diff --git a/src/Init/Meta.lean b/src/Init/Meta.lean index 146b068820c3..4b25a919e8c8 100644 --- a/src/Init/Meta.lean +++ b/src/Init/Meta.lean @@ -374,6 +374,9 @@ partial def structEq : Syntax → Syntax → Bool instance : BEq Lean.Syntax := ⟨structEq⟩ instance : BEq (Lean.TSyntax k) := ⟨(·.raw == ·.raw)⟩ +/-- +Finds the first `SourceInfo` from the back of `stx` or `none` if no `SourceInfo` can be found. +-/ partial def getTailInfo? : Syntax → Option SourceInfo | atom info _ => info | ident info .. => info @@ -382,14 +385,39 @@ partial def getTailInfo? : Syntax → Option SourceInfo | node info _ _ => info | _ => none +/-- +Finds the first `SourceInfo` from the back of `stx` or `SourceInfo.none` +if no `SourceInfo` can be found. +-/ def getTailInfo (stx : Syntax) : SourceInfo := stx.getTailInfo?.getD SourceInfo.none +/-- +Finds the trailing size of the first `SourceInfo` from the back of `stx`. +If no `SourceInfo` can be found or the first `SourceInfo` from the back of `stx` contains no +trailing whitespace, the result is `0`. +-/ def getTrailingSize (stx : Syntax) : Nat := match stx.getTailInfo? with | some (SourceInfo.original (trailing := trailing) ..) => trailing.bsize | _ => 0 +/-- +Finds the trailing whitespace substring of the first `SourceInfo` from the back of `stx`. +If no `SourceInfo` can be found or the first `SourceInfo` from the back of `stx` contains +no trailing whitespace, the result is `none`. +-/ +def getTrailing? (stx : Syntax) : Option Substring := + stx.getTailInfo.getTrailing? + +/-- +Finds the tail position of the trailing whitespace of the first `SourceInfo` from the back of `stx`. +If no `SourceInfo` can be found or the first `SourceInfo` from the back of `stx` contains +no trailing whitespace and lacks a tail position, the result is `none`. +-/ +def getTrailingTailPos? (stx : Syntax) (canonicalOnly := false) : Option String.Pos := + stx.getTailInfo.getTrailingTailPos? canonicalOnly + /-- Return substring of original input covering `stx`. Result is meaningful only if all involved `SourceInfo.original`s refer to the same string (as is the case after parsing). -/ diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index 39618b16d1ca..35fd5c7ca442 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -3654,7 +3654,8 @@ namespace SourceInfo /-- Gets the position information from a `SourceInfo`, if available. -If `originalOnly` is true, then `.synthetic` syntax will also return `none`. +If `canonicalOnly` is true, then `.synthetic` syntax with `canonical := false` +will also return `none`. -/ def getPos? (info : SourceInfo) (canonicalOnly := false) : Option String.Pos := match info, canonicalOnly with @@ -3665,7 +3666,8 @@ def getPos? (info : SourceInfo) (canonicalOnly := false) : Option String.Pos := /-- Gets the end position information from a `SourceInfo`, if available. -If `originalOnly` is true, then `.synthetic` syntax will also return `none`. +If `canonicalOnly` is true, then `.synthetic` syntax with `canonical := false` +will also return `none`. -/ def getTailPos? (info : SourceInfo) (canonicalOnly := false) : Option String.Pos := match info, canonicalOnly with @@ -3674,6 +3676,24 @@ def getTailPos? (info : SourceInfo) (canonicalOnly := false) : Option String.Pos | synthetic (endPos := endPos) .., false => some endPos | _, _ => none +/-- +Gets the substring representing the trailing whitespace of a `SourceInfo`, if available. +-/ +def getTrailing? (info : SourceInfo) : Option Substring := + match info with + | original (trailing := trailing) .. => some trailing + | _ => none + +/-- +Gets the end position information of the trailing whitespace of a `SourceInfo`, if available. +If `canonicalOnly` is true, then `.synthetic` syntax with `canonical := false` +will also return `none`. +-/ +def getTrailingTailPos? (info : SourceInfo) (canonicalOnly := false) : Option String.Pos := + match info.getTrailing? with + | some trailing => some trailing.stopPos + | none => info.getTailPos? canonicalOnly + end SourceInfo /-- @@ -3972,7 +3992,6 @@ position information. def getPos? (stx : Syntax) (canonicalOnly := false) : Option String.Pos := stx.getHeadInfo.getPos? canonicalOnly - /-- Get the ending position of the syntax, if possible. If `canonicalOnly` is true, non-canonical `synthetic` nodes are treated as not carrying diff --git a/src/Lean/Data/Lsp/Basic.lean b/src/Lean/Data/Lsp/Basic.lean index 2403101173c3..55fb5dd8ab85 100644 --- a/src/Lean/Data/Lsp/Basic.lean +++ b/src/Lean/Data/Lsp/Basic.lean @@ -365,6 +365,7 @@ structure TextDocumentRegistrationOptions where inductive MarkupKind where | plaintext | markdown + deriving DecidableEq, Hashable instance : FromJson MarkupKind := ⟨fun | str "plaintext" => Except.ok MarkupKind.plaintext @@ -378,7 +379,7 @@ instance : ToJson MarkupKind := ⟨fun structure MarkupContent where kind : MarkupKind value : String - deriving ToJson, FromJson + deriving ToJson, FromJson, DecidableEq, Hashable /-- Reference to the progress of some in-flight piece of work. diff --git a/src/Lean/Data/Lsp/LanguageFeatures.lean b/src/Lean/Data/Lsp/LanguageFeatures.lean index f5c46a13d68d..038d200e4d86 100644 --- a/src/Lean/Data/Lsp/LanguageFeatures.lean +++ b/src/Lean/Data/Lsp/LanguageFeatures.lean @@ -25,7 +25,7 @@ inductive CompletionItemKind where | unit | value | enum | keyword | snippet | color | file | reference | folder | enumMember | constant | struct | event | operator | typeParameter - deriving Inhabited, DecidableEq, Repr + deriving Inhabited, DecidableEq, Repr, Hashable instance : ToJson CompletionItemKind where toJson a := toJson (a.toCtorIdx + 1) @@ -39,11 +39,11 @@ structure InsertReplaceEdit where newText : String insert : Range replace : Range - deriving FromJson, ToJson + deriving FromJson, ToJson, BEq, Hashable inductive CompletionItemTag where | deprecated - deriving Inhabited, DecidableEq, Repr + deriving Inhabited, DecidableEq, Repr, Hashable instance : ToJson CompletionItemTag where toJson t := toJson (t.toCtorIdx + 1) @@ -73,7 +73,7 @@ structure CompletionItem where commitCharacters? : string[] command? : Command -/ - deriving FromJson, ToJson, Inhabited + deriving FromJson, ToJson, Inhabited, BEq, Hashable structure CompletionList where isIncomplete : Bool diff --git a/src/Lean/Elab/App.lean b/src/Lean/Elab/App.lean index e68ba7362b94..eaad3a6053df 100644 --- a/src/Lean/Elab/App.lean +++ b/src/Lean/Elab/App.lean @@ -1399,8 +1399,8 @@ private def elabAppLValsAux (namedArgs : Array NamedArg) (args : Array Arg) (exp let rec loop : Expr → List LVal → TermElabM Expr | f, [] => elabAppArgs f namedArgs args expectedType? explicit ellipsis | f, lval::lvals => do - if let LVal.fieldName (fullRef := fullRef) .. := lval then - addDotCompletionInfo fullRef f expectedType? + if let LVal.fieldName (ref := ref) .. := lval then + addDotCompletionInfo ref f expectedType? let hasArgs := !namedArgs.isEmpty || !args.isEmpty let (f, lvalRes) ← resolveLVal f lval hasArgs match lvalRes with @@ -1650,6 +1650,14 @@ private def getSuccesses (candidates : Array (TermElabResult Expr)) : TermElabM -/ private def mergeFailures (failures : Array (TermElabResult Expr)) : TermElabM α := do let exs := failures.map fun | .error ex _ => ex | _ => unreachable! + let trees := failures.map (fun | .error _ s => s.meta.core.infoState.trees | _ => unreachable!) + |>.filterMap (·[0]?) + -- Retain partial `InfoTree` subtrees in an `.ofChoiceInfo` node in case of multiple failures. + -- This ensures that the language server still has `Info` to work with when multiple overloaded + -- elaborators fail. + withInfoContext (mkInfo := pure <| .ofChoiceInfo { elaborator := .anonymous, stx := ← getRef }) do + for tree in trees do + pushInfoTree tree throwErrorWithNestedErrors "overloaded" exs private def elabAppAux (f : Syntax) (namedArgs : Array NamedArg) (args : Array Arg) (ellipsis : Bool) (expectedType? : Option Expr) : TermElabM Expr := do diff --git a/src/Lean/Elab/BuiltinTerm.lean b/src/Lean/Elab/BuiltinTerm.lean index f9a225b9b681..a25106ed2885 100644 --- a/src/Lean/Elab/BuiltinTerm.lean +++ b/src/Lean/Elab/BuiltinTerm.lean @@ -42,16 +42,15 @@ private def elabOptLevel (stx : Syntax) : TermElabM Level := @[builtin_term_elab «completion»] def elabCompletion : TermElab := fun stx expectedType? => do /- `ident.` is ambiguous in Lean, we may try to be completing a declaration name or access a "field". -/ if stx[0].isIdent then - /- If we can elaborate the identifier successfully, we assume it is a dot-completion. Otherwise, we treat it as - identifier completion with a dangling `.`. - Recall that the server falls back to identifier completion when dot-completion fails. -/ + -- Add both an `id` and a `dot` `CompletionInfo` and have the language server figure out which + -- one to use. + addCompletionInfo <| CompletionInfo.id stx stx[0].getId (danglingDot := true) (← getLCtx) expectedType? let s ← saveState try let e ← elabTerm stx[0] none addDotCompletionInfo stx e expectedType? catch _ => s.restore - addCompletionInfo <| CompletionInfo.id stx stx[0].getId (danglingDot := true) (← getLCtx) expectedType? throwErrorAt stx[1] "invalid field notation, identifier or numeral expected" else elabPipeCompletion stx expectedType? @@ -328,7 +327,7 @@ private def mkSilentAnnotationIfHole (e : Expr) : TermElabM Expr := do @[builtin_term_elab withAnnotateTerm] def elabWithAnnotateTerm : TermElab := fun stx expectedType? => do match stx with | `(with_annotate_term $stx $e) => - withInfoContext' stx (elabTerm e expectedType?) (mkTermInfo .anonymous (expectedType? := expectedType?) stx) + withTermInfoContext' .anonymous stx (expectedType? := expectedType?) (elabTerm e expectedType?) | _ => throwUnsupportedSyntax private unsafe def evalFilePathUnsafe (stx : Syntax) : TermElabM System.FilePath := diff --git a/src/Lean/Elab/Extra.lean b/src/Lean/Elab/Extra.lean index 38372ebae9da..c4d792c7c38e 100644 --- a/src/Lean/Elab/Extra.lean +++ b/src/Lean/Elab/Extra.lean @@ -327,15 +327,18 @@ private def toExprCore (t : Tree) : TermElabM Expr := do | .term _ trees e => modifyInfoState (fun s => { s with trees := s.trees ++ trees }); return e | .binop ref kind f lhs rhs => - withRef ref <| withInfoContext' ref (mkInfo := mkTermInfo .anonymous ref) do - mkBinOp (kind == .lazy) f (← toExprCore lhs) (← toExprCore rhs) + withRef ref <| + withTermInfoContext' .anonymous ref do + mkBinOp (kind == .lazy) f (← toExprCore lhs) (← toExprCore rhs) | .unop ref f arg => - withRef ref <| withInfoContext' ref (mkInfo := mkTermInfo .anonymous ref) do - mkUnOp f (← toExprCore arg) + withRef ref <| + withTermInfoContext' .anonymous ref do + mkUnOp f (← toExprCore arg) | .macroExpansion macroName stx stx' nested => - withRef stx <| withInfoContext' stx (mkInfo := mkTermInfo macroName stx) do - withMacroExpansion stx stx' do - toExprCore nested + withRef stx <| + withTermInfoContext' macroName stx <| + withMacroExpansion stx stx' <| + toExprCore nested /-- Auxiliary function to decide whether we should coerce `f`'s argument to `maxType` or not. diff --git a/src/Lean/Elab/InfoTree/Main.lean b/src/Lean/Elab/InfoTree/Main.lean index 5ba988886f3b..eb129079d45e 100644 --- a/src/Lean/Elab/InfoTree/Main.lean +++ b/src/Lean/Elab/InfoTree/Main.lean @@ -139,12 +139,16 @@ def TermInfo.runMetaM (info : TermInfo) (ctx : ContextInfo) (x : MetaM α) : IO def TermInfo.format (ctx : ContextInfo) (info : TermInfo) : IO Format := do info.runMetaM ctx do - let ty : Format ← try - Meta.ppExpr (← Meta.inferType info.expr) - catch _ => - pure "" + let ty : Format ← + try + Meta.ppExpr (← Meta.inferType info.expr) + catch _ => + pure "" return f!"{← Meta.ppExpr info.expr} {if info.isBinder then "(isBinder := true) " else ""}: {ty} @ {formatElabInfo ctx info.toElabInfo}" +def PartialTermInfo.format (ctx : ContextInfo) (info : PartialTermInfo) : Format := + f!"Partial term @ {formatElabInfo ctx info.toElabInfo}" + def CompletionInfo.format (ctx : ContextInfo) (info : CompletionInfo) : IO Format := match info with | .dot i (expectedType? := expectedType?) .. => return f!"[.] {← i.format ctx} : {expectedType?}" @@ -191,9 +195,13 @@ def FieldRedeclInfo.format (ctx : ContextInfo) (info : FieldRedeclInfo) : Format def OmissionInfo.format (ctx : ContextInfo) (info : OmissionInfo) : IO Format := do return f!"Omission @ {← TermInfo.format ctx info.toTermInfo}\nReason: {info.reason}" +def ChoiceInfo.format (ctx : ContextInfo) (info : ChoiceInfo) : Format := + f!"Choice @ {formatElabInfo ctx info.toElabInfo}" + def Info.format (ctx : ContextInfo) : Info → IO Format | ofTacticInfo i => i.format ctx | ofTermInfo i => i.format ctx + | ofPartialTermInfo i => pure <| i.format ctx | ofCommandInfo i => i.format ctx | ofMacroExpansionInfo i => i.format ctx | ofOptionInfo i => i.format ctx @@ -204,10 +212,12 @@ def Info.format (ctx : ContextInfo) : Info → IO Format | ofFVarAliasInfo i => pure <| i.format | ofFieldRedeclInfo i => pure <| i.format ctx | ofOmissionInfo i => i.format ctx + | ofChoiceInfo i => pure <| i.format ctx def Info.toElabInfo? : Info → Option ElabInfo | ofTacticInfo i => some i.toElabInfo | ofTermInfo i => some i.toElabInfo + | ofPartialTermInfo i => some i.toElabInfo | ofCommandInfo i => some i.toElabInfo | ofMacroExpansionInfo _ => none | ofOptionInfo _ => none @@ -218,6 +228,7 @@ def Info.toElabInfo? : Info → Option ElabInfo | ofFVarAliasInfo _ => none | ofFieldRedeclInfo _ => none | ofOmissionInfo i => some i.toElabInfo + | ofChoiceInfo i => some i.toElabInfo /-- Helper function for propagating the tactic metavariable context to its children nodes. @@ -311,24 +322,36 @@ def realizeGlobalNameWithInfos (ref : Syntax) (id : Name) : CoreM (List (Name × addConstInfo ref n return ns -/-- Use this to descend a node on the infotree that is being built. +/-- +Adds a node containing the `InfoTree`s generated by `x` to the `InfoTree`s in `m`. + +If `x` succeeds and `mkInfo` yields an `Info`, the `InfoTree`s of `x` become subtrees of a node +containing the `Info` produced by `mkInfo`, which is then added to the `InfoTree`s in `m`. +If `x` succeeds and `mkInfo` yields an `MVarId`, the `InfoTree`s of `x` are discarded and a `hole` +node is added to the `InfoTree`s in `m`. +If `x` fails, the `InfoTree`s of `x` become subtrees of a node containing the `Info` produced by +`mkInfoOnError`, which is then added to the `InfoTree`s in `m`. -It saves the current list of trees `t₀` and resets it and then runs `x >>= mkInfo`, producing either an `i : Info` or a hole id. -Running `x >>= mkInfo` will modify the trees state and produce a new list of trees `t₁`. -In the `i : Info` case, `t₁` become the children of a node `node i t₁` that is appended to `t₀`. - -/ -def withInfoContext' [MonadFinally m] (x : m α) (mkInfo : α → m (Sum Info MVarId)) : m α := do +The `InfoTree`s in `m` are reset before `x` is executed and restored with the addition of a new tree +after `x` is executed. +-/ +def withInfoContext' + [MonadFinally m] + (x : m α) + (mkInfo : α → m (Sum Info MVarId)) + (mkInfoOnError : m Info) : + m α := do if (← getInfoState).enabled then let treesSaved ← getResetInfoTrees Prod.fst <$> MonadFinally.tryFinally' x fun a? => do - match a? with - | none => modifyInfoTrees fun _ => treesSaved - | some a => - let info ← mkInfo a - modifyInfoTrees fun trees => - match info with - | Sum.inl info => treesSaved.push <| InfoTree.node info trees - | Sum.inr mvarId => treesSaved.push <| InfoTree.hole mvarId + let info ← do + match a? with + | none => pure <| .inl <| ← mkInfoOnError + | some a => mkInfo a + modifyInfoTrees fun trees => + match info with + | Sum.inl info => treesSaved.push <| InfoTree.node info trees + | Sum.inr mvarId => treesSaved.push <| InfoTree.hole mvarId else x diff --git a/src/Lean/Elab/InfoTree/Types.lean b/src/Lean/Elab/InfoTree/Types.lean index 02a189bb28bb..5ecdd7dde97d 100644 --- a/src/Lean/Elab/InfoTree/Types.lean +++ b/src/Lean/Elab/InfoTree/Types.lean @@ -70,6 +70,18 @@ structure TermInfo extends ElabInfo where isBinder : Bool := false deriving Inhabited +/-- +Used instead of `TermInfo` when a term couldn't successfully be elaborated, +and so there is no complete expression available. + +The main purpose of `PartialTermInfo` is to ensure that the sub-`InfoTree`s of a failed elaborator +are retained so that they can still be used in the language server. +-/ +structure PartialTermInfo extends ElabInfo where + lctx : LocalContext -- The local context when the term was elaborated. + expectedType? : Option Expr + deriving Inhabited + structure CommandInfo extends ElabInfo where deriving Inhabited @@ -79,7 +91,7 @@ inductive CompletionInfo where | dot (termInfo : TermInfo) (expectedType? : Option Expr) | id (stx : Syntax) (id : Name) (danglingDot : Bool) (lctx : LocalContext) (expectedType? : Option Expr) | dotId (stx : Syntax) (id : Name) (lctx : LocalContext) (expectedType? : Option Expr) - | fieldId (stx : Syntax) (id : Name) (lctx : LocalContext) (structName : Name) + | fieldId (stx : Syntax) (id : Option Name) (lctx : LocalContext) (structName : Name) | namespaceId (stx : Syntax) | option (stx : Syntax) | endSection (stx : Syntax) (scopeNames : List String) @@ -165,10 +177,18 @@ regular delaboration settings. structure OmissionInfo extends TermInfo where reason : String +/-- +Indicates that all overloaded elaborators failed. The subtrees of a `ChoiceInfo` node are the +partial `InfoTree`s of those failed elaborators. Retaining these partial `InfoTree`s helps +the language server provide interactivity even when all overloaded elaborators failed. +-/ +structure ChoiceInfo extends ElabInfo where + /-- Header information for a node in `InfoTree`. -/ inductive Info where | ofTacticInfo (i : TacticInfo) | ofTermInfo (i : TermInfo) + | ofPartialTermInfo (i : PartialTermInfo) | ofCommandInfo (i : CommandInfo) | ofMacroExpansionInfo (i : MacroExpansionInfo) | ofOptionInfo (i : OptionInfo) @@ -179,6 +199,7 @@ inductive Info where | ofFVarAliasInfo (i : FVarAliasInfo) | ofFieldRedeclInfo (i : FieldRedeclInfo) | ofOmissionInfo (i : OmissionInfo) + | ofChoiceInfo (i : ChoiceInfo) deriving Inhabited /-- The InfoTree is a structure that is generated during elaboration and used diff --git a/src/Lean/Elab/LetRec.lean b/src/Lean/Elab/LetRec.lean index 03c2adc4977c..22663d005548 100644 --- a/src/Lean/Elab/LetRec.lean +++ b/src/Lean/Elab/LetRec.lean @@ -90,9 +90,12 @@ private def elabLetRecDeclValues (view : LetRecView) : TermElabM (Array Expr) := for i in [0:view.binderIds.size] do addLocalVarInfo view.binderIds[i]! xs[i]! withDeclName view.declName do - withInfoContext' view.valStx (mkInfo := (pure <| .inl <| mkBodyInfo view.valStx ·)) do - let value ← elabTermEnsuringType view.valStx type - mkLambdaFVars xs value + withInfoContext' view.valStx + (mkInfo := (pure <| .inl <| mkBodyInfo view.valStx ·)) + (mkInfoOnError := (pure <| mkBodyInfo view.valStx none)) + do + let value ← elabTermEnsuringType view.valStx type + mkLambdaFVars xs value private def registerLetRecsToLift (views : Array LetRecDeclView) (fvars : Array Expr) (values : Array Expr) : TermElabM Unit := do let letRecsToLiftCurr := (← get).letRecsToLift diff --git a/src/Lean/Elab/Match.lean b/src/Lean/Elab/Match.lean index e9e1085a9493..3fb9c33c86ce 100644 --- a/src/Lean/Elab/Match.lean +++ b/src/Lean/Elab/Match.lean @@ -644,7 +644,7 @@ where if inaccessible? p |>.isSome then return mkMData k (← withReader (fun _ => true) (go b)) else if let some (stx, p) := patternWithRef? p then - Elab.withInfoContext' (go p) fun p => do + Elab.withInfoContext' (go p) (mkInfoOnError := mkPartialTermInfo .anonymous stx) fun p => do /- If `p` is a free variable and we are not inside of an "inaccessible" pattern, this `p` is a binder. -/ mkTermInfo Name.anonymous stx p (isBinder := p.isFVar && !(← read)) else diff --git a/src/Lean/Elab/MutualDef.lean b/src/Lean/Elab/MutualDef.lean index d31248655bb0..55b057eb6bce 100644 --- a/src/Lean/Elab/MutualDef.lean +++ b/src/Lean/Elab/MutualDef.lean @@ -283,7 +283,7 @@ private partial def withFunLocalDecls {α} (headers : Array DefViewElabHeader) ( loop 0 #[] private def expandWhereStructInst : Macro - | `(Parser.Command.whereStructInst|where $[$decls:letDecl];* $[$whereDecls?:whereDecls]?) => do + | whereStx@`(Parser.Command.whereStructInst|where%$whereTk $[$decls:letDecl];* $[$whereDecls?:whereDecls]?) => do let letIdDecls ← decls.mapM fun stx => match stx with | `(letDecl|$_decl:letPatDecl) => Macro.throwErrorAt stx "patterns are not allowed here" | `(letDecl|$decl:letEqnsDecl) => expandLetEqnsDecl decl (useExplicit := false) @@ -300,7 +300,30 @@ private def expandWhereStructInst : Macro `(structInstField|$id:ident := $val) | stx@`(letIdDecl|_ $_* $[: $_]? := $_) => Macro.throwErrorAt stx "'_' is not allowed here" | _ => Macro.throwUnsupported + + let startOfStructureTkInfo : SourceInfo := + match whereTk.getPos? with + | some pos => .synthetic pos ⟨pos.byteIdx + 1⟩ true + | none => .none + -- Position the closing `}` at the end of the trailing whitespace of `where $[$_:letDecl];*`. + -- We need an accurate range of the generated structure instance in the generated `TermInfo` + -- so that we can determine the expected type in structure field completion. + let structureStxTailInfo := + whereStx[1].getTailInfo? + <|> whereStx[0].getTailInfo? + let endOfStructureTkInfo : SourceInfo := + match structureStxTailInfo with + | some (SourceInfo.original _ _ trailing _) => + let tokenPos := trailing.str.prev trailing.stopPos + let tokenEndPos := trailing.stopPos + .synthetic tokenPos tokenEndPos true + | _ => .none + let body ← `(structInst| { $structInstFields,* }) + let body := body.raw.setInfo <| + match startOfStructureTkInfo.getPos?, endOfStructureTkInfo.getTailPos? with + | some startPos, some endPos => .synthetic startPos endPos true + | _, _ => .none match whereDecls? with | some whereDecls => expandWhereDecls whereDecls body | none => return body @@ -417,12 +440,15 @@ private def elabFunValues (headers : Array DefViewElabHeader) (vars : Array Expr -- Store instantiated body in info tree for the benefit of the unused variables linter -- and other metaprograms that may want to inspect it without paying for the instantiation -- again - withInfoContext' valStx (mkInfo := (pure <| .inl <| mkBodyInfo valStx ·)) do - -- synthesize mvars here to force the top-level tactic block (if any) to run - let val ← elabTermEnsuringType valStx type <* synthesizeSyntheticMVarsNoPostponing - -- NOTE: without this `instantiatedMVars`, `mkLambdaFVars` may leave around a redex that - -- leads to more section variables being included than necessary - instantiateMVarsProfiling val + withInfoContext' valStx + (mkInfo := (pure <| .inl <| mkBodyInfo valStx ·)) + (mkInfoOnError := (pure <| mkBodyInfo valStx none)) + do + -- synthesize mvars here to force the top-level tactic block (if any) to run + let val ← elabTermEnsuringType valStx type <* synthesizeSyntheticMVarsNoPostponing + -- NOTE: without this `instantiatedMVars`, `mkLambdaFVars` may leave around a redex that + -- leads to more section variables being included than necessary + instantiateMVarsProfiling val let val ← mkLambdaFVars xs val if linter.unusedSectionVars.get (← getOptions) && !header.type.hasSorry && !val.hasSorry then let unusedVars ← vars.filterMapM fun var => do diff --git a/src/Lean/Elab/StructInst.lean b/src/Lean/Elab/StructInst.lean index bbfd5d58a009..5f0ad77beafe 100644 --- a/src/Lean/Elab/StructInst.lean +++ b/src/Lean/Elab/StructInst.lean @@ -144,7 +144,7 @@ private def getStructSource (structStx : Syntax) : TermElabM Source := ``` -/ private def isModifyOp? (stx : Syntax) : TermElabM (Option Syntax) := do - let s? ← stx[2].getSepArgs.foldlM (init := none) fun s? arg => do + let s? ← stx[2][0].getSepArgs.foldlM (init := none) fun s? arg => do /- arg is of the form `structInstFieldAbbrev <|> structInstField` -/ if arg.getKind == ``Lean.Parser.Term.structInstField then /- Remark: the syntax for `structInstField` is @@ -199,7 +199,7 @@ private def elabModifyOp (stx modifyOp : Syntax) (sources : Array ExplicitSource let valField := modifyOp.setArg 0 <| mkNode ``Parser.Term.structInstLVal #[valFirst, valRest] let valSource := mkSourcesWithSyntax #[s] let val := stx.setArg 1 valSource - let val := val.setArg 2 <| mkNullNode #[valField] + let val := val.setArg 2 <| mkNode ``Parser.Term.structInstFields #[mkNullNode #[valField]] trace[Elab.struct.modifyOp] "{stx}\nval: {val}" cont val @@ -359,7 +359,7 @@ private def mkStructView (stx : Syntax) (structName : Name) (source : Source) : /- Recall that `stx` is of the form ``` leading_parser "{" >> optional (atomic (sepBy1 termParser ", " >> " with ")) - >> sepByIndent (structInstFieldAbbrev <|> structInstField) ... + >> structInstFields (sepByIndent (structInstFieldAbbrev <|> structInstField) ...) >> optional ".." >> optional (" : " >> termParser) >> " }" @@ -367,7 +367,7 @@ private def mkStructView (stx : Syntax) (structName : Name) (source : Source) : This method assumes that `structInstFieldAbbrev` had already been expanded. -/ - let fields ← stx[2].getSepArgs.toList.mapM fun fieldStx => do + let fields ← stx[2][0].getSepArgs.toList.mapM fun fieldStx => do let val := fieldStx[2] let first ← toFieldLHS fieldStx[0][0] let rest ← fieldStx[0][1].getArgs.toList.mapM toFieldLHS @@ -511,7 +511,9 @@ mutual let valStx := s.ref -- construct substructure syntax using s.ref as template let valStx := valStx.setArg 4 mkNullNode -- erase optional expected type let args := substructFields.toArray.map (·.toSyntax) - let valStx := valStx.setArg 2 (mkNullNode <| mkSepArray args (mkAtom ",")) + let fieldsStx := mkNode ``Parser.Term.structInstFields + #[mkNullNode <| mkSepArray args (mkAtom ",")] + let valStx := valStx.setArg 2 fieldsStx let valStx ← updateSource valStx return { field with lhs := [field.lhs.head!], val := FieldVal.term valStx } /-- diff --git a/src/Lean/Elab/Tactic/Config.lean b/src/Lean/Elab/Tactic/Config.lean index 160b8cb0928c..9961bf24b270 100644 --- a/src/Lean/Elab/Tactic/Config.lean +++ b/src/Lean/Elab/Tactic/Config.lean @@ -114,6 +114,13 @@ private def elabConfig (recover : Bool) (structName : Name) (items : Array Confi let e ← Term.withSynthesize <| Term.elabTermEnsuringType stx (mkConst structName) instantiateMVars e +section +-- We automatically disable the following option for `macro`s but the subsequent `def` both contains +-- a quotation and is called only by `macro`s, so we disable the option for it manually. Note that +-- we can't use `in` as it is parsed as a single command and so the option would not influence the +-- parser. +set_option internal.parseQuotWithCurrentStage false + private def mkConfigElaborator (doc? : Option (TSyntax ``Parser.Command.docComment)) (elabName type monadName : Ident) (adapt recover : Term) : MacroM (TSyntax `command) := do @@ -148,6 +155,8 @@ private def mkConfigElaborator throwError msg go) +end + /-! `declare_config_elab elabName TypeName` declares a function `elabName : Syntax → TacticM TypeName` that elaborates a tactic configuration. diff --git a/src/Lean/Elab/Term.lean b/src/Lean/Elab/Term.lean index 62e688d21515..97f2873d512f 100644 --- a/src/Lean/Elab/Term.lean +++ b/src/Lean/Elab/Term.lean @@ -1298,13 +1298,20 @@ def isTacticOrPostponedHole? (e : Expr) : TermElabM (Option MVarId) := do | _ => return none | _ => pure none -def mkTermInfo (elaborator : Name) (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none) (lctx? : Option LocalContext := none) (isBinder := false) : TermElabM (Sum Info MVarId) := do +def mkTermInfo (elaborator : Name) (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none) + (lctx? : Option LocalContext := none) (isBinder := false) : + TermElabM (Sum Info MVarId) := do match (← isTacticOrPostponedHole? e) with | some mvarId => return Sum.inr mvarId | none => let e := removeSaveInfoAnnotation e return Sum.inl <| Info.ofTermInfo { elaborator, lctx := lctx?.getD (← getLCtx), expr := e, stx, expectedType?, isBinder } +def mkPartialTermInfo (elaborator : Name) (stx : Syntax) (expectedType? : Option Expr := none) + (lctx? : Option LocalContext := none) : + TermElabM Info := do + return Info.ofPartialTermInfo { elaborator, lctx := lctx?.getD (← getLCtx), stx, expectedType? } + /-- Pushes a new leaf node to the info tree associating the expression `e` to the syntax `stx`. As a result, when the user hovers over `stx` they will see the type of `e`, and if `e` @@ -1326,41 +1333,54 @@ def addTermInfo (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none) if (← read).inPattern && !force then return mkPatternWithRef e stx else - withInfoContext' (pure ()) (fun _ => mkTermInfo elaborator stx e expectedType? lctx? isBinder) |> discard + discard <| withInfoContext' + (pure ()) + (fun _ => mkTermInfo elaborator stx e expectedType? lctx? isBinder) + (mkPartialTermInfo elaborator stx expectedType? lctx?) return e def addTermInfo' (stx : Syntax) (e : Expr) (expectedType? : Option Expr := none) (lctx? : Option LocalContext := none) (elaborator := Name.anonymous) (isBinder := false) : TermElabM Unit := discard <| addTermInfo stx e expectedType? lctx? elaborator isBinder -def withInfoContext' (stx : Syntax) (x : TermElabM Expr) (mkInfo : Expr → TermElabM (Sum Info MVarId)) : TermElabM Expr := do +def withInfoContext' (stx : Syntax) (x : TermElabM Expr) + (mkInfo : Expr → TermElabM (Sum Info MVarId)) (mkInfoOnError : TermElabM Info) : + TermElabM Expr := do if (← read).inPattern then let e ← x return mkPatternWithRef e stx else - Elab.withInfoContext' x mkInfo + Elab.withInfoContext' x mkInfo mkInfoOnError /-- Info node capturing `def/let rec` bodies, used by the unused variables linter. -/ structure BodyInfo where - /-- The body as a fully elaborated term. -/ - value : Expr + /-- The body as a fully elaborated term. `none` if the body failed to elaborate. -/ + value? : Option Expr deriving TypeName /-- Creates an `Info.ofCustomInfo` node backed by a `BodyInfo`. -/ -def mkBodyInfo (stx : Syntax) (value : Expr) : Info := - .ofCustomInfo { stx, value := .mk { value : BodyInfo } } +def mkBodyInfo (stx : Syntax) (value? : Option Expr) : Info := + .ofCustomInfo { stx, value := .mk { value? : BodyInfo } } /-- Extracts a `BodyInfo` custom info. -/ def getBodyInfo? : Info → Option BodyInfo | .ofCustomInfo { value, .. } => value.get? BodyInfo | _ => none +def withTermInfoContext' (elaborator : Name) (stx : Syntax) (x : TermElabM Expr) + (expectedType? : Option Expr := none) (lctx? : Option LocalContext := none) + (isBinder : Bool := false) : + TermElabM Expr := + withInfoContext' stx x + (mkTermInfo elaborator stx (expectedType? := expectedType?) (lctx? := lctx?) (isBinder := isBinder)) + (mkPartialTermInfo elaborator stx (expectedType? := expectedType?) (lctx? := lctx?)) + /-- Postpone the elaboration of `stx`, return a metavariable that acts as a placeholder, and ensures the info tree is updated and a hole id is introduced. When `stx` is elaborated, new info nodes are created and attached to the new hole id in the info tree. -/ def postponeElabTerm (stx : Syntax) (expectedType? : Option Expr) : TermElabM Expr := do - withInfoContext' stx (mkInfo := mkTermInfo .anonymous (expectedType? := expectedType?) stx) do + withTermInfoContext' .anonymous stx (expectedType? := expectedType?) do postponeElabTermCore stx expectedType? /-- @@ -1372,7 +1392,7 @@ private def elabUsingElabFnsAux (s : SavedState) (stx : Syntax) (expectedType? : | (elabFn::elabFns) => try -- record elaborator in info tree, but only when not backtracking to other elaborators (outer `try`) - withInfoContext' stx (mkInfo := mkTermInfo elabFn.declName (expectedType? := expectedType?) stx) + withTermInfoContext' elabFn.declName stx (expectedType? := expectedType?) (try elabFn.value stx expectedType? catch ex => match ex with @@ -1755,7 +1775,7 @@ private partial def elabTermAux (expectedType? : Option Expr) (catchExPostpone : let result ← match (← liftMacroM (expandMacroImpl? env stx)) with | some (decl, stxNew?) => let stxNew ← liftMacroM <| liftExcept stxNew? - withInfoContext' stx (mkInfo := mkTermInfo decl (expectedType? := expectedType?) stx) <| + withTermInfoContext' decl stx (expectedType? := expectedType?) <| withMacroExpansion stx stxNew <| withRef stxNew <| elabTermAux expectedType? catchExPostpone implicitLambda stxNew diff --git a/src/Lean/Linter/UnusedVariables.lean b/src/Lean/Linter/UnusedVariables.lean index 5cefd948e5c7..74bcb7260e5b 100644 --- a/src/Lean/Linter/UnusedVariables.lean +++ b/src/Lean/Linter/UnusedVariables.lean @@ -393,16 +393,17 @@ where | .ofCustomInfo ti => if !linter.unusedVariables.analyzeTactics.get ci.options then if let some bodyInfo := ti.value.get? Elab.Term.BodyInfo then - -- the body is the only `Expr` we will analyze in this case - -- NOTE: we include it even if no tactics are present as at least for parameters we want - -- to lint only truly unused binders - let (e, _) := instantiateMVarsCore ci.mctx bodyInfo.value - modify fun s => { s with - assignments := s.assignments.push (.insert {} ⟨.anonymous⟩ e) } - let tacticsPresent := children.any (·.findInfo? (· matches .ofTacticInfo ..) |>.isSome) - withReader (· || tacticsPresent) do - go children.toArray ci - return false + if let some value := bodyInfo.value? then + -- the body is the only `Expr` we will analyze in this case + -- NOTE: we include it even if no tactics are present as at least for parameters we want + -- to lint only truly unused binders + let (e, _) := instantiateMVarsCore ci.mctx value + modify fun s => { s with + assignments := s.assignments.push (.insert {} ⟨.anonymous⟩ e) } + let tacticsPresent := children.any (·.findInfo? (· matches .ofTacticInfo ..) |>.isSome) + withReader (· || tacticsPresent) do + go children.toArray ci + return false | .ofTermInfo ti => if ignored then return true match ti.expr with diff --git a/src/Lean/Parser/Command.lean b/src/Lean/Parser/Command.lean index 2eab73ec025f..825cab538e04 100644 --- a/src/Lean/Parser/Command.lean +++ b/src/Lean/Parser/Command.lean @@ -137,7 +137,7 @@ def declValEqns := leading_parser def whereStructField := leading_parser Term.letDecl def whereStructInst := leading_parser - ppIndent ppSpace >> "where" >> sepByIndent (ppGroup whereStructField) "; " (allowTrailingSep := true) >> + ppIndent ppSpace >> "where" >> Term.structInstFields (sepByIndent (ppGroup whereStructField) "; " (allowTrailingSep := true)) >> optional Term.whereDecls /-- `declVal` matches the right-hand side of a declaration, one of: * `:= expr` (a "simple declaration") diff --git a/src/Lean/Parser/Term.lean b/src/Lean/Parser/Term.lean index 15d4cc1a0d30..66b9a4d5d27f 100644 --- a/src/Lean/Parser/Term.lean +++ b/src/Lean/Parser/Term.lean @@ -281,6 +281,12 @@ def structInstFieldAbbrev := leading_parser atomic (ident >> notFollowedBy ("." <|> ":=" <|> symbol "[") "invalid field abbreviation") def optEllipsis := leading_parser optional " .." +/- +Tags the structure instance field syntax with a `Lean.Parser.Term.structInstFields` syntax node. +This node is used to enable structure instance field completion in the whitespace +of a structure instance notation. +-/ +def structInstFields (p : Parser) : Parser := node `Lean.Parser.Term.structInstFields p /-- Structure instance. `{ x := e, ... }` assigns `e` to field `x`, which may be inherited. If `e` is itself a variable called `x`, it can be elided: @@ -292,7 +298,7 @@ The structure type can be specified if not inferable: -/ @[builtin_term_parser] def structInst := leading_parser "{ " >> withoutPosition (optional (atomic (sepBy1 termParser ", " >> " with ")) - >> sepByIndent (structInstFieldAbbrev <|> structInstField) ", " (allowTrailingSep := true) + >> structInstFields (sepByIndent (structInstFieldAbbrev <|> structInstField) ", " (allowTrailingSep := true)) >> optEllipsis >> optional (" : " >> termParser)) >> " }" def typeSpec := leading_parser " : " >> termParser @@ -984,6 +990,7 @@ builtin_initialize register_parser_alias bracketedBinder register_parser_alias attrKind register_parser_alias optSemicolon + register_parser_alias structInstFields end Parser end Lean diff --git a/src/Lean/Server/Completion.lean b/src/Lean/Server/Completion.lean index 0312a117145c..aec3001948f7 100644 --- a/src/Lean/Server/Completion.lean +++ b/src/Lean/Server/Completion.lean @@ -1,1021 +1,57 @@ /- Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura +Authors: Leonardo de Moura, Marc Huisinga -/ prelude -import Lean.Environment -import Lean.Parser.Term -import Lean.Data.FuzzyMatching -import Lean.Data.Lsp.LanguageFeatures -import Lean.Data.Lsp.Capabilities -import Lean.Data.Lsp.Utf16 -import Lean.Meta.CompletionName -import Lean.Meta.Tactic.Apply -import Lean.Meta.Match.MatcherInfo -import Lean.Elab.Tactic.Doc -import Lean.Server.InfoUtils -import Lean.Parser.Extension -import Lean.Server.FileSource -import Lean.Server.CompletionItemData - -private partial def Lean.Server.Completion.consumeImplicitPrefix (e : Expr) (k : Expr → MetaM α) : MetaM α := do - match e with - | Expr.forallE n d b c => - -- We do not consume instance implicit arguments because the user probably wants be aware of this dependency - if c == .implicit then - Meta.withLocalDecl n c d fun arg => - consumeImplicitPrefix (b.instantiate1 arg) k - else - k e - | _ => k e - -namespace Lean.Lsp - -/-- -Identifier that is sent from the server to the client as part of the `CompletionItem.data?` field. -Needed to resolve the `CompletionItem` when the client sends a `completionItem/resolve` request -for that item, again containing the `data?` field provided by the server. --/ -inductive CompletionIdentifier where - | const (declName : Name) - | fvar (id : FVarId) - deriving FromJson, ToJson - -/-- -`CompletionItemData` that also contains a `CompletionIdentifier`. -See the documentation of`CompletionItemData` and `CompletionIdentifier`. --/ -structure CompletionItemDataWithId extends CompletionItemData where - id? : Option CompletionIdentifier - deriving FromJson, ToJson - -/-- -Fills the `CompletionItem.detail?` field of `item` using the pretty-printed type identified by `id`. --/ -def CompletionItem.resolve - (item : CompletionItem) - (id : CompletionIdentifier) - : MetaM CompletionItem := do - let env ← getEnv - let lctx ← getLCtx - let mut item := item - - if item.detail?.isNone then - let type? := match id with - | .const declName => - env.find? declName |>.map ConstantInfo.type - | .fvar id => - lctx.find? id |>.map LocalDecl.type - let detail? ← type?.mapM fun type => - Lean.Server.Completion.consumeImplicitPrefix type fun typeWithoutImplicits => - return toString (← Meta.ppExpr typeWithoutImplicits) - item := { item with detail? := detail? } - - return item - -end Lean.Lsp +import Lean.Server.Completion.CompletionCollectors namespace Lean.Server.Completion open Lsp open Elab -open Meta -open FuzzyMatching - -abbrev EligibleHeaderDecls := Std.HashMap Name ConstantInfo - -/-- Cached header declarations for which `allowCompletion headerEnv decl` is true. -/ -builtin_initialize eligibleHeaderDeclsRef : IO.Ref (Option EligibleHeaderDecls) ← - IO.mkRef none -/-- -Returns the declarations in the header for which `allowCompletion env decl` is true, caching them -if not already cached. --/ -def getEligibleHeaderDecls (env : Environment) : IO EligibleHeaderDecls := do - eligibleHeaderDeclsRef.modifyGet fun - | some eligibleHeaderDecls => (eligibleHeaderDecls, some eligibleHeaderDecls) - | none => - let (_, eligibleHeaderDecls) := - StateT.run (m := Id) (s := {}) do - -- `map₁` are the header decls - env.constants.map₁.forM fun declName c => do - modify fun eligibleHeaderDecls => - if allowCompletion env declName then - eligibleHeaderDecls.insert declName c - else - eligibleHeaderDecls - (eligibleHeaderDecls, some eligibleHeaderDecls) - -/-- Iterate over all declarations that are allowed in completion results. -/ -private def forEligibleDeclsM [Monad m] [MonadEnv m] [MonadLiftT (ST IO.RealWorld) m] - [MonadLiftT IO m] (f : Name → ConstantInfo → m PUnit) : m PUnit := do - let env ← getEnv - (← getEligibleHeaderDecls env).forM f - -- map₂ are exactly the local decls - env.constants.map₂.forM fun name c => do - if allowCompletion env name then - f name c - -/-- Checks whether this declaration can appear in completion results. -/ -private def allowCompletion (eligibleHeaderDecls : EligibleHeaderDecls) (env : Environment) - (declName : Name) : Bool := - eligibleHeaderDecls.contains declName || - env.constants.map₂.contains declName && Lean.Meta.allowCompletion env declName +private def filterDuplicateCompletionItems + (items : Array ScoredCompletionItem) + : Array ScoredCompletionItem := + let duplicationGroups := items.groupByKey fun s => ( + s.item.label, + s.item.textEdit?, + s.item.detail?, + s.item.kind?, + s.item.tags?, + s.item.documentation?, + ) + duplicationGroups.map (fun _ duplicateItems => duplicateItems.getMax? (·.score < ·.score) |>.get!) + |>.valuesArray /-- Sorts `items` descendingly according to their score and ascendingly according to their label for equal scores. -/ -private def sortCompletionItems (items : Array (CompletionItem × Float)) : Array CompletionItem := - let items := items.qsort fun (i1, s1) (i2, s2) => +private def sortCompletionItems (items : Array ScoredCompletionItem) : Array CompletionItem := + let items := items.qsort fun ⟨i1, s1⟩ ⟨i2, s2⟩ => if s1 != s2 then s1 > s2 else i1.label.map (·.toLower) < i2.label.map (·.toLower) items.map (·.1) -/-- Intermediate state while completions are being computed. -/ -structure State where - /-- All completion items and their fuzzy match scores so far. -/ - items : Array (CompletionItem × Float) := #[] - -/-- -Monad used for completion computation that allows modifying a completion `State` and reading -`CompletionParams`. --/ -abbrev M := OptionT $ ReaderT CompletionParams $ StateRefT State MetaM - -/-- Adds a new completion item to the state in `M`. -/ -private def addItem - (item : CompletionItem) - (score : Float) - (id? : Option CompletionIdentifier := none) - : M Unit := do - let params ← read - let data := { params, id? : CompletionItemDataWithId } - let item := { item with data? := toJson data } - modify fun s => { s with items := s.items.push (item, score) } - -/-- -Adds a new completion item with the given `label`, `id`, `kind` and `score` to the state in `M`. -Computes the doc string from the environment if available. --/ -private def addUnresolvedCompletionItem - (label : Name) - (id : CompletionIdentifier) - (kind : CompletionItemKind) - (score : Float) - : M Unit := do - let env ← getEnv - let (docStringPrefix?, tags?) := Id.run do - let .const declName := id - | (none, none) - let some param := Linter.deprecatedAttr.getParam? env declName - | (none, none) - let docstringPrefix := - if let some text := param.text? then - text - else if let some newName := param.newName? then - s!"`{declName}` has been deprecated, use `{newName}` instead." - else - s!"`{declName}` has been deprecated." - (some docstringPrefix, some #[CompletionItemTag.deprecated]) - let docString? ← do - let .const declName := id - | pure none - findDocString? env declName - let doc? := do - let docValue ← - match docStringPrefix?, docString? with - | none, none => none - | some docStringPrefix, none => docStringPrefix - | none, docString => docString - | some docStringPrefix, some docString => s!"{docStringPrefix}\n\n{docString}" - pure { value := docValue , kind := MarkupKind.markdown : MarkupContent } - let item := { label := label.toString, kind? := kind, documentation? := doc?, tags?} - addItem item score id - -private def getCompletionKindForDecl (constInfo : ConstantInfo) : M CompletionItemKind := do - let env ← getEnv - if constInfo.isCtor then - return CompletionItemKind.constructor - else if constInfo.isInductive then - if isClass env constInfo.name then - return CompletionItemKind.class - else if (← isEnumType constInfo.name) then - return CompletionItemKind.enum - else - return CompletionItemKind.struct - else if constInfo.isTheorem then - return CompletionItemKind.event - else if (← isProjectionFn constInfo.name) then - return CompletionItemKind.field - else - let isFunction : Bool ← withTheReader Core.Context ({ · with maxHeartbeats := 0 }) do - return (← whnf constInfo.type).isForall - if isFunction then - return CompletionItemKind.function - else - return CompletionItemKind.constant - -private def addUnresolvedCompletionItemForDecl (label : Name) (declName : Name) (score : Float) : M Unit := do - if let some c := (← getEnv).find? declName then - addUnresolvedCompletionItem label (.const declName) (← getCompletionKindForDecl c) score - -private def addKeywordCompletionItem (keyword : String) (score : Float) : M Unit := do - let item := { label := keyword, detail? := "keyword", documentation? := none, kind? := CompletionItemKind.keyword } - addItem item score - -private def addNamespaceCompletionItem (ns : Name) (score : Float) : M Unit := do - let item := { label := ns.toString, detail? := "namespace", documentation? := none, kind? := CompletionItemKind.module } - addItem item score - -private def runM (params : CompletionParams) (ctx : ContextInfo) (lctx : LocalContext) (x : M Unit) - : IO (Option CompletionList) := - ctx.runMetaM lctx do - match (← x.run |>.run params |>.run {}) with - | (none, _) => return none - | (some _, s) => - return some { items := sortCompletionItems s.items, isIncomplete := true } - -private def matchAtomic (id : Name) (declName : Name) : Option Float := - match id, declName with - | .str .anonymous s₁, .str .anonymous s₂ => fuzzyMatchScoreWithThreshold? s₁ s₂ - | _, _ => none - -private def normPrivateName? (declName : Name) : MetaM (Option Name) := do - match privateToUserName? declName with - | none => return declName - | some userName => - if mkPrivateName (← getEnv) userName == declName then - return userName - else - return none - -/-- - Return the auto-completion label if `id` can be auto completed using `declName` assuming namespace `ns` is open. - This function only succeeds with atomic labels. BTW, it seems most clients only use the last part. - - Remark: `danglingDot == true` when the completion point is an identifier followed by `.`. --/ -private def matchDecl? (ns : Name) (id : Name) (danglingDot : Bool) (declName : Name) : MetaM (Option (Name × Float)) := do - let some declName ← normPrivateName? declName - | return none - if !ns.isPrefixOf declName then - return none - let declName := declName.replacePrefix ns Name.anonymous - if danglingDot then - -- If the input is `id.` and `declName` is of the form `id.atomic`, complete with `atomicName` - if id.isPrefixOf declName then - let declName := declName.replacePrefix id Name.anonymous - if declName.isAtomic && !declName.isAnonymous then - return some (declName, 1) - else if let (.str p₁ s₁, .str p₂ s₂) := (id, declName) then - if p₁ == p₂ then - -- If the namespaces agree, fuzzy-match on the trailing part - return fuzzyMatchScoreWithThreshold? s₁ s₂ |>.map (.mkSimple s₂, ·) - else if p₁.isAnonymous then - -- If `id` is namespace-less, also fuzzy-match declaration names in arbitrary namespaces - -- (but don't match the namespace itself). - -- Penalize score by component length of added namespace. - return fuzzyMatchScoreWithThreshold? s₁ s₂ |>.map (declName, · / (p₂.getNumParts + 1).toFloat) - return none - -/-- - Truncate the given identifier and make sure it has length `≤ newLength`. - This function assumes `id` does not contain `Name.num` constructors. --/ -private partial def truncate (id : Name) (newLen : Nat) : Name := - let rec go (id : Name) : Name × Nat := - match id with - | Name.anonymous => (id, 0) - | Name.num .. => unreachable! - | .str p s => - let (p', len) := go p - if len + 1 >= newLen then - (p', len) - else - let optDot := if p.isAnonymous then 0 else 1 - let len' := len + optDot + s.length - if len' ≤ newLen then - (id, len') - else - (Name.mkStr p (s.extract 0 ⟨newLen - optDot - len⟩), newLen) - (go id).1 - -inductive HoverInfo where - | after - | inside (delta : Nat) - -def matchNamespace (ns : Name) (nsFragment : Name) (danglingDot : Bool) : Option Float := - if danglingDot then - if nsFragment != ns && nsFragment.isPrefixOf ns then - some 1 - else - none - else - match ns, nsFragment with - | .str p₁ s₁, .str p₂ s₂ => - if p₁ == p₂ then fuzzyMatchScoreWithThreshold? s₂ s₁ else none - | _, _ => none - -def completeNamespaces (ctx : ContextInfo) (id : Name) (danglingDot : Bool) : M Unit := do - let env ← getEnv - let add (ns : Name) (ns' : Name) (score : Float) : M Unit := - if danglingDot then - addNamespaceCompletionItem (ns.replacePrefix (ns' ++ id) Name.anonymous) score - else - addNamespaceCompletionItem (ns.replacePrefix ns' Name.anonymous) score - env.getNamespaceSet |>.forM fun ns => do - unless ns.isInternal || env.contains ns do -- Ignore internal and namespaces that are also declaration names - for openDecl in ctx.openDecls do - match openDecl with - | OpenDecl.simple ns' _ => - if let some score := matchNamespace ns (ns' ++ id) danglingDot then - add ns ns' score - return () - | _ => pure () - -- use current namespace - let rec visitNamespaces (ns' : Name) : M Unit := do - if let some score := matchNamespace ns (ns' ++ id) danglingDot then - add ns ns' score - else - match ns' with - | Name.str p .. => visitNamespaces p - | _ => return () - visitNamespaces ctx.currNamespace - -private def idCompletionCore - (ctx : ContextInfo) - (stx : Syntax) - (id : Name) - (hoverInfo : HoverInfo) - (danglingDot : Bool) - : M Unit := do - let mut id := id - if id.hasMacroScopes then - if stx.getHeadInfo matches .original .. then - id := id.eraseMacroScopes - else - -- Identifier is synthetic and has macro scopes => no completions - -- Erasing the macro scopes does not make sense in this case because the identifier name - -- is some random synthetic string. - return - let mut danglingDot := danglingDot - if let HoverInfo.inside delta := hoverInfo then - id := truncate id delta - danglingDot := false - if id.isAtomic then - -- search for matches in the local context - for localDecl in (← getLCtx) do - if let some score := matchAtomic id localDecl.userName then - addUnresolvedCompletionItem localDecl.userName (.fvar localDecl.fvarId) (kind := CompletionItemKind.variable) score - -- search for matches in the environment - let env ← getEnv - forEligibleDeclsM fun declName c => do - let bestMatch? ← (·.2) <$> StateT.run (s := none) do - let matchUsingNamespace (ns : Name) : StateT (Option (Name × Float)) M Unit := do - let some (label, score) ← matchDecl? ns id danglingDot declName - | return - modify fun - | none => - some (label, score) - | some (bestLabel, bestScore) => - -- for open namespaces `A` and `A.B` and a decl `A.B.c`, pick the decl `c` over `B.c` - if label.isSuffixOf bestLabel then - some (label, score) - else - some (bestLabel, bestScore) - let rec visitNamespaces (ns : Name) : StateT (Option (Name × Float)) M Unit := do - let Name.str p .. := ns - | return () - matchUsingNamespace ns - visitNamespaces p - -- use current namespace - visitNamespaces ctx.currNamespace - -- use open decls - for openDecl in ctx.openDecls do - let OpenDecl.simple ns exs := openDecl - | pure () - if exs.contains declName then - continue - matchUsingNamespace ns - matchUsingNamespace Name.anonymous - if let some (bestLabel, bestScore) := bestMatch? then - addUnresolvedCompletionItem bestLabel (.const declName) (← getCompletionKindForDecl c) bestScore - -- Recall that aliases may not be atomic and include the namespace where they were created. - let matchAlias (ns : Name) (alias : Name) : Option Float := - if ns.isPrefixOf alias then - matchAtomic id (alias.replacePrefix ns Name.anonymous) - else - none - let eligibleHeaderDecls ← getEligibleHeaderDecls env - -- Auxiliary function for `alias` - let addAlias (alias : Name) (declNames : List Name) (score : Float) : M Unit := do - declNames.forM fun declName => do - if allowCompletion eligibleHeaderDecls env declName then - addUnresolvedCompletionItemForDecl (.mkSimple alias.getString!) declName score - -- search explicitly open `ids` - for openDecl in ctx.openDecls do - match openDecl with - | OpenDecl.explicit openedId resolvedId => - if allowCompletion eligibleHeaderDecls env resolvedId then - if let some score := matchAtomic id openedId then - addUnresolvedCompletionItemForDecl (.mkSimple openedId.getString!) resolvedId score - | OpenDecl.simple ns _ => - getAliasState env |>.forM fun alias declNames => do - if let some score := matchAlias ns alias then - addAlias alias declNames score - -- search for aliases - getAliasState env |>.forM fun alias declNames => do - -- use current namespace - let rec searchAlias (ns : Name) : M Unit := do - if let some score := matchAlias ns alias then - addAlias alias declNames score - else - match ns with - | Name.str p .. => searchAlias p - | _ => return () - searchAlias ctx.currNamespace - -- Search keywords - if let .str .anonymous s := id then - let keywords := Parser.getTokenTable env - for keyword in keywords.findPrefix s do - if let some score := fuzzyMatchScoreWithThreshold? s keyword then - addKeywordCompletionItem keyword score - -- Search namespaces - completeNamespaces ctx id danglingDot - -private def idCompletion - (params : CompletionParams) - (ctx : ContextInfo) - (lctx : LocalContext) - (stx : Syntax) - (id : Name) - (hoverInfo : HoverInfo) - (danglingDot : Bool) - : IO (Option CompletionList) := - runM params ctx lctx do - idCompletionCore ctx stx id hoverInfo danglingDot - -private def unfoldeDefinitionGuarded? (e : Expr) : MetaM (Option Expr) := - try unfoldDefinition? e catch _ => pure none - -/-- Return `true` if `e` is a `declName`-application, or can be unfolded (delta-reduced) to one. -/ -private partial def isDefEqToAppOf (e : Expr) (declName : Name) : MetaM Bool := do - let isConstOf := match e.getAppFn with - | .const name .. => (privateToUserName? name).getD name == declName - | _ => false - if isConstOf then - return true - let some e ← unfoldeDefinitionGuarded? e | return false - isDefEqToAppOf e declName - -private def isDotCompletionMethod (typeName : Name) (info : ConstantInfo) : MetaM Bool := - forallTelescopeReducing info.type fun xs _ => do - for x in xs do - let localDecl ← x.fvarId!.getDecl - let type := localDecl.type.consumeMData - if (← isDefEqToAppOf type typeName) then - return true - return false - -/-- -Checks whether the expected type of `info.type` can be reduced to an application of `typeName`. --/ -private def isDotIdCompletionMethod (typeName : Name) (info : ConstantInfo) : MetaM Bool := do - forallTelescopeReducing info.type fun _ type => - isDefEqToAppOf type.consumeMData typeName - -/-- -Converts `n` to `Name.anonymous` if `n` is a private prefix (see `Lean.isPrivatePrefix`). --/ -private def stripPrivatePrefix (n : Name) : Name := - match n with - | .num _ 0 => if isPrivatePrefix n then .anonymous else n - | _ => n - -/-- -Compares `n₁` and `n₂` modulo private prefixes. Similar to `Name.cmp` but ignores all -private prefixes in both names. -Necessary because the namespaces of private names do not contain private prefixes. --/ -partial def cmpModPrivate (n₁ n₂ : Name) : Ordering := - let n₁ := stripPrivatePrefix n₁ - let n₂ := stripPrivatePrefix n₂ - match n₁, n₂ with - | .anonymous, .anonymous => Ordering.eq - | .anonymous, _ => Ordering.lt - | _, .anonymous => Ordering.gt - | .num p₁ i₁, .num p₂ i₂ => - match compare i₁ i₂ with - | Ordering.eq => cmpModPrivate p₁ p₂ - | ord => ord - | .num _ _, .str _ _ => Ordering.lt - | .str _ _, .num _ _ => Ordering.gt - | .str p₁ n₁, .str p₂ n₂ => - match compare n₁ n₂ with - | Ordering.eq => cmpModPrivate p₁ p₂ - | ord => ord - -/-- -`NameSet` where names are compared according to `cmpModPrivate`. -Helps speed up dot completion because it allows us to look up names without first having to -strip the private prefix from deep in the name, letting us reject most names without -having to scan the full name first. --/ -def NameSetModPrivate := RBTree Name cmpModPrivate - -/-- - Given a type, try to extract relevant type names for dot notation field completion. - We extract the type name, parent struct names, and unfold the type. - The process mimics the dot notation elaboration procedure at `App.lean` -/ -private partial def getDotCompletionTypeNames (type : Expr) : MetaM NameSetModPrivate := - return (← visit type |>.run RBTree.empty).2 -where - visit (type : Expr) : StateRefT NameSetModPrivate MetaM Unit := do - let .const typeName _ := type.getAppFn | return () - modify fun s => s.insert typeName - if isStructure (← getEnv) typeName then - for parentName in (← getAllParentStructures typeName) do - modify fun s => s.insert parentName - let some type ← unfoldeDefinitionGuarded? type | return () - visit type - -private def dotCompletion - (params : CompletionParams) - (ctx : ContextInfo) - (info : TermInfo) - (hoverInfo : HoverInfo) - : IO (Option CompletionList) := - runM params ctx info.lctx do - let nameSet ← try - getDotCompletionTypeNames (← instantiateMVars (← inferType info.expr)) - catch _ => - pure RBTree.empty - - if nameSet.isEmpty then - let stx := info.stx - if stx.isIdent then - idCompletionCore ctx stx stx.getId hoverInfo (danglingDot := false) - else if stx.getKind == ``Lean.Parser.Term.completion && stx[0].isIdent then - -- TODO: truncation when there is a dangling dot - idCompletionCore ctx stx stx[0].getId HoverInfo.after (danglingDot := true) - else - failure - return - - forEligibleDeclsM fun declName c => do - let unnormedTypeName := declName.getPrefix - if ! nameSet.contains unnormedTypeName then - return - let some declName ← normPrivateName? declName - | return - let typeName := declName.getPrefix - if ! (← isDotCompletionMethod typeName c) then - return - let completionKind ← getCompletionKindForDecl c - addUnresolvedCompletionItem (.mkSimple c.name.getString!) (.const c.name) (kind := completionKind) 1 - -private def dotIdCompletion - (params : CompletionParams) - (ctx : ContextInfo) - (lctx : LocalContext) - (id : Name) - (expectedType? : Option Expr) - : IO (Option CompletionList) := - runM params ctx lctx do - let some expectedType := expectedType? - | return () - - let resultTypeFn := (← instantiateMVars expectedType).cleanupAnnotations.getAppFn.cleanupAnnotations - let .const .. := resultTypeFn - | return () - - let nameSet ← try - getDotCompletionTypeNames resultTypeFn - catch _ => - pure RBTree.empty - - forEligibleDeclsM fun declName c => do - let unnormedTypeName := declName.getPrefix - if ! nameSet.contains unnormedTypeName then - return - - let some declName ← normPrivateName? declName - | return - - let typeName := declName.getPrefix - if ! (← isDotIdCompletionMethod typeName c) then - return - - let completionKind ← getCompletionKindForDecl c - if id.isAnonymous then - -- We're completing a lone dot => offer all decls of the type - addUnresolvedCompletionItem (.mkSimple c.name.getString!) (.const c.name) completionKind 1 - return - - let some (label, score) ← matchDecl? typeName id (danglingDot := false) declName | pure () - addUnresolvedCompletionItem label (.const c.name) completionKind score - -private def fieldIdCompletion - (params : CompletionParams) - (ctx : ContextInfo) - (lctx : LocalContext) - (id : Name) - (structName : Name) - : IO (Option CompletionList) := - runM params ctx lctx do - let idStr := id.toString - let fieldNames := getStructureFieldsFlattened (← getEnv) structName (includeSubobjectFields := false) - for fieldName in fieldNames do - let .str _ fieldName := fieldName | continue - let some score := fuzzyMatchScoreWithThreshold? idStr fieldName | continue - let item := { label := fieldName, detail? := "field", documentation? := none, kind? := CompletionItemKind.field } - addItem item score - -private def optionCompletion - (params : CompletionParams) - (ctx : ContextInfo) - (stx : Syntax) - (caps : ClientCapabilities) - : IO (Option CompletionList) := - ctx.runMetaM {} do - let (partialName, trailingDot) := - -- `stx` is from `"set_option" >> ident` - match stx[1].getSubstring? (withLeading := false) (withTrailing := false) with - | none => ("", false) -- the `ident` is `missing`, list all options - | some ss => - if !ss.str.atEnd ss.stopPos && ss.str.get ss.stopPos == '.' then - -- include trailing dot, which is not parsed by `ident` - (ss.toString ++ ".", true) - else - (ss.toString, false) - -- HACK(WN): unfold the type so ForIn works - let (decls : RBMap _ _ _) ← getOptionDecls - let opts ← getOptions - let mut items := #[] - for ⟨name, decl⟩ in decls do - if let some score := fuzzyMatchScoreWithThreshold? partialName name.toString then - let textEdit := - if !caps.textDocument?.any (·.completion?.any (·.completionItem?.any (·.insertReplaceSupport?.any (·)))) then - none -- InsertReplaceEdit not supported by client - else if let some ⟨start, stop⟩ := stx[1].getRange? then - let stop := if trailingDot then stop + ' ' else stop - let range := ⟨ctx.fileMap.utf8PosToLspPos start, ctx.fileMap.utf8PosToLspPos stop⟩ - some { newText := name.toString, insert := range, replace := range : InsertReplaceEdit } - else - none - items := items.push - ({ label := name.toString - detail? := s!"({opts.get name decl.defValue}), {decl.descr}" - documentation? := none, - kind? := CompletionItemKind.property -- TODO: investigate whether this is the best kind for options. - textEdit? := textEdit - data? := toJson { params, id? := none : CompletionItemDataWithId } }, score) - return some { items := sortCompletionItems items, isIncomplete := true } - -private def tacticCompletion (params : CompletionParams) (ctx : ContextInfo) - : IO (Option CompletionList) := ctx.runMetaM .empty do - let allTacticDocs ← Tactic.Doc.allTacticDocs - let items : Array (CompletionItem × Float) := allTacticDocs.map fun tacticDoc => - ({ - label := tacticDoc.userName - detail? := none - documentation? := tacticDoc.docString.map fun docString => - { value := docString, kind := MarkupKind.markdown : MarkupContent } - kind? := CompletionItemKind.keyword - data? := toJson { params, id? := none : CompletionItemDataWithId } - }, 1) - return some { items := sortCompletionItems items, isIncomplete := true } - -private def findBest? - (infoTree : InfoTree) - (gt : α → α → Bool) - (f : ContextInfo → Info → PersistentArray InfoTree → Option α) - : Option α := - infoTree.visitM (m := Id) (postNode := choose) |>.join -where - choose - (ctx : ContextInfo) - (info : Info) - (cs : PersistentArray InfoTree) - (childValues : List (Option (Option α))) - : Option α := - let bestChildValue := childValues.map (·.join) |>.foldl (init := none) fun v best => - if isBetter v best then - v - else - best - if let some v := f ctx info cs then - if isBetter v bestChildValue then - v - else - bestChildValue - else - bestChildValue - isBetter (a b : Option α) : Bool := - match a, b with - | none, none => false - | some _, none => true - | none, some _ => false - | some a, some b => gt a b - -/-- -If there are `Info`s that contain `hoverPos` and have a nonempty `LocalContext`, -yields the closest one of those `Info`s. -Otherwise, yields the closest `Info` that contains `hoverPos` and has an empty `LocalContext`. --/ -private def findClosestInfoWithLocalContextAt? - (hoverPos : String.Pos) - (infoTree : InfoTree) - : Option (ContextInfo × Info) := - findBest? infoTree isBetter fun ctx info _ => - if info.occursInOrOnBoundary hoverPos then - (ctx, info) - else - none -where - isBetter (a b : ContextInfo × Info) : Bool := - let (_, ia) := a - let (_, ib) := b - if !ia.lctx.isEmpty && ib.lctx.isEmpty then - true - else if ia.lctx.isEmpty && !ib.lctx.isEmpty then - false - else if ia.isSmaller ib then - true - else if ib.isSmaller ia then - false - else - false - -private def findSyntheticIdentifierCompletion? - (hoverPos : String.Pos) - (infoTree : InfoTree) - : Option (HoverInfo × ContextInfo × CompletionInfo) := do - let some (ctx, info) := findClosestInfoWithLocalContextAt? hoverPos infoTree - | none - let some stack := info.stx.findStack? (·.getRange?.any (·.contains hoverPos (includeStop := true))) - | none - let stack := stack.dropWhile fun (stx, _) => !(stx matches `($_:ident) || stx matches `($_:ident.)) - let some (stx, _) := stack.head? - | none - let isDotIdCompletion := stack.any fun (stx, _) => stx matches `(.$_:ident) - if isDotIdCompletion then - -- An identifier completion is never useful in a dotId completion context. - none - let some (id, danglingDot) := - match stx with - | `($id:ident) => some (id.getId, false) - | `($id:ident.) => some (id.getId, true) - | _ => none - | none - let tailPos := stx.getTailPos?.get! - let hoverInfo := - if hoverPos < tailPos then - HoverInfo.inside (tailPos - hoverPos).byteIdx - else - HoverInfo.after - some (hoverInfo, ctx, .id stx id danglingDot info.lctx none) - -private partial def getIndentationAmount (fileMap : FileMap) (line : Nat) : Nat := Id.run do - let lineStartPos := fileMap.lineStart line - let lineEndPos := fileMap.lineStart (line + 1) - let mut it : String.Iterator := ⟨fileMap.source, lineStartPos⟩ - let mut indentationAmount := 0 - while it.pos < lineEndPos do - let c := it.curr - if c = ' ' || c = '\t' then - indentationAmount := indentationAmount + 1 - else - break - it := it.next - return indentationAmount - -private partial def isSyntheticTacticCompletion - (fileMap : FileMap) - (hoverPos : String.Pos) - (cmdStx : Syntax) - : Bool := Id.run do - let hoverFilePos := fileMap.toPosition hoverPos - let mut hoverLineIndentation := getIndentationAmount fileMap hoverFilePos.line - if hoverFilePos.column < hoverLineIndentation then - -- Ignore trailing whitespace after the cursor - hoverLineIndentation := hoverFilePos.column - go hoverFilePos hoverLineIndentation cmdStx 0 -where - go - (hoverFilePos : Position) - (hoverLineIndentation : Nat) - (stx : Syntax) - (leadingWs : Nat) - : Bool := Id.run do - match stx.getPos?, stx.getTailPos? with - | some startPos, some endPos => - let isCursorInCompletionRange := - startPos.byteIdx - leadingWs <= hoverPos.byteIdx - && hoverPos.byteIdx <= endPos.byteIdx + stx.getTrailingSize - if ! isCursorInCompletionRange then - return false - let mut wsBeforeArg := leadingWs - for arg in stx.getArgs do - if go hoverFilePos hoverLineIndentation arg wsBeforeArg then - return true - -- We must account for the whitespace before an argument because the syntax nodes we use - -- to identify tactic blocks only start *after* the whitespace following a `by`, and we - -- want to provide tactic completions in that whitespace as well. - -- This method of computing whitespace assumes that there are no syntax nodes without tokens - -- after `by` and before the first proper tactic syntax. - wsBeforeArg := arg.getTrailingSize - return isCompletionInEmptyTacticBlock stx - || isCompletionAfterSemicolon stx - || isCompletionOnTacticBlockIndentation hoverFilePos hoverLineIndentation stx - | _, _ => - -- Empty tactic blocks typically lack ranges since they do not contain any tokens. - -- We do not perform more precise range checking in this case because we assume that empty - -- tactic blocks always occur within other syntax with ranges that let us narrow down the - -- search to the degree that we can be sure that the cursor is indeed in this empty tactic - -- block. - return isCompletionInEmptyTacticBlock stx - - isCompletionOnTacticBlockIndentation - (hoverFilePos : Position) - (hoverLineIndentation : Nat) - (stx : Syntax) - : Bool := Id.run do - let isCursorInIndentation := hoverFilePos.column <= hoverLineIndentation - if ! isCursorInIndentation then - -- Do not trigger tactic completion at the end of a properly indented tactic block line since - -- that line might already have entered term mode by that point. - return false - let some tacticsNode := getTacticsNode? stx - | return false - let some firstTacticPos := tacticsNode.getPos? - | return false - let firstTacticLine := fileMap.toPosition firstTacticPos |>.line - let firstTacticIndentation := getIndentationAmount fileMap firstTacticLine - -- This ensures that we do not accidentally provide tactic completions in a term mode proof - - -- tactic completions are only provided at the same indentation level as the other tactics in - -- that tactic block. - let isCursorInTacticBlock := hoverLineIndentation == firstTacticIndentation - return isCursorInProperWhitespace && isCursorInTacticBlock - - isCompletionAfterSemicolon (stx : Syntax) : Bool := Id.run do - let some tacticsNode := getTacticsNode? stx - | return false - let tactics := tacticsNode.getArgs - -- We want to provide completions in the case of `skip;`, so the cursor must only be on - -- whitespace, not in proper whitespace. - return isCursorOnWhitspace && tactics.any fun tactic => Id.run do - let some tailPos := tactic.getTailPos? - | return false - let isCursorAfterSemicolon := - tactic.isToken ";" - && tailPos.byteIdx <= hoverPos.byteIdx - && hoverPos.byteIdx <= tailPos.byteIdx + tactic.getTrailingSize - return isCursorAfterSemicolon - - getTacticsNode? (stx : Syntax) : Option Syntax := - if stx.getKind == `Lean.Parser.Tactic.tacticSeq1Indented then - some stx[0] - else if stx.getKind == `Lean.Parser.Tactic.tacticSeqBracketed then - some stx[1] - else - none - - isCompletionInEmptyTacticBlock (stx : Syntax) : Bool := - isCursorInProperWhitespace && isEmptyTacticBlock stx - - isCursorOnWhitspace : Bool := - fileMap.source.atEnd hoverPos || (fileMap.source.get hoverPos).isWhitespace - - isCursorInProperWhitespace : Bool := - (fileMap.source.atEnd hoverPos || (fileMap.source.get hoverPos).isWhitespace) - && (fileMap.source.get (hoverPos - ⟨1⟩)).isWhitespace - - isEmptyTacticBlock (stx : Syntax) : Bool := - stx.getKind == `Lean.Parser.Tactic.tacticSeq && isEmpty stx - || stx.getKind == `Lean.Parser.Tactic.tacticSeq1Indented && isEmpty stx - || stx.getKind == `Lean.Parser.Tactic.tacticSeqBracketed && isEmpty stx[1] - - isEmpty : Syntax → Bool - | .missing => true - | .ident .. => false - | .atom .. => false - | .node _ _ args => args.all isEmpty - -private partial def findOutermostContextInfo? (i : InfoTree) : Option ContextInfo := - go i -where - go (i : InfoTree) : Option ContextInfo := do - match i with - | .context ctx i => - match ctx with - | .commandCtx ctxInfo => - some { ctxInfo with } - | _ => - -- This shouldn't happen (see the `PartialContextInfo` docstring), - -- but let's continue searching regardless - go i - | .node _ cs => - cs.findSome? go - | .hole .. => - none - -private def findSyntheticTacticCompletion? - (fileMap : FileMap) - (hoverPos : String.Pos) - (cmdStx : Syntax) - (infoTree : InfoTree) - : Option (HoverInfo × ContextInfo × CompletionInfo) := do - let ctx ← findOutermostContextInfo? infoTree - if ! isSyntheticTacticCompletion fileMap hoverPos cmdStx then - none - -- Neither `HoverInfo` nor the syntax in `.tactic` are important for tactic completion. - return (HoverInfo.after, ctx, .tactic .missing) - -private def findCompletionInfoAt? - (fileMap : FileMap) - (hoverPos : String.Pos) - (cmdStx : Syntax) - (infoTree : InfoTree) - : Option (HoverInfo × ContextInfo × CompletionInfo) := - let ⟨hoverLine, _⟩ := fileMap.toPosition hoverPos - match infoTree.foldInfo (init := none) (choose hoverLine) with - | some (hoverInfo, ctx, Info.ofCompletionInfo info) => - some (hoverInfo, ctx, info) - | _ => - findSyntheticTacticCompletion? fileMap hoverPos cmdStx infoTree <|> - findSyntheticIdentifierCompletion? hoverPos infoTree -where - choose - (hoverLine : Nat) - (ctx : ContextInfo) - (info : Info) - (best? : Option (HoverInfo × ContextInfo × Info)) - : Option (HoverInfo × ContextInfo × Info) := - if !info.isCompletion then - best? - else if info.occursInOrOnBoundary hoverPos then - let headPos := info.pos?.get! - let tailPos := info.tailPos?.get! - let hoverInfo := - if hoverPos < tailPos then - HoverInfo.inside (hoverPos - headPos).byteIdx - else - HoverInfo.after - let ⟨headPosLine, _⟩ := fileMap.toPosition headPos - let ⟨tailPosLine, _⟩ := fileMap.toPosition info.tailPos?.get! - if headPosLine != hoverLine || headPosLine != tailPosLine then - best? - else match best? with - | none => (hoverInfo, ctx, info) - | some (_, _, best) => - if isBetter info best then - (hoverInfo, ctx, info) - else - best? - else - best? - - isBetter : Info → Info → Bool - | i₁@(.ofCompletionInfo ci₁), i₂@(.ofCompletionInfo ci₂) => - -- Use the smallest info available and prefer non-id completion over id completions as a - -- tie-breaker. - -- This is necessary because the elaborator sometimes generates both for the same range. - -- If two infos are equivalent, always prefer the first one. - if i₁.isSmaller i₂ then - true - else if i₂.isSmaller i₁ then - false - else if !(ci₁ matches .id ..) && ci₂ matches .id .. then - true - else if ci₁ matches .id .. && !(ci₂ matches .id ..) then - false - else - true - | .ofCompletionInfo _, _ => true - | _, .ofCompletionInfo _ => false - | _, _ => true - /-- Assigns the `CompletionItem.sortText?` for all items in `completions` according to their order in `completions`. This is necessary because clients will use their own sort order if the server does not set it. -/ -private def assignSortTexts (completions : CompletionList) : CompletionList := Id.run do - if completions.items.isEmpty then - return completions - let items := completions.items.mapIdx fun i item => +private def assignSortTexts (completions : Array CompletionItem) : Array CompletionItem := Id.run do + if completions.isEmpty then + return #[] + let items := completions.mapIdx fun i item => { item with sortText? := toString i } let maxDigits := items[items.size - 1]!.sortText?.get!.length let items := items.map fun item => let sortText := item.sortText?.get! let pad := List.replicate (maxDigits - sortText.length) '0' |>.asString { item with sortText? := pad ++ sortText } - { completions with items := items } + items partial def find? (params : CompletionParams) @@ -1024,40 +60,37 @@ partial def find? (cmdStx : Syntax) (infoTree : InfoTree) (caps : ClientCapabilities) - : IO (Option CompletionList) := do - let some (hoverInfo, ctx, info) := findCompletionInfoAt? fileMap hoverPos cmdStx infoTree - | return none - let completionList? ← - match info with - | .dot info .. => - dotCompletion params ctx info hoverInfo - | .id stx id danglingDot lctx .. => - idCompletion params ctx lctx stx id hoverInfo danglingDot - | .dotId _ id lctx expectedType? => - dotIdCompletion params ctx lctx id expectedType? - | .fieldId _ id lctx structName => - fieldIdCompletion params ctx lctx id structName - | .option stx => - optionCompletion params ctx stx caps - | .tactic .. => - tacticCompletion params ctx - | _ => return none - return completionList?.map assignSortTexts + : IO CompletionList := do + let prioritizedPartitions := findPrioritizedCompletionPartitionsAt fileMap hoverPos cmdStx infoTree + let mut allCompletions := #[] + for partition in prioritizedPartitions do + for (i, completionInfoPos) in partition do + let completions : Array ScoredCompletionItem ← + match i.info with + | .id stx id danglingDot lctx .. => + idCompletion params completionInfoPos i.ctx lctx stx id i.hoverInfo danglingDot + | .dot info .. => + dotCompletion params completionInfoPos i.ctx info + | .dotId _ id lctx expectedType? => + dotIdCompletion params completionInfoPos i.ctx lctx id expectedType? + | .fieldId _ id lctx structName => + fieldIdCompletion params completionInfoPos i.ctx lctx id structName + | .option stx => + optionCompletion params completionInfoPos i.ctx stx caps + | .tactic .. => + tacticCompletion params completionInfoPos i.ctx + | _ => + pure #[] + allCompletions := allCompletions ++ completions + if ! allCompletions.isEmpty then + -- Stop accumulating completions with lower priority if we found completions for a higher + -- priority. + break -/-- -Fills the `CompletionItem.detail?` field of `item` using the pretty-printed type identified by `id` -in the context found at `hoverPos` in `infoTree`. --/ -def resolveCompletionItem? - (fileMap : FileMap) - (hoverPos : String.Pos) - (cmdStx : Syntax) - (infoTree : InfoTree) - (item : CompletionItem) - (id : CompletionIdentifier) - : IO CompletionItem := do - let some (_, ctx, info) := findCompletionInfoAt? fileMap hoverPos cmdStx infoTree - | return item - ctx.runMetaM info.lctx (item.resolve id) + let finalCompletions := allCompletions + |> filterDuplicateCompletionItems + |> sortCompletionItems + |> assignSortTexts + return { items := finalCompletions, isIncomplete := true } end Lean.Server.Completion diff --git a/src/Lean/Server/Completion/CompletionCollectors.lean b/src/Lean/Server/Completion/CompletionCollectors.lean new file mode 100644 index 000000000000..7743a8a2d29d --- /dev/null +++ b/src/Lean/Server/Completion/CompletionCollectors.lean @@ -0,0 +1,606 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Marc Huisinga +-/ +prelude +import Lean.Data.FuzzyMatching +import Lean.Elab.Tactic.Doc +import Lean.Server.Completion.CompletionResolution +import Lean.Server.Completion.EligibleHeaderDecls + +namespace Lean.Server.Completion +open Elab +open Lean.Lsp +open Meta +open FuzzyMatching + +section Infrastructure + + structure ScoredCompletionItem where + item : CompletionItem + score : Float + deriving Inhabited + + private structure Context where + params : CompletionParams + completionInfoPos : Nat + + + /-- Intermediate state while completions are being computed. -/ + private structure State where + /-- All completion items and their fuzzy match scores so far. -/ + items : Array ScoredCompletionItem := #[] + + /-- + Monad used for completion computation that allows modifying a completion `State` and reading + `CompletionParams`. + -/ + private abbrev M := ReaderT Context $ StateRefT State MetaM + + /-- Adds a new completion item to the state in `M`. -/ + private def addItem + (item : CompletionItem) + (score : Float) + (id? : Option CompletionIdentifier := none) + : M Unit := do + let ctx ← read + let data := { + params := ctx.params, + cPos := ctx.completionInfoPos, + id? + : ResolvableCompletionItemData + } + let item := { item with data? := toJson data } + modify fun s => { s with items := s.items.push ⟨item, score⟩ } + + /-- + Adds a new completion item with the given `label`, `id`, `kind` and `score` to the state in `M`. + Computes the doc string from the environment if available. + -/ + private def addUnresolvedCompletionItem + (label : Name) + (id : CompletionIdentifier) + (kind : CompletionItemKind) + (score : Float) + : M Unit := do + let env ← getEnv + let (docStringPrefix?, tags?) := Id.run do + let .const declName := id + | (none, none) + let some param := Linter.deprecatedAttr.getParam? env declName + | (none, none) + let docstringPrefix := + if let some text := param.text? then + text + else if let some newName := param.newName? then + s!"`{declName}` has been deprecated, use `{newName}` instead." + else + s!"`{declName}` has been deprecated." + (some docstringPrefix, some #[CompletionItemTag.deprecated]) + let docString? ← do + let .const declName := id + | pure none + findDocString? env declName + let doc? := do + let docValue ← + match docStringPrefix?, docString? with + | none, none => none + | some docStringPrefix, none => docStringPrefix + | none, docString => docString + | some docStringPrefix, some docString => s!"{docStringPrefix}\n\n{docString}" + pure { value := docValue , kind := MarkupKind.markdown : MarkupContent } + let item := { label := label.toString, kind? := kind, documentation? := doc?, tags?} + addItem item score id + + private def getCompletionKindForDecl (constInfo : ConstantInfo) : M CompletionItemKind := do + let env ← getEnv + if constInfo.isCtor then + return CompletionItemKind.constructor + else if constInfo.isInductive then + if isClass env constInfo.name then + return CompletionItemKind.class + else if (← isEnumType constInfo.name) then + return CompletionItemKind.enum + else + return CompletionItemKind.struct + else if constInfo.isTheorem then + return CompletionItemKind.event + else if (← isProjectionFn constInfo.name) then + return CompletionItemKind.field + else + let isFunction : Bool ← withTheReader Core.Context ({ · with maxHeartbeats := 0 }) do + return (← whnf constInfo.type).isForall + if isFunction then + return CompletionItemKind.function + else + return CompletionItemKind.constant + + private def addUnresolvedCompletionItemForDecl (label : Name) (declName : Name) (score : Float) : M Unit := do + if let some c := (← getEnv).find? declName then + addUnresolvedCompletionItem label (.const declName) (← getCompletionKindForDecl c) score + + private def addKeywordCompletionItem (keyword : String) (score : Float) : M Unit := do + let item := { label := keyword, detail? := "keyword", documentation? := none, kind? := CompletionItemKind.keyword } + addItem item score + + private def addNamespaceCompletionItem (ns : Name) (score : Float) : M Unit := do + let item := { label := ns.toString, detail? := "namespace", documentation? := none, kind? := CompletionItemKind.module } + addItem item score + + private def runM + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (lctx : LocalContext) + (x : M Unit) + : IO (Array ScoredCompletionItem) := + ctx.runMetaM lctx do + let (_, s) ← x.run ⟨params, completionInfoPos⟩ |>.run {} + return s.items + +end Infrastructure + +section Utils + + private def normPrivateName? (declName : Name) : MetaM (Option Name) := do + match privateToUserName? declName with + | none => return declName + | some userName => + if mkPrivateName (← getEnv) userName == declName then + return userName + else + return none + + /-- + Return the auto-completion label if `id` can be auto completed using `declName` assuming namespace `ns` is open. + This function only succeeds with atomic labels. BTW, it seems most clients only use the last part. + + Remark: `danglingDot == true` when the completion point is an identifier followed by `.`. + -/ + private def matchDecl? (ns : Name) (id : Name) (danglingDot : Bool) (declName : Name) : MetaM (Option (Name × Float)) := do + let some declName ← normPrivateName? declName + | return none + if !ns.isPrefixOf declName then + return none + let declName := declName.replacePrefix ns Name.anonymous + if danglingDot then + -- If the input is `id.` and `declName` is of the form `id.atomic`, complete with `atomicName` + if id.isPrefixOf declName then + let declName := declName.replacePrefix id Name.anonymous + if declName.isAtomic && !declName.isAnonymous then + return some (declName, 1) + else if let (.str p₁ s₁, .str p₂ s₂) := (id, declName) then + if p₁ == p₂ then + -- If the namespaces agree, fuzzy-match on the trailing part + return fuzzyMatchScoreWithThreshold? s₁ s₂ |>.map (.mkSimple s₂, ·) + else if p₁.isAnonymous then + -- If `id` is namespace-less, also fuzzy-match declaration names in arbitrary namespaces + -- (but don't match the namespace itself). + -- Penalize score by component length of added namespace. + return fuzzyMatchScoreWithThreshold? s₁ s₂ |>.map (declName, · / (p₂.getNumParts + 1).toFloat) + return none + +end Utils + +section IdCompletionUtils + + private def matchAtomic (id : Name) (declName : Name) : Option Float := + match id, declName with + | .str .anonymous s₁, .str .anonymous s₂ => fuzzyMatchScoreWithThreshold? s₁ s₂ + | _, _ => none + + /-- + Truncate the given identifier and make sure it has length `≤ newLength`. + This function assumes `id` does not contain `Name.num` constructors. + -/ + private partial def truncate (id : Name) (newLen : Nat) : Name := + let rec go (id : Name) : Name × Nat := + match id with + | Name.anonymous => (id, 0) + | Name.num .. => unreachable! + | .str p s => + let (p', len) := go p + if len + 1 >= newLen then + (p', len) + else + let optDot := if p.isAnonymous then 0 else 1 + let len' := len + optDot + s.length + if len' ≤ newLen then + (id, len') + else + (Name.mkStr p (s.extract 0 ⟨newLen - optDot - len⟩), newLen) + (go id).1 + + def matchNamespace (ns : Name) (nsFragment : Name) (danglingDot : Bool) : Option Float := + if danglingDot then + if nsFragment != ns && nsFragment.isPrefixOf ns then + some 1 + else + none + else + match ns, nsFragment with + | .str p₁ s₁, .str p₂ s₂ => + if p₁ == p₂ then fuzzyMatchScoreWithThreshold? s₂ s₁ else none + | _, _ => none + + def completeNamespaces (ctx : ContextInfo) (id : Name) (danglingDot : Bool) : M Unit := do + let env ← getEnv + let add (ns : Name) (ns' : Name) (score : Float) : M Unit := + if danglingDot then + addNamespaceCompletionItem (ns.replacePrefix (ns' ++ id) Name.anonymous) score + else + addNamespaceCompletionItem (ns.replacePrefix ns' Name.anonymous) score + env.getNamespaceSet |>.forM fun ns => do + unless ns.isInternal || env.contains ns do -- Ignore internal and namespaces that are also declaration names + for openDecl in ctx.openDecls do + match openDecl with + | OpenDecl.simple ns' _ => + if let some score := matchNamespace ns (ns' ++ id) danglingDot then + add ns ns' score + return () + | _ => pure () + -- use current namespace + let rec visitNamespaces (ns' : Name) : M Unit := do + if let some score := matchNamespace ns (ns' ++ id) danglingDot then + add ns ns' score + else + match ns' with + | Name.str p .. => visitNamespaces p + | _ => return () + visitNamespaces ctx.currNamespace + +end IdCompletionUtils + +section DotCompletionUtils + + private def unfoldeDefinitionGuarded? (e : Expr) : MetaM (Option Expr) := + try unfoldDefinition? e catch _ => pure none + + /-- Return `true` if `e` is a `declName`-application, or can be unfolded (delta-reduced) to one. -/ + private partial def isDefEqToAppOf (e : Expr) (declName : Name) : MetaM Bool := do + let isConstOf := match e.getAppFn with + | .const name .. => (privateToUserName? name).getD name == declName + | _ => false + if isConstOf then + return true + let some e ← unfoldeDefinitionGuarded? e | return false + isDefEqToAppOf e declName + + private def isDotCompletionMethod (typeName : Name) (info : ConstantInfo) : MetaM Bool := + forallTelescopeReducing info.type fun xs _ => do + for x in xs do + let localDecl ← x.fvarId!.getDecl + let type := localDecl.type.consumeMData + if (← isDefEqToAppOf type typeName) then + return true + return false + + /-- + Checks whether the expected type of `info.type` can be reduced to an application of `typeName`. + -/ + private def isDotIdCompletionMethod (typeName : Name) (info : ConstantInfo) : MetaM Bool := do + forallTelescopeReducing info.type fun _ type => + isDefEqToAppOf type.consumeMData typeName + + /-- + Converts `n` to `Name.anonymous` if `n` is a private prefix (see `Lean.isPrivatePrefix`). + -/ + private def stripPrivatePrefix (n : Name) : Name := + match n with + | .num _ 0 => if isPrivatePrefix n then .anonymous else n + | _ => n + + /-- + Compares `n₁` and `n₂` modulo private prefixes. Similar to `Name.cmp` but ignores all + private prefixes in both names. + Necessary because the namespaces of private names do not contain private prefixes. + -/ + private partial def cmpModPrivate (n₁ n₂ : Name) : Ordering := + let n₁ := stripPrivatePrefix n₁ + let n₂ := stripPrivatePrefix n₂ + match n₁, n₂ with + | .anonymous, .anonymous => Ordering.eq + | .anonymous, _ => Ordering.lt + | _, .anonymous => Ordering.gt + | .num p₁ i₁, .num p₂ i₂ => + match compare i₁ i₂ with + | Ordering.eq => cmpModPrivate p₁ p₂ + | ord => ord + | .num _ _, .str _ _ => Ordering.lt + | .str _ _, .num _ _ => Ordering.gt + | .str p₁ n₁, .str p₂ n₂ => + match compare n₁ n₂ with + | Ordering.eq => cmpModPrivate p₁ p₂ + | ord => ord + + /-- + `NameSet` where names are compared according to `cmpModPrivate`. + Helps speed up dot completion because it allows us to look up names without first having to + strip the private prefix from deep in the name, letting us reject most names without + having to scan the full name first. + -/ + private def NameSetModPrivate := RBTree Name cmpModPrivate + + /-- + Given a type, try to extract relevant type names for dot notation field completion. + We extract the type name, parent struct names, and unfold the type. + The process mimics the dot notation elaboration procedure at `App.lean` -/ + private partial def getDotCompletionTypeNames (type : Expr) : MetaM NameSetModPrivate := + return (← visit type |>.run RBTree.empty).2 + where + visit (type : Expr) : StateRefT NameSetModPrivate MetaM Unit := do + let .const typeName _ := type.getAppFn | return () + modify fun s => s.insert typeName + if isStructure (← getEnv) typeName then + for parentName in (← getAllParentStructures typeName) do + modify fun s => s.insert parentName + let some type ← unfoldeDefinitionGuarded? type | return () + visit type + +end DotCompletionUtils + +private def idCompletionCore + (ctx : ContextInfo) + (stx : Syntax) + (id : Name) + (hoverInfo : HoverInfo) + (danglingDot : Bool) + : M Unit := do + let mut id := id + if id.hasMacroScopes then + if stx.getHeadInfo matches .original .. then + id := id.eraseMacroScopes + else + -- Identifier is synthetic and has macro scopes => no completions + -- Erasing the macro scopes does not make sense in this case because the identifier name + -- is some random synthetic string. + return + let mut danglingDot := danglingDot + if let HoverInfo.inside delta := hoverInfo then + id := truncate id delta + danglingDot := false + if id.isAtomic then + -- search for matches in the local context + for localDecl in (← getLCtx) do + if let some score := matchAtomic id localDecl.userName then + addUnresolvedCompletionItem localDecl.userName (.fvar localDecl.fvarId) (kind := CompletionItemKind.variable) score + -- search for matches in the environment + let env ← getEnv + forEligibleDeclsM fun declName c => do + let bestMatch? ← (·.2) <$> StateT.run (s := none) do + let matchUsingNamespace (ns : Name) : StateT (Option (Name × Float)) M Unit := do + let some (label, score) ← matchDecl? ns id danglingDot declName + | return + modify fun + | none => + some (label, score) + | some (bestLabel, bestScore) => + -- for open namespaces `A` and `A.B` and a decl `A.B.c`, pick the decl `c` over `B.c` + if label.isSuffixOf bestLabel then + some (label, score) + else + some (bestLabel, bestScore) + let rec visitNamespaces (ns : Name) : StateT (Option (Name × Float)) M Unit := do + let Name.str p .. := ns + | return () + matchUsingNamespace ns + visitNamespaces p + -- use current namespace + visitNamespaces ctx.currNamespace + -- use open decls + for openDecl in ctx.openDecls do + let OpenDecl.simple ns exs := openDecl + | pure () + if exs.contains declName then + continue + matchUsingNamespace ns + matchUsingNamespace Name.anonymous + if let some (bestLabel, bestScore) := bestMatch? then + addUnresolvedCompletionItem bestLabel (.const declName) (← getCompletionKindForDecl c) bestScore + -- Recall that aliases may not be atomic and include the namespace where they were created. + let matchAlias (ns : Name) (alias : Name) : Option Float := + if ns.isPrefixOf alias then + matchAtomic id (alias.replacePrefix ns Name.anonymous) + else + none + let eligibleHeaderDecls ← getEligibleHeaderDecls env + -- Auxiliary function for `alias` + let addAlias (alias : Name) (declNames : List Name) (score : Float) : M Unit := do + declNames.forM fun declName => do + if allowCompletion eligibleHeaderDecls env declName then + addUnresolvedCompletionItemForDecl (.mkSimple alias.getString!) declName score + -- search explicitly open `ids` + for openDecl in ctx.openDecls do + match openDecl with + | OpenDecl.explicit openedId resolvedId => + if allowCompletion eligibleHeaderDecls env resolvedId then + if let some score := matchAtomic id openedId then + addUnresolvedCompletionItemForDecl (.mkSimple openedId.getString!) resolvedId score + | OpenDecl.simple ns _ => + getAliasState env |>.forM fun alias declNames => do + if let some score := matchAlias ns alias then + addAlias alias declNames score + -- search for aliases + getAliasState env |>.forM fun alias declNames => do + -- use current namespace + let rec searchAlias (ns : Name) : M Unit := do + if let some score := matchAlias ns alias then + addAlias alias declNames score + else + match ns with + | Name.str p .. => searchAlias p + | _ => return () + searchAlias ctx.currNamespace + -- Search keywords + if let .str .anonymous s := id then + let keywords := Parser.getTokenTable env + for keyword in keywords.findPrefix s do + if let some score := fuzzyMatchScoreWithThreshold? s keyword then + addKeywordCompletionItem keyword score + -- Search namespaces + completeNamespaces ctx id danglingDot + +def idCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (lctx : LocalContext) + (stx : Syntax) + (id : Name) + (hoverInfo : HoverInfo) + (danglingDot : Bool) + : IO (Array ScoredCompletionItem) := + runM params completionInfoPos ctx lctx do + idCompletionCore ctx stx id hoverInfo danglingDot + +def dotCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (info : TermInfo) + : IO (Array ScoredCompletionItem) := + runM params completionInfoPos ctx info.lctx do + let nameSet ← try + getDotCompletionTypeNames (← instantiateMVars (← inferType info.expr)) + catch _ => + pure RBTree.empty + if nameSet.isEmpty then + return + + forEligibleDeclsM fun declName c => do + let unnormedTypeName := declName.getPrefix + if ! nameSet.contains unnormedTypeName then + return + let some declName ← normPrivateName? declName + | return + let typeName := declName.getPrefix + if ! (← isDotCompletionMethod typeName c) then + return + let completionKind ← getCompletionKindForDecl c + addUnresolvedCompletionItem (.mkSimple c.name.getString!) (.const c.name) (kind := completionKind) 1 + +def dotIdCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (lctx : LocalContext) + (id : Name) + (expectedType? : Option Expr) + : IO (Array ScoredCompletionItem) := + runM params completionInfoPos ctx lctx do + let some expectedType := expectedType? + | return () + + let resultTypeFn := (← instantiateMVars expectedType).cleanupAnnotations.getAppFn.cleanupAnnotations + let .const .. := resultTypeFn + | return () + + let nameSet ← try + getDotCompletionTypeNames resultTypeFn + catch _ => + pure RBTree.empty + + forEligibleDeclsM fun declName c => do + let unnormedTypeName := declName.getPrefix + if ! nameSet.contains unnormedTypeName then + return + + let some declName ← normPrivateName? declName + | return + + let typeName := declName.getPrefix + if ! (← isDotIdCompletionMethod typeName c) then + return + + let completionKind ← getCompletionKindForDecl c + if id.isAnonymous then + -- We're completing a lone dot => offer all decls of the type + addUnresolvedCompletionItem (.mkSimple c.name.getString!) (.const c.name) completionKind 1 + return + + let some (label, score) ← matchDecl? typeName id (danglingDot := false) declName | pure () + addUnresolvedCompletionItem label (.const c.name) completionKind score + +def fieldIdCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (lctx : LocalContext) + (id : Option Name) + (structName : Name) + : IO (Array ScoredCompletionItem) := + runM params completionInfoPos ctx lctx do + let idStr := id.map (·.toString) |>.getD "" + let fieldNames := getStructureFieldsFlattened (← getEnv) structName (includeSubobjectFields := false) + for fieldName in fieldNames do + let .str _ fieldName := fieldName | continue + let some score := fuzzyMatchScoreWithThreshold? idStr fieldName | continue + let item := { label := fieldName, detail? := "field", documentation? := none, kind? := CompletionItemKind.field } + addItem item score + +def optionCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + (stx : Syntax) + (caps : ClientCapabilities) + : IO (Array ScoredCompletionItem) := + ctx.runMetaM {} do + let (partialName, trailingDot) := + -- `stx` is from `"set_option" >> ident` + match stx[1].getSubstring? (withLeading := false) (withTrailing := false) with + | none => ("", false) -- the `ident` is `missing`, list all options + | some ss => + if !ss.str.atEnd ss.stopPos && ss.str.get ss.stopPos == '.' then + -- include trailing dot, which is not parsed by `ident` + (ss.toString ++ ".", true) + else + (ss.toString, false) + -- HACK(WN): unfold the type so ForIn works + let (decls : RBMap _ _ _) ← getOptionDecls + let opts ← getOptions + let mut items := #[] + for ⟨name, decl⟩ in decls do + if let some score := fuzzyMatchScoreWithThreshold? partialName name.toString then + let textEdit := + if !caps.textDocument?.any (·.completion?.any (·.completionItem?.any (·.insertReplaceSupport?.any (·)))) then + none -- InsertReplaceEdit not supported by client + else if let some ⟨start, stop⟩ := stx[1].getRange? then + let stop := if trailingDot then stop + ' ' else stop + let range := ⟨ctx.fileMap.utf8PosToLspPos start, ctx.fileMap.utf8PosToLspPos stop⟩ + some { newText := name.toString, insert := range, replace := range : InsertReplaceEdit } + else + none + items := items.push ⟨{ + label := name.toString + detail? := s!"({opts.get name decl.defValue}), {decl.descr}" + documentation? := none, + kind? := CompletionItemKind.property -- TODO: investigate whether this is the best kind for options. + textEdit? := textEdit + data? := toJson { + params, + cPos := completionInfoPos, + id? := none : ResolvableCompletionItemData + } + }, score⟩ + return items + +def tacticCompletion + (params : CompletionParams) + (completionInfoPos : Nat) + (ctx : ContextInfo) + : IO (Array ScoredCompletionItem) := ctx.runMetaM .empty do + let allTacticDocs ← Tactic.Doc.allTacticDocs + let items : Array ScoredCompletionItem := allTacticDocs.map fun tacticDoc => + ⟨{ + label := tacticDoc.userName + detail? := none + documentation? := tacticDoc.docString.map fun docString => + { value := docString, kind := MarkupKind.markdown : MarkupContent } + kind? := CompletionItemKind.keyword + data? := toJson { params, cPos := completionInfoPos, id? := none : ResolvableCompletionItemData } + }, 1⟩ + return items + +end Lean.Server.Completion diff --git a/src/Lean/Server/Completion/CompletionInfoSelection.lean b/src/Lean/Server/Completion/CompletionInfoSelection.lean new file mode 100644 index 000000000000..0ce20cab48bf --- /dev/null +++ b/src/Lean/Server/Completion/CompletionInfoSelection.lean @@ -0,0 +1,131 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Marc Huisinga +-/ +prelude +import Lean.Server.Completion.SyntheticCompletion + +namespace Lean.Server.Completion +open Elab + +private def filterDuplicateCompletionInfos + (infos : Array ContextualizedCompletionInfo) + : Array ContextualizedCompletionInfo := Id.run do + -- We don't expect there to be too many duplicate completion infos, + -- so it's fine if this is quadratic (we don't need to implement `Hashable` / `LT` this way). + let mut deduplicatedInfos : Array ContextualizedCompletionInfo := #[] + for i in infos do + if deduplicatedInfos.any (fun di => eq di.info i.info) then + continue + deduplicatedInfos := deduplicatedInfos.push i + deduplicatedInfos +where + eq : CompletionInfo → CompletionInfo → Bool + | .dot ti₁ .., .dot ti₂ .. => + ti₁.stx.eqWithInfo ti₂.stx && ti₁.expr == ti₂.expr + | .id stx₁ id₁ .., .id stx₂ id₂ .. => + stx₁.eqWithInfo stx₂ && id₁ == id₂ + | .dotId stx₁ id₁ .., .id stx₂ id₂ .. => + stx₁.eqWithInfo stx₂ && id₁ == id₂ + | .fieldId stx₁ id₁? _ structName₁, .fieldId stx₂ id₂? _ structName₂ => + stx₁.eqWithInfo stx₂ && id₁? == id₂? && structName₁ == structName₂ + | .namespaceId stx₁, .namespaceId stx₂ => + stx₁.eqWithInfo stx₂ + | .option stx₁, .option stx₂ => + stx₁.eqWithInfo stx₂ + | .endSection stx₁ scopeNames₁, .endSection stx₂ scopeNames₂ => + stx₁.eqWithInfo stx₂ && scopeNames₁ == scopeNames₂ + | .tactic stx₁, .tactic stx₂ => + stx₁.eqWithInfo stx₂ + | _, _ => + false + +def findCompletionInfosAt + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + : Array ContextualizedCompletionInfo := Id.run do + let ⟨hoverLine, _⟩ := fileMap.toPosition hoverPos + let mut completionInfoCandidates := infoTree.foldInfo (init := #[]) (go hoverLine) + if completionInfoCandidates.isEmpty then + completionInfoCandidates := findSyntheticCompletions fileMap hoverPos cmdStx infoTree + return filterDuplicateCompletionInfos completionInfoCandidates +where + go + (hoverLine : Nat) + (ctx : ContextInfo) + (info : Info) + (best : Array ContextualizedCompletionInfo) + : Array ContextualizedCompletionInfo := Id.run do + let .ofCompletionInfo completionInfo := info + | return best + if ! info.occursInOrOnBoundary hoverPos then + return best + let headPos := info.pos?.get! + let tailPos := info.tailPos?.get! + let hoverInfo := + if hoverPos < tailPos then + HoverInfo.inside (hoverPos - headPos).byteIdx + else + HoverInfo.after + let ⟨headPosLine, _⟩ := fileMap.toPosition headPos + let ⟨tailPosLine, _⟩ := fileMap.toPosition info.tailPos?.get! + if headPosLine != hoverLine || headPosLine != tailPosLine then + return best + return best.push { hoverInfo, ctx, info := completionInfo } + +private def computePrioritizedCompletionPartitions + (items : Array (ContextualizedCompletionInfo × Nat)) + : Array (Array (ContextualizedCompletionInfo × Nat)) := + let partitions := items.groupByKey fun (i, _) => + let isId := i.info matches .id .. + let size? := Info.ofCompletionInfo i.info |>.size? + (isId, size?) + -- Sort partitions so that non-id completions infos come before id completion infos and + -- within those two groups, smaller sizes come before larger sizes. + let partitionsByPriority := partitions.toArray.qsort + fun ((isId₁, size₁?), _) ((isId₂, size₂?), _) => + match size₁?, size₂? with + | some _, none => true + | none, some _ => false + | _, _ => + match isId₁, isId₂ with + | false, true => true + | true, false => false + | _, _ => Id.run do + let some size₁ := size₁? + | return false + let some size₂ := size₂? + | return false + return size₁ < size₂ + partitionsByPriority.map (·.2) + +/-- +Finds all `CompletionInfo`s (both from the `InfoTree` and synthetic ones), prioritizes them, +arranges them in partitions of `CompletionInfo`s with the same priority and sorts these partitions +so that `CompletionInfo`s with the highest priority come first. +The returned `CompletionInfo`s are also tagged with their index in `findCompletionInfosAt` so that +when resolving a `CompletionItem`, we can reconstruct which `CompletionInfo` it was created from. + +In general, the `InfoTree` may contain multiple different `CompletionInfo`s covering `hoverPos`, +and so we need to decide which of these `CompletionInfo`s we want to use to show completions to the +user. We choose priorities by the following rules: +- Synthetic completions have the lowest priority since they are only intended as a backup. +- Non-identifier completions have the highest priority since they tend to be much more helpful than + identifier completions when available since there are typically way too many of the latter. +- Within the three groups [non-id completions, id completions, synthetic completions], + `CompletionInfo`s with a smaller range are considered to be better. +-/ +def findPrioritizedCompletionPartitionsAt + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + : Array (Array (ContextualizedCompletionInfo × Nat)) := + findCompletionInfosAt fileMap hoverPos cmdStx infoTree + |>.zipWithIndex + |> computePrioritizedCompletionPartitions + +end Lean.Server.Completion diff --git a/src/Lean/Server/CompletionItemData.lean b/src/Lean/Server/Completion/CompletionItemData.lean similarity index 100% rename from src/Lean/Server/CompletionItemData.lean rename to src/Lean/Server/Completion/CompletionItemData.lean diff --git a/src/Lean/Server/Completion/CompletionResolution.lean b/src/Lean/Server/Completion/CompletionResolution.lean new file mode 100644 index 000000000000..48e4302c3d87 --- /dev/null +++ b/src/Lean/Server/Completion/CompletionResolution.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Marc Huisinga +-/ +prelude +import Lean.Server.Completion.CompletionItemData +import Lean.Server.Completion.CompletionInfoSelection + +namespace Lean.Lsp + +/-- +Identifier that is sent from the server to the client as part of the `CompletionItem.data?` field. +Needed to resolve the `CompletionItem` when the client sends a `completionItem/resolve` request +for that item, again containing the `data?` field provided by the server. +-/ +inductive CompletionIdentifier where + | const (declName : Name) + | fvar (id : FVarId) + deriving FromJson, ToJson + +/-- +`CompletionItemData` that contains additional information to identify the item +in order to resolve it. +-/ +structure ResolvableCompletionItemData extends CompletionItemData where + /-- Position of the completion info that this completion item was created from. -/ + cPos : Nat + id? : Option CompletionIdentifier + deriving FromJson, ToJson + +private partial def consumeImplicitPrefix (e : Expr) (k : Expr → MetaM α) : MetaM α := do + match e with + | Expr.forallE n d b c => + -- We do not consume instance implicit arguments because the user probably wants be aware of this dependency + if c == .implicit then + Meta.withLocalDecl n c d fun arg => + consumeImplicitPrefix (b.instantiate1 arg) k + else + k e + | _ => k e + +/-- +Fills the `CompletionItem.detail?` field of `item` using the pretty-printed type identified by `id`. +-/ +def CompletionItem.resolve + (item : CompletionItem) + (id : CompletionIdentifier) + : MetaM CompletionItem := do + let env ← getEnv + let lctx ← getLCtx + let mut item := item + + if item.detail?.isNone then + let type? := match id with + | .const declName => + env.find? declName |>.map ConstantInfo.type + | .fvar id => + lctx.find? id |>.map LocalDecl.type + let detail? ← type?.mapM fun type => + consumeImplicitPrefix type fun typeWithoutImplicits => + return toString (← Meta.ppExpr typeWithoutImplicits) + item := { item with detail? := detail? } + + return item + +end Lean.Lsp + +namespace Lean.Server.Completion +open Lean.Lsp +open Elab + +/-- +Fills the `CompletionItem.detail?` field of `item` using the pretty-printed type identified by `id` +in the context found at `hoverPos` in `infoTree`. +-/ +def resolveCompletionItem? + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + (item : CompletionItem) + (id : CompletionIdentifier) + (completionInfoPos : Nat) + : IO CompletionItem := do + let completionInfos := findCompletionInfosAt fileMap hoverPos cmdStx infoTree + let some i := completionInfos.get? completionInfoPos + | return item + i.ctx.runMetaM i.info.lctx (item.resolve id) + +end Lean.Server.Completion diff --git a/src/Lean/Server/Completion/CompletionUtils.lean b/src/Lean/Server/Completion/CompletionUtils.lean new file mode 100644 index 000000000000..67fcfe57b055 --- /dev/null +++ b/src/Lean/Server/Completion/CompletionUtils.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Marc Huisinga +-/ +prelude +import Init.Prelude +import Lean.Elab.InfoTree.Types + +namespace Lean.Server.Completion +open Elab + +inductive HoverInfo : Type where + | after + | inside (delta : Nat) + +structure ContextualizedCompletionInfo where + hoverInfo : HoverInfo + ctx : ContextInfo + info : CompletionInfo + +end Lean.Server.Completion diff --git a/src/Lean/Server/Completion/EligibleHeaderDecls.lean b/src/Lean/Server/Completion/EligibleHeaderDecls.lean new file mode 100644 index 000000000000..56eed2f16ced --- /dev/null +++ b/src/Lean/Server/Completion/EligibleHeaderDecls.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Marc Huisinga +-/ +prelude +import Lean.Meta.CompletionName + +namespace Lean.Server.Completion +open Meta + +abbrev EligibleHeaderDecls := Std.HashMap Name ConstantInfo + +/-- Cached header declarations for which `allowCompletion headerEnv decl` is true. -/ +builtin_initialize eligibleHeaderDeclsRef : IO.Ref (Option EligibleHeaderDecls) ← + IO.mkRef none + +/-- +Returns the declarations in the header for which `allowCompletion env decl` is true, caching them +if not already cached. +-/ +def getEligibleHeaderDecls (env : Environment) : IO EligibleHeaderDecls := do + eligibleHeaderDeclsRef.modifyGet fun + | some eligibleHeaderDecls => (eligibleHeaderDecls, some eligibleHeaderDecls) + | none => + let (_, eligibleHeaderDecls) := + StateT.run (m := Id) (s := {}) do + -- `map₁` are the header decls + env.constants.map₁.forM fun declName c => do + modify fun eligibleHeaderDecls => + if allowCompletion env declName then + eligibleHeaderDecls.insert declName c + else + eligibleHeaderDecls + (eligibleHeaderDecls, some eligibleHeaderDecls) + +/-- Iterate over all declarations that are allowed in completion results. -/ +def forEligibleDeclsM [Monad m] [MonadEnv m] [MonadLiftT (ST IO.RealWorld) m] + [MonadLiftT IO m] (f : Name → ConstantInfo → m PUnit) : m PUnit := do + let env ← getEnv + (← getEligibleHeaderDecls env).forM f + -- map₂ are exactly the local decls + env.constants.map₂.forM fun name c => do + if allowCompletion env name then + f name c + +/-- Checks whether this declaration can appear in completion results. -/ +def allowCompletion (eligibleHeaderDecls : EligibleHeaderDecls) (env : Environment) + (declName : Name) : Bool := + eligibleHeaderDecls.contains declName || + env.constants.map₂.contains declName && Lean.Meta.allowCompletion env declName + +end Lean.Server.Completion diff --git a/src/Lean/Server/ImportCompletion.lean b/src/Lean/Server/Completion/ImportCompletion.lean similarity index 98% rename from src/Lean/Server/ImportCompletion.lean rename to src/Lean/Server/Completion/ImportCompletion.lean index 1e096a2152e7..702c2445ca0b 100644 --- a/src/Lean/Server/ImportCompletion.lean +++ b/src/Lean/Server/Completion/ImportCompletion.lean @@ -4,13 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Marc Huisinga -/ prelude -import Lean.Data.Name import Lean.Data.NameTrie -import Lean.Data.Lsp.Utf16 -import Lean.Data.Lsp.LanguageFeatures import Lean.Util.Paths import Lean.Util.LakePath -import Lean.Server.CompletionItemData +import Lean.Server.Completion.CompletionItemData namespace ImportCompletion diff --git a/src/Lean/Server/Completion/SyntheticCompletion.lean b/src/Lean/Server/Completion/SyntheticCompletion.lean new file mode 100644 index 000000000000..86336eb3b21c --- /dev/null +++ b/src/Lean/Server/Completion/SyntheticCompletion.lean @@ -0,0 +1,396 @@ +/- +Copyright (c) 2024 Lean FRO, LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Marc Huisinga +-/ +prelude +import Lean.Server.InfoUtils +import Lean.Server.Completion.CompletionUtils + +namespace Lean.Server.Completion +open Elab + +private def findBest? + (infoTree : InfoTree) + (gt : α → α → Bool) + (f : ContextInfo → Info → PersistentArray InfoTree → Option α) + : Option α := + infoTree.visitM (m := Id) (postNode := choose) |>.join +where + choose + (ctx : ContextInfo) + (info : Info) + (cs : PersistentArray InfoTree) + (childValues : List (Option (Option α))) + : Option α := + let bestChildValue := childValues.map (·.join) |>.foldl (init := none) fun v best => + if isBetter v best then + v + else + best + if let some v := f ctx info cs then + if isBetter v bestChildValue then + v + else + bestChildValue + else + bestChildValue + isBetter (a b : Option α) : Bool := + match a, b with + | none, none => false + | some _, none => true + | none, some _ => false + | some a, some b => gt a b + +/-- +If there are `Info`s that contain `hoverPos` and have a nonempty `LocalContext`, +yields the closest one of those `Info`s. +Otherwise, yields the closest `Info` that contains `hoverPos` and has an empty `LocalContext`. +-/ +private def findClosestInfoWithLocalContextAt? + (hoverPos : String.Pos) + (infoTree : InfoTree) + : Option (ContextInfo × Info) := + findBest? infoTree isBetter fun ctx info _ => + if info.occursInOrOnBoundary hoverPos then + (ctx, info) + else + none +where + isBetter (a b : ContextInfo × Info) : Bool := + let (_, ia) := a + let (_, ib) := b + if !ia.lctx.isEmpty && ib.lctx.isEmpty then + true + else if ia.lctx.isEmpty && !ib.lctx.isEmpty then + false + else if ia.isSmaller ib then + true + else if ib.isSmaller ia then + false + else + false + +private def findSyntheticIdentifierCompletion? + (hoverPos : String.Pos) + (infoTree : InfoTree) + : Option ContextualizedCompletionInfo := do + let some (ctx, info) := findClosestInfoWithLocalContextAt? hoverPos infoTree + | none + let some stack := info.stx.findStack? (·.getRange?.any (·.contains hoverPos (includeStop := true))) + | none + let stack := stack.dropWhile fun (stx, _) => !(stx matches `($_:ident) || stx matches `($_:ident.)) + let some (stx, _) := stack.head? + | none + let isDotIdCompletion := stack.any fun (stx, _) => stx matches `(.$_:ident) + if isDotIdCompletion then + -- An identifier completion is never useful in a dotId completion context. + none + let some (id, danglingDot) := + match stx with + | `($id:ident) => some (id.getId, false) + | `($id:ident.) => some (id.getId, true) + | _ => none + | none + let tailPos := stx.getTailPos?.get! + let hoverInfo := + if hoverPos < tailPos then + HoverInfo.inside (tailPos - hoverPos).byteIdx + else + HoverInfo.after + some { hoverInfo, ctx, info := .id stx id danglingDot info.lctx none } + +private partial def getIndentationAmount (fileMap : FileMap) (line : Nat) : Nat := Id.run do + let lineStartPos := fileMap.lineStart line + let lineEndPos := fileMap.lineStart (line + 1) + let mut it : String.Iterator := ⟨fileMap.source, lineStartPos⟩ + let mut indentationAmount := 0 + while it.pos < lineEndPos do + let c := it.curr + if c = ' ' || c = '\t' then + indentationAmount := indentationAmount + 1 + else + break + it := it.next + return indentationAmount + +private partial def isCursorOnWhitespace (fileMap : FileMap) (hoverPos : String.Pos) : Bool := + fileMap.source.atEnd hoverPos || (fileMap.source.get hoverPos).isWhitespace + +private partial def isCursorInProperWhitespace (fileMap : FileMap) (hoverPos : String.Pos) : Bool := + (fileMap.source.atEnd hoverPos || (fileMap.source.get hoverPos).isWhitespace) + && (fileMap.source.get (hoverPos - ⟨1⟩)).isWhitespace + +private partial def isSyntheticTacticCompletion + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + : Bool := Id.run do + let hoverFilePos := fileMap.toPosition hoverPos + let mut hoverLineIndentation := getIndentationAmount fileMap hoverFilePos.line + if hoverFilePos.column < hoverLineIndentation then + -- Ignore trailing whitespace after the cursor + hoverLineIndentation := hoverFilePos.column + go hoverFilePos hoverLineIndentation cmdStx 0 +where + go + (hoverFilePos : Position) + (hoverLineIndentation : Nat) + (stx : Syntax) + (leadingWs : Nat) + : Bool := Id.run do + match stx.getPos?, stx.getTailPos? with + | some startPos, some endPos => + let isCursorInCompletionRange := + startPos.byteIdx - leadingWs <= hoverPos.byteIdx + && hoverPos.byteIdx <= endPos.byteIdx + stx.getTrailingSize + if ! isCursorInCompletionRange then + return false + let mut wsBeforeArg := leadingWs + for arg in stx.getArgs do + if go hoverFilePos hoverLineIndentation arg wsBeforeArg then + return true + -- We must account for the whitespace before an argument because the syntax nodes we use + -- to identify tactic blocks only start *after* the whitespace following a `by`, and we + -- want to provide tactic completions in that whitespace as well. + -- This method of computing whitespace assumes that there are no syntax nodes without tokens + -- after `by` and before the first proper tactic syntax. + wsBeforeArg := arg.getTrailingSize + return isCompletionInEmptyTacticBlock stx + || isCompletionAfterSemicolon stx + || isCompletionOnTacticBlockIndentation hoverFilePos hoverLineIndentation stx + | _, _ => + -- Empty tactic blocks typically lack ranges since they do not contain any tokens. + -- We do not perform more precise range checking in this case because we assume that empty + -- tactic blocks always occur within other syntax with ranges that let us narrow down the + -- search to the degree that we can be sure that the cursor is indeed in this empty tactic + -- block. + return isCompletionInEmptyTacticBlock stx + + isCompletionOnTacticBlockIndentation + (hoverFilePos : Position) + (hoverLineIndentation : Nat) + (stx : Syntax) + : Bool := Id.run do + let isCursorInIndentation := hoverFilePos.column <= hoverLineIndentation + if ! isCursorInIndentation then + -- Do not trigger tactic completion at the end of a properly indented tactic block line since + -- that line might already have entered term mode by that point. + return false + let some tacticsNode := getTacticsNode? stx + | return false + let some firstTacticPos := tacticsNode.getPos? + | return false + let firstTacticLine := fileMap.toPosition firstTacticPos |>.line + let firstTacticIndentation := getIndentationAmount fileMap firstTacticLine + -- This ensures that we do not accidentally provide tactic completions in a term mode proof - + -- tactic completions are only provided at the same indentation level as the other tactics in + -- that tactic block. + let isCursorInTacticBlock := hoverLineIndentation == firstTacticIndentation + return isCursorInProperWhitespace fileMap hoverPos && isCursorInTacticBlock + + isCompletionAfterSemicolon (stx : Syntax) : Bool := Id.run do + let some tacticsNode := getTacticsNode? stx + | return false + let tactics := tacticsNode.getArgs + -- We want to provide completions in the case of `skip;`, so the cursor must only be on + -- whitespace, not in proper whitespace. + return isCursorOnWhitespace fileMap hoverPos && tactics.any fun tactic => Id.run do + let some tailPos := tactic.getTailPos? + | return false + let isCursorAfterSemicolon := + tactic.isToken ";" + && tailPos.byteIdx <= hoverPos.byteIdx + && hoverPos.byteIdx <= tailPos.byteIdx + tactic.getTrailingSize + return isCursorAfterSemicolon + + getTacticsNode? (stx : Syntax) : Option Syntax := + if stx.getKind == ``Parser.Tactic.tacticSeq1Indented then + some stx[0] + else if stx.getKind == ``Parser.Tactic.tacticSeqBracketed then + some stx[1] + else + none + + isCompletionInEmptyTacticBlock (stx : Syntax) : Bool := + isCursorInProperWhitespace fileMap hoverPos && isEmptyTacticBlock stx + + isEmptyTacticBlock (stx : Syntax) : Bool := + stx.getKind == ``Parser.Tactic.tacticSeq && isEmpty stx + || stx.getKind == ``Parser.Tactic.tacticSeq1Indented && isEmpty stx + || stx.getKind == ``Parser.Tactic.tacticSeqBracketed && isEmpty stx[1] + + isEmpty : Syntax → Bool + | .missing => true + | .ident .. => false + | .atom .. => false + | .node _ _ args => args.all isEmpty + +private partial def findOutermostContextInfo? (i : InfoTree) : Option ContextInfo := + go i +where + go (i : InfoTree) : Option ContextInfo := do + match i with + | .context ctx i => + match ctx with + | .commandCtx ctxInfo => + some { ctxInfo with } + | _ => + -- This shouldn't happen (see the `PartialContextInfo` docstring), + -- but let's continue searching regardless + go i + | .node _ cs => + cs.findSome? go + | .hole .. => + none + +private def findSyntheticTacticCompletion? + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + : Option ContextualizedCompletionInfo := do + let ctx ← findOutermostContextInfo? infoTree + if ! isSyntheticTacticCompletion fileMap hoverPos cmdStx then + none + -- Neither `HoverInfo` nor the syntax in `.tactic` are important for tactic completion. + return { hoverInfo := HoverInfo.after, ctx, info := .tactic .missing } + +private def findExpectedTypeAt (infoTree : InfoTree) (hoverPos : String.Pos) : Option (ContextInfo × Expr) := do + let (ctx, .ofTermInfo i) ← infoTree.smallestInfo? fun i => Id.run do + let some pos := i.pos? + | return false + let some tailPos := i.tailPos? + | return false + let .ofTermInfo ti := i + | return false + return ti.expectedType?.isSome && pos <= hoverPos && hoverPos <= tailPos + | none + (ctx, i.expectedType?.get!) + +private partial def foldWithLeadingToken [Inhabited α] + (f : α → Option Syntax → Syntax → α) + (init : α) + (stx : Syntax) + : α := + let (_, r) := go none init stx + r +where + go [Inhabited α] (leadingToken? : Option Syntax) (acc : α) (stx : Syntax) : Option Syntax × α := + let acc := f acc leadingToken? stx + match stx with + | .missing => (none, acc) + | .atom .. => (stx, acc) + | .ident .. => (stx, acc) + | .node _ _ args => Id.run do + let mut acc := acc + let mut lastToken? := none + for arg in args do + let (lastToken'?, acc') := go (lastToken? <|> leadingToken?) acc arg + lastToken? := lastToken'? <|> lastToken? + acc := acc' + return (lastToken?, acc) + +private def findWithLeadingToken? + (p : Option Syntax → Syntax → Bool) + (stx : Syntax) + : Option Syntax := + foldWithLeadingToken (stx := stx) (init := none) fun foundStx? leadingToken? stx => + match foundStx? with + | some foundStx => foundStx + | none => + if p leadingToken? stx then + some stx + else + none + +private def isSyntheticStructFieldCompletion + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + : Bool := Id.run do + let isCursorOnWhitespace := isCursorOnWhitespace fileMap hoverPos + let isCursorInProperWhitespace := isCursorInProperWhitespace fileMap hoverPos + if ! isCursorOnWhitespace then + return false + + let hoverFilePos := fileMap.toPosition hoverPos + let mut hoverLineIndentation := getIndentationAmount fileMap hoverFilePos.line + if hoverFilePos.column < hoverLineIndentation then + -- Ignore trailing whitespace after the cursor + hoverLineIndentation := hoverFilePos.column + + return Option.isSome <| findWithLeadingToken? (stx := cmdStx) fun leadingToken? stx => Id.run do + let some leadingToken := leadingToken? + | return false + + if stx.getKind != ``Parser.Term.structInstFields then + return false + + let fieldsAndSeps := stx[0].getArgs + let some outerBoundsStart := leadingToken.getTailPos? (canonicalOnly := true) + | return false + let some outerBoundsStop := + stx.getTrailingTailPos? (canonicalOnly := true) + <|> leadingToken.getTrailingTailPos? (canonicalOnly := true) + | return false + let outerBounds : String.Range := ⟨outerBoundsStart, outerBoundsStop⟩ + + let isCompletionInEmptyBlock := + fieldsAndSeps.isEmpty && outerBounds.contains hoverPos (includeStop := true) + if isCompletionInEmptyBlock then + return true + + let isCompletionAfterSep := fieldsAndSeps.zipWithIndex.any fun (fieldOrSep, i) => Id.run do + if i % 2 == 0 || !fieldOrSep.isAtom then + return false + let sep := fieldOrSep + let some sepTailPos := sep.getTailPos? + | return false + return sepTailPos <= hoverPos + && hoverPos.byteIdx <= sepTailPos.byteIdx + sep.getTrailingSize + if isCompletionAfterSep then + return true + + let isCompletionOnIndentation := Id.run do + if ! isCursorInProperWhitespace then + return false + let isCursorInIndentation := hoverFilePos.column <= hoverLineIndentation + if ! isCursorInIndentation then + return false + let some firstFieldPos := stx.getPos? + | return false + let firstFieldLine := fileMap.toPosition firstFieldPos |>.line + let firstFieldIndentation := getIndentationAmount fileMap firstFieldLine + let isCursorInBlock := hoverLineIndentation == firstFieldIndentation + return isCursorInBlock + return isCompletionOnIndentation + +private def findSyntheticFieldCompletion? + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + : Option ContextualizedCompletionInfo := do + if ! isSyntheticStructFieldCompletion fileMap hoverPos cmdStx then + none + let (ctx, expectedType) ← findExpectedTypeAt infoTree hoverPos + let .const typeName _ := expectedType.getAppFn + | none + if ! isStructure ctx.env typeName then + none + return { hoverInfo := HoverInfo.after, ctx, info := .fieldId .missing none .empty typeName } + +def findSyntheticCompletions + (fileMap : FileMap) + (hoverPos : String.Pos) + (cmdStx : Syntax) + (infoTree : InfoTree) + : Array ContextualizedCompletionInfo := + let syntheticCompletionData? : Option ContextualizedCompletionInfo := + findSyntheticTacticCompletion? fileMap hoverPos cmdStx infoTree <|> + findSyntheticFieldCompletion? fileMap hoverPos cmdStx infoTree <|> + findSyntheticIdentifierCompletion? hoverPos infoTree + syntheticCompletionData?.map (#[·]) |>.getD #[] + +end Lean.Server.Completion diff --git a/src/Lean/Server/FileWorker.lean b/src/Lean/Server/FileWorker.lean index a9b7f4d8885e..a6a4dccbc36d 100644 --- a/src/Lean/Server/FileWorker.lean +++ b/src/Lean/Server/FileWorker.lean @@ -28,7 +28,7 @@ import Lean.Server.FileWorker.WidgetRequests import Lean.Server.FileWorker.SetupFile import Lean.Server.Rpc.Basic import Lean.Widget.InteractiveDiagnostic -import Lean.Server.ImportCompletion +import Lean.Server.Completion.ImportCompletion /-! For general server architecture, see `README.md`. For details of IPC communication, see `Watchdog.lean`. diff --git a/src/Lean/Server/FileWorker/RequestHandling.lean b/src/Lean/Server/FileWorker/RequestHandling.lean index 57418b4fa1cf..424b475ad287 100644 --- a/src/Lean/Server/FileWorker/RequestHandling.lean +++ b/src/Lean/Server/FileWorker/RequestHandling.lean @@ -46,10 +46,11 @@ def handleCompletion (p : CompletionParams) mapTask (findCompletionCmdDataAtPos doc pos) fun cmdData? => do let some (cmdStx, infoTree) := cmdData? -- work around https://github.com/microsoft/vscode/issues/155738 - | return { items := #[{label := "-"}], isIncomplete := true } - if let some r ← Completion.find? p doc.meta.text pos cmdStx infoTree caps then - return r - return { items := #[ ], isIncomplete := true } + | return { + items := #[{label := "-", data? := toJson { params := p : Lean.Lsp.CompletionItemData }}], + isIncomplete := true + } + Completion.find? p doc.meta.text pos cmdStx infoTree caps /-- Handles `completionItem/resolve` requests that are sent by the client after the user selects @@ -62,7 +63,7 @@ def handleCompletionItemResolve (item : CompletionItem) : RequestM (RequestTask CompletionItem) := do let doc ← readDoc let text := doc.meta.text - let some (data : CompletionItemDataWithId) := item.data?.bind fun data => (fromJson? data).toOption + let some (data : ResolvableCompletionItemData) := item.data?.bind fun data => (fromJson? data).toOption | return .pure item let some id := data.id? | return .pure item @@ -70,7 +71,7 @@ def handleCompletionItemResolve (item : CompletionItem) mapTask (findCompletionCmdDataAtPos doc pos) fun cmdData? => do let some (cmdStx, infoTree) := cmdData? | return item - Completion.resolveCompletionItem? text pos cmdStx infoTree item id + Completion.resolveCompletionItem? text pos cmdStx infoTree item id data.cPos open Elab in def handleHover (p : HoverParams) diff --git a/src/Lean/Server/InfoUtils.lean b/src/Lean/Server/InfoUtils.lean index 2f3fba396994..b642d7384911 100644 --- a/src/Lean/Server/InfoUtils.lean +++ b/src/Lean/Server/InfoUtils.lean @@ -136,6 +136,7 @@ def InfoTree.getCompletionInfos (infoTree : InfoTree) : Array (ContextInfo × Co def Info.stx : Info → Syntax | ofTacticInfo i => i.stx | ofTermInfo i => i.stx + | ofPartialTermInfo i => i.stx | ofCommandInfo i => i.stx | ofMacroExpansionInfo i => i.stx | ofOptionInfo i => i.stx @@ -146,6 +147,7 @@ def Info.stx : Info → Syntax | ofFVarAliasInfo _ => .missing | ofFieldRedeclInfo i => i.stx | ofOmissionInfo i => i.stx + | ofChoiceInfo i => i.stx def Info.lctx : Info → LocalContext | .ofTermInfo i => i.lctx diff --git a/stage0/src/runtime/io.cpp b/stage0/src/runtime/io.cpp index afeea75c92ff..44f18cc58733 100644 --- a/stage0/src/runtime/io.cpp +++ b/stage0/src/runtime/io.cpp @@ -5,6 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Sebastian Ullrich */ #if defined(LEAN_WINDOWS) +#include #include #include #define NOMINMAX // prevent ntdef.h from defining min/max macros @@ -630,6 +631,176 @@ extern "C" LEAN_EXPORT obj_res lean_io_prim_handle_put_str(b_obj_arg h, b_obj_ar } } +/* Std.Time.Timestamp.now : IO Timestamp */ +extern "C" LEAN_EXPORT obj_res lean_get_current_time(obj_arg /* w */) { + using namespace std::chrono; + + std::chrono::system_clock::time_point now = std::chrono::system_clock::now(); + long long timestamp = std::chrono::duration_cast(now.time_since_epoch()).count(); + + long long secs = timestamp / 1000000000; + long long nano = timestamp % 1000000000; + + lean_object *lean_ts = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_ts, 0, lean_int64_to_int(secs)); + lean_ctor_set(lean_ts, 1, lean_int64_to_int(nano)); + + return lean_io_result_mk_ok(lean_ts); +} + +/* Std.Time.Database.Windows.getNextTransition : @&String -> Int64 -> Bool -> IO (Option (Int64 × TimeZone)) */ +extern "C" LEAN_EXPORT obj_res lean_windows_get_next_transition(b_obj_arg timezone_str, uint64_t tm_obj, uint8 default_time, obj_arg /* w */) { +#if defined(LEAN_WINDOWS) + UErrorCode status = U_ZERO_ERROR; + const char* dst_name_id = lean_string_cstr(timezone_str); + + UChar tzID[256]; + u_strFromUTF8(tzID, sizeof(tzID) / sizeof(tzID[0]), NULL, dst_name_id, strlen(dst_name_id), &status); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read identifier"))); + } + + UCalendar *cal = ucal_open(tzID, -1, NULL, UCAL_GREGORIAN, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); + } + + int64_t tm = 0; + + if (!default_time) { + int64_t timestamp_secs = (int64_t)tm_obj; + + ucal_setMillis(cal, timestamp_secs * 1000, &status); + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); + } + + UDate nextTransition; + if (!ucal_getTimeZoneTransitionDate(cal, UCAL_TZ_TRANSITION_NEXT, &nextTransition, &status)) { + ucal_close(cal); + return io_result_mk_ok(mk_option_none()); + } + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get next transation"))); + } + + tm = (int64_t)(nextTransition / 1000.0); + } + + int32_t dst_offset = ucal_get(cal, UCAL_DST_OFFSET, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get dst_offset"))); + } + + int is_dst = dst_offset != 0; + + int32_t tzIDLength = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_DST : UCAL_STANDARD, "en_US", tzID, 32, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to timezone identifier"))); + } + + char dst_name[256]; + int32_t dst_name_len; + u_strToUTF8(dst_name, sizeof(dst_name), &dst_name_len, tzID, tzIDLength, &status); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to convert DST name to UTF-8"))); + } + + UChar display_name[32]; + int32_t display_name_len = ucal_getTimeZoneDisplayName(cal, is_dst ? UCAL_SHORT_DST : UCAL_SHORT_STANDARD, "en_US", display_name, 32, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to read abbreaviation"))); + } + + char display_name_str[256]; + int32_t display_name_str_len; + u_strToUTF8(display_name_str, sizeof(display_name_str), &display_name_str_len, display_name, display_name_len, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get abbreviation to cstr"))); + } + + int32_t zone_offset = ucal_get(cal, UCAL_ZONE_OFFSET, &status); + zone_offset += dst_offset; + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get zone_offset"))); + } + + ucal_close(cal); + + int offset_seconds = zone_offset / 1000; + + lean_object *lean_tz = lean_alloc_ctor(0, 3, 1); + lean_ctor_set(lean_tz, 0, lean_int_to_int(offset_seconds)); + lean_ctor_set(lean_tz, 1, lean_mk_string_from_bytes_unchecked(dst_name, dst_name_len)); + lean_ctor_set(lean_tz, 2, lean_mk_string_from_bytes_unchecked(display_name_str, display_name_str_len)); + lean_ctor_set_uint8(lean_tz, sizeof(void*)*3, is_dst); + + lean_object *lean_pair = lean_alloc_ctor(0, 2, 0); + lean_ctor_set(lean_pair, 0, lean_box_uint64((uint64_t)tm)); + lean_ctor_set(lean_pair, 1, lean_tz); + + return lean_io_result_mk_ok(mk_option_some(lean_pair)); +#else + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone, its windows only."))); +#endif +} + +/* Std.Time.Database.Windows.getLocalTimeZoneIdentifierAt : Int64 → IO String */ +extern "C" LEAN_EXPORT obj_res lean_get_windows_local_timezone_id_at(uint64_t tm_obj, obj_arg /* w */) { +#if defined(LEAN_WINDOWS) + UErrorCode status = U_ZERO_ERROR; + UCalendar* cal = ucal_open(NULL, -1, NULL, UCAL_GREGORIAN, &status); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to open calendar"))); + } + + int64_t timestamp_secs = (int64_t)tm_obj; + ucal_setMillis(cal, timestamp_secs * 1000, &status); + + if (U_FAILURE(status)) { + ucal_close(cal); + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to set calendar time"))); + } + + UChar tzId[256]; + int32_t tzIdLength = ucal_getTimeZoneID(cal, tzId, sizeof(tzId) / sizeof(tzId[0]), &status); + ucal_close(cal); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to get timezone ID"))); + } + + char tzIdStr[256]; + u_strToUTF8(tzIdStr, sizeof(tzIdStr), NULL, tzId, tzIdLength, &status); + + if (U_FAILURE(status)) { + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("failed to convert timezone ID to UTF-8"))); + } + + return lean_io_result_mk_ok(lean_mk_ascii_string_unchecked(tzIdStr)); +#else + return lean_io_result_mk_error(lean_decode_io_error(EINVAL, mk_string("timezone retrieval is Windows-only"))); +#endif +} + /* monoMsNow : BaseIO Nat */ extern "C" LEAN_EXPORT obj_res lean_io_mono_ms_now(obj_arg /* w */) { static_assert(sizeof(std::chrono::milliseconds::rep) <= sizeof(uint64), "size of std::chrono::nanoseconds::rep may not exceed 64"); diff --git a/stage0/stdlib/Init/Data/Array/Attach.c b/stage0/stdlib/Init/Data/Array/Attach.c index 338de863c19b..f91080e85d3d 100644 --- a/stage0/stdlib/Init/Data/Array/Attach.c +++ b/stage0/stdlib/Init/Data/Array/Attach.c @@ -13,22 +13,112 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Array_pmap(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_attachWithImpl___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Array_pmap___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Array_unattach___spec__1___rarg(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_attach(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Array_unattach___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_pmapImpl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_attach___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_pmap___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_attach___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_attachWithImpl(lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_attachWithImpl___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_unattach___rarg(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_unattach(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_pmapImpl___rarg(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Array_unattach___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Array_pmap___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Array_pmap___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_List_reverse___rarg(x_3); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_1); +x_8 = lean_apply_2(x_1, x_6, lean_box(0)); +lean_ctor_set(x_2, 1, x_3); +lean_ctor_set(x_2, 0, x_8); +{ +lean_object* _tmp_1 = x_7; +lean_object* _tmp_2 = x_2; +x_2 = _tmp_1; +x_3 = _tmp_2; +} +goto _start; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); +lean_inc(x_1); +x_12 = lean_apply_2(x_1, x_10, lean_box(0)); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +x_2 = x_11; +x_3 = x_13; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Array_pmap___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_List_mapTR_loop___at_Array_pmap___spec__1___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_pmap___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_array_to_list(x_2); +x_5 = lean_box(0); +x_6 = l_List_mapTR_loop___at_Array_pmap___spec__1___rarg(x_1, x_4, x_5); +x_7 = lean_array_mk(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_pmap(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_pmap___rarg), 3, 0); +return x_4; +} +} LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_attachWithImpl___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -77,6 +167,71 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +lean_inc(x_1); +x_11 = lean_apply_2(x_1, x_6, lean_box(0)); +x_12 = lean_array_uset(x_8, x_3, x_11); +x_3 = x_10; +x_4 = x_12; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg___boxed), 4, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_pmapImpl___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_array_size(x_2); +x_5 = 0; +x_6 = l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg(x_1, x_4, x_5, x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Attach_0__Array_pmapImpl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Attach_0__Array_pmapImpl___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at___private_Init_Data_Array_Attach_0__Array_pmapImpl___spec__1___rarg(x_1, x_5, x_6, x_4); +return x_7; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Array_unattach___spec__1___rarg(size_t x_1, size_t x_2, lean_object* x_3) { _start: { diff --git a/stage0/stdlib/Init/Data/List/Nat/Range.c b/stage0/stdlib/Init/Data/List/Nat/Range.c index bfca59b415a5..4c205b2fc5cd 100644 --- a/stage0/stdlib/Init/Data/List/Nat/Range.c +++ b/stage0/stdlib/Init/Data/List/Nat/Range.c @@ -13,44 +13,44 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8; -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378_; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_627_; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373_; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13; -static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20; -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_632_; -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1() { +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3; +static lean_object* l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1() { _start: { lean_object* x_1; @@ -58,7 +58,7 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2() { _start: { lean_object* x_1; @@ -66,7 +66,7 @@ x_1 = lean_mk_string_unchecked("Parser", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3() { _start: { lean_object* x_1; @@ -74,7 +74,7 @@ x_1 = lean_mk_string_unchecked("Tactic", 6, 6); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4() { _start: { lean_object* x_1; @@ -82,19 +82,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3; -x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3; +x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -103,7 +103,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7() { _start: { lean_object* x_1; @@ -111,19 +111,19 @@ x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3; -x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3; +x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9() { _start: { lean_object* x_1; @@ -131,17 +131,17 @@ x_1 = lean_mk_string_unchecked("null", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11() { _start: { lean_object* x_1; @@ -149,41 +149,41 @@ x_1 = lean_mk_string_unchecked("simp", 4, 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3; -x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3; +x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15() { _start: { lean_object* x_1; @@ -191,25 +191,25 @@ x_1 = lean_mk_string_unchecked("optConfig", 9, 9); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3; -x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3; +x_4 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -217,23 +217,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -241,63 +241,63 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -305,23 +305,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -329,23 +329,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -353,23 +353,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6; -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5; -x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30; +x_2 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5; +x_3 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -377,19 +377,19 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378_() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_632_() { +static lean_object* _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_627_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31; +x_1 = l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31; return x_1; } } @@ -418,72 +418,72 @@ lean_dec_ref(res); res = initialize_Init_Data_List_Erase(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__1); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__2); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__3); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__4); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__5); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__6); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__7); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__8); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__9); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__10); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__11); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__12); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__13); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__14); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__15); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__16); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__17); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__18); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__19); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__20); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__21); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__22); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__23); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__24); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__25); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__26); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__27); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__28); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__29); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__30); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378____closed__31); -l___auto____x40_Init_Data_List_Nat_Range___hyg_378_ = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_378_(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_378_); -l___auto____x40_Init_Data_List_Nat_Range___hyg_632_ = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_632_(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_632_); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__1); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__2); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__3); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__4); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__5); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__6); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__7); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__8); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__9); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__10); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__11); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__12); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__13); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__14); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__15); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__16); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__17); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__18); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__19); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__20); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__21); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__22); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__23); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__24); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__25); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__26); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__27); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__28); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__29); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__30); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31 = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373____closed__31); +l___auto____x40_Init_Data_List_Nat_Range___hyg_373_ = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_373_(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_373_); +l___auto____x40_Init_Data_List_Nat_Range___hyg_627_ = _init_l___auto____x40_Init_Data_List_Nat_Range___hyg_627_(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Nat_Range___hyg_627_); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/List/Sort/Basic.c b/stage0/stdlib/Init/Data/List/Sort/Basic.c index d0a58e7d392c..b4eee2123d83 100644 --- a/stage0/stdlib/Init/Data/List/Sort/Basic.c +++ b/stage0/stdlib/Init/Data/List/Sort/Basic.c @@ -37,6 +37,7 @@ static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed LEAN_EXPORT lean_object* l_List_mergeSort___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__51; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__23; +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_812_; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__41; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__12; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__13; @@ -88,7 +89,6 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__4; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__28; lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_818_; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__31; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__27; static lean_object* l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__1; @@ -995,7 +995,7 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Sort_Basic___hyg_818_() { +static lean_object* _init_l___auto____x40_Init_Data_List_Sort_Basic___hyg_812_() { _start: { lean_object* x_1; @@ -1365,8 +1365,8 @@ l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__59 = _init_l___aut lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Basic___hyg_12____closed__59); l___auto____x40_Init_Data_List_Sort_Basic___hyg_12_ = _init_l___auto____x40_Init_Data_List_Sort_Basic___hyg_12_(); lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Basic___hyg_12_); -l___auto____x40_Init_Data_List_Sort_Basic___hyg_818_ = _init_l___auto____x40_Init_Data_List_Sort_Basic___hyg_818_(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Basic___hyg_818_); +l___auto____x40_Init_Data_List_Sort_Basic___hyg_812_ = _init_l___auto____x40_Init_Data_List_Sort_Basic___hyg_812_(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Basic___hyg_812_); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/List/Sort/Impl.c b/stage0/stdlib/Init/Data/List/Sort/Impl.c index db4715f696cb..151b88a17180 100644 --- a/stage0/stdlib/Init/Data/List/Sort/Impl.c +++ b/stage0/stdlib/Init/Data/List/Sort/Impl.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Sort_Impl___hyg_2003_; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_List_MergeSort_Internal_splitRevAt_go___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_MergeSort_Internal_mergeSortTR_u2082_run_x27(lean_object*); @@ -62,6 +61,7 @@ LEAN_EXPORT lean_object* l_List_MergeSort_Internal_mergeSortTR_u2082_run___rarg_ LEAN_EXPORT lean_object* l_List_MergeSort_Internal_splitRevInTwo_x27___rarg___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127____closed__28; LEAN_EXPORT lean_object* l_List_MergeSort_Internal_mergeSortTR_u2082(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Data_List_Sort_Impl___hyg_1991_; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127____closed__58; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -1379,7 +1379,7 @@ lean_dec(x_1); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Data_List_Sort_Impl___hyg_2003_() { +static lean_object* _init_l___auto____x40_Init_Data_List_Sort_Impl___hyg_1991_() { _start: { lean_object* x_1; @@ -1789,8 +1789,8 @@ l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127____closed__59 = _init_l___au lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127____closed__59); l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127_ = _init_l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127_(); lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Impl___hyg_1127_); -l___auto____x40_Init_Data_List_Sort_Impl___hyg_2003_ = _init_l___auto____x40_Init_Data_List_Sort_Impl___hyg_2003_(); -lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Impl___hyg_2003_); +l___auto____x40_Init_Data_List_Sort_Impl___hyg_1991_ = _init_l___auto____x40_Init_Data_List_Sort_Impl___hyg_1991_(); +lean_mark_persistent(l___auto____x40_Init_Data_List_Sort_Impl___hyg_1991_); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/Nat/Linear.c b/stage0/stdlib/Init/Data/Nat/Linear.c index 864ea2902f96..197603f6c0b2 100644 --- a/stage0/stdlib/Init/Data/Nat/Linear.c +++ b/stage0/stdlib/Init/Data/Nat/Linear.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.Nat.Linear -// Imports: Init.ByCases Init.Data.Prod +// Imports: Init.ByCases Init.Data.Prod Init.Data.RArray #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,12 +14,12 @@ extern "C" { #endif uint8_t l_Nat_blt(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isNum_x3f_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_monomialToExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Expr_toNormPoly(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_mul_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Expr_toPoly_go___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_hugeFuel; LEAN_EXPORT uint8_t l_Nat_Linear_Poly_isNonZero(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_insert(lean_object*, lean_object*, lean_object*); @@ -27,6 +27,7 @@ LEAN_EXPORT lean_object* l_Nat_Linear_PolyCnstr_isUnsat___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isNum_x3f_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_combineAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Expr_denote_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_RArray_getImpl___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_cancelAux_match__2_splitter(lean_object*); static lean_object* l_Nat_Linear_Poly_isNum_x3f___closed__1; LEAN_EXPORT lean_object* l_Nat_Linear_Poly_denote(lean_object*, lean_object*); @@ -34,7 +35,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isZ LEAN_EXPORT lean_object* l_Nat_Linear_instInhabitedExpr; LEAN_EXPORT lean_object* l_Nat_Linear_Poly_isNum_x3f(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_mul___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_denote_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Certificate_combineHyps_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Expr_denote___boxed(lean_object*, lean_object*); @@ -45,7 +45,6 @@ LEAN_EXPORT uint8_t l_Nat_Linear_PolyCnstr_isUnsat(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_cancel(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote_go___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Expr_toPoly_go(lean_object*, lean_object*, lean_object*); static lean_object* l_Nat_Linear_instInhabitedExpr___closed__1; LEAN_EXPORT lean_object* l_Nat_Linear_Poly_mul_go___boxed(lean_object*, lean_object*); @@ -56,7 +55,6 @@ static lean_object* l_Nat_Linear_Certificate_combine___closed__1; LEAN_EXPORT lean_object* l_Nat_Linear_fixedVar; LEAN_EXPORT lean_object* l_Nat_Linear_Poly_norm_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_toExpr_go(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_PolyCnstr_isValid___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isZero_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_cancelAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -66,13 +64,14 @@ LEAN_EXPORT lean_object* l_Nat_Linear_Poly_isNonZero___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Certificate_combineHyps_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Expr_denote_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_cancelAux_match__1_splitter(lean_object*); -LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_PolyCnstr_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Nat_Linear_Poly_isZero(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_norm(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Certificate_combine(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isZero_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_isNum_x3f_match__1_splitter(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_cancelAux_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -83,8 +82,8 @@ LEAN_EXPORT lean_object* l_Nat_elimOffset___rarg(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Expr_toPoly___boxed(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Poly_combineAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_Poly_denote_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_PolyCnstr_toExpr(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_ExprCnstr_toPoly(lean_object*); LEAN_EXPORT lean_object* l_Nat_Linear_Certificate_combineHyps(lean_object*, lean_object*); static lean_object* l_Nat_Linear_instBEqPolyCnstr___closed__1; @@ -116,51 +115,6 @@ x_1 = lean_unsigned_to_nat(100000000u); return x_1; } } -LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote_go(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -lean_dec(x_2); -x_3 = lean_unsigned_to_nat(0u); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_eq(x_2, x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_unsigned_to_nat(1u); -x_9 = lean_nat_sub(x_2, x_8); -lean_dec(x_2); -x_1 = x_5; -x_2 = x_9; -goto _start; -} -else -{ -lean_dec(x_2); -lean_inc(x_4); -return x_4; -} -} -} -} -LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote_go___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Nat_Linear_Var_denote_go(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote(lean_object* x_1, lean_object* x_2) { _start: { @@ -170,13 +124,12 @@ x_4 = lean_nat_dec_eq(x_2, x_3); if (x_4 == 0) { lean_object* x_5; -x_5 = l_Nat_Linear_Var_denote_go(x_1, x_2); +x_5 = l_Lean_RArray_getImpl___rarg(x_1, x_2); return x_5; } else { lean_object* x_6; -lean_dec(x_2); x_6 = lean_unsigned_to_nat(1u); return x_6; } @@ -187,6 +140,7 @@ LEAN_EXPORT lean_object* l_Nat_Linear_Var_denote___boxed(lean_object* x_1, lean_ { lean_object* x_3; x_3 = l_Nat_Linear_Var_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } @@ -218,15 +172,12 @@ case 0: lean_object* x_3; x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -lean_dec(x_2); return x_3; } case 1: { lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); x_5 = l_Nat_Linear_Var_denote(x_1, x_4); return x_5; } @@ -234,10 +185,7 @@ case 2: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_dec(x_2); x_8 = l_Nat_Linear_Expr_denote(x_1, x_6); x_9 = l_Nat_Linear_Expr_denote(x_1, x_7); x_10 = lean_nat_add(x_8, x_9); @@ -249,27 +197,19 @@ case 3: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); -lean_inc(x_12); -lean_dec(x_2); x_13 = l_Nat_Linear_Expr_denote(x_1, x_12); x_14 = lean_nat_mul(x_11, x_13); lean_dec(x_13); -lean_dec(x_11); return x_14; } default: { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_2, 0); -lean_inc(x_15); x_16 = lean_ctor_get(x_2, 1); -lean_inc(x_16); -lean_dec(x_2); x_17 = l_Nat_Linear_Expr_denote(x_1, x_15); x_18 = lean_nat_mul(x_17, x_16); -lean_dec(x_16); lean_dec(x_17); return x_18; } @@ -281,6 +221,7 @@ LEAN_EXPORT lean_object* l_Nat_Linear_Expr_denote___boxed(lean_object* x_1, lean { lean_object* x_3; x_3 = l_Nat_Linear_Expr_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } @@ -298,19 +239,12 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_dec(x_2); x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -lean_dec(x_4); x_8 = l_Nat_Linear_Var_denote(x_1, x_7); x_9 = lean_nat_mul(x_6, x_8); lean_dec(x_8); -lean_dec(x_6); x_10 = l_Nat_Linear_Poly_denote(x_1, x_5); x_11 = lean_nat_add(x_9, x_10); lean_dec(x_10); @@ -324,6 +258,7 @@ LEAN_EXPORT lean_object* l_Nat_Linear_Poly_denote___boxed(lean_object* x_1, lean { lean_object* x_3; x_3 = l_Nat_Linear_Poly_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } @@ -2455,7 +2390,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2520,7 +2455,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -2575,7 +2510,7 @@ return x_10; else { uint8_t x_11; -x_11 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1(x_4, x_7); +x_11 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1(x_4, x_7); if (x_11 == 0) { uint8_t x_12; @@ -2585,29 +2520,29 @@ return x_12; else { uint8_t x_13; -x_13 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1(x_5, x_8); +x_13 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1(x_5, x_8); return x_13; } } } } } -LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____spec__1(x_1, x_2); +x_3 = l_List_beq___at___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466_(x_1, x_2); +x_3 = l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2618,7 +2553,7 @@ static lean_object* _init_l_Nat_Linear_instBEqPolyCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1466____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Data_Nat_Linear_0__Nat_Linear_beqPolyCnstr____x40_Init_Data_Nat_Linear___hyg_1382____boxed), 2, 0); return x_1; } } @@ -3762,6 +3697,7 @@ return x_6; } lean_object* initialize_Init_ByCases(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Prod(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_RArray(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_Nat_Linear(uint8_t builtin, lean_object* w) { lean_object * res; @@ -3773,6 +3709,9 @@ lean_dec_ref(res); res = initialize_Init_Data_Prod(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_RArray(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Nat_Linear_fixedVar = _init_l_Nat_Linear_fixedVar(); lean_mark_persistent(l_Nat_Linear_fixedVar); l_Nat_Linear_instInhabitedExpr___closed__1 = _init_l_Nat_Linear_instInhabitedExpr___closed__1(); diff --git a/stage0/stdlib/Init/Data/Nat/SOM.c b/stage0/stdlib/Init/Data/Nat/SOM.c index 8614e8809ffb..1f423defede2 100644 --- a/stage0/stdlib/Init/Data/Nat/SOM.c +++ b/stage0/stdlib/Init/Data/Nat/SOM.c @@ -78,15 +78,12 @@ case 0: lean_object* x_3; x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); -lean_dec(x_2); return x_3; } case 1: { lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); x_5 = l_Nat_Linear_Var_denote(x_1, x_4); return x_5; } @@ -94,10 +91,7 @@ case 2: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_dec(x_2); x_8 = l_Nat_SOM_Expr_denote(x_1, x_6); x_9 = l_Nat_SOM_Expr_denote(x_1, x_7); x_10 = lean_nat_add(x_8, x_9); @@ -109,10 +103,7 @@ return x_10; { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); -lean_inc(x_12); -lean_dec(x_2); x_13 = l_Nat_SOM_Expr_denote(x_1, x_11); x_14 = l_Nat_SOM_Expr_denote(x_1, x_12); x_15 = lean_nat_mul(x_13, x_14); @@ -128,6 +119,7 @@ LEAN_EXPORT lean_object* l_Nat_SOM_Expr_denote___boxed(lean_object* x_1, lean_ob { lean_object* x_3; x_3 = l_Nat_SOM_Expr_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } @@ -145,10 +137,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_dec(x_2); x_6 = l_Nat_Linear_Var_denote(x_1, x_4); x_7 = l_Nat_SOM_Mon_denote(x_1, x_5); x_8 = lean_nat_mul(x_6, x_7); @@ -163,6 +152,7 @@ LEAN_EXPORT lean_object* l_Nat_SOM_Mon_denote___boxed(lean_object* x_1, lean_obj { lean_object* x_3; x_3 = l_Nat_SOM_Mon_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } @@ -413,19 +403,12 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_dec(x_2); x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -lean_dec(x_4); x_8 = l_Nat_SOM_Mon_denote(x_1, x_7); x_9 = lean_nat_mul(x_6, x_8); lean_dec(x_8); -lean_dec(x_6); x_10 = l_Nat_SOM_Poly_denote(x_1, x_5); x_11 = lean_nat_add(x_9, x_10); lean_dec(x_10); @@ -439,6 +422,7 @@ LEAN_EXPORT lean_object* l_Nat_SOM_Poly_denote___boxed(lean_object* x_1, lean_ob { lean_object* x_3; x_3 = l_Nat_SOM_Poly_denote(x_1, x_2); +lean_dec(x_2); lean_dec(x_1); return x_3; } diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index c324380528df..7983c749d46d 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStr1(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__7; @@ -29,13 +30,13 @@ LEAN_EXPORT lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object static lean_object* l_Lean_quoteNameMk___closed__7; LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2852____spec__1___boxed(lean_object*, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1; static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__2; static lean_object* l_Lean_mkGroupNode___closed__2; LEAN_EXPORT lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___boxed(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStringGap(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNameLit___closed__2; @@ -49,7 +50,6 @@ static lean_object* l_Lean_origin___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__9; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__8; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__4; @@ -60,9 +60,9 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___rarg(lean_obj LEAN_EXPORT lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__11; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__63; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__16; LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStr1___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__6; @@ -71,21 +71,19 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__6; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18; lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___closed__1; uint32_t lean_string_utf8_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getSepElems___rarg(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1; @@ -104,23 +102,20 @@ static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; LEAN_EXPORT lean_object* l_Lean_TSyntax_getId(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; LEAN_EXPORT lean_object* l_Lean_githash; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24; static lean_object* l_Lean_versionString___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getName(lean_object*); LEAN_EXPORT uint8_t l_String_anyAux___at_Lean_Name_isInaccessibleUserName___spec__1(uint32_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84; lean_object* lean_uint32_to_nat(uint32_t); LEAN_EXPORT lean_object* l_Lean_Syntax_instRepr; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4; LEAN_EXPORT lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__13; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52; @@ -131,9 +126,9 @@ LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape___boxed(lean_ob LEAN_EXPORT lean_object* l_Array_filterSepElemsM___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__4; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30; static lean_object* l_Lean_Name_reprPrec___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__48; LEAN_EXPORT lean_object* lean_name_append_after(lean_object*, lean_object*); @@ -142,9 +137,9 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_ static lean_object* l_Lean_TSyntax_getScientific___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__10; static lean_object* l_Lean_termEval__prio_____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__15; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18; static lean_object* l_Lean_termEval__prec_____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -156,44 +151,45 @@ LEAN_EXPORT lean_object* l_Lean_Name_toString___boxed(lean_object*, lean_object* lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17527_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__16; LEAN_EXPORT lean_object* l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__4; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__3; static lean_object* l_Lean_Option_hasQuote___rarg___closed__4; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_location; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkOptionalNode___boxed(lean_object*); static lean_object* l_Lean_evalPrec___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__7; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNumLit___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; LEAN_EXPORT lean_object* l_Lean_TSyntax_getString(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Parser_Tactic_getConfigItems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_structEq___boxed(lean_object*, lean_object*); static lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__2; extern lean_object* l_Lean_reservedMacroScope; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7; lean_object* l_String_quote(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__11; static lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); @@ -201,7 +197,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__6; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_origin; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20; lean_object* l_Lean_TSyntaxArray_rawImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_hasQuote(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_curr(lean_object*); @@ -211,13 +206,17 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3; LEAN_EXPORT lean_object* l_Lean_mkIdentFrom___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__1; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkHole___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; @@ -227,10 +226,11 @@ extern lean_object* l_Lean_Parser_Tactic_simpStar; LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__73; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__79; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom___boxed(lean_object*, lean_object*, lean_object*); @@ -240,8 +240,10 @@ static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__18; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm(lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; static lean_object* l_Lean_version_specialDesc___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__19; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Syntax_findAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -251,24 +253,24 @@ static lean_object* l_Lean_Syntax_instReprTSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(lean_object*); static lean_object* l_Lean_Syntax_SepArray_ofElems___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isLetterLike___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__4; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__6; static lean_object* l_Lean_termEval__prio_____closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__6; LEAN_EXPORT lean_object* l_Lean_Name_escapePart___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279_(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39; static lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArray___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionSepArray(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux___boxed(lean_object*, lean_object*); @@ -277,11 +279,9 @@ static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53; LEAN_EXPORT uint8_t l_Lean_isGreek(uint32_t); LEAN_EXPORT uint8_t l_Lean_isIdBeginEscape(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArith; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems___boxed(lean_object*, lean_object*); @@ -289,14 +289,12 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_instQuoteBoolMkStr1___closed__1; extern lean_object* l_Lean_Parser_Tactic_optConfig; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16; LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19; LEAN_EXPORT uint8_t l_Lean_Syntax_instBEqTSyntax___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6; lean_object* lean_version_get_patch(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -304,17 +302,13 @@ LEAN_EXPORT lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t) LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteTermMkStr1; static lean_object* l_Lean_termEval__prec_____closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__6; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfoAux(lean_object*, lean_object*); @@ -326,6 +320,7 @@ static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go(lean_object*); lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__48; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__61; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__2; @@ -340,6 +335,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_versionString___closed__6; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__6; @@ -348,21 +344,23 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__10; lean_object* l___private_Init_Data_Repr_0__reprSourceInfo____x40_Init_Data_Repr___hyg_2444_(lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailing_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__1; static lean_object* l_Lean_toolchain___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object*); static lean_object* l_Lean_termEval__prio_____closed__5; static lean_object* l_Lean_termEval__prec_____closed__9; LEAN_EXPORT uint8_t l_Lean_Syntax_isAtom(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_instReprEtaStructMode; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNatLit(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__3; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__19; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27; LEAN_EXPORT uint8_t l_Lean_version_isRelease; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm___boxed(lean_object*); @@ -377,17 +375,18 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__8; static lean_object* l_Lean_Syntax_instBEq___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__71; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__7; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4; static lean_object* l_Lean_Parser_Tactic_getConfigItems___closed__4; LEAN_EXPORT lean_object* l_String_anyAux___at_Lean_Name_escapePart___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_getConfigItems(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_discharger; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__4; @@ -395,29 +394,29 @@ static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__14; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionTSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__5; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape(uint8_t, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrio___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18; LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f(lean_object*); static lean_object* l_Lean_versionStringCore___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19; LEAN_EXPORT lean_object* l_Lean_Syntax_splitNameLit(lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f___boxed(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17; LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instReprPreresolved; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKind___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_getConfigItems___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168_(lean_object*, lean_object*); @@ -429,14 +428,14 @@ LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Le static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__12; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__3; static lean_object* l_Lean_versionStringCore___closed__3; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263_(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__4; static lean_object* l_Lean_Syntax_mkCharLit___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_mkOptionalNode___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__16; @@ -446,12 +445,15 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__63; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__55; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_noConfusion(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__45; LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStr1(uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__18; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__4; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_instReprConfig; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit(lean_object*); @@ -472,9 +474,9 @@ static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKind___rarg(lean_object*); static lean_object* l_Array_getSepElems___rarg___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19; LEAN_EXPORT uint8_t l_Lean_Name_toStringWithSep___lambda__1(lean_object*); static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__3; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25; @@ -484,12 +486,11 @@ lean_object* lean_string_utf8_next(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__14; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; LEAN_EXPORT lean_object* l_Lean_Meta_ApplyNewGoals_noConfusion___rarg___lambda__1___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_instReprTransparencyMode; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__3; static lean_object* l_Lean_version_patch___closed__1; @@ -497,6 +498,7 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1___boxed(lean_object static lean_object* l_Lean_instQuoteBoolMkStr1___closed__2; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8; lean_object* l_Char_toUpper(uint32_t); LEAN_EXPORT lean_object* l_List_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____spec__4(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; @@ -506,9 +508,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__7; LEAN_EXPORT lean_object* l_Lean_instQuoteNatNumLitKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9; static lean_object* l_Lean_termEval__prec_____closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast(lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object*); @@ -523,10 +523,12 @@ static lean_object* l_Lean_quoteNameMk___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm___boxed(lean_object*); LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray___closed__1; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14; LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexLitAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__12; @@ -539,13 +541,12 @@ static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__4; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__13; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__20; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__2; LEAN_EXPORT uint32_t l_Lean_TSyntax_getChar(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__7; @@ -554,10 +555,12 @@ LEAN_EXPORT lean_object* l_Lean_versionString; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__9; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____spec__5(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23; static lean_object* l_Lean_termEval__prio_____closed__9; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__2(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; extern lean_object* l_Lean_Parser_Tactic_simpLemma; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -566,14 +569,16 @@ LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___boxed(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__58; LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16581_(lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_versionString___closed__2; static lean_object* l_Lean_Name_reprPrec___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__65; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13; static lean_object* l_Lean_Syntax_getHead_x3f___closed__3; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot(lean_object*, lean_object*, lean_object*, lean_object*); @@ -586,10 +591,9 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_ static lean_object* l_Lean_termEval__prec_____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; lean_object* l_outOfBounds___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16; LEAN_EXPORT lean_object* l_Lean_quoteNameMk(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16; LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__10; @@ -607,11 +611,11 @@ static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69; uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Parser_Tactic_getConfigItems___spec__1(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__4; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__9; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit___boxed(lean_object*); static lean_object* l_Lean_versionStringCore___closed__4; @@ -620,13 +624,12 @@ LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__3; lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_getGithash___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28; lean_object* l_Lean_TSyntaxArray_mkImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNumLit___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17219_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__22; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__4; @@ -634,23 +637,20 @@ lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__9; static lean_object* l_Lean_termEval__prio_____closed__1; LEAN_EXPORT uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeDepTermMkIdentIdent(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__3; static lean_object* l_Lean_Option_hasQuote___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux(lean_object*); lean_object* l_Char_quote(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Lean_Name_getRoot(lean_object*); @@ -666,6 +666,7 @@ static lean_object* l_Lean_mkOptionalNode___closed__1; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__12; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___at_Lean_Syntax_setTailInfoAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__18; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; @@ -673,8 +674,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instReprConfig__1; static lean_object* l_Lean_Meta_instReprConfig__1___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__1; LEAN_EXPORT lean_object* l_Lean_instQuoteCharCharLitKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); @@ -686,13 +685,12 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_hasArgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArray(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____boxed(lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__5; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isIdEndEscape___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqPreresolved; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___boxed(lean_object*); @@ -704,24 +702,25 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_TSyntax_expandIn LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__9; lean_object* l_Lean_extractMacroScopes(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrio(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__57; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__20; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3; static lean_object* l_Lean_quoteNameMk___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_copyHeadTailInfoFrom___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkCharLit___closed__2; @@ -739,28 +738,30 @@ static lean_object* l_Lean_TSyntax_getScientific___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__80; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10; LEAN_EXPORT lean_object* l_Lean_instQuoteOfCoeHTCTTSyntaxConsSyntaxNodeKindNil___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_version_minor; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10; lean_object* lean_string_utf8_prev(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkOptionalNode(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7; static lean_object* l_Lean_toolchain___closed__8; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__8; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__1; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28; static lean_object* l_Lean_mkSepArray___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_isIdBeginEscape___boxed(lean_object*); @@ -772,8 +773,10 @@ static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61; LEAN_EXPORT lean_object* l_Array_mapSepElemsM(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__7; lean_object* lean_version_get_special_desc(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5; static lean_object* l_Lean_Option_hasQuote___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeScientificLitTerm(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__8; @@ -784,6 +787,7 @@ LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStr1(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteNameMkStr1(lean_object*); LEAN_EXPORT lean_object* lean_mk_syntax_ident(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__41; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__56; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46; @@ -800,12 +804,9 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__44; LEAN_EXPORT lean_object* l_Lean_Syntax_isToken___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo(lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold; @@ -819,8 +820,8 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Name_isInaccessibleUserName___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17511_(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8; lean_object* lean_get_githash(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__23; static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__1; @@ -830,6 +831,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_mkSepArray___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__16; LEAN_EXPORT lean_object* l_Lean_version_major; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__8; @@ -838,19 +840,22 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkScientificLit___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, uint8_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm___boxed(lean_object*); static lean_object* l_Lean_mkHole___closed__2; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__25; +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingTailPos_x3f___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__9; static lean_object* l_Lean_versionString___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2; LEAN_EXPORT lean_object* l_Array_mapSepElems(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_simpErase; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43; lean_object* l_Substring_takeWhileAux___at_Substring_trimLeft___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2852____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCharLit(uint32_t, lean_object*); @@ -864,41 +869,40 @@ lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_ lean_object* l_Substring_nextn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* lean_name_append_index_after(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13; LEAN_EXPORT uint8_t l_String_anyAux___at_Lean_Name_escapePart___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__76; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__24; static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Syntax_findAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_quoteNameMk___closed__8; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715_(uint8_t, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__1; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__7; LEAN_EXPORT lean_object* l_Lean_TSyntax_getId___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object*); static lean_object* l_Lean_Syntax_getHead_x3f___closed__2; +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; static lean_object* l_Lean_Name_reprPrec___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__3(lean_object*, lean_object*); static lean_object* l_Lean_toolchain___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__74; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15; LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Name_hasNum(lean_object*); static lean_object* l_Lean_Syntax_mkStrLit___closed__2; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16597_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___at_Lean_Syntax_setHeadInfoAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__5; static lean_object* l_Lean_quoteNameMk___closed__2; @@ -911,12 +915,14 @@ LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1(lean_object*, static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__7; static lean_object* l_Lean_toolchain___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7; LEAN_EXPORT lean_object* l_Lean_TSyntax_getDocString___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteCharCharLitKind(uint32_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f(lean_object*); LEAN_EXPORT uint32_t l_Lean_idEndEscape; @@ -928,32 +934,31 @@ LEAN_EXPORT uint8_t l_Lean_isNumericSubscript(uint32_t); LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_replacePrefix___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__13; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99; static lean_object* l_Lean_Name_reprPrec___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__16; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____spec__1(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__15; LEAN_EXPORT uint8_t l_Lean_Name_toString_maybePseudoSyntax(lean_object*); static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__14; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); uint8_t lean_internal_is_stage0(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__28; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getScientific(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20; lean_object* lean_string_length(lean_object*); @@ -970,12 +975,13 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__15; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5; LEAN_EXPORT lean_object* l_Lean_instQuoteOfCoeHTCTTSyntaxConsSyntaxNodeKindNil(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__4; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArray___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__54; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28; @@ -984,16 +990,18 @@ LEAN_EXPORT uint8_t l_Lean_isSubScriptAlnum(uint32_t); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__15; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68; LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); lean_object* l_Substring_takeRightWhileAux___at_Substring_trimRight___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_getTrailingTailPos_x3f(lean_object*, uint8_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__42; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__6; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__14; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__9; @@ -1011,18 +1019,18 @@ LEAN_EXPORT lean_object* l_Lean_instQuoteOfCoeHTCTTSyntaxConsSyntaxNodeKindNil__ static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__40; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17860_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllKind; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__12; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28; static lean_object* l_Lean_toolchain___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__38; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__4; @@ -1033,7 +1041,6 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuo static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__19; static lean_object* l_Lean_Syntax_getHead_x3f___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__60; static uint8_t l_Lean_version_isRelease___closed__1; @@ -1041,16 +1048,17 @@ static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__1; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24; static lean_object* l_Lean_versionString___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_findAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKindNil___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); uint8_t l_String_isPrefixOf(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__21; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17876_(lean_object*, lean_object*, lean_object*); lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__76; LEAN_EXPORT lean_object* l_Lean_Name_escapePart(lean_object*, uint8_t); @@ -1064,13 +1072,12 @@ LEAN_EXPORT lean_object* l_String_toName(lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__21; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; static lean_object* l_Lean_quoteNameMk___closed__1; LEAN_EXPORT lean_object* l_Lean_NameGenerator_mkChild(lean_object*); LEAN_EXPORT uint8_t l_Lean_Syntax_isNone(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32; static lean_object* l_Lean_Parser_Tactic_getConfigItems___closed__5; @@ -1078,17 +1085,15 @@ LEAN_EXPORT lean_object* l_Lean_Name_getRoot___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__14; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_eraseSuffix_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9; LEAN_EXPORT uint8_t l_Lean_isIdEndEscape(uint32_t); LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22; LEAN_EXPORT lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getName___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1104,7 +1109,6 @@ static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__7; static lean_object* l_Lean_mkCIdentFrom___closed__2; lean_object* l_Array_mkArray1___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__59; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23; lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__75; @@ -1112,7 +1116,9 @@ static lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailing_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__18; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArith; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__6; LEAN_EXPORT lean_object* l_Lean_Name_instDecidableEq___boxed(lean_object*, lean_object*); @@ -1126,16 +1132,20 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArray___boxed(lean_ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__50; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__22; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28; static lean_object* l_Lean_termEval__prec_____closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionTSepArray___boxed(lean_object*, lean_object*); static lean_object* l_Lean_termEval__prec_____closed__5; static lean_object* l_Lean_githash___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15; LEAN_EXPORT uint8_t l_Lean_Meta_Occurrences_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105_(lean_object*, lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__4; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux(lean_object*); @@ -1144,7 +1154,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__13; LEAN_EXPORT uint8_t lean_is_inaccessible_user_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_next(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38; static lean_object* l_Lean_versionStringCore___closed__7; static lean_object* l_Lean_Syntax_instCoeOutSepArrayArray___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__29; @@ -1153,16 +1162,17 @@ static lean_object* l_Lean_instQuoteBoolMkStr1___closed__5; static lean_object* l_Lean_instQuoteNameMkStr1___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__4; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKindNil___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19; lean_object* lean_array_mk(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__6; +lean_object* l_Lean_SourceInfo_getTrailing_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticErw______; LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStr1___boxed(lean_object*); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__1; @@ -1177,6 +1187,7 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Arr static lean_object* l_Lean_versionStringCore___closed__5; static lean_object* l_Lean_Meta_instReprEtaStructMode___closed__1; uint8_t l_Substring_beq(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__18; @@ -1195,21 +1206,18 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArray(lean_object*) LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831_(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc___boxed(lean_object*); static lean_object* l_Lean_mkOptionalNode___closed__3; size_t lean_array_size(lean_object*); static lean_object* l_Lean_termEval__prec_____closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__53; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__22; @@ -1222,17 +1230,13 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__17; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_getConfigItems___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__51; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCApp(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__4; LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStr1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdent(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; static lean_object* l_Lean_quoteNameMk___closed__4; LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKindNil(lean_object*, lean_object*); @@ -1243,6 +1247,7 @@ lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__21; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__4; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__8; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731_(uint8_t, lean_object*); static lean_object* l_Lean_quoteNameMk___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getChar___boxed(lean_object*); @@ -1259,7 +1264,6 @@ static lean_object* l_Lean_termEval__prec_____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__72; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_mkSepArray___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prec____1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6; lean_object* lean_nat_pred(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__10; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); @@ -1267,23 +1271,24 @@ lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Parser_Tactic_getConfigItems___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__17; lean_object* l_panic___at___private_Init_Prelude_0__Lean_assembleParts___spec__1(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__34; LEAN_EXPORT lean_object* l_Lean_Meta_Occurrences_isAll___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__10; static lean_object* l_Lean_instQuoteTermMkStr1___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instReprTSyntax___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9; static lean_object* l_Lean_mkGroupNode___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__49; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93; @@ -1294,8 +1299,10 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeArraySepArray(lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__1; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__33; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___at_Array_mapSepElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3; LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStr1(lean_object*); static lean_object* l_Lean_termEval__prio_____closed__2; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1___closed__1; @@ -1328,38 +1335,36 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; LEAN_EXPORT lean_object* l_Array_getSepElems(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__70; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17203_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termEval__prio_____closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__78; LEAN_EXPORT lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__12; LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2852_(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAutoUnfold; static lean_object* l_Lean_termEval__prio_____closed__4; static lean_object* l_Lean_Option_hasQuote___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_mkHole___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__5; LEAN_EXPORT lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14; LEAN_EXPORT lean_object* l_Lean_Name_modifyBase(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__1; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__1___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_appendConfig(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17; LEAN_EXPORT uint8_t l_Lean_isIdFirst(uint32_t); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; LEAN_EXPORT lean_object* l_Lean_version_specialDesc; LEAN_EXPORT lean_object* l_String_anyAux___at_Lean_Name_escapePart___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60; lean_object* l_String_toSubstring_x27(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__7; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__23; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__27; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__8; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__3; @@ -1370,6 +1375,7 @@ LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_version_get_major(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__43; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel___boxed(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -1380,9 +1386,11 @@ static lean_object* l_Lean_termEval__prec_____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_isIdFirst___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__16; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__39; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17; static lean_object* l_Lean_Meta_instReprConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold; @@ -1397,24 +1405,23 @@ static lean_object* l_Lean_toolchain___closed__1; lean_object* l_Char_ofNat(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getPatch___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4; static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__3; LEAN_EXPORT uint8_t l_Lean_isIdRest(uint32_t); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKind___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__11; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__10; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__3; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__5; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2168____closed__6; static lean_object* l_Lean_versionStringCore___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo___boxed(lean_object*); @@ -1423,7 +1430,6 @@ LEAN_EXPORT lean_object* l_Lean_termEval__prec__; static lean_object* l_Lean_Syntax_mkApp___closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStr4Nil___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30; static lean_object* l_Lean_versionStringCore___closed__8; static lean_object* l_Lean_Parser_Tactic_tacticErw_________closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object* x_1) { @@ -7709,6 +7715,46 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailing_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Syntax_getTailInfo(x_1); +x_3 = l_Lean_SourceInfo_getTrailing_x3f(x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailing_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Syntax_getTrailing_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingTailPos_x3f(lean_object* x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_Syntax_getTailInfo(x_1); +x_4 = l_Lean_SourceInfo_getTrailingTailPos_x3f(x_3, x_2); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingTailPos_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Lean_Syntax_getTrailingTailPos_x3f(x_1, x_3); +lean_dec(x_1); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object* x_1, uint8_t x_2, uint8_t x_3) { _start: { @@ -14780,7 +14826,7 @@ static lean_object* _init_l_List_foldr___at_Substring_toName___spec__1___closed_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_List_foldr___at_Substring_toName___spec__1___closed__1; x_2 = l_List_foldr___at_Substring_toName___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(946u); +x_3 = lean_unsigned_to_nat(974u); x_4 = lean_unsigned_to_nat(10u); x_5 = l_List_foldr___at_Substring_toName___spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -20574,7 +20620,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1() { _start: { lean_object* x_1; @@ -20582,33 +20628,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.TransparencyMode.all", 30, 30); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20616,23 +20662,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20640,7 +20686,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7() { _start: { lean_object* x_1; @@ -20648,33 +20694,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.TransparencyMode.default", 34, 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20682,23 +20728,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20706,7 +20752,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13() { _start: { lean_object* x_1; @@ -20714,33 +20760,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.TransparencyMode.reducible", 36, 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20748,23 +20794,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20772,7 +20818,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19() { _start: { lean_object* x_1; @@ -20780,33 +20826,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.TransparencyMode.instances", 36, 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20814,23 +20860,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20838,7 +20884,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -20850,14 +20896,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -20870,14 +20916,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -20890,14 +20936,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -20910,14 +20956,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24; +x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -20925,13 +20971,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -20940,7 +20986,7 @@ static lean_object* _init_l_Lean_Meta_instReprTransparencyMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____boxed), 2, 0); return x_1; } } @@ -20952,7 +20998,7 @@ x_1 = l_Lean_Meta_instReprTransparencyMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1() { _start: { lean_object* x_1; @@ -20960,33 +21006,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.EtaStructMode.all", 27, 27); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -20994,23 +21040,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21018,7 +21064,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7() { _start: { lean_object* x_1; @@ -21026,33 +21072,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.EtaStructMode.notClasses", 34, 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21060,23 +21106,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21084,7 +21130,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13() { _start: { lean_object* x_1; @@ -21092,33 +21138,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.EtaStructMode.none", 28, 28); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21126,23 +21172,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21150,7 +21196,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -21162,14 +21208,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -21182,14 +21228,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -21202,14 +21248,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -21217,13 +21263,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -21232,7 +21278,7 @@ static lean_object* _init_l_Lean_Meta_instReprEtaStructMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____boxed), 2, 0); return x_1; } } @@ -21244,7 +21290,7 @@ x_1 = l_Lean_Meta_instReprEtaStructMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1() { _start: { lean_object* x_1; @@ -21252,33 +21298,33 @@ x_1 = lean_mk_string_unchecked("zeta", 4, 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21286,7 +21332,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -21295,7 +21341,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6() { _start: { lean_object* x_1; @@ -21303,17 +21349,17 @@ x_1 = lean_mk_string_unchecked("beta", 4, 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8() { _start: { lean_object* x_1; @@ -21321,17 +21367,17 @@ x_1 = lean_mk_string_unchecked("eta", 3, 3); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10() { _start: { lean_object* x_1; @@ -21339,17 +21385,17 @@ x_1 = lean_mk_string_unchecked("etaStruct", 9, 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -21358,7 +21404,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13() { _start: { lean_object* x_1; @@ -21366,17 +21412,17 @@ x_1 = lean_mk_string_unchecked("iota", 4, 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15() { _start: { lean_object* x_1; @@ -21384,17 +21430,17 @@ x_1 = lean_mk_string_unchecked("proj", 4, 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17() { _start: { lean_object* x_1; @@ -21402,17 +21448,17 @@ x_1 = lean_mk_string_unchecked("decide", 6, 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -21421,7 +21467,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20() { _start: { lean_object* x_1; @@ -21429,17 +21475,17 @@ x_1 = lean_mk_string_unchecked("autoUnfold", 10, 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -21448,7 +21494,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23() { _start: { lean_object* x_1; @@ -21456,17 +21502,17 @@ x_1 = lean_mk_string_unchecked("failIfUnchanged", 15, 15); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -21475,7 +21521,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26() { _start: { lean_object* x_1; @@ -21483,17 +21529,17 @@ x_1 = lean_mk_string_unchecked("unfoldPartialApp", 16, 16); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -21502,7 +21548,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29() { _start: { lean_object* x_1; @@ -21510,17 +21556,17 @@ x_1 = lean_mk_string_unchecked("zetaDelta", 9, 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31() { _start: { lean_object* x_1; @@ -21528,17 +21574,17 @@ x_1 = lean_mk_string_unchecked("index", 5, 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -21547,7 +21593,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -21557,23 +21603,23 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21581,7 +21627,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -21591,23 +21637,23 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21615,7 +21661,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; @@ -21624,8 +21670,8 @@ x_4 = lean_ctor_get_uint8(x_1, 1); x_5 = lean_ctor_get_uint8(x_1, 2); x_6 = lean_ctor_get_uint8(x_1, 3); x_7 = lean_unsigned_to_nat(0u); -x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715_(x_6, x_7); -x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12; +x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731_(x_6, x_7); +x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12; x_10 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -21644,28 +21690,28 @@ x_20 = lean_ctor_get_uint8(x_1, 11); if (x_3 == 0) { lean_object* x_173; -x_173 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_173 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_21 = x_173; goto block_172; } else { lean_object* x_174; -x_174 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_174 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_21 = x_174; goto block_172; } block_172: { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_22 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5; +x_22 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5; x_23 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); x_24 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_24, 0, x_23); lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_11); -x_25 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4; +x_25 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4; x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_24); @@ -21677,7 +21723,7 @@ x_29 = lean_box(1); x_30 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7; +x_31 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7; x_32 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); @@ -21688,14 +21734,14 @@ lean_ctor_set(x_34, 1, x_33); if (x_4 == 0) { lean_object* x_170; -x_170 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_170 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_35 = x_170; goto block_169; } else { lean_object* x_171; -x_171 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_171 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_35 = x_171; goto block_169; } @@ -21717,7 +21763,7 @@ lean_ctor_set(x_39, 1, x_27); x_40 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_40, 0, x_39); lean_ctor_set(x_40, 1, x_29); -x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9; +x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9; x_42 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -21727,14 +21773,14 @@ lean_ctor_set(x_43, 1, x_33); if (x_5 == 0) { lean_object* x_167; -x_167 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_167 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_44 = x_167; goto block_166; } else { lean_object* x_168; -x_168 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_168 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_44 = x_168; goto block_166; } @@ -21757,7 +21803,7 @@ lean_ctor_set(x_49, 1, x_27); x_50 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_29); -x_51 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11; +x_51 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11; x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -21773,7 +21819,7 @@ lean_ctor_set(x_55, 1, x_27); x_56 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_56, 0, x_55); lean_ctor_set(x_56, 1, x_29); -x_57 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14; +x_57 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14; x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -21783,14 +21829,14 @@ lean_ctor_set(x_59, 1, x_33); if (x_13 == 0) { lean_object* x_164; -x_164 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_164 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_60 = x_164; goto block_163; } else { lean_object* x_165; -x_165 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_165 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_60 = x_165; goto block_163; } @@ -21812,7 +21858,7 @@ lean_ctor_set(x_64, 1, x_27); x_65 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_65, 0, x_64); lean_ctor_set(x_65, 1, x_29); -x_66 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16; +x_66 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16; x_67 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_67, 0, x_65); lean_ctor_set(x_67, 1, x_66); @@ -21822,14 +21868,14 @@ lean_ctor_set(x_68, 1, x_33); if (x_14 == 0) { lean_object* x_161; -x_161 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_161 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_69 = x_161; goto block_160; } else { lean_object* x_162; -x_162 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_162 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_69 = x_162; goto block_160; } @@ -21851,7 +21897,7 @@ lean_ctor_set(x_73, 1, x_27); x_74 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_74, 0, x_73); lean_ctor_set(x_74, 1, x_29); -x_75 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18; +x_75 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18; x_76 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_76, 0, x_74); lean_ctor_set(x_76, 1, x_75); @@ -21861,21 +21907,21 @@ lean_ctor_set(x_77, 1, x_33); if (x_15 == 0) { lean_object* x_158; -x_158 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_158 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_78 = x_158; goto block_157; } else { lean_object* x_159; -x_159 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_159 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_78 = x_159; goto block_157; } block_157: { lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_79 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19; +x_79 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19; x_80 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_78); @@ -21891,7 +21937,7 @@ lean_ctor_set(x_83, 1, x_27); x_84 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_84, 0, x_83); lean_ctor_set(x_84, 1, x_29); -x_85 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21; +x_85 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21; x_86 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_86, 0, x_84); lean_ctor_set(x_86, 1, x_85); @@ -21901,21 +21947,21 @@ lean_ctor_set(x_87, 1, x_33); if (x_16 == 0) { lean_object* x_155; -x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_88 = x_155; goto block_154; } else { lean_object* x_156; -x_156 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_156 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_88 = x_156; goto block_154; } block_154: { lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_89 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22; +x_89 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22; x_90 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_90, 0, x_89); lean_ctor_set(x_90, 1, x_88); @@ -21931,7 +21977,7 @@ lean_ctor_set(x_93, 1, x_27); x_94 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_94, 0, x_93); lean_ctor_set(x_94, 1, x_29); -x_95 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24; +x_95 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24; x_96 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_96, 0, x_94); lean_ctor_set(x_96, 1, x_95); @@ -21941,21 +21987,21 @@ lean_ctor_set(x_97, 1, x_33); if (x_17 == 0) { lean_object* x_152; -x_152 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_152 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_98 = x_152; goto block_151; } else { lean_object* x_153; -x_153 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_153 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_98 = x_153; goto block_151; } block_151: { lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_99 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25; +x_99 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25; x_100 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_100, 0, x_99); lean_ctor_set(x_100, 1, x_98); @@ -21971,7 +22017,7 @@ lean_ctor_set(x_103, 1, x_27); x_104 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_104, 0, x_103); lean_ctor_set(x_104, 1, x_29); -x_105 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27; +x_105 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27; x_106 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_106, 0, x_104); lean_ctor_set(x_106, 1, x_105); @@ -21981,21 +22027,21 @@ lean_ctor_set(x_107, 1, x_33); if (x_18 == 0) { lean_object* x_149; -x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_108 = x_149; goto block_148; } else { lean_object* x_150; -x_150 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_150 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_108 = x_150; goto block_148; } block_148: { lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_109 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28; +x_109 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28; x_110 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_108); @@ -22011,7 +22057,7 @@ lean_ctor_set(x_113, 1, x_27); x_114 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_114, 0, x_113); lean_ctor_set(x_114, 1, x_29); -x_115 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30; +x_115 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30; x_116 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_116, 0, x_114); lean_ctor_set(x_116, 1, x_115); @@ -22021,14 +22067,14 @@ lean_ctor_set(x_117, 1, x_33); if (x_19 == 0) { lean_object* x_146; -x_146 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_146 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_118 = x_146; goto block_145; } else { lean_object* x_147; -x_147 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_147 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_118 = x_147; goto block_145; } @@ -22050,7 +22096,7 @@ lean_ctor_set(x_122, 1, x_27); x_123 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_123, 0, x_122); lean_ctor_set(x_123, 1, x_29); -x_124 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32; +x_124 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32; x_125 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_125, 0, x_123); lean_ctor_set(x_125, 1, x_124); @@ -22060,7 +22106,7 @@ lean_ctor_set(x_126, 1, x_33); if (x_20 == 0) { lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_127 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36; +x_127 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36; x_128 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_128, 0, x_126); lean_ctor_set(x_128, 1, x_127); @@ -22084,7 +22130,7 @@ return x_135; else { lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_136 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39; +x_136 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39; x_137 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_137, 0, x_126); lean_ctor_set(x_137, 1, x_136); @@ -22117,11 +22163,11 @@ return x_144; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -22131,7 +22177,7 @@ static lean_object* _init_l_Lean_Meta_instReprConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____boxed), 2, 0); return x_1; } } @@ -22143,7 +22189,7 @@ x_1 = l_Lean_Meta_instReprConfig___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1() { _start: { lean_object* x_1; @@ -22151,33 +22197,33 @@ x_1 = lean_mk_string_unchecked("maxSteps", 8, 8); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2408____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22185,7 +22231,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -22194,7 +22240,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6() { _start: { lean_object* x_1; @@ -22202,17 +22248,17 @@ x_1 = lean_mk_string_unchecked("maxDischargeDepth", 17, 17); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -22221,7 +22267,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9() { _start: { lean_object* x_1; @@ -22229,17 +22275,17 @@ x_1 = lean_mk_string_unchecked("contextual", 10, 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11() { _start: { lean_object* x_1; @@ -22247,17 +22293,17 @@ x_1 = lean_mk_string_unchecked("memoize", 7, 7); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -22266,7 +22312,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14() { _start: { lean_object* x_1; @@ -22274,17 +22320,17 @@ x_1 = lean_mk_string_unchecked("singlePass", 10, 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16() { _start: { lean_object* x_1; @@ -22292,17 +22338,17 @@ x_1 = lean_mk_string_unchecked("arith", 5, 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18() { _start: { lean_object* x_1; @@ -22310,17 +22356,17 @@ x_1 = lean_mk_string_unchecked("dsimp", 5, 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20() { _start: { lean_object* x_1; @@ -22328,17 +22374,17 @@ x_1 = lean_mk_string_unchecked("ground", 6, 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22() { _start: { lean_object* x_1; @@ -22346,17 +22392,17 @@ x_1 = lean_mk_string_unchecked("implicitDefEqProofs", 19, 19); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -22365,23 +22411,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -22389,23 +22435,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -22413,7 +22459,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; lean_object* x_56; @@ -22422,7 +22468,7 @@ lean_inc(x_3); x_4 = l___private_Init_Data_Repr_0__Nat_reprFast(x_3); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5; +x_6 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -22430,7 +22476,7 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4; +x_10 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -22442,7 +22488,7 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7; +x_16 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -22455,7 +22501,7 @@ lean_inc(x_20); x_21 = l___private_Init_Data_Repr_0__Nat_reprFast(x_20); x_22 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_22, 0, x_21); -x_23 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8; x_24 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); @@ -22471,7 +22517,7 @@ lean_ctor_set(x_27, 1, x_12); x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_14); -x_29 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10; +x_29 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10; x_30 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); @@ -22486,8 +22532,8 @@ x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 4); x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 5); x_38 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 6); x_39 = lean_unsigned_to_nat(0u); -x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715_(x_38, x_39); -x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12; +x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731_(x_38, x_39); +x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12; x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_40); @@ -22510,21 +22556,21 @@ lean_dec(x_1); if (x_32 == 0) { lean_object* x_290; -x_290 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_290 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_56 = x_290; goto block_289; } else { lean_object* x_291; -x_291 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_291 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_56 = x_291; goto block_289; } block_289: { lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_57 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22; +x_57 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22; x_58 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_56); @@ -22540,7 +22586,7 @@ lean_ctor_set(x_61, 1, x_12); x_62 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_62, 1, x_14); -x_63 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12; +x_63 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12; x_64 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_64, 0, x_62); lean_ctor_set(x_64, 1, x_63); @@ -22550,21 +22596,21 @@ lean_ctor_set(x_65, 1, x_18); if (x_33 == 0) { lean_object* x_287; -x_287 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_287 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_66 = x_287; goto block_286; } else { lean_object* x_288; -x_288 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_288 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_66 = x_288; goto block_286; } block_286: { lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_67 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13; +x_67 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13; x_68 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_68, 0, x_67); lean_ctor_set(x_68, 1, x_66); @@ -22580,7 +22626,7 @@ lean_ctor_set(x_71, 1, x_12); x_72 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_72, 0, x_71); lean_ctor_set(x_72, 1, x_14); -x_73 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15; +x_73 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15; x_74 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_74, 0, x_72); lean_ctor_set(x_74, 1, x_73); @@ -22590,14 +22636,14 @@ lean_ctor_set(x_75, 1, x_18); if (x_34 == 0) { lean_object* x_284; -x_284 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_284 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_76 = x_284; goto block_283; } else { lean_object* x_285; -x_285 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_285 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_76 = x_285; goto block_283; } @@ -22619,7 +22665,7 @@ lean_ctor_set(x_80, 1, x_12); x_81 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_14); -x_82 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2; +x_82 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2; x_83 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_83, 0, x_81); lean_ctor_set(x_83, 1, x_82); @@ -22629,21 +22675,21 @@ lean_ctor_set(x_84, 1, x_18); if (x_35 == 0) { lean_object* x_281; -x_281 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_281 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_85 = x_281; goto block_280; } else { lean_object* x_282; -x_282 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_282 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_85 = x_282; goto block_280; } block_280: { lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_86 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5; +x_86 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5; x_87 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_87, 0, x_86); lean_ctor_set(x_87, 1, x_85); @@ -22659,7 +22705,7 @@ lean_ctor_set(x_90, 1, x_12); x_91 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_91, 0, x_90); lean_ctor_set(x_91, 1, x_14); -x_92 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7; +x_92 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7; x_93 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_93, 0, x_91); lean_ctor_set(x_93, 1, x_92); @@ -22669,14 +22715,14 @@ lean_ctor_set(x_94, 1, x_18); if (x_36 == 0) { lean_object* x_278; -x_278 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_278 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_95 = x_278; goto block_277; } else { lean_object* x_279; -x_279 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_279 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_95 = x_279; goto block_277; } @@ -22698,7 +22744,7 @@ lean_ctor_set(x_99, 1, x_12); x_100 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_100, 0, x_99); lean_ctor_set(x_100, 1, x_14); -x_101 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9; +x_101 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9; x_102 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_102, 0, x_100); lean_ctor_set(x_102, 1, x_101); @@ -22708,14 +22754,14 @@ lean_ctor_set(x_103, 1, x_18); if (x_37 == 0) { lean_object* x_275; -x_275 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_275 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_104 = x_275; goto block_274; } else { lean_object* x_276; -x_276 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_276 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_104 = x_276; goto block_274; } @@ -22738,7 +22784,7 @@ lean_ctor_set(x_109, 1, x_12); x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_14); -x_111 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11; +x_111 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11; x_112 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_112, 0, x_110); lean_ctor_set(x_112, 1, x_111); @@ -22754,7 +22800,7 @@ lean_ctor_set(x_115, 1, x_12); x_116 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_116, 0, x_115); lean_ctor_set(x_116, 1, x_14); -x_117 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14; +x_117 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14; x_118 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_118, 0, x_116); lean_ctor_set(x_118, 1, x_117); @@ -22764,14 +22810,14 @@ lean_ctor_set(x_119, 1, x_18); if (x_44 == 0) { lean_object* x_272; -x_272 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_272 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_120 = x_272; goto block_271; } else { lean_object* x_273; -x_273 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_273 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_120 = x_273; goto block_271; } @@ -22793,7 +22839,7 @@ lean_ctor_set(x_124, 1, x_12); x_125 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_125, 0, x_124); lean_ctor_set(x_125, 1, x_14); -x_126 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16; +x_126 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16; x_127 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_127, 0, x_125); lean_ctor_set(x_127, 1, x_126); @@ -22803,14 +22849,14 @@ lean_ctor_set(x_128, 1, x_18); if (x_45 == 0) { lean_object* x_269; -x_269 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_269 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_129 = x_269; goto block_268; } else { lean_object* x_270; -x_270 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_270 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_129 = x_270; goto block_268; } @@ -22832,7 +22878,7 @@ lean_ctor_set(x_133, 1, x_12); x_134 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_134, 0, x_133); lean_ctor_set(x_134, 1, x_14); -x_135 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18; +x_135 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18; x_136 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_136, 0, x_134); lean_ctor_set(x_136, 1, x_135); @@ -22842,21 +22888,21 @@ lean_ctor_set(x_137, 1, x_18); if (x_46 == 0) { lean_object* x_266; -x_266 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_266 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_138 = x_266; goto block_265; } else { lean_object* x_267; -x_267 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_267 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_138 = x_267; goto block_265; } block_265: { lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_139 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19; +x_139 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19; x_140 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_140, 0, x_139); lean_ctor_set(x_140, 1, x_138); @@ -22872,7 +22918,7 @@ lean_ctor_set(x_143, 1, x_12); x_144 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_144, 0, x_143); lean_ctor_set(x_144, 1, x_14); -x_145 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17; +x_145 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17; x_146 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_146, 0, x_144); lean_ctor_set(x_146, 1, x_145); @@ -22882,21 +22928,21 @@ lean_ctor_set(x_147, 1, x_18); if (x_47 == 0) { lean_object* x_263; -x_263 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_263 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_148 = x_263; goto block_262; } else { lean_object* x_264; -x_264 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_264 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_148 = x_264; goto block_262; } block_262: { lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33; +x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33; x_150 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_150, 0, x_149); lean_ctor_set(x_150, 1, x_148); @@ -22912,7 +22958,7 @@ lean_ctor_set(x_153, 1, x_12); x_154 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_154, 0, x_153); lean_ctor_set(x_154, 1, x_14); -x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21; +x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21; x_156 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_156, 0, x_154); lean_ctor_set(x_156, 1, x_155); @@ -22922,14 +22968,14 @@ lean_ctor_set(x_157, 1, x_18); if (x_48 == 0) { lean_object* x_260; -x_260 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_260 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_158 = x_260; goto block_259; } else { lean_object* x_261; -x_261 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_261 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_158 = x_261; goto block_259; } @@ -22951,7 +22997,7 @@ lean_ctor_set(x_162, 1, x_12); x_163 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_163, 0, x_162); lean_ctor_set(x_163, 1, x_14); -x_164 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19; +x_164 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19; x_165 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_165, 0, x_163); lean_ctor_set(x_165, 1, x_164); @@ -22961,14 +23007,14 @@ lean_ctor_set(x_166, 1, x_18); if (x_49 == 0) { lean_object* x_257; -x_257 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_257 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_167 = x_257; goto block_256; } else { lean_object* x_258; -x_258 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_258 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_167 = x_258; goto block_256; } @@ -22990,7 +23036,7 @@ lean_ctor_set(x_171, 1, x_12); x_172 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_172, 0, x_171); lean_ctor_set(x_172, 1, x_14); -x_173 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24; +x_173 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24; x_174 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_174, 0, x_172); lean_ctor_set(x_174, 1, x_173); @@ -23000,21 +23046,21 @@ lean_ctor_set(x_175, 1, x_18); if (x_50 == 0) { lean_object* x_254; -x_254 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_254 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_176 = x_254; goto block_253; } else { lean_object* x_255; -x_255 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_255 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_176 = x_255; goto block_253; } block_253: { lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_177 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25; +x_177 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25; x_178 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_178, 0, x_177); lean_ctor_set(x_178, 1, x_176); @@ -23030,7 +23076,7 @@ lean_ctor_set(x_181, 1, x_12); x_182 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_182, 0, x_181); lean_ctor_set(x_182, 1, x_14); -x_183 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21; +x_183 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21; x_184 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_184, 0, x_182); lean_ctor_set(x_184, 1, x_183); @@ -23040,14 +23086,14 @@ lean_ctor_set(x_185, 1, x_18); if (x_51 == 0) { lean_object* x_251; -x_251 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_251 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_186 = x_251; goto block_250; } else { lean_object* x_252; -x_252 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_252 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_186 = x_252; goto block_250; } @@ -23069,7 +23115,7 @@ lean_ctor_set(x_190, 1, x_12); x_191 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_191, 0, x_190); lean_ctor_set(x_191, 1, x_14); -x_192 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27; +x_192 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27; x_193 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_193, 0, x_191); lean_ctor_set(x_193, 1, x_192); @@ -23079,21 +23125,21 @@ lean_ctor_set(x_194, 1, x_18); if (x_52 == 0) { lean_object* x_248; -x_248 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_248 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_195 = x_248; goto block_247; } else { lean_object* x_249; -x_249 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_249 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_195 = x_249; goto block_247; } block_247: { lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_196 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28; +x_196 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28; x_197 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_197, 0, x_196); lean_ctor_set(x_197, 1, x_195); @@ -23109,7 +23155,7 @@ lean_ctor_set(x_200, 1, x_12); x_201 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_201, 0, x_200); lean_ctor_set(x_201, 1, x_14); -x_202 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30; +x_202 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30; x_203 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_203, 0, x_201); lean_ctor_set(x_203, 1, x_202); @@ -23119,14 +23165,14 @@ lean_ctor_set(x_204, 1, x_18); if (x_53 == 0) { lean_object* x_245; -x_245 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_245 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_205 = x_245; goto block_244; } else { lean_object* x_246; -x_246 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_246 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_205 = x_246; goto block_244; } @@ -23148,7 +23194,7 @@ lean_ctor_set(x_209, 1, x_12); x_210 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_210, 0, x_209); lean_ctor_set(x_210, 1, x_14); -x_211 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32; +x_211 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32; x_212 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_212, 0, x_210); lean_ctor_set(x_212, 1, x_211); @@ -23158,14 +23204,14 @@ lean_ctor_set(x_213, 1, x_18); if (x_54 == 0) { lean_object* x_242; -x_242 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34; +x_242 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34; x_214 = x_242; goto block_241; } else { lean_object* x_243; -x_243 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37; +x_243 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37; x_214 = x_243; goto block_241; } @@ -23187,7 +23233,7 @@ lean_ctor_set(x_218, 1, x_12); x_219 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_219, 0, x_218); lean_ctor_set(x_219, 1, x_14); -x_220 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23; +x_220 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23; x_221 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_221, 0, x_219); lean_ctor_set(x_221, 1, x_220); @@ -23197,7 +23243,7 @@ lean_ctor_set(x_222, 1, x_18); if (x_55 == 0) { lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_223 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26; +x_223 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26; x_224 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_224, 0, x_222); lean_ctor_set(x_224, 1, x_223); @@ -23221,7 +23267,7 @@ return x_231; else { lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; -x_232 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28; +x_232 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28; x_233 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_233, 0, x_222); lean_ctor_set(x_233, 1, x_232); @@ -23261,11 +23307,11 @@ return x_240; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -23274,7 +23320,7 @@ static lean_object* _init_l_Lean_Meta_instReprConfig__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____boxed), 2, 0); return x_1; } } @@ -24492,7 +24538,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -27503,7 +27549,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_2 = l_String_toSubstring_x27(x_1); return x_2; } @@ -27513,7 +27559,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -27525,7 +27571,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; x_2 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Lean_Parser_Tactic_getConfigItems___closed__1; -x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -29885,26 +29931,26 @@ x_1 = l_Lean_Parser_Tactic_simpAutoUnfold___closed__25; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; @@ -29923,12 +29969,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -30118,81 +30164,81 @@ x_1 = l_Lean_Parser_Tactic_simpArith___closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; x_2 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Lean_Parser_Tactic_getConfigItems___closed__1; -x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17; +x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; @@ -30211,12 +30257,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -30255,10 +30301,10 @@ x_26 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_15, x_17, x_22, x_24); x_27 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__5; lean_inc(x_6); x_28 = l_Lean_Syntax_node1(x_6, x_27, x_26); -x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4; +x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4; x_30 = l_Lean_addMacroScope(x_8, x_29, x_7); -x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3; -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7; +x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7; lean_inc(x_6); x_33 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_33, 0, x_6); @@ -30426,7 +30472,7 @@ x_1 = l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16581_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16597_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; @@ -30445,12 +30491,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -30489,11 +30535,11 @@ x_26 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_15, x_17, x_22, x_24); x_27 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__5; lean_inc(x_6); x_28 = l_Lean_Syntax_node1(x_6, x_27, x_26); -x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; +x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; lean_inc(x_7); lean_inc(x_8); x_30 = l_Lean_addMacroScope(x_8, x_29, x_7); -x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; +x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; lean_inc(x_6); x_32 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_32, 0, x_6); @@ -30508,10 +30554,10 @@ lean_inc(x_6); x_33 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_32, x_17, x_22, x_24); lean_inc(x_6); x_34 = l_Lean_Syntax_node1(x_6, x_27, x_33); -x_35 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4; +x_35 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4; x_36 = l_Lean_addMacroScope(x_8, x_35, x_7); -x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3; -x_38 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7; +x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3; +x_38 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7; lean_inc(x_6); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_6); @@ -30721,7 +30767,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1() { _start: { lean_object* x_1; @@ -30729,7 +30775,7 @@ x_1 = lean_mk_string_unchecked("simp_all", 8, 8); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; @@ -30748,12 +30794,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -30795,7 +30841,7 @@ x_33 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac x_34 = l_Lean_Syntax_setKind(x_1, x_33); x_35 = lean_unsigned_to_nat(0u); x_36 = l_Lean_Syntax_getArg(x_34, x_35); -x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1; +x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1; x_38 = 1; x_39 = l_Lean_mkAtomFrom(x_36, x_37, x_38); lean_dec(x_36); @@ -30929,7 +30975,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArith___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17203_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17219_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; @@ -30948,12 +30994,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -30992,10 +31038,10 @@ x_26 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_15, x_17, x_22, x_24); x_27 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__5; lean_inc(x_6); x_28 = l_Lean_Syntax_node1(x_6, x_27, x_26); -x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4; +x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4; x_30 = l_Lean_addMacroScope(x_8, x_29, x_7); -x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3; -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7; +x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7; lean_inc(x_6); x_33 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_33, 0, x_6); @@ -31015,7 +31061,7 @@ x_40 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac x_41 = l_Lean_Syntax_setKind(x_1, x_40); x_42 = lean_unsigned_to_nat(0u); x_43 = l_Lean_Syntax_getArg(x_41, x_42); -x_44 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1; +x_44 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1; x_45 = 1; x_46 = l_Lean_mkAtomFrom(x_43, x_44, x_45); lean_dec(x_43); @@ -31149,7 +31195,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17511_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17527_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; @@ -31168,12 +31214,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -31212,11 +31258,11 @@ x_26 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_15, x_17, x_22, x_24); x_27 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw________1___closed__5; lean_inc(x_6); x_28 = l_Lean_Syntax_node1(x_6, x_27, x_26); -x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; +x_29 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; lean_inc(x_7); lean_inc(x_8); x_30 = l_Lean_addMacroScope(x_8, x_29, x_7); -x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; +x_31 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; lean_inc(x_6); x_32 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_32, 0, x_6); @@ -31231,10 +31277,10 @@ lean_inc(x_6); x_33 = l_Lean_Syntax_node5(x_6, x_25, x_10, x_32, x_17, x_22, x_24); lean_inc(x_6); x_34 = l_Lean_Syntax_node1(x_6, x_27, x_33); -x_35 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4; +x_35 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4; x_36 = l_Lean_addMacroScope(x_8, x_35, x_7); -x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3; -x_38 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7; +x_37 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3; +x_38 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7; lean_inc(x_6); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_6); @@ -31254,7 +31300,7 @@ x_46 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac x_47 = l_Lean_Syntax_setKind(x_1, x_46); x_48 = lean_unsigned_to_nat(0u); x_49 = l_Lean_Syntax_getArg(x_47, x_48); -x_50 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1; +x_50 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1; x_51 = 1; x_52 = l_Lean_mkAtomFrom(x_49, x_50, x_51); lean_dec(x_49); @@ -31402,7 +31448,7 @@ x_1 = l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17860_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17876_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; @@ -31421,12 +31467,12 @@ lean_inc(x_6); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2; +x_11 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2; lean_inc(x_7); lean_inc(x_8); x_12 = l_Lean_addMacroScope(x_8, x_11, x_7); x_13 = lean_box(0); -x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1; +x_14 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1; lean_inc(x_6); x_15 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_15, 0, x_6); @@ -31468,7 +31514,7 @@ x_33 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tac x_34 = l_Lean_Syntax_setKind(x_1, x_33); x_35 = lean_unsigned_to_nat(0u); x_36 = l_Lean_Syntax_getArg(x_34, x_35); -x_37 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18; +x_37 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18; x_38 = 1; x_39 = l_Lean_mkAtomFrom(x_36, x_37, x_38); lean_dec(x_36); @@ -32055,236 +32101,236 @@ l_Lean_TSyntax_expandInterpolatedStr___closed__4 = _init_l_Lean_TSyntax_expandIn lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__4); l_Lean_TSyntax_expandInterpolatedStr___closed__5 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__5(); lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11567____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11583____closed__24); l_Lean_Meta_instReprTransparencyMode___closed__1 = _init_l_Lean_Meta_instReprTransparencyMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprTransparencyMode___closed__1); l_Lean_Meta_instReprTransparencyMode = _init_l_Lean_Meta_instReprTransparencyMode(); lean_mark_persistent(l_Lean_Meta_instReprTransparencyMode); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11715____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11731____closed__18); l_Lean_Meta_instReprEtaStructMode___closed__1 = _init_l_Lean_Meta_instReprEtaStructMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprEtaStructMode___closed__1); l_Lean_Meta_instReprEtaStructMode = _init_l_Lean_Meta_instReprEtaStructMode(); lean_mark_persistent(l_Lean_Meta_instReprEtaStructMode); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__24); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__25); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__26); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__27); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__28); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__29); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__30); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__31); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__32); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__33); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__34); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__35); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__36); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__37); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__38); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11831____closed__39); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__25); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__26); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__27); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__28); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__29); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__30); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__31); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__32); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__33); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__34); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__35); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__36); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__37); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__38); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11847____closed__39); l_Lean_Meta_instReprConfig___closed__1 = _init_l_Lean_Meta_instReprConfig___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprConfig___closed__1); l_Lean_Meta_instReprConfig = _init_l_Lean_Meta_instReprConfig(); lean_mark_persistent(l_Lean_Meta_instReprConfig); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__24); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__25); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__26); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__27); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12089____closed__28); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__25); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__26); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__27); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_12105____closed__28); l_Lean_Meta_instReprConfig__1___closed__1 = _init_l_Lean_Meta_instReprConfig__1___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprConfig__1___closed__1); l_Lean_Meta_instReprConfig__1 = _init_l_Lean_Meta_instReprConfig__1(); @@ -32931,10 +32977,10 @@ l_Lean_Parser_Tactic_simpAutoUnfold___closed__25 = _init_l_Lean_Parser_Tactic_si lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold___closed__25); l_Lean_Parser_Tactic_simpAutoUnfold = _init_l_Lean_Parser_Tactic_simpAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15990____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16006____closed__2); l_Lean_Parser_Tactic_simpArith___closed__1 = _init_l_Lean_Parser_Tactic_simpArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__1); l_Lean_Parser_Tactic_simpArith___closed__2 = _init_l_Lean_Parser_Tactic_simpArith___closed__2(); @@ -32957,20 +33003,20 @@ l_Lean_Parser_Tactic_simpArith___closed__10 = _init_l_Lean_Parser_Tactic_simpAri lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__10); l_Lean_Parser_Tactic_simpArith = _init_l_Lean_Parser_Tactic_simpArith(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16263____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16279____closed__7); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2(); @@ -33021,8 +33067,8 @@ l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13 = _init_l_Lean_Parser_Tactic lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13); l_Lean_Parser_Tactic_simpAllAutoUnfold = _init_l_Lean_Parser_Tactic_simpAllAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16940____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16956____closed__1); l_Lean_Parser_Tactic_simpAllArith___closed__1 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllArith___closed__1); l_Lean_Parser_Tactic_simpAllArith___closed__2 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__2(); diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index de6403f69857..70fc01b2ca05 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -296,6 +296,7 @@ LEAN_EXPORT lean_object* l_instLEFin___boxed(lean_object*); static lean_object* l_Array_mkArray4___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Macro_throwError(lean_object*); LEAN_EXPORT lean_object* l_instDecidableEqUInt32___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailingTailPos_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_namedPattern(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_get___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Macro_instMonadQuotationMacroM___lambda__1(lean_object*, lean_object*); @@ -647,6 +648,7 @@ LEAN_EXPORT lean_object* l_Array_mkArray8___rarg(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Eq_ndrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_interpolatedStrKind; LEAN_EXPORT lean_object* l_BitVec_ofNatLt___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailingTailPos_x3f(lean_object*, uint8_t); static lean_object* l_instInhabitedSubstring___closed__2; static lean_object* l_instHashableString___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntaxArray_mkImpl___boxed(lean_object*); @@ -757,6 +759,7 @@ LEAN_EXPORT lean_object* l_instPowOfHomogeneousPow___rarg(lean_object*, lean_obj LEAN_EXPORT lean_object* l_EStateM_seqRight(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EStateM_get___rarg(lean_object*); LEAN_EXPORT lean_object* l_instInhabitedReaderT(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailing_x3f(lean_object*); static lean_object* l_EStateM_instMonadStateOf___closed__3; LEAN_EXPORT lean_object* l_EStateM_adaptExcept(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instLTPos; @@ -772,6 +775,7 @@ LEAN_EXPORT lean_object* l_EStateM_map___rarg(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_node8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instInhabitedForall(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailing_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_UInt16_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_instMonadLiftT___rarg___boxed(lean_object*); @@ -7864,6 +7868,87 @@ lean_dec(x_1); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailing_x3f(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 2); +lean_inc(x_2); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailing_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_SourceInfo_getTrailing_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailingTailPos_x3f(lean_object* x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_SourceInfo_getTrailing_x3f(x_1); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = l_Lean_SourceInfo_getTailPos_x3f(x_1, x_2); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_6, 2); +lean_inc(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = lean_ctor_get(x_8, 2); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_SourceInfo_getTrailingTailPos_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Lean_SourceInfo_getTrailingTailPos_x3f(x_1, x_3); +lean_dec(x_1); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_node1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { diff --git a/stage0/stdlib/Lake/DSL/DeclUtil.c b/stage0/stdlib/Lake/DSL/DeclUtil.c index 583bf5acefb3..6c8d49ca13d2 100644 --- a/stage0/stdlib/Lake/DSL/DeclUtil.c +++ b/stage0/stdlib/Lake/DSL/DeclUtil.c @@ -41,6 +41,7 @@ static lean_object* l_Lake_DSL_identOrStr___closed__9; static lean_object* l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__3; LEAN_EXPORT lean_object* l_Lake_DSL_elabConfigDecl___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_DSL_expandAttrs___closed__6; +static lean_object* l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; LEAN_EXPORT lean_object* l_Lake_DSL_elabConfigDecl___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lake_DSL_identOrStr___closed__1; @@ -306,6 +307,7 @@ static lean_object* l_Lake_DSL_bracketedSimpleBinder___closed__6; static lean_object* l_Lake_DSL_structVal___closed__3; static lean_object* l_Lake_DSL_structDeclSig___closed__5; LEAN_EXPORT lean_object* l_Lake_DSL_elabConfigDecl___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lake_DSL_elabConfigDecl___lambda__2___closed__25; LEAN_EXPORT lean_object* l_Lake_DSL_elabConfigDecl___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_DSL_bracketedSimpleBinder___closed__5; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -2583,279 +2585,281 @@ return x_28; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; x_29 = lean_unsigned_to_nat(0u); x_30 = l_Lean_Syntax_getArg(x_13, x_29); x_31 = lean_unsigned_to_nat(2u); x_32 = l_Lean_Syntax_getArg(x_13, x_31); x_33 = l_Lean_Syntax_getId(x_30); -x_34 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; -lean_inc(x_1); lean_inc(x_33); -x_35 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_35, 0, x_13); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_34); -lean_ctor_set(x_35, 3, x_1); -x_36 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_35, x_8, x_9, x_10); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_38 = lean_ctor_get(x_36, 1); -x_39 = lean_ctor_get(x_36, 0); -lean_dec(x_39); -x_40 = lean_st_ref_get(x_9, x_38); -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; +lean_inc(x_1); +x_36 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_36, 0, x_13); +lean_ctor_set(x_36, 1, x_34); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_36, 3, x_1); +x_37 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_36, x_8, x_9, x_10); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_ctor_get(x_40, 0); -x_43 = lean_ctor_get(x_40, 1); -x_44 = lean_ctor_get(x_42, 0); -lean_inc(x_44); -lean_dec(x_42); +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_39 = lean_ctor_get(x_37, 1); +x_40 = lean_ctor_get(x_37, 0); +lean_dec(x_40); +x_41 = lean_st_ref_get(x_9, x_39); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_41, 1); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); lean_inc(x_1); -x_45 = l_Lean_findField_x3f(x_44, x_1, x_33); -lean_dec(x_44); -if (lean_obj_tag(x_45) == 0) +x_46 = l_Lean_findField_x3f(x_45, x_1, x_33); +lean_dec(x_45); +if (lean_obj_tag(x_46) == 0) { -uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_dec(x_32); -x_46 = 0; +x_47 = 0; lean_inc(x_1); -x_47 = l_Lean_MessageData_ofConstName(x_1, x_46); -x_48 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -lean_ctor_set_tag(x_40, 7); -lean_ctor_set(x_40, 1, x_47); -lean_ctor_set(x_40, 0, x_48); -x_49 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_49); -lean_ctor_set(x_36, 0, x_40); -x_50 = l_Lean_MessageData_ofName(x_33); -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_36); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -x_54 = 1; +x_48 = l_Lean_MessageData_ofConstName(x_1, x_47); +x_49 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +lean_ctor_set_tag(x_41, 7); +lean_ctor_set(x_41, 1, x_48); +lean_ctor_set(x_41, 0, x_49); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_37, 7); +lean_ctor_set(x_37, 1, x_50); +lean_ctor_set(x_37, 0, x_41); +x_51 = l_Lean_MessageData_ofName(x_33); +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_37); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +x_55 = 1; lean_inc(x_8); -x_55 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_53, x_54, x_8, x_9, x_43); +x_56 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_54, x_55, x_8, x_9, x_44); lean_dec(x_30); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_7); -x_14 = x_57; -x_15 = x_56; +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_7); +x_14 = x_58; +x_15 = x_57; goto block_20; } else { -uint8_t x_58; -lean_free_object(x_36); -x_58 = !lean_is_exclusive(x_45); -if (x_58 == 0) +uint8_t x_59; +lean_free_object(x_37); +x_59 = !lean_is_exclusive(x_46); +if (x_59 == 0) { -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_45, 0); -lean_dec(x_59); -lean_ctor_set(x_40, 1, x_32); -lean_ctor_set(x_40, 0, x_30); -x_60 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_40); -lean_ctor_set(x_45, 0, x_60); -x_14 = x_45; -x_15 = x_43; +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_46, 0); +lean_dec(x_60); +lean_ctor_set(x_41, 1, x_32); +lean_ctor_set(x_41, 0, x_30); +x_61 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_41); +lean_ctor_set(x_46, 0, x_61); +x_14 = x_46; +x_15 = x_44; goto block_20; } else { -lean_object* x_61; lean_object* x_62; -lean_dec(x_45); -lean_ctor_set(x_40, 1, x_32); -lean_ctor_set(x_40, 0, x_30); -x_61 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_40); -x_62 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_62, 0, x_61); -x_14 = x_62; -x_15 = x_43; +lean_object* x_62; lean_object* x_63; +lean_dec(x_46); +lean_ctor_set(x_41, 1, x_32); +lean_ctor_set(x_41, 0, x_30); +x_62 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_41); +x_63 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_63, 0, x_62); +x_14 = x_63; +x_15 = x_44; goto block_20; } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_40, 0); -x_64 = lean_ctor_get(x_40, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_40); -x_65 = lean_ctor_get(x_63, 0); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_41, 0); +x_65 = lean_ctor_get(x_41, 1); lean_inc(x_65); -lean_dec(x_63); +lean_inc(x_64); +lean_dec(x_41); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +lean_dec(x_64); lean_inc(x_1); -x_66 = l_Lean_findField_x3f(x_65, x_1, x_33); -lean_dec(x_65); -if (lean_obj_tag(x_66) == 0) +x_67 = l_Lean_findField_x3f(x_66, x_1, x_33); +lean_dec(x_66); +if (lean_obj_tag(x_67) == 0) { -uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_dec(x_32); -x_67 = 0; +x_68 = 0; lean_inc(x_1); -x_68 = l_Lean_MessageData_ofConstName(x_1, x_67); -x_69 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_68); -x_71 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_36, 7); -lean_ctor_set(x_36, 1, x_71); -lean_ctor_set(x_36, 0, x_70); -x_72 = l_Lean_MessageData_ofName(x_33); -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_36); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_73); -lean_ctor_set(x_75, 1, x_74); -x_76 = 1; +x_69 = l_Lean_MessageData_ofConstName(x_1, x_68); +x_70 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_69); +x_72 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_37, 7); +lean_ctor_set(x_37, 1, x_72); +lean_ctor_set(x_37, 0, x_71); +x_73 = l_Lean_MessageData_ofName(x_33); +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_37); +lean_ctor_set(x_74, 1, x_73); +x_75 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +x_77 = 1; lean_inc(x_8); -x_77 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_75, x_76, x_8, x_9, x_64); +x_78 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_76, x_77, x_8, x_9, x_65); lean_dec(x_30); -x_78 = lean_ctor_get(x_77, 1); -lean_inc(x_78); -lean_dec(x_77); -x_79 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_79, 0, x_7); -x_14 = x_79; -x_15 = x_78; +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_7); +x_14 = x_80; +x_15 = x_79; goto block_20; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_free_object(x_36); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - x_80 = x_66; +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_free_object(x_37); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + x_81 = x_67; } else { - lean_dec_ref(x_66); - x_80 = lean_box(0); + lean_dec_ref(x_67); + x_81 = lean_box(0); } -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_30); -lean_ctor_set(x_81, 1, x_32); -x_82 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_81); -if (lean_is_scalar(x_80)) { - x_83 = lean_alloc_ctor(1, 1, 0); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_30); +lean_ctor_set(x_82, 1, x_32); +x_83 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_82); +if (lean_is_scalar(x_81)) { + x_84 = lean_alloc_ctor(1, 1, 0); } else { - x_83 = x_80; + x_84 = x_81; } -lean_ctor_set(x_83, 0, x_82); -x_14 = x_83; -x_15 = x_64; +lean_ctor_set(x_84, 0, x_83); +x_14 = x_84; +x_15 = x_65; goto block_20; } } } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_84 = lean_ctor_get(x_36, 1); -lean_inc(x_84); -lean_dec(x_36); -x_85 = lean_st_ref_get(x_9, x_84); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_85 = lean_ctor_get(x_37, 1); +lean_inc(x_85); +lean_dec(x_37); +x_86 = lean_st_ref_get(x_9, x_85); +x_87 = lean_ctor_get(x_86, 0); lean_inc(x_87); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_88 = x_85; +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_89 = x_86; } else { - lean_dec_ref(x_85); - x_88 = lean_box(0); + lean_dec_ref(x_86); + x_89 = lean_box(0); } -x_89 = lean_ctor_get(x_86, 0); -lean_inc(x_89); -lean_dec(x_86); +x_90 = lean_ctor_get(x_87, 0); +lean_inc(x_90); +lean_dec(x_87); lean_inc(x_1); -x_90 = l_Lean_findField_x3f(x_89, x_1, x_33); -lean_dec(x_89); -if (lean_obj_tag(x_90) == 0) +x_91 = l_Lean_findField_x3f(x_90, x_1, x_33); +lean_dec(x_90); +if (lean_obj_tag(x_91) == 0) { -uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_dec(x_32); -x_91 = 0; +x_92 = 0; lean_inc(x_1); -x_92 = l_Lean_MessageData_ofConstName(x_1, x_91); -x_93 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -if (lean_is_scalar(x_88)) { - x_94 = lean_alloc_ctor(7, 2, 0); -} else { - x_94 = x_88; - lean_ctor_set_tag(x_94, 7); -} -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_92); -x_95 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -x_96 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean_MessageData_ofName(x_33); -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -x_99 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -x_101 = 1; +x_93 = l_Lean_MessageData_ofConstName(x_1, x_92); +x_94 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +if (lean_is_scalar(x_89)) { + x_95 = lean_alloc_ctor(7, 2, 0); +} else { + x_95 = x_89; + lean_ctor_set_tag(x_95, 7); +} +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_93); +x_96 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +x_97 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_MessageData_ofName(x_33); +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_101 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +x_102 = 1; lean_inc(x_8); -x_102 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_100, x_101, x_8, x_9, x_87); +x_103 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_30, x_101, x_102, x_8, x_9, x_88); lean_dec(x_30); -x_103 = lean_ctor_get(x_102, 1); -lean_inc(x_103); -lean_dec(x_102); -x_104 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_104, 0, x_7); -x_14 = x_104; -x_15 = x_103; +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +lean_dec(x_103); +x_105 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_105, 0, x_7); +x_14 = x_105; +x_15 = x_104; goto block_20; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - x_105 = x_90; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + x_106 = x_91; } else { - lean_dec_ref(x_90); - x_105 = lean_box(0); + lean_dec_ref(x_91); + x_106 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_106 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_89)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_106 = x_88; + x_107 = x_89; } -lean_ctor_set(x_106, 0, x_30); -lean_ctor_set(x_106, 1, x_32); -x_107 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_106); -if (lean_is_scalar(x_105)) { - x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_30); +lean_ctor_set(x_107, 1, x_32); +x_108 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_7, x_33, x_107); +if (lean_is_scalar(x_106)) { + x_109 = lean_alloc_ctor(1, 1, 0); } else { - x_108 = x_105; + x_109 = x_106; } -lean_ctor_set(x_108, 0, x_107); -x_14 = x_108; -x_15 = x_87; +lean_ctor_set(x_109, 0, x_108); +x_14 = x_109; +x_15 = x_88; goto block_20; } } @@ -3318,279 +3322,281 @@ return x_30; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; x_31 = lean_unsigned_to_nat(0u); x_32 = l_Lean_Syntax_getArg(x_16, x_31); x_33 = lean_unsigned_to_nat(2u); x_34 = l_Lean_Syntax_getArg(x_16, x_33); x_35 = l_Lean_Syntax_getId(x_32); -x_36 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; -lean_inc(x_1); lean_inc(x_35); -x_37 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_37, 0, x_16); -lean_ctor_set(x_37, 1, x_35); -lean_ctor_set(x_37, 2, x_36); -lean_ctor_set(x_37, 3, x_1); -x_38 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_37, x_11, x_12, x_13); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; +lean_inc(x_1); +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_16); +lean_ctor_set(x_38, 1, x_36); +lean_ctor_set(x_38, 2, x_37); +lean_ctor_set(x_38, 3, x_1); +x_39 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_38, x_11, x_12, x_13); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_40 = lean_ctor_get(x_38, 1); -x_41 = lean_ctor_get(x_38, 0); -lean_dec(x_41); -x_42 = lean_st_ref_get(x_12, x_40); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_39, 1); +x_42 = lean_ctor_get(x_39, 0); +lean_dec(x_42); +x_43 = lean_st_ref_get(x_12, x_41); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_ctor_get(x_42, 1); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -lean_dec(x_44); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_ctor_get(x_43, 1); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); lean_inc(x_1); -x_47 = l_Lean_findField_x3f(x_46, x_1, x_35); -lean_dec(x_46); -if (lean_obj_tag(x_47) == 0) +x_48 = l_Lean_findField_x3f(x_47, x_1, x_35); +lean_dec(x_47); +if (lean_obj_tag(x_48) == 0) { -uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_dec(x_34); -x_48 = 0; +x_49 = 0; lean_inc(x_1); -x_49 = l_Lean_MessageData_ofConstName(x_1, x_48); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -lean_ctor_set_tag(x_42, 7); -lean_ctor_set(x_42, 1, x_49); -lean_ctor_set(x_42, 0, x_50); -x_51 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_51); -lean_ctor_set(x_38, 0, x_42); -x_52 = l_Lean_MessageData_ofName(x_35); -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_38); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -x_56 = 1; +x_50 = l_Lean_MessageData_ofConstName(x_1, x_49); +x_51 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +lean_ctor_set_tag(x_43, 7); +lean_ctor_set(x_43, 1, x_50); +lean_ctor_set(x_43, 0, x_51); +x_52 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_52); +lean_ctor_set(x_39, 0, x_43); +x_53 = l_Lean_MessageData_ofName(x_35); +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_39); +lean_ctor_set(x_54, 1, x_53); +x_55 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = 1; lean_inc(x_11); -x_57 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_55, x_56, x_11, x_12, x_45); +x_58 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_56, x_57, x_11, x_12, x_46); lean_dec(x_32); -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_10); -x_17 = x_59; -x_18 = x_58; +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_10); +x_17 = x_60; +x_18 = x_59; goto block_23; } else { -uint8_t x_60; -lean_free_object(x_38); -x_60 = !lean_is_exclusive(x_47); -if (x_60 == 0) +uint8_t x_61; +lean_free_object(x_39); +x_61 = !lean_is_exclusive(x_48); +if (x_61 == 0) { -lean_object* x_61; lean_object* x_62; -x_61 = lean_ctor_get(x_47, 0); -lean_dec(x_61); -lean_ctor_set(x_42, 1, x_34); -lean_ctor_set(x_42, 0, x_32); -x_62 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_42); -lean_ctor_set(x_47, 0, x_62); -x_17 = x_47; -x_18 = x_45; +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_48, 0); +lean_dec(x_62); +lean_ctor_set(x_43, 1, x_34); +lean_ctor_set(x_43, 0, x_32); +x_63 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_43); +lean_ctor_set(x_48, 0, x_63); +x_17 = x_48; +x_18 = x_46; goto block_23; } else { -lean_object* x_63; lean_object* x_64; -lean_dec(x_47); -lean_ctor_set(x_42, 1, x_34); -lean_ctor_set(x_42, 0, x_32); -x_63 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_42); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -x_17 = x_64; -x_18 = x_45; +lean_object* x_64; lean_object* x_65; +lean_dec(x_48); +lean_ctor_set(x_43, 1, x_34); +lean_ctor_set(x_43, 0, x_32); +x_64 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_43); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_17 = x_65; +x_18 = x_46; goto block_23; } } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = lean_ctor_get(x_42, 0); -x_66 = lean_ctor_get(x_42, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_42); -x_67 = lean_ctor_get(x_65, 0); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_43, 0); +x_67 = lean_ctor_get(x_43, 1); lean_inc(x_67); -lean_dec(x_65); +lean_inc(x_66); +lean_dec(x_43); +x_68 = lean_ctor_get(x_66, 0); +lean_inc(x_68); +lean_dec(x_66); lean_inc(x_1); -x_68 = l_Lean_findField_x3f(x_67, x_1, x_35); -lean_dec(x_67); -if (lean_obj_tag(x_68) == 0) +x_69 = l_Lean_findField_x3f(x_68, x_1, x_35); +lean_dec(x_68); +if (lean_obj_tag(x_69) == 0) { -uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_dec(x_34); -x_69 = 0; +x_70 = 0; lean_inc(x_1); -x_70 = l_Lean_MessageData_ofConstName(x_1, x_69); -x_71 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_70); -x_73 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_38, 7); -lean_ctor_set(x_38, 1, x_73); -lean_ctor_set(x_38, 0, x_72); -x_74 = l_Lean_MessageData_ofName(x_35); -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_38); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -x_78 = 1; +x_71 = l_Lean_MessageData_ofConstName(x_1, x_70); +x_72 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +x_73 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_39, 7); +lean_ctor_set(x_39, 1, x_74); +lean_ctor_set(x_39, 0, x_73); +x_75 = l_Lean_MessageData_ofName(x_35); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_39); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = 1; lean_inc(x_11); -x_79 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_77, x_78, x_11, x_12, x_66); +x_80 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_78, x_79, x_11, x_12, x_67); lean_dec(x_32); -x_80 = lean_ctor_get(x_79, 1); -lean_inc(x_80); -lean_dec(x_79); -x_81 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_81, 0, x_10); -x_17 = x_81; -x_18 = x_80; +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_82, 0, x_10); +x_17 = x_82; +x_18 = x_81; goto block_23; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_free_object(x_38); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - x_82 = x_68; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_free_object(x_39); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + x_83 = x_69; } else { - lean_dec_ref(x_68); - x_82 = lean_box(0); + lean_dec_ref(x_69); + x_83 = lean_box(0); } -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_32); -lean_ctor_set(x_83, 1, x_34); -x_84 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_83); -if (lean_is_scalar(x_82)) { - x_85 = lean_alloc_ctor(1, 1, 0); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_32); +lean_ctor_set(x_84, 1, x_34); +x_85 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_84); +if (lean_is_scalar(x_83)) { + x_86 = lean_alloc_ctor(1, 1, 0); } else { - x_85 = x_82; + x_86 = x_83; } -lean_ctor_set(x_85, 0, x_84); -x_17 = x_85; -x_18 = x_66; +lean_ctor_set(x_86, 0, x_85); +x_17 = x_86; +x_18 = x_67; goto block_23; } } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_86 = lean_ctor_get(x_38, 1); -lean_inc(x_86); -lean_dec(x_38); -x_87 = lean_st_ref_get(x_12, x_86); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_87 = lean_ctor_get(x_39, 1); +lean_inc(x_87); +lean_dec(x_39); +x_88 = lean_st_ref_get(x_12, x_87); +x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_90 = x_87; +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_91 = x_88; } else { - lean_dec_ref(x_87); - x_90 = lean_box(0); + lean_dec_ref(x_88); + x_91 = lean_box(0); } -x_91 = lean_ctor_get(x_88, 0); -lean_inc(x_91); -lean_dec(x_88); +x_92 = lean_ctor_get(x_89, 0); +lean_inc(x_92); +lean_dec(x_89); lean_inc(x_1); -x_92 = l_Lean_findField_x3f(x_91, x_1, x_35); -lean_dec(x_91); -if (lean_obj_tag(x_92) == 0) +x_93 = l_Lean_findField_x3f(x_92, x_1, x_35); +lean_dec(x_92); +if (lean_obj_tag(x_93) == 0) { -uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +uint8_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_dec(x_34); -x_93 = 0; +x_94 = 0; lean_inc(x_1); -x_94 = l_Lean_MessageData_ofConstName(x_1, x_93); -x_95 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -if (lean_is_scalar(x_90)) { - x_96 = lean_alloc_ctor(7, 2, 0); +x_95 = l_Lean_MessageData_ofConstName(x_1, x_94); +x_96 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +if (lean_is_scalar(x_91)) { + x_97 = lean_alloc_ctor(7, 2, 0); } else { - x_96 = x_90; - lean_ctor_set_tag(x_96, 7); + x_97 = x_91; + lean_ctor_set_tag(x_97, 7); } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -x_97 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -x_99 = l_Lean_MessageData_ofName(x_35); -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_102 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -x_103 = 1; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Lean_MessageData_ofName(x_35); +x_101 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +x_104 = 1; lean_inc(x_11); -x_104 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_102, x_103, x_11, x_12, x_89); +x_105 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_32, x_103, x_104, x_11, x_12, x_90); lean_dec(x_32); -x_105 = lean_ctor_get(x_104, 1); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_10); -x_17 = x_106; -x_18 = x_105; +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_10); +x_17 = x_107; +x_18 = x_106; goto block_23; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - x_107 = x_92; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + x_108 = x_93; } else { - lean_dec_ref(x_92); - x_107 = lean_box(0); + lean_dec_ref(x_93); + x_108 = lean_box(0); } -if (lean_is_scalar(x_90)) { - x_108 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_91)) { + x_109 = lean_alloc_ctor(0, 2, 0); } else { - x_108 = x_90; + x_109 = x_91; } -lean_ctor_set(x_108, 0, x_32); -lean_ctor_set(x_108, 1, x_34); -x_109 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_108); -if (lean_is_scalar(x_107)) { - x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_109, 0, x_32); +lean_ctor_set(x_109, 1, x_34); +x_110 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_10, x_35, x_109); +if (lean_is_scalar(x_108)) { + x_111 = lean_alloc_ctor(1, 1, 0); } else { - x_110 = x_107; + x_111 = x_108; } -lean_ctor_set(x_110, 0, x_109); -x_17 = x_110; -x_18 = x_89; +lean_ctor_set(x_111, 0, x_110); +x_17 = x_111; +x_18 = x_90; goto block_23; } } @@ -4051,279 +4057,281 @@ return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; x_30 = lean_unsigned_to_nat(0u); x_31 = l_Lean_Syntax_getArg(x_14, x_30); x_32 = lean_unsigned_to_nat(2u); x_33 = l_Lean_Syntax_getArg(x_14, x_32); x_34 = l_Lean_Syntax_getId(x_31); -x_35 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; -lean_inc(x_1); lean_inc(x_34); -x_36 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_36, 0, x_14); -lean_ctor_set(x_36, 1, x_34); -lean_ctor_set(x_36, 2, x_35); -lean_ctor_set(x_36, 3, x_1); -x_37 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_36, x_9, x_10, x_11); -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_34); +x_36 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__8; +lean_inc(x_1); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_14); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_36); +lean_ctor_set(x_37, 3, x_1); +x_38 = l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2(x_37, x_9, x_10, x_11); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_39 = lean_ctor_get(x_37, 1); -x_40 = lean_ctor_get(x_37, 0); -lean_dec(x_40); -x_41 = lean_st_ref_get(x_10, x_39); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = lean_ctor_get(x_38, 1); +x_41 = lean_ctor_get(x_38, 0); +lean_dec(x_41); +x_42 = lean_st_ref_get(x_10, x_40); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_41, 0); -x_44 = lean_ctor_get(x_41, 1); -x_45 = lean_ctor_get(x_43, 0); -lean_inc(x_45); -lean_dec(x_43); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_42, 1); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +lean_dec(x_44); lean_inc(x_1); -x_46 = l_Lean_findField_x3f(x_45, x_1, x_34); -lean_dec(x_45); -if (lean_obj_tag(x_46) == 0) +x_47 = l_Lean_findField_x3f(x_46, x_1, x_34); +lean_dec(x_46); +if (lean_obj_tag(x_47) == 0) { -uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_dec(x_33); -x_47 = 0; +x_48 = 0; lean_inc(x_1); -x_48 = l_Lean_MessageData_ofConstName(x_1, x_47); -x_49 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -lean_ctor_set_tag(x_41, 7); -lean_ctor_set(x_41, 1, x_48); -lean_ctor_set(x_41, 0, x_49); -x_50 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_37, 7); -lean_ctor_set(x_37, 1, x_50); -lean_ctor_set(x_37, 0, x_41); -x_51 = l_Lean_MessageData_ofName(x_34); -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_37); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -x_55 = 1; +x_49 = l_Lean_MessageData_ofConstName(x_1, x_48); +x_50 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +lean_ctor_set_tag(x_42, 7); +lean_ctor_set(x_42, 1, x_49); +lean_ctor_set(x_42, 0, x_50); +x_51 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_38, 7); +lean_ctor_set(x_38, 1, x_51); +lean_ctor_set(x_38, 0, x_42); +x_52 = l_Lean_MessageData_ofName(x_34); +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_38); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = 1; lean_inc(x_9); -x_56 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_54, x_55, x_9, x_10, x_44); +x_57 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_55, x_56, x_9, x_10, x_45); lean_dec(x_31); -x_57 = lean_ctor_get(x_56, 1); -lean_inc(x_57); -lean_dec(x_56); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_8); -x_15 = x_58; -x_16 = x_57; +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_8); +x_15 = x_59; +x_16 = x_58; goto block_21; } else { -uint8_t x_59; -lean_free_object(x_37); -x_59 = !lean_is_exclusive(x_46); -if (x_59 == 0) +uint8_t x_60; +lean_free_object(x_38); +x_60 = !lean_is_exclusive(x_47); +if (x_60 == 0) { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_46, 0); -lean_dec(x_60); -lean_ctor_set(x_41, 1, x_33); -lean_ctor_set(x_41, 0, x_31); -x_61 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_41); -lean_ctor_set(x_46, 0, x_61); -x_15 = x_46; -x_16 = x_44; +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_47, 0); +lean_dec(x_61); +lean_ctor_set(x_42, 1, x_33); +lean_ctor_set(x_42, 0, x_31); +x_62 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_42); +lean_ctor_set(x_47, 0, x_62); +x_15 = x_47; +x_16 = x_45; goto block_21; } else { -lean_object* x_62; lean_object* x_63; -lean_dec(x_46); -lean_ctor_set(x_41, 1, x_33); -lean_ctor_set(x_41, 0, x_31); -x_62 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_41); -x_63 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_63, 0, x_62); -x_15 = x_63; -x_16 = x_44; +lean_object* x_63; lean_object* x_64; +lean_dec(x_47); +lean_ctor_set(x_42, 1, x_33); +lean_ctor_set(x_42, 0, x_31); +x_63 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_42); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_15 = x_64; +x_16 = x_45; goto block_21; } } } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = lean_ctor_get(x_41, 0); -x_65 = lean_ctor_get(x_41, 1); -lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_41); -x_66 = lean_ctor_get(x_64, 0); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_42, 0); +x_66 = lean_ctor_get(x_42, 1); lean_inc(x_66); -lean_dec(x_64); +lean_inc(x_65); +lean_dec(x_42); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +lean_dec(x_65); lean_inc(x_1); -x_67 = l_Lean_findField_x3f(x_66, x_1, x_34); -lean_dec(x_66); -if (lean_obj_tag(x_67) == 0) +x_68 = l_Lean_findField_x3f(x_67, x_1, x_34); +lean_dec(x_67); +if (lean_obj_tag(x_68) == 0) { -uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_33); -x_68 = 0; +x_69 = 0; lean_inc(x_1); -x_69 = l_Lean_MessageData_ofConstName(x_1, x_68); -x_70 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -x_71 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_69); -x_72 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -lean_ctor_set_tag(x_37, 7); -lean_ctor_set(x_37, 1, x_72); -lean_ctor_set(x_37, 0, x_71); -x_73 = l_Lean_MessageData_ofName(x_34); -x_74 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_74, 0, x_37); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -x_77 = 1; +x_70 = l_Lean_MessageData_ofConstName(x_1, x_69); +x_71 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +x_72 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_70); +x_73 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +lean_ctor_set_tag(x_38, 7); +lean_ctor_set(x_38, 1, x_73); +lean_ctor_set(x_38, 0, x_72); +x_74 = l_Lean_MessageData_ofName(x_34); +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_38); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +x_78 = 1; lean_inc(x_9); -x_78 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_76, x_77, x_9, x_10, x_65); +x_79 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_77, x_78, x_9, x_10, x_66); lean_dec(x_31); -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_8); -x_15 = x_80; -x_16 = x_79; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_81, 0, x_8); +x_15 = x_81; +x_16 = x_80; goto block_21; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_free_object(x_37); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - x_81 = x_67; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_free_object(x_38); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + x_82 = x_68; } else { - lean_dec_ref(x_67); - x_81 = lean_box(0); + lean_dec_ref(x_68); + x_82 = lean_box(0); } -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_31); -lean_ctor_set(x_82, 1, x_33); -x_83 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_82); -if (lean_is_scalar(x_81)) { - x_84 = lean_alloc_ctor(1, 1, 0); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_31); +lean_ctor_set(x_83, 1, x_33); +x_84 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_83); +if (lean_is_scalar(x_82)) { + x_85 = lean_alloc_ctor(1, 1, 0); } else { - x_84 = x_81; + x_85 = x_82; } -lean_ctor_set(x_84, 0, x_83); -x_15 = x_84; -x_16 = x_65; +lean_ctor_set(x_85, 0, x_84); +x_15 = x_85; +x_16 = x_66; goto block_21; } } } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_85 = lean_ctor_get(x_37, 1); -lean_inc(x_85); -lean_dec(x_37); -x_86 = lean_st_ref_get(x_10, x_85); -x_87 = lean_ctor_get(x_86, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_86, 1); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_86 = lean_ctor_get(x_38, 1); +lean_inc(x_86); +lean_dec(x_38); +x_87 = lean_st_ref_get(x_10, x_86); +x_88 = lean_ctor_get(x_87, 0); lean_inc(x_88); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_89 = x_86; +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; } else { - lean_dec_ref(x_86); - x_89 = lean_box(0); + lean_dec_ref(x_87); + x_90 = lean_box(0); } -x_90 = lean_ctor_get(x_87, 0); -lean_inc(x_90); -lean_dec(x_87); +x_91 = lean_ctor_get(x_88, 0); +lean_inc(x_91); +lean_dec(x_88); lean_inc(x_1); -x_91 = l_Lean_findField_x3f(x_90, x_1, x_34); -lean_dec(x_90); -if (lean_obj_tag(x_91) == 0) +x_92 = l_Lean_findField_x3f(x_91, x_1, x_34); +lean_dec(x_91); +if (lean_obj_tag(x_92) == 0) { -uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_dec(x_33); -x_92 = 0; +x_93 = 0; lean_inc(x_1); -x_93 = l_Lean_MessageData_ofConstName(x_1, x_92); -x_94 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; -if (lean_is_scalar(x_89)) { - x_95 = lean_alloc_ctor(7, 2, 0); +x_94 = l_Lean_MessageData_ofConstName(x_1, x_93); +x_95 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__10; +if (lean_is_scalar(x_90)) { + x_96 = lean_alloc_ctor(7, 2, 0); } else { - x_95 = x_89; - lean_ctor_set_tag(x_95, 7); + x_96 = x_90; + lean_ctor_set_tag(x_96, 7); } -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_93); -x_96 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; -x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_MessageData_ofName(x_34); -x_99 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; -x_101 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -x_102 = 1; +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +x_97 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__12; +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_MessageData_ofName(x_34); +x_100 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +x_101 = l_Array_forIn_x27Unsafe_loop___at_Lake_DSL_elabConfigDecl___spec__3___closed__14; +x_102 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +x_103 = 1; lean_inc(x_9); -x_103 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_101, x_102, x_9, x_10, x_88); +x_104 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_31, x_102, x_103, x_9, x_10, x_89); lean_dec(x_31); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -lean_dec(x_103); -x_105 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_105, 0, x_8); -x_15 = x_105; -x_16 = x_104; +x_105 = lean_ctor_get(x_104, 1); +lean_inc(x_105); +lean_dec(x_104); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_8); +x_15 = x_106; +x_16 = x_105; goto block_21; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - x_106 = x_91; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + x_107 = x_92; } else { - lean_dec_ref(x_91); - x_106 = lean_box(0); + lean_dec_ref(x_92); + x_107 = lean_box(0); } -if (lean_is_scalar(x_89)) { - x_107 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_90)) { + x_108 = lean_alloc_ctor(0, 2, 0); } else { - x_107 = x_89; + x_108 = x_90; } -lean_ctor_set(x_107, 0, x_31); -lean_ctor_set(x_107, 1, x_33); -x_108 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_107); -if (lean_is_scalar(x_106)) { - x_109 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_31); +lean_ctor_set(x_108, 1, x_33); +x_109 = l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(x_8, x_34, x_108); +if (lean_is_scalar(x_107)) { + x_110 = lean_alloc_ctor(1, 1, 0); } else { - x_109 = x_106; + x_110 = x_107; } -lean_ctor_set(x_109, 0, x_108); -x_15 = x_109; -x_16 = x_88; +lean_ctor_set(x_110, 0, x_109); +x_15 = x_110; +x_16 = x_89; goto block_21; } } @@ -4552,23 +4560,43 @@ static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked(",", 1, 1); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lake_DSL_expandAttrs___closed__2; +x_2 = l_Lake_DSL_expandAttrs___closed__3; +x_3 = l_Lake_DSL_expandAttrs___closed__4; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; +x_2 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__5() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -4576,19 +4604,19 @@ x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__6() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_expandAttrs___closed__4; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__7() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -4596,19 +4624,19 @@ x_1 = lean_mk_string_unchecked("declaration", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__8() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_simpleDeclSig___closed__7; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__7; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__9() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__11() { _start: { lean_object* x_1; @@ -4616,19 +4644,19 @@ x_1 = lean_mk_string_unchecked("declModifiers", 13, 13); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__10() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_simpleDeclSig___closed__7; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__9; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__11() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__13() { _start: { lean_object* x_1; @@ -4636,7 +4664,7 @@ x_1 = lean_mk_string_unchecked("@[", 2, 2); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__12() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__14() { _start: { lean_object* x_1; @@ -4644,7 +4672,7 @@ x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__13() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__15() { _start: { lean_object* x_1; @@ -4652,19 +4680,19 @@ x_1 = lean_mk_string_unchecked("abbrev", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__14() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_simpleDeclSig___closed__7; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__15() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__17() { _start: { lean_object* x_1; @@ -4672,19 +4700,19 @@ x_1 = lean_mk_string_unchecked("optDeclSig", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__16() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_simpleDeclSig___closed__7; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__17() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__19() { _start: { lean_object* x_1; @@ -4692,7 +4720,7 @@ x_1 = lean_mk_string_unchecked("Termination", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__18() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__20() { _start: { lean_object* x_1; @@ -4700,19 +4728,19 @@ x_1 = lean_mk_string_unchecked("suffix", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__19() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; -x_3 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__17; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +x_3 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__20() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -4721,7 +4749,7 @@ x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__21() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__23() { _start: { lean_object* x_1; @@ -4729,19 +4757,19 @@ x_1 = lean_mk_string_unchecked("declId", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__22() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lake_DSL_expandAttrs___closed__2; x_2 = l_Lake_DSL_expandAttrs___closed__3; x_3 = l_Lake_DSL_simpleDeclSig___closed__7; -x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +x_4 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__23; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__23() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -4755,12 +4783,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__24() { +static lean_object* _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__23; +x_2 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__25; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -4807,45 +4835,45 @@ if (lean_obj_tag(x_8) == 0) lean_free_object(x_25); if (lean_obj_tag(x_9) == 0) { -lean_object* x_676; uint8_t x_677; -x_676 = l_Lean_Elab_Command_getRef(x_12, x_13, x_28); -x_677 = !lean_is_exclusive(x_676); -if (x_677 == 0) +lean_object* x_686; uint8_t x_687; +x_686 = l_Lean_Elab_Command_getRef(x_12, x_13, x_28); +x_687 = !lean_is_exclusive(x_686); +if (x_687 == 0) { -lean_object* x_678; lean_object* x_679; lean_object* x_680; -x_678 = lean_ctor_get(x_676, 0); -x_679 = lean_ctor_get(x_676, 1); +lean_object* x_688; lean_object* x_689; lean_object* x_690; +x_688 = lean_ctor_get(x_686, 0); +x_689 = lean_ctor_get(x_686, 1); lean_inc(x_13); lean_inc(x_12); -x_680 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_678, x_29, x_12, x_13, x_679); -lean_dec(x_678); -if (lean_obj_tag(x_680) == 0) -{ -lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; -x_681 = lean_ctor_get(x_680, 0); -lean_inc(x_681); -x_682 = lean_ctor_get(x_680, 1); -lean_inc(x_682); -lean_dec(x_680); -x_683 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -lean_ctor_set_tag(x_676, 1); -lean_ctor_set(x_676, 1, x_683); -lean_ctor_set(x_676, 0, x_681); -x_684 = lean_array_mk(x_676); -x_685 = lean_box(2); -x_686 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_687 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_687, 0, x_685); -lean_ctor_set(x_687, 1, x_686); -lean_ctor_set(x_687, 2, x_684); -x_31 = x_687; -x_32 = x_682; -goto block_675; -} -else -{ -uint8_t x_688; -lean_free_object(x_676); +x_690 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_688, x_29, x_12, x_13, x_689); +lean_dec(x_688); +if (lean_obj_tag(x_690) == 0) +{ +lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; +x_691 = lean_ctor_get(x_690, 0); +lean_inc(x_691); +x_692 = lean_ctor_get(x_690, 1); +lean_inc(x_692); +lean_dec(x_690); +x_693 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +lean_ctor_set_tag(x_686, 1); +lean_ctor_set(x_686, 1, x_693); +lean_ctor_set(x_686, 0, x_691); +x_694 = lean_array_mk(x_686); +x_695 = lean_box(2); +x_696 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_697 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_697, 0, x_695); +lean_ctor_set(x_697, 1, x_696); +lean_ctor_set(x_697, 2, x_694); +x_31 = x_697; +x_32 = x_692; +goto block_685; +} +else +{ +uint8_t x_698; +lean_free_object(x_686); lean_dec(x_30); lean_dec(x_23); lean_dec(x_13); @@ -4853,64 +4881,64 @@ lean_dec(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_688 = !lean_is_exclusive(x_680); -if (x_688 == 0) +x_698 = !lean_is_exclusive(x_690); +if (x_698 == 0) { -return x_680; +return x_690; } else { -lean_object* x_689; lean_object* x_690; lean_object* x_691; -x_689 = lean_ctor_get(x_680, 0); -x_690 = lean_ctor_get(x_680, 1); -lean_inc(x_690); -lean_inc(x_689); -lean_dec(x_680); -x_691 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_691, 0, x_689); -lean_ctor_set(x_691, 1, x_690); -return x_691; +lean_object* x_699; lean_object* x_700; lean_object* x_701; +x_699 = lean_ctor_get(x_690, 0); +x_700 = lean_ctor_get(x_690, 1); +lean_inc(x_700); +lean_inc(x_699); +lean_dec(x_690); +x_701 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_701, 0, x_699); +lean_ctor_set(x_701, 1, x_700); +return x_701; } } } else { -lean_object* x_692; lean_object* x_693; lean_object* x_694; -x_692 = lean_ctor_get(x_676, 0); -x_693 = lean_ctor_get(x_676, 1); -lean_inc(x_693); -lean_inc(x_692); -lean_dec(x_676); +lean_object* x_702; lean_object* x_703; lean_object* x_704; +x_702 = lean_ctor_get(x_686, 0); +x_703 = lean_ctor_get(x_686, 1); +lean_inc(x_703); +lean_inc(x_702); +lean_dec(x_686); lean_inc(x_13); lean_inc(x_12); -x_694 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_692, x_29, x_12, x_13, x_693); -lean_dec(x_692); -if (lean_obj_tag(x_694) == 0) -{ -lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; -x_695 = lean_ctor_get(x_694, 0); -lean_inc(x_695); -x_696 = lean_ctor_get(x_694, 1); -lean_inc(x_696); -lean_dec(x_694); -x_697 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -x_698 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_698, 0, x_695); -lean_ctor_set(x_698, 1, x_697); -x_699 = lean_array_mk(x_698); -x_700 = lean_box(2); -x_701 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_702 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_702, 0, x_700); -lean_ctor_set(x_702, 1, x_701); -lean_ctor_set(x_702, 2, x_699); -x_31 = x_702; -x_32 = x_696; -goto block_675; +x_704 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_702, x_29, x_12, x_13, x_703); +lean_dec(x_702); +if (lean_obj_tag(x_704) == 0) +{ +lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; +x_705 = lean_ctor_get(x_704, 0); +lean_inc(x_705); +x_706 = lean_ctor_get(x_704, 1); +lean_inc(x_706); +lean_dec(x_704); +x_707 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +x_708 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_708, 0, x_705); +lean_ctor_set(x_708, 1, x_707); +x_709 = lean_array_mk(x_708); +x_710 = lean_box(2); +x_711 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_712 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_712, 0, x_710); +lean_ctor_set(x_712, 1, x_711); +lean_ctor_set(x_712, 2, x_709); +x_31 = x_712; +x_32 = x_706; +goto block_685; } else { -lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; +lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_dec(x_30); lean_dec(x_23); lean_dec(x_13); @@ -4918,77 +4946,77 @@ lean_dec(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_703 = lean_ctor_get(x_694, 0); -lean_inc(x_703); -x_704 = lean_ctor_get(x_694, 1); -lean_inc(x_704); -if (lean_is_exclusive(x_694)) { - lean_ctor_release(x_694, 0); - lean_ctor_release(x_694, 1); - x_705 = x_694; +x_713 = lean_ctor_get(x_704, 0); +lean_inc(x_713); +x_714 = lean_ctor_get(x_704, 1); +lean_inc(x_714); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_715 = x_704; } else { - lean_dec_ref(x_694); - x_705 = lean_box(0); + lean_dec_ref(x_704); + x_715 = lean_box(0); } -if (lean_is_scalar(x_705)) { - x_706 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_715)) { + x_716 = lean_alloc_ctor(1, 2, 0); } else { - x_706 = x_705; + x_716 = x_715; } -lean_ctor_set(x_706, 0, x_703); -lean_ctor_set(x_706, 1, x_704); -return x_706; +lean_ctor_set(x_716, 0, x_713); +lean_ctor_set(x_716, 1, x_714); +return x_716; } } } else { -lean_object* x_707; lean_object* x_708; uint8_t x_709; -x_707 = lean_ctor_get(x_9, 0); -lean_inc(x_707); -lean_dec(x_9); -x_708 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_707, x_29, x_12, x_13, x_28); -x_709 = !lean_is_exclusive(x_708); -if (x_709 == 0) -{ -lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; -x_710 = lean_ctor_get(x_708, 1); -x_711 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -lean_ctor_set_tag(x_708, 1); -lean_ctor_set(x_708, 1, x_711); -x_712 = lean_array_mk(x_708); -x_713 = lean_box(2); -x_714 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_715 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_715, 0, x_713); -lean_ctor_set(x_715, 1, x_714); -lean_ctor_set(x_715, 2, x_712); -x_31 = x_715; -x_32 = x_710; -goto block_675; -} -else -{ -lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; -x_716 = lean_ctor_get(x_708, 0); -x_717 = lean_ctor_get(x_708, 1); +lean_object* x_717; lean_object* x_718; uint8_t x_719; +x_717 = lean_ctor_get(x_9, 0); lean_inc(x_717); -lean_inc(x_716); -lean_dec(x_708); -x_718 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -x_719 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_719, 0, x_716); -lean_ctor_set(x_719, 1, x_718); -x_720 = lean_array_mk(x_719); -x_721 = lean_box(2); -x_722 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_723 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_723, 0, x_721); -lean_ctor_set(x_723, 1, x_722); -lean_ctor_set(x_723, 2, x_720); -x_31 = x_723; -x_32 = x_717; -goto block_675; +lean_dec(x_9); +x_718 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_717, x_29, x_12, x_13, x_28); +x_719 = !lean_is_exclusive(x_718); +if (x_719 == 0) +{ +lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; +x_720 = lean_ctor_get(x_718, 1); +x_721 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +lean_ctor_set_tag(x_718, 1); +lean_ctor_set(x_718, 1, x_721); +x_722 = lean_array_mk(x_718); +x_723 = lean_box(2); +x_724 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_725 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_725, 0, x_723); +lean_ctor_set(x_725, 1, x_724); +lean_ctor_set(x_725, 2, x_722); +x_31 = x_725; +x_32 = x_720; +goto block_685; +} +else +{ +lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; +x_726 = lean_ctor_get(x_718, 0); +x_727 = lean_ctor_get(x_718, 1); +lean_inc(x_727); +lean_inc(x_726); +lean_dec(x_718); +x_728 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +x_729 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_729, 0, x_726); +lean_ctor_set(x_729, 1, x_728); +x_730 = lean_array_mk(x_729); +x_731 = lean_box(2); +x_732 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_733 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_733, 0, x_731); +lean_ctor_set(x_733, 1, x_732); +lean_ctor_set(x_733, 2, x_730); +x_31 = x_733; +x_32 = x_727; +goto block_685; } } } @@ -4996,49 +5024,49 @@ else { if (lean_obj_tag(x_9) == 0) { -lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; -x_724 = lean_ctor_get(x_8, 0); -x_725 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -lean_inc(x_724); +lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; +x_734 = lean_ctor_get(x_8, 0); +x_735 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +lean_inc(x_734); lean_ctor_set_tag(x_25, 1); -lean_ctor_set(x_25, 1, x_725); -lean_ctor_set(x_25, 0, x_724); -x_726 = lean_array_mk(x_25); -x_727 = lean_box(2); -x_728 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_729 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_729, 0, x_727); -lean_ctor_set(x_729, 1, x_728); -lean_ctor_set(x_729, 2, x_726); -x_31 = x_729; +lean_ctor_set(x_25, 1, x_735); +lean_ctor_set(x_25, 0, x_734); +x_736 = lean_array_mk(x_25); +x_737 = lean_box(2); +x_738 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_739 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_739, 0, x_737); +lean_ctor_set(x_739, 1, x_738); +lean_ctor_set(x_739, 2, x_736); +x_31 = x_739; x_32 = x_28; -goto block_675; +goto block_685; } else { -lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; -x_730 = lean_ctor_get(x_8, 0); -x_731 = lean_ctor_get(x_9, 0); -lean_inc(x_731); +lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; +x_740 = lean_ctor_get(x_8, 0); +x_741 = lean_ctor_get(x_9, 0); +lean_inc(x_741); lean_dec(x_9); -x_732 = l_Lean_mkIdentFrom(x_730, x_731, x_29); -x_733 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_742 = l_Lean_mkIdentFrom(x_740, x_741, x_29); +x_743 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; lean_ctor_set_tag(x_25, 1); -lean_ctor_set(x_25, 1, x_733); -lean_ctor_set(x_25, 0, x_732); -x_734 = lean_array_mk(x_25); -x_735 = lean_box(2); -x_736 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_737 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_737, 0, x_735); -lean_ctor_set(x_737, 1, x_736); -lean_ctor_set(x_737, 2, x_734); -x_31 = x_737; +lean_ctor_set(x_25, 1, x_743); +lean_ctor_set(x_25, 0, x_742); +x_744 = lean_array_mk(x_25); +x_745 = lean_box(2); +x_746 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_747 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_747, 0, x_745); +lean_ctor_set(x_747, 1, x_746); +lean_ctor_set(x_747, 2, x_744); +x_31 = x_747; x_32 = x_28; -goto block_675; +goto block_685; } } -block_675: +block_685: { lean_object* x_33; uint8_t x_34; x_33 = l_Lean_Elab_Command_getRef(x_12, x_13, x_32); @@ -5103,7 +5131,7 @@ x_58 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_56); x_59 = !lean_is_exclusive(x_58); if (x_59 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; x_60 = lean_ctor_get(x_58, 1); x_61 = lean_ctor_get(x_58, 0); lean_dec(x_61); @@ -5121,7 +5149,7 @@ lean_ctor_set(x_65, 1, x_63); lean_ctor_set(x_65, 2, x_64); x_66 = lean_array_size(x_23); x_67 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_66, x_17, x_23); -x_68 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +x_68 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; x_69 = l_Lean_mkSepArray(x_67, x_68); lean_dec(x_67); x_70 = l_Array_append___rarg(x_64, x_69); @@ -5131,2269 +5159,2287 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_53); lean_ctor_set(x_71, 1, x_63); lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_53); +x_73 = l_Lean_Syntax_node1(x_53, x_72, x_71); +x_74 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; lean_inc(x_65); lean_inc(x_53); -x_73 = l_Lean_Syntax_node1(x_53, x_72, x_65); -x_74 = l_Lake_DSL_structVal___closed__18; +x_75 = l_Lean_Syntax_node1(x_53, x_74, x_65); +x_76 = l_Lake_DSL_structVal___closed__18; lean_inc(x_53); lean_ctor_set_tag(x_54, 2); -lean_ctor_set(x_54, 1, x_74); +lean_ctor_set(x_54, 1, x_76); lean_ctor_set(x_54, 0, x_53); -x_75 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +x_77 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; lean_inc(x_65); -x_76 = l_Lean_Syntax_node6(x_53, x_75, x_58, x_65, x_71, x_73, x_65, x_54); -x_77 = l_Lean_Elab_Command_getRef(x_12, x_13, x_60); -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_79 = lean_ctor_get(x_77, 0); -x_80 = lean_ctor_get(x_77, 1); -x_81 = l_Lean_SourceInfo_fromRef(x_79, x_29); -lean_dec(x_79); -x_82 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_80); -x_83 = !lean_is_exclusive(x_82); -if (x_83 == 0) +x_78 = l_Lean_Syntax_node6(x_53, x_77, x_58, x_65, x_73, x_75, x_65, x_54); +x_79 = l_Lean_Elab_Command_getRef(x_12, x_13, x_60); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_84 = lean_ctor_get(x_82, 1); -x_85 = lean_ctor_get(x_82, 0); -lean_dec(x_85); -x_86 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_84); -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_81 = lean_ctor_get(x_79, 0); +x_82 = lean_ctor_get(x_79, 1); +x_83 = l_Lean_SourceInfo_fromRef(x_81, x_29); +lean_dec(x_81); +x_84 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_82); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_88 = lean_ctor_get(x_86, 1); -x_89 = lean_ctor_get(x_86, 0); -lean_dec(x_89); -x_90 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -lean_ctor_set_tag(x_86, 2); -lean_ctor_set(x_86, 1, x_90); -lean_ctor_set(x_86, 0, x_81); -x_91 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_92 = l_Lean_Syntax_SepArray_ofElems(x_91, x_4); -x_93 = l_Array_append___rarg(x_64, x_92); -lean_dec(x_92); -lean_inc(x_81); -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_81); -lean_ctor_set(x_94, 1, x_63); -lean_ctor_set(x_94, 2, x_93); -x_95 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -lean_ctor_set_tag(x_82, 2); -lean_ctor_set(x_82, 1, x_95); -lean_ctor_set(x_82, 0, x_81); -x_96 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_97 = l_Lean_Syntax_node3(x_81, x_96, x_86, x_94, x_82); -lean_inc(x_81); -x_98 = l_Lean_Syntax_node1(x_81, x_63, x_97); -lean_inc(x_81); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_81); -lean_ctor_set(x_99, 1, x_63); -lean_ctor_set(x_99, 2, x_64); -x_100 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_100); -lean_ctor_set(x_77, 0, x_81); -x_101 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_84, 1); +x_87 = lean_ctor_get(x_84, 0); +lean_dec(x_87); +x_88 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_86); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +lean_ctor_set_tag(x_88, 2); +lean_ctor_set(x_88, 1, x_92); +lean_ctor_set(x_88, 0, x_83); +x_93 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_94 = l_Lean_Syntax_SepArray_ofElems(x_93, x_4); +x_95 = l_Array_append___rarg(x_64, x_94); +lean_dec(x_94); +lean_inc(x_83); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_83); +lean_ctor_set(x_96, 1, x_63); +lean_ctor_set(x_96, 2, x_95); +x_97 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +lean_ctor_set_tag(x_84, 2); +lean_ctor_set(x_84, 1, x_97); +lean_ctor_set(x_84, 0, x_83); +x_98 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_99 = l_Lean_Syntax_node3(x_83, x_98, x_88, x_96, x_84); +lean_inc(x_83); +x_100 = l_Lean_Syntax_node1(x_83, x_63, x_99); +lean_inc(x_83); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_83); +lean_ctor_set(x_101, 1, x_63); +lean_ctor_set(x_101, 2, x_64); +x_102 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_102); +lean_ctor_set(x_79, 0, x_83); +x_103 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_101); -lean_ctor_set(x_49, 0, x_81); -x_102 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_103 = l_Lean_Syntax_node2(x_81, x_102, x_49, x_30); -lean_inc(x_81); -x_104 = l_Lean_Syntax_node1(x_81, x_63, x_103); -x_105 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_99); -lean_inc(x_81); -x_106 = l_Lean_Syntax_node2(x_81, x_105, x_99, x_104); -x_107 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_103); +lean_ctor_set(x_49, 0, x_83); +x_104 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_105 = l_Lean_Syntax_node2(x_83, x_104, x_49, x_30); +lean_inc(x_83); +x_106 = l_Lean_Syntax_node1(x_83, x_63, x_105); +x_107 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_101); +lean_inc(x_83); +x_108 = l_Lean_Syntax_node2(x_83, x_107, x_101, x_106); +x_109 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_107); -lean_ctor_set(x_33, 0, x_81); -x_108 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_99, 2); -lean_inc(x_81); -x_109 = l_Lean_Syntax_node2(x_81, x_108, x_99, x_99); +lean_ctor_set(x_33, 1, x_109); +lean_ctor_set(x_33, 0, x_83); +x_110 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_101, 2); +lean_inc(x_83); +x_111 = l_Lean_Syntax_node2(x_83, x_110, x_101, x_101); if (lean_obj_tag(x_7) == 0) { -x_110 = x_64; -goto block_137; +x_112 = x_64; +goto block_139; } else { -lean_object* x_138; lean_object* x_139; -x_138 = lean_ctor_get(x_7, 0); -lean_inc(x_138); +lean_object* x_140; lean_object* x_141; +x_140 = lean_ctor_get(x_7, 0); +lean_inc(x_140); lean_dec(x_7); -x_139 = l_Array_mkArray1___rarg(x_138); -x_110 = x_139; -goto block_137; -} -block_137: -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_111 = l_Array_append___rarg(x_64, x_110); -lean_dec(x_110); -lean_inc(x_81); -x_112 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_112, 0, x_81); -lean_ctor_set(x_112, 1, x_63); -lean_ctor_set(x_112, 2, x_111); -x_113 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_99, 3); -lean_inc(x_81); -x_114 = l_Lean_Syntax_node6(x_81, x_113, x_112, x_98, x_99, x_99, x_99, x_99); +x_141 = l_Array_mkArray1___rarg(x_140); +x_112 = x_141; +goto block_139; +} +block_139: +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = l_Array_append___rarg(x_64, x_112); +lean_dec(x_112); +lean_inc(x_83); +x_114 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_114, 0, x_83); +lean_ctor_set(x_114, 1, x_63); +lean_ctor_set(x_114, 2, x_113); +x_115 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_101, 3); +lean_inc(x_83); +x_116 = l_Lean_Syntax_node6(x_83, x_115, x_114, x_100, x_101, x_101, x_101, x_101); if (lean_obj_tag(x_5) == 0) { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_115 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_81); -lean_ctor_set(x_116, 1, x_63); -lean_ctor_set(x_116, 2, x_115); -x_117 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_118 = l_Lean_Syntax_node4(x_81, x_117, x_33, x_76, x_109, x_116); -x_119 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_120 = l_Lean_Syntax_node4(x_81, x_119, x_77, x_31, x_106, x_118); -x_121 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_122 = l_Lean_Syntax_node2(x_81, x_121, x_114, x_120); -lean_inc(x_122); -x_123 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_123, 0, x_122); -x_124 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_122, x_123, x_12, x_13, x_88); -return x_124; -} -else -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_125 = lean_ctor_get(x_5, 0); -lean_inc(x_125); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_117 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_83); +lean_ctor_set(x_118, 1, x_63); +lean_ctor_set(x_118, 2, x_117); +x_119 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_120 = l_Lean_Syntax_node4(x_83, x_119, x_33, x_78, x_111, x_118); +x_121 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_122 = l_Lean_Syntax_node4(x_83, x_121, x_79, x_31, x_108, x_120); +x_123 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_124 = l_Lean_Syntax_node2(x_83, x_123, x_116, x_122); +lean_inc(x_124); +x_125 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_125, 0, x_124); +x_126 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_124, x_125, x_12, x_13, x_90); +return x_126; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_127 = lean_ctor_get(x_5, 0); +lean_inc(x_127); lean_dec(x_5); -x_126 = l_Array_mkArray1___rarg(x_125); -x_127 = l_Array_append___rarg(x_64, x_126); -lean_dec(x_126); -lean_inc(x_81); -x_128 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_128, 0, x_81); -lean_ctor_set(x_128, 1, x_63); -lean_ctor_set(x_128, 2, x_127); -x_129 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_130 = l_Lean_Syntax_node4(x_81, x_129, x_33, x_76, x_109, x_128); -x_131 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_132 = l_Lean_Syntax_node4(x_81, x_131, x_77, x_31, x_106, x_130); -x_133 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_134 = l_Lean_Syntax_node2(x_81, x_133, x_114, x_132); -lean_inc(x_134); -x_135 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_135, 0, x_134); -x_136 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_134, x_135, x_12, x_13, x_88); -return x_136; -} -} -} -else -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_140 = lean_ctor_get(x_86, 1); -lean_inc(x_140); -lean_dec(x_86); -x_141 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -x_142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_142, 0, x_81); -lean_ctor_set(x_142, 1, x_141); -x_143 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_144 = l_Lean_Syntax_SepArray_ofElems(x_143, x_4); -x_145 = l_Array_append___rarg(x_64, x_144); -lean_dec(x_144); -lean_inc(x_81); -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_81); -lean_ctor_set(x_146, 1, x_63); -lean_ctor_set(x_146, 2, x_145); -x_147 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -lean_ctor_set_tag(x_82, 2); -lean_ctor_set(x_82, 1, x_147); -lean_ctor_set(x_82, 0, x_81); -x_148 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_149 = l_Lean_Syntax_node3(x_81, x_148, x_142, x_146, x_82); -lean_inc(x_81); -x_150 = l_Lean_Syntax_node1(x_81, x_63, x_149); -lean_inc(x_81); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_81); -lean_ctor_set(x_151, 1, x_63); -lean_ctor_set(x_151, 2, x_64); -x_152 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_152); -lean_ctor_set(x_77, 0, x_81); -x_153 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); +x_128 = l_Array_mkArray1___rarg(x_127); +x_129 = l_Array_append___rarg(x_64, x_128); +lean_dec(x_128); +lean_inc(x_83); +x_130 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_130, 0, x_83); +lean_ctor_set(x_130, 1, x_63); +lean_ctor_set(x_130, 2, x_129); +x_131 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_132 = l_Lean_Syntax_node4(x_83, x_131, x_33, x_78, x_111, x_130); +x_133 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_134 = l_Lean_Syntax_node4(x_83, x_133, x_79, x_31, x_108, x_132); +x_135 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_136 = l_Lean_Syntax_node2(x_83, x_135, x_116, x_134); +lean_inc(x_136); +x_137 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_137, 0, x_136); +x_138 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_136, x_137, x_12, x_13, x_90); +return x_138; +} +} +} +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_142 = lean_ctor_get(x_88, 1); +lean_inc(x_142); +lean_dec(x_88); +x_143 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +x_144 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_144, 0, x_83); +lean_ctor_set(x_144, 1, x_143); +x_145 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_146 = l_Lean_Syntax_SepArray_ofElems(x_145, x_4); +x_147 = l_Array_append___rarg(x_64, x_146); +lean_dec(x_146); +lean_inc(x_83); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_83); +lean_ctor_set(x_148, 1, x_63); +lean_ctor_set(x_148, 2, x_147); +x_149 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +lean_ctor_set_tag(x_84, 2); +lean_ctor_set(x_84, 1, x_149); +lean_ctor_set(x_84, 0, x_83); +x_150 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_151 = l_Lean_Syntax_node3(x_83, x_150, x_144, x_148, x_84); +lean_inc(x_83); +x_152 = l_Lean_Syntax_node1(x_83, x_63, x_151); +lean_inc(x_83); +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_83); +lean_ctor_set(x_153, 1, x_63); +lean_ctor_set(x_153, 2, x_64); +x_154 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_154); +lean_ctor_set(x_79, 0, x_83); +x_155 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_153); -lean_ctor_set(x_49, 0, x_81); -x_154 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_155 = l_Lean_Syntax_node2(x_81, x_154, x_49, x_30); -lean_inc(x_81); -x_156 = l_Lean_Syntax_node1(x_81, x_63, x_155); -x_157 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_151); -lean_inc(x_81); -x_158 = l_Lean_Syntax_node2(x_81, x_157, x_151, x_156); -x_159 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_155); +lean_ctor_set(x_49, 0, x_83); +x_156 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_157 = l_Lean_Syntax_node2(x_83, x_156, x_49, x_30); +lean_inc(x_83); +x_158 = l_Lean_Syntax_node1(x_83, x_63, x_157); +x_159 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_153); +lean_inc(x_83); +x_160 = l_Lean_Syntax_node2(x_83, x_159, x_153, x_158); +x_161 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_159); -lean_ctor_set(x_33, 0, x_81); -x_160 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_151, 2); -lean_inc(x_81); -x_161 = l_Lean_Syntax_node2(x_81, x_160, x_151, x_151); +lean_ctor_set(x_33, 1, x_161); +lean_ctor_set(x_33, 0, x_83); +x_162 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_153, 2); +lean_inc(x_83); +x_163 = l_Lean_Syntax_node2(x_83, x_162, x_153, x_153); if (lean_obj_tag(x_7) == 0) { -x_162 = x_64; -goto block_189; +x_164 = x_64; +goto block_191; } else { -lean_object* x_190; lean_object* x_191; -x_190 = lean_ctor_get(x_7, 0); -lean_inc(x_190); +lean_object* x_192; lean_object* x_193; +x_192 = lean_ctor_get(x_7, 0); +lean_inc(x_192); lean_dec(x_7); -x_191 = l_Array_mkArray1___rarg(x_190); -x_162 = x_191; -goto block_189; -} -block_189: -{ -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_163 = l_Array_append___rarg(x_64, x_162); -lean_dec(x_162); -lean_inc(x_81); -x_164 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_164, 0, x_81); -lean_ctor_set(x_164, 1, x_63); -lean_ctor_set(x_164, 2, x_163); -x_165 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_151, 3); -lean_inc(x_81); -x_166 = l_Lean_Syntax_node6(x_81, x_165, x_164, x_150, x_151, x_151, x_151, x_151); +x_193 = l_Array_mkArray1___rarg(x_192); +x_164 = x_193; +goto block_191; +} +block_191: +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_165 = l_Array_append___rarg(x_64, x_164); +lean_dec(x_164); +lean_inc(x_83); +x_166 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_166, 0, x_83); +lean_ctor_set(x_166, 1, x_63); +lean_ctor_set(x_166, 2, x_165); +x_167 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_153, 3); +lean_inc(x_83); +x_168 = l_Lean_Syntax_node6(x_83, x_167, x_166, x_152, x_153, x_153, x_153, x_153); if (lean_obj_tag(x_5) == 0) { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_167 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_81); -lean_ctor_set(x_168, 1, x_63); -lean_ctor_set(x_168, 2, x_167); -x_169 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_170 = l_Lean_Syntax_node4(x_81, x_169, x_33, x_76, x_161, x_168); -x_171 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_172 = l_Lean_Syntax_node4(x_81, x_171, x_77, x_31, x_158, x_170); -x_173 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_174 = l_Lean_Syntax_node2(x_81, x_173, x_166, x_172); -lean_inc(x_174); -x_175 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_175, 0, x_174); -x_176 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_174, x_175, x_12, x_13, x_140); -return x_176; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_169 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_83); +lean_ctor_set(x_170, 1, x_63); +lean_ctor_set(x_170, 2, x_169); +x_171 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_172 = l_Lean_Syntax_node4(x_83, x_171, x_33, x_78, x_163, x_170); +x_173 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_174 = l_Lean_Syntax_node4(x_83, x_173, x_79, x_31, x_160, x_172); +x_175 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_176 = l_Lean_Syntax_node2(x_83, x_175, x_168, x_174); +lean_inc(x_176); +x_177 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_177, 0, x_176); +x_178 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_176, x_177, x_12, x_13, x_142); +return x_178; } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_177 = lean_ctor_get(x_5, 0); -lean_inc(x_177); +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_179 = lean_ctor_get(x_5, 0); +lean_inc(x_179); lean_dec(x_5); -x_178 = l_Array_mkArray1___rarg(x_177); -x_179 = l_Array_append___rarg(x_64, x_178); -lean_dec(x_178); -lean_inc(x_81); -x_180 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_180, 0, x_81); -lean_ctor_set(x_180, 1, x_63); -lean_ctor_set(x_180, 2, x_179); -x_181 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_182 = l_Lean_Syntax_node4(x_81, x_181, x_33, x_76, x_161, x_180); -x_183 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_184 = l_Lean_Syntax_node4(x_81, x_183, x_77, x_31, x_158, x_182); -x_185 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_186 = l_Lean_Syntax_node2(x_81, x_185, x_166, x_184); -lean_inc(x_186); -x_187 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_187, 0, x_186); -x_188 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_186, x_187, x_12, x_13, x_140); -return x_188; -} -} -} -} -else -{ -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -x_192 = lean_ctor_get(x_82, 1); -lean_inc(x_192); -lean_dec(x_82); -x_193 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_192); -x_194 = lean_ctor_get(x_193, 1); +x_180 = l_Array_mkArray1___rarg(x_179); +x_181 = l_Array_append___rarg(x_64, x_180); +lean_dec(x_180); +lean_inc(x_83); +x_182 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_182, 0, x_83); +lean_ctor_set(x_182, 1, x_63); +lean_ctor_set(x_182, 2, x_181); +x_183 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_184 = l_Lean_Syntax_node4(x_83, x_183, x_33, x_78, x_163, x_182); +x_185 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_186 = l_Lean_Syntax_node4(x_83, x_185, x_79, x_31, x_160, x_184); +x_187 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_188 = l_Lean_Syntax_node2(x_83, x_187, x_168, x_186); +lean_inc(x_188); +x_189 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_189, 0, x_188); +x_190 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_188, x_189, x_12, x_13, x_142); +return x_190; +} +} +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_194 = lean_ctor_get(x_84, 1); lean_inc(x_194); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - x_195 = x_193; +lean_dec(x_84); +x_195 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_194); +x_196 = lean_ctor_get(x_195, 1); +lean_inc(x_196); +if (lean_is_exclusive(x_195)) { + lean_ctor_release(x_195, 0); + lean_ctor_release(x_195, 1); + x_197 = x_195; } else { - lean_dec_ref(x_193); - x_195 = lean_box(0); + lean_dec_ref(x_195); + x_197 = lean_box(0); } -x_196 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -if (lean_is_scalar(x_195)) { - x_197 = lean_alloc_ctor(2, 2, 0); +x_198 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +if (lean_is_scalar(x_197)) { + x_199 = lean_alloc_ctor(2, 2, 0); } else { - x_197 = x_195; - lean_ctor_set_tag(x_197, 2); -} -lean_ctor_set(x_197, 0, x_81); -lean_ctor_set(x_197, 1, x_196); -x_198 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_199 = l_Lean_Syntax_SepArray_ofElems(x_198, x_4); -x_200 = l_Array_append___rarg(x_64, x_199); -lean_dec(x_199); -lean_inc(x_81); -x_201 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_201, 0, x_81); -lean_ctor_set(x_201, 1, x_63); -lean_ctor_set(x_201, 2, x_200); -x_202 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -x_203 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_203, 0, x_81); -lean_ctor_set(x_203, 1, x_202); -x_204 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_205 = l_Lean_Syntax_node3(x_81, x_204, x_197, x_201, x_203); -lean_inc(x_81); -x_206 = l_Lean_Syntax_node1(x_81, x_63, x_205); -lean_inc(x_81); -x_207 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_207, 0, x_81); -lean_ctor_set(x_207, 1, x_63); -lean_ctor_set(x_207, 2, x_64); -x_208 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_208); -lean_ctor_set(x_77, 0, x_81); -x_209 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); + x_199 = x_197; + lean_ctor_set_tag(x_199, 2); +} +lean_ctor_set(x_199, 0, x_83); +lean_ctor_set(x_199, 1, x_198); +x_200 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_201 = l_Lean_Syntax_SepArray_ofElems(x_200, x_4); +x_202 = l_Array_append___rarg(x_64, x_201); +lean_dec(x_201); +lean_inc(x_83); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_83); +lean_ctor_set(x_203, 1, x_63); +lean_ctor_set(x_203, 2, x_202); +x_204 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_83); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_207 = l_Lean_Syntax_node3(x_83, x_206, x_199, x_203, x_205); +lean_inc(x_83); +x_208 = l_Lean_Syntax_node1(x_83, x_63, x_207); +lean_inc(x_83); +x_209 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_209, 0, x_83); +lean_ctor_set(x_209, 1, x_63); +lean_ctor_set(x_209, 2, x_64); +x_210 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_210); +lean_ctor_set(x_79, 0, x_83); +x_211 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_209); -lean_ctor_set(x_49, 0, x_81); -x_210 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_211 = l_Lean_Syntax_node2(x_81, x_210, x_49, x_30); -lean_inc(x_81); -x_212 = l_Lean_Syntax_node1(x_81, x_63, x_211); -x_213 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_207); -lean_inc(x_81); -x_214 = l_Lean_Syntax_node2(x_81, x_213, x_207, x_212); -x_215 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_211); +lean_ctor_set(x_49, 0, x_83); +x_212 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_213 = l_Lean_Syntax_node2(x_83, x_212, x_49, x_30); +lean_inc(x_83); +x_214 = l_Lean_Syntax_node1(x_83, x_63, x_213); +x_215 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_209); +lean_inc(x_83); +x_216 = l_Lean_Syntax_node2(x_83, x_215, x_209, x_214); +x_217 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_215); -lean_ctor_set(x_33, 0, x_81); -x_216 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_207, 2); -lean_inc(x_81); -x_217 = l_Lean_Syntax_node2(x_81, x_216, x_207, x_207); +lean_ctor_set(x_33, 1, x_217); +lean_ctor_set(x_33, 0, x_83); +x_218 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_209, 2); +lean_inc(x_83); +x_219 = l_Lean_Syntax_node2(x_83, x_218, x_209, x_209); if (lean_obj_tag(x_7) == 0) { -x_218 = x_64; -goto block_245; +x_220 = x_64; +goto block_247; } else { -lean_object* x_246; lean_object* x_247; -x_246 = lean_ctor_get(x_7, 0); -lean_inc(x_246); +lean_object* x_248; lean_object* x_249; +x_248 = lean_ctor_get(x_7, 0); +lean_inc(x_248); lean_dec(x_7); -x_247 = l_Array_mkArray1___rarg(x_246); -x_218 = x_247; -goto block_245; -} -block_245: -{ -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_219 = l_Array_append___rarg(x_64, x_218); -lean_dec(x_218); -lean_inc(x_81); -x_220 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_220, 0, x_81); -lean_ctor_set(x_220, 1, x_63); -lean_ctor_set(x_220, 2, x_219); -x_221 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_207, 3); -lean_inc(x_81); -x_222 = l_Lean_Syntax_node6(x_81, x_221, x_220, x_206, x_207, x_207, x_207, x_207); +x_249 = l_Array_mkArray1___rarg(x_248); +x_220 = x_249; +goto block_247; +} +block_247: +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_221 = l_Array_append___rarg(x_64, x_220); +lean_dec(x_220); +lean_inc(x_83); +x_222 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_222, 0, x_83); +lean_ctor_set(x_222, 1, x_63); +lean_ctor_set(x_222, 2, x_221); +x_223 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_209, 3); +lean_inc(x_83); +x_224 = l_Lean_Syntax_node6(x_83, x_223, x_222, x_208, x_209, x_209, x_209, x_209); if (lean_obj_tag(x_5) == 0) { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_223 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_224, 0, x_81); -lean_ctor_set(x_224, 1, x_63); -lean_ctor_set(x_224, 2, x_223); -x_225 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_226 = l_Lean_Syntax_node4(x_81, x_225, x_33, x_76, x_217, x_224); -x_227 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_228 = l_Lean_Syntax_node4(x_81, x_227, x_77, x_31, x_214, x_226); -x_229 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_230 = l_Lean_Syntax_node2(x_81, x_229, x_222, x_228); -lean_inc(x_230); -x_231 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_231, 0, x_230); -x_232 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_230, x_231, x_12, x_13, x_194); -return x_232; -} -else -{ -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_233 = lean_ctor_get(x_5, 0); -lean_inc(x_233); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_225 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_83); +lean_ctor_set(x_226, 1, x_63); +lean_ctor_set(x_226, 2, x_225); +x_227 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_228 = l_Lean_Syntax_node4(x_83, x_227, x_33, x_78, x_219, x_226); +x_229 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_230 = l_Lean_Syntax_node4(x_83, x_229, x_79, x_31, x_216, x_228); +x_231 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_232 = l_Lean_Syntax_node2(x_83, x_231, x_224, x_230); +lean_inc(x_232); +x_233 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_233, 0, x_232); +x_234 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_232, x_233, x_12, x_13, x_196); +return x_234; +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_235 = lean_ctor_get(x_5, 0); +lean_inc(x_235); lean_dec(x_5); -x_234 = l_Array_mkArray1___rarg(x_233); -x_235 = l_Array_append___rarg(x_64, x_234); -lean_dec(x_234); -lean_inc(x_81); -x_236 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_236, 0, x_81); -lean_ctor_set(x_236, 1, x_63); -lean_ctor_set(x_236, 2, x_235); -x_237 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_81); -x_238 = l_Lean_Syntax_node4(x_81, x_237, x_33, x_76, x_217, x_236); -x_239 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_240 = l_Lean_Syntax_node4(x_81, x_239, x_77, x_31, x_214, x_238); -x_241 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_242 = l_Lean_Syntax_node2(x_81, x_241, x_222, x_240); -lean_inc(x_242); -x_243 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_243, 0, x_242); -x_244 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_242, x_243, x_12, x_13, x_194); -return x_244; +x_236 = l_Array_mkArray1___rarg(x_235); +x_237 = l_Array_append___rarg(x_64, x_236); +lean_dec(x_236); +lean_inc(x_83); +x_238 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_238, 0, x_83); +lean_ctor_set(x_238, 1, x_63); +lean_ctor_set(x_238, 2, x_237); +x_239 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_83); +x_240 = l_Lean_Syntax_node4(x_83, x_239, x_33, x_78, x_219, x_238); +x_241 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_242 = l_Lean_Syntax_node4(x_83, x_241, x_79, x_31, x_216, x_240); +x_243 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_244 = l_Lean_Syntax_node2(x_83, x_243, x_224, x_242); +lean_inc(x_244); +x_245 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_245, 0, x_244); +x_246 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_244, x_245, x_12, x_13, x_196); +return x_246; } } } } else { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; -x_248 = lean_ctor_get(x_77, 0); -x_249 = lean_ctor_get(x_77, 1); -lean_inc(x_249); -lean_inc(x_248); -lean_dec(x_77); -x_250 = l_Lean_SourceInfo_fromRef(x_248, x_29); -lean_dec(x_248); -x_251 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_249); -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_253 = x_251; +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_250 = lean_ctor_get(x_79, 0); +x_251 = lean_ctor_get(x_79, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_79); +x_252 = l_Lean_SourceInfo_fromRef(x_250, x_29); +lean_dec(x_250); +x_253 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_251); +x_254 = lean_ctor_get(x_253, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_253)) { + lean_ctor_release(x_253, 0); + lean_ctor_release(x_253, 1); + x_255 = x_253; } else { - lean_dec_ref(x_251); - x_253 = lean_box(0); + lean_dec_ref(x_253); + x_255 = lean_box(0); } -x_254 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_252); -x_255 = lean_ctor_get(x_254, 1); -lean_inc(x_255); -if (lean_is_exclusive(x_254)) { - lean_ctor_release(x_254, 0); - lean_ctor_release(x_254, 1); - x_256 = x_254; +x_256 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_254); +x_257 = lean_ctor_get(x_256, 1); +lean_inc(x_257); +if (lean_is_exclusive(x_256)) { + lean_ctor_release(x_256, 0); + lean_ctor_release(x_256, 1); + x_258 = x_256; } else { - lean_dec_ref(x_254); - x_256 = lean_box(0); + lean_dec_ref(x_256); + x_258 = lean_box(0); } -x_257 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_250); -if (lean_is_scalar(x_256)) { - x_258 = lean_alloc_ctor(2, 2, 0); +x_259 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_252); +if (lean_is_scalar(x_258)) { + x_260 = lean_alloc_ctor(2, 2, 0); } else { - x_258 = x_256; - lean_ctor_set_tag(x_258, 2); -} -lean_ctor_set(x_258, 0, x_250); -lean_ctor_set(x_258, 1, x_257); -x_259 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_260 = l_Lean_Syntax_SepArray_ofElems(x_259, x_4); -x_261 = l_Array_append___rarg(x_64, x_260); -lean_dec(x_260); -lean_inc(x_250); -x_262 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_262, 0, x_250); -lean_ctor_set(x_262, 1, x_63); -lean_ctor_set(x_262, 2, x_261); -x_263 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_250); -if (lean_is_scalar(x_253)) { - x_264 = lean_alloc_ctor(2, 2, 0); + x_260 = x_258; + lean_ctor_set_tag(x_260, 2); +} +lean_ctor_set(x_260, 0, x_252); +lean_ctor_set(x_260, 1, x_259); +x_261 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_262 = l_Lean_Syntax_SepArray_ofElems(x_261, x_4); +x_263 = l_Array_append___rarg(x_64, x_262); +lean_dec(x_262); +lean_inc(x_252); +x_264 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_264, 0, x_252); +lean_ctor_set(x_264, 1, x_63); +lean_ctor_set(x_264, 2, x_263); +x_265 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_252); +if (lean_is_scalar(x_255)) { + x_266 = lean_alloc_ctor(2, 2, 0); } else { - x_264 = x_253; - lean_ctor_set_tag(x_264, 2); + x_266 = x_255; + lean_ctor_set_tag(x_266, 2); } -lean_ctor_set(x_264, 0, x_250); -lean_ctor_set(x_264, 1, x_263); -x_265 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_250); -x_266 = l_Lean_Syntax_node3(x_250, x_265, x_258, x_262, x_264); -lean_inc(x_250); -x_267 = l_Lean_Syntax_node1(x_250, x_63, x_266); -lean_inc(x_250); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_250); -lean_ctor_set(x_268, 1, x_63); -lean_ctor_set(x_268, 2, x_64); -x_269 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_250); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_250); -lean_ctor_set(x_270, 1, x_269); -x_271 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_250); +lean_ctor_set(x_266, 0, x_252); +lean_ctor_set(x_266, 1, x_265); +x_267 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_252); +x_268 = l_Lean_Syntax_node3(x_252, x_267, x_260, x_264, x_266); +lean_inc(x_252); +x_269 = l_Lean_Syntax_node1(x_252, x_63, x_268); +lean_inc(x_252); +x_270 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_270, 0, x_252); +lean_ctor_set(x_270, 1, x_63); +lean_ctor_set(x_270, 2, x_64); +x_271 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_252); +x_272 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_272, 0, x_252); +lean_ctor_set(x_272, 1, x_271); +x_273 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_252); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_271); -lean_ctor_set(x_49, 0, x_250); -x_272 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_250); -x_273 = l_Lean_Syntax_node2(x_250, x_272, x_49, x_30); -lean_inc(x_250); -x_274 = l_Lean_Syntax_node1(x_250, x_63, x_273); -x_275 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_268); -lean_inc(x_250); -x_276 = l_Lean_Syntax_node2(x_250, x_275, x_268, x_274); -x_277 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_250); +lean_ctor_set(x_49, 1, x_273); +lean_ctor_set(x_49, 0, x_252); +x_274 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_252); +x_275 = l_Lean_Syntax_node2(x_252, x_274, x_49, x_30); +lean_inc(x_252); +x_276 = l_Lean_Syntax_node1(x_252, x_63, x_275); +x_277 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_270); +lean_inc(x_252); +x_278 = l_Lean_Syntax_node2(x_252, x_277, x_270, x_276); +x_279 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_252); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_277); -lean_ctor_set(x_33, 0, x_250); -x_278 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_268, 2); -lean_inc(x_250); -x_279 = l_Lean_Syntax_node2(x_250, x_278, x_268, x_268); +lean_ctor_set(x_33, 1, x_279); +lean_ctor_set(x_33, 0, x_252); +x_280 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_270, 2); +lean_inc(x_252); +x_281 = l_Lean_Syntax_node2(x_252, x_280, x_270, x_270); if (lean_obj_tag(x_7) == 0) { -x_280 = x_64; -goto block_307; +x_282 = x_64; +goto block_309; } else { -lean_object* x_308; lean_object* x_309; -x_308 = lean_ctor_get(x_7, 0); -lean_inc(x_308); +lean_object* x_310; lean_object* x_311; +x_310 = lean_ctor_get(x_7, 0); +lean_inc(x_310); lean_dec(x_7); -x_309 = l_Array_mkArray1___rarg(x_308); -x_280 = x_309; -goto block_307; +x_311 = l_Array_mkArray1___rarg(x_310); +x_282 = x_311; +goto block_309; } -block_307: +block_309: { -lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_281 = l_Array_append___rarg(x_64, x_280); -lean_dec(x_280); -lean_inc(x_250); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_250); -lean_ctor_set(x_282, 1, x_63); -lean_ctor_set(x_282, 2, x_281); -x_283 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_268, 3); -lean_inc(x_250); -x_284 = l_Lean_Syntax_node6(x_250, x_283, x_282, x_267, x_268, x_268, x_268, x_268); +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_283 = l_Array_append___rarg(x_64, x_282); +lean_dec(x_282); +lean_inc(x_252); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_252); +lean_ctor_set(x_284, 1, x_63); +lean_ctor_set(x_284, 2, x_283); +x_285 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_270, 3); +lean_inc(x_252); +x_286 = l_Lean_Syntax_node6(x_252, x_285, x_284, x_269, x_270, x_270, x_270, x_270); if (lean_obj_tag(x_5) == 0) { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; -x_285 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_250); -x_286 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_286, 0, x_250); -lean_ctor_set(x_286, 1, x_63); -lean_ctor_set(x_286, 2, x_285); -x_287 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_250); -x_288 = l_Lean_Syntax_node4(x_250, x_287, x_33, x_76, x_279, x_286); -x_289 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_250); -x_290 = l_Lean_Syntax_node4(x_250, x_289, x_270, x_31, x_276, x_288); -x_291 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_292 = l_Lean_Syntax_node2(x_250, x_291, x_284, x_290); -lean_inc(x_292); -x_293 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_293, 0, x_292); -x_294 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_292, x_293, x_12, x_13, x_255); -return x_294; +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_287 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_252); +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_252); +lean_ctor_set(x_288, 1, x_63); +lean_ctor_set(x_288, 2, x_287); +x_289 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_252); +x_290 = l_Lean_Syntax_node4(x_252, x_289, x_33, x_78, x_281, x_288); +x_291 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_252); +x_292 = l_Lean_Syntax_node4(x_252, x_291, x_272, x_31, x_278, x_290); +x_293 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_294 = l_Lean_Syntax_node2(x_252, x_293, x_286, x_292); +lean_inc(x_294); +x_295 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_295, 0, x_294); +x_296 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_294, x_295, x_12, x_13, x_257); +return x_296; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_295 = lean_ctor_get(x_5, 0); -lean_inc(x_295); +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; +x_297 = lean_ctor_get(x_5, 0); +lean_inc(x_297); lean_dec(x_5); -x_296 = l_Array_mkArray1___rarg(x_295); -x_297 = l_Array_append___rarg(x_64, x_296); -lean_dec(x_296); -lean_inc(x_250); -x_298 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_298, 0, x_250); -lean_ctor_set(x_298, 1, x_63); -lean_ctor_set(x_298, 2, x_297); -x_299 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_250); -x_300 = l_Lean_Syntax_node4(x_250, x_299, x_33, x_76, x_279, x_298); -x_301 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_250); -x_302 = l_Lean_Syntax_node4(x_250, x_301, x_270, x_31, x_276, x_300); -x_303 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_304 = l_Lean_Syntax_node2(x_250, x_303, x_284, x_302); -lean_inc(x_304); -x_305 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_305, 0, x_304); -x_306 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_304, x_305, x_12, x_13, x_255); -return x_306; +x_298 = l_Array_mkArray1___rarg(x_297); +x_299 = l_Array_append___rarg(x_64, x_298); +lean_dec(x_298); +lean_inc(x_252); +x_300 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_300, 0, x_252); +lean_ctor_set(x_300, 1, x_63); +lean_ctor_set(x_300, 2, x_299); +x_301 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_252); +x_302 = l_Lean_Syntax_node4(x_252, x_301, x_33, x_78, x_281, x_300); +x_303 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_252); +x_304 = l_Lean_Syntax_node4(x_252, x_303, x_272, x_31, x_278, x_302); +x_305 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_306 = l_Lean_Syntax_node2(x_252, x_305, x_286, x_304); +lean_inc(x_306); +x_307 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_307, 0, x_306); +x_308 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_306, x_307, x_12, x_13, x_257); +return x_308; } } } } else { -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; size_t x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; -x_310 = lean_ctor_get(x_58, 1); -lean_inc(x_310); +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; size_t x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; +x_312 = lean_ctor_get(x_58, 1); +lean_inc(x_312); lean_dec(x_58); -x_311 = l_Lake_DSL_structVal___closed__3; +x_313 = l_Lake_DSL_structVal___closed__3; lean_inc(x_53); -x_312 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_312, 0, x_53); -lean_ctor_set(x_312, 1, x_311); -x_313 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_314 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +x_314 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_314, 0, x_53); +lean_ctor_set(x_314, 1, x_313); +x_315 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_316 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_53); -x_315 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_315, 0, x_53); -lean_ctor_set(x_315, 1, x_313); -lean_ctor_set(x_315, 2, x_314); -x_316 = lean_array_size(x_23); -x_317 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_316, x_17, x_23); -x_318 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_319 = l_Lean_mkSepArray(x_317, x_318); -lean_dec(x_317); -x_320 = l_Array_append___rarg(x_314, x_319); +x_317 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_317, 0, x_53); +lean_ctor_set(x_317, 1, x_315); +lean_ctor_set(x_317, 2, x_316); +x_318 = lean_array_size(x_23); +x_319 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_318, x_17, x_23); +x_320 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_321 = l_Lean_mkSepArray(x_319, x_320); lean_dec(x_319); +x_322 = l_Array_append___rarg(x_316, x_321); +lean_dec(x_321); lean_inc(x_53); -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_53); -lean_ctor_set(x_321, 1, x_313); -lean_ctor_set(x_321, 2, x_320); -x_322 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_315); +x_323 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_323, 0, x_53); +lean_ctor_set(x_323, 1, x_315); +lean_ctor_set(x_323, 2, x_322); +x_324 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; lean_inc(x_53); -x_323 = l_Lean_Syntax_node1(x_53, x_322, x_315); -x_324 = l_Lake_DSL_structVal___closed__18; +x_325 = l_Lean_Syntax_node1(x_53, x_324, x_323); +x_326 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_317); +lean_inc(x_53); +x_327 = l_Lean_Syntax_node1(x_53, x_326, x_317); +x_328 = l_Lake_DSL_structVal___closed__18; lean_inc(x_53); lean_ctor_set_tag(x_54, 2); -lean_ctor_set(x_54, 1, x_324); +lean_ctor_set(x_54, 1, x_328); lean_ctor_set(x_54, 0, x_53); -x_325 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_315); -x_326 = l_Lean_Syntax_node6(x_53, x_325, x_312, x_315, x_321, x_323, x_315, x_54); -x_327 = l_Lean_Elab_Command_getRef(x_12, x_13, x_310); -x_328 = lean_ctor_get(x_327, 0); -lean_inc(x_328); -x_329 = lean_ctor_get(x_327, 1); -lean_inc(x_329); -if (lean_is_exclusive(x_327)) { - lean_ctor_release(x_327, 0); - lean_ctor_release(x_327, 1); - x_330 = x_327; -} else { - lean_dec_ref(x_327); - x_330 = lean_box(0); -} -x_331 = l_Lean_SourceInfo_fromRef(x_328, x_29); -lean_dec(x_328); -x_332 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_329); -x_333 = lean_ctor_get(x_332, 1); +x_329 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_317); +x_330 = l_Lean_Syntax_node6(x_53, x_329, x_314, x_317, x_325, x_327, x_317, x_54); +x_331 = l_Lean_Elab_Command_getRef(x_12, x_13, x_312); +x_332 = lean_ctor_get(x_331, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_331, 1); lean_inc(x_333); -if (lean_is_exclusive(x_332)) { - lean_ctor_release(x_332, 0); - lean_ctor_release(x_332, 1); - x_334 = x_332; +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_334 = x_331; } else { - lean_dec_ref(x_332); + lean_dec_ref(x_331); x_334 = lean_box(0); } -x_335 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_333); -x_336 = lean_ctor_get(x_335, 1); -lean_inc(x_336); -if (lean_is_exclusive(x_335)) { - lean_ctor_release(x_335, 0); - lean_ctor_release(x_335, 1); - x_337 = x_335; +x_335 = l_Lean_SourceInfo_fromRef(x_332, x_29); +lean_dec(x_332); +x_336 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_333); +x_337 = lean_ctor_get(x_336, 1); +lean_inc(x_337); +if (lean_is_exclusive(x_336)) { + lean_ctor_release(x_336, 0); + lean_ctor_release(x_336, 1); + x_338 = x_336; } else { - lean_dec_ref(x_335); - x_337 = lean_box(0); -} -x_338 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_331); -if (lean_is_scalar(x_337)) { - x_339 = lean_alloc_ctor(2, 2, 0); -} else { - x_339 = x_337; - lean_ctor_set_tag(x_339, 2); -} -lean_ctor_set(x_339, 0, x_331); -lean_ctor_set(x_339, 1, x_338); -x_340 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_341 = l_Lean_Syntax_SepArray_ofElems(x_340, x_4); -x_342 = l_Array_append___rarg(x_314, x_341); -lean_dec(x_341); -lean_inc(x_331); -x_343 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_343, 0, x_331); -lean_ctor_set(x_343, 1, x_313); -lean_ctor_set(x_343, 2, x_342); -x_344 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_331); + lean_dec_ref(x_336); + x_338 = lean_box(0); +} +x_339 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_337); +x_340 = lean_ctor_get(x_339, 1); +lean_inc(x_340); +if (lean_is_exclusive(x_339)) { + lean_ctor_release(x_339, 0); + lean_ctor_release(x_339, 1); + x_341 = x_339; +} else { + lean_dec_ref(x_339); + x_341 = lean_box(0); +} +x_342 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_335); +if (lean_is_scalar(x_341)) { + x_343 = lean_alloc_ctor(2, 2, 0); +} else { + x_343 = x_341; + lean_ctor_set_tag(x_343, 2); +} +lean_ctor_set(x_343, 0, x_335); +lean_ctor_set(x_343, 1, x_342); +x_344 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_345 = l_Lean_Syntax_SepArray_ofElems(x_344, x_4); +x_346 = l_Array_append___rarg(x_316, x_345); +lean_dec(x_345); +lean_inc(x_335); +x_347 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_347, 0, x_335); +lean_ctor_set(x_347, 1, x_315); +lean_ctor_set(x_347, 2, x_346); +x_348 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_335); +if (lean_is_scalar(x_338)) { + x_349 = lean_alloc_ctor(2, 2, 0); +} else { + x_349 = x_338; + lean_ctor_set_tag(x_349, 2); +} +lean_ctor_set(x_349, 0, x_335); +lean_ctor_set(x_349, 1, x_348); +x_350 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_335); +x_351 = l_Lean_Syntax_node3(x_335, x_350, x_343, x_347, x_349); +lean_inc(x_335); +x_352 = l_Lean_Syntax_node1(x_335, x_315, x_351); +lean_inc(x_335); +x_353 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_353, 0, x_335); +lean_ctor_set(x_353, 1, x_315); +lean_ctor_set(x_353, 2, x_316); +x_354 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_335); if (lean_is_scalar(x_334)) { - x_345 = lean_alloc_ctor(2, 2, 0); -} else { - x_345 = x_334; - lean_ctor_set_tag(x_345, 2); -} -lean_ctor_set(x_345, 0, x_331); -lean_ctor_set(x_345, 1, x_344); -x_346 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_331); -x_347 = l_Lean_Syntax_node3(x_331, x_346, x_339, x_343, x_345); -lean_inc(x_331); -x_348 = l_Lean_Syntax_node1(x_331, x_313, x_347); -lean_inc(x_331); -x_349 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_349, 0, x_331); -lean_ctor_set(x_349, 1, x_313); -lean_ctor_set(x_349, 2, x_314); -x_350 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_331); -if (lean_is_scalar(x_330)) { - x_351 = lean_alloc_ctor(2, 2, 0); -} else { - x_351 = x_330; - lean_ctor_set_tag(x_351, 2); -} -lean_ctor_set(x_351, 0, x_331); -lean_ctor_set(x_351, 1, x_350); -x_352 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_331); + x_355 = lean_alloc_ctor(2, 2, 0); +} else { + x_355 = x_334; + lean_ctor_set_tag(x_355, 2); +} +lean_ctor_set(x_355, 0, x_335); +lean_ctor_set(x_355, 1, x_354); +x_356 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_335); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_352); -lean_ctor_set(x_49, 0, x_331); -x_353 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_331); -x_354 = l_Lean_Syntax_node2(x_331, x_353, x_49, x_30); -lean_inc(x_331); -x_355 = l_Lean_Syntax_node1(x_331, x_313, x_354); -x_356 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_349); -lean_inc(x_331); -x_357 = l_Lean_Syntax_node2(x_331, x_356, x_349, x_355); -x_358 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_331); +lean_ctor_set(x_49, 1, x_356); +lean_ctor_set(x_49, 0, x_335); +x_357 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_335); +x_358 = l_Lean_Syntax_node2(x_335, x_357, x_49, x_30); +lean_inc(x_335); +x_359 = l_Lean_Syntax_node1(x_335, x_315, x_358); +x_360 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_353); +lean_inc(x_335); +x_361 = l_Lean_Syntax_node2(x_335, x_360, x_353, x_359); +x_362 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_335); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_358); -lean_ctor_set(x_33, 0, x_331); -x_359 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_349, 2); -lean_inc(x_331); -x_360 = l_Lean_Syntax_node2(x_331, x_359, x_349, x_349); +lean_ctor_set(x_33, 1, x_362); +lean_ctor_set(x_33, 0, x_335); +x_363 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_353, 2); +lean_inc(x_335); +x_364 = l_Lean_Syntax_node2(x_335, x_363, x_353, x_353); if (lean_obj_tag(x_7) == 0) { -x_361 = x_314; -goto block_388; +x_365 = x_316; +goto block_392; } else { -lean_object* x_389; lean_object* x_390; -x_389 = lean_ctor_get(x_7, 0); -lean_inc(x_389); +lean_object* x_393; lean_object* x_394; +x_393 = lean_ctor_get(x_7, 0); +lean_inc(x_393); lean_dec(x_7); -x_390 = l_Array_mkArray1___rarg(x_389); -x_361 = x_390; -goto block_388; +x_394 = l_Array_mkArray1___rarg(x_393); +x_365 = x_394; +goto block_392; } -block_388: -{ -lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; -x_362 = l_Array_append___rarg(x_314, x_361); -lean_dec(x_361); -lean_inc(x_331); -x_363 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_363, 0, x_331); -lean_ctor_set(x_363, 1, x_313); -lean_ctor_set(x_363, 2, x_362); -x_364 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_349, 3); -lean_inc(x_331); -x_365 = l_Lean_Syntax_node6(x_331, x_364, x_363, x_348, x_349, x_349, x_349, x_349); -if (lean_obj_tag(x_5) == 0) +block_392: { -lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -x_366 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_331); +lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_366 = l_Array_append___rarg(x_316, x_365); +lean_dec(x_365); +lean_inc(x_335); x_367 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_367, 0, x_331); -lean_ctor_set(x_367, 1, x_313); +lean_ctor_set(x_367, 0, x_335); +lean_ctor_set(x_367, 1, x_315); lean_ctor_set(x_367, 2, x_366); -x_368 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_331); -x_369 = l_Lean_Syntax_node4(x_331, x_368, x_33, x_326, x_360, x_367); -x_370 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_331); -x_371 = l_Lean_Syntax_node4(x_331, x_370, x_351, x_31, x_357, x_369); -x_372 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_373 = l_Lean_Syntax_node2(x_331, x_372, x_365, x_371); -lean_inc(x_373); -x_374 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_374, 0, x_373); -x_375 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_373, x_374, x_12, x_13, x_336); -return x_375; +x_368 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_353, 3); +lean_inc(x_335); +x_369 = l_Lean_Syntax_node6(x_335, x_368, x_367, x_352, x_353, x_353, x_353, x_353); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; +x_370 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_335); +x_371 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_371, 0, x_335); +lean_ctor_set(x_371, 1, x_315); +lean_ctor_set(x_371, 2, x_370); +x_372 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_335); +x_373 = l_Lean_Syntax_node4(x_335, x_372, x_33, x_330, x_364, x_371); +x_374 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_335); +x_375 = l_Lean_Syntax_node4(x_335, x_374, x_355, x_31, x_361, x_373); +x_376 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_377 = l_Lean_Syntax_node2(x_335, x_376, x_369, x_375); +lean_inc(x_377); +x_378 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_378, 0, x_377); +x_379 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_377, x_378, x_12, x_13, x_340); +return x_379; } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; -x_376 = lean_ctor_get(x_5, 0); -lean_inc(x_376); +lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; +x_380 = lean_ctor_get(x_5, 0); +lean_inc(x_380); lean_dec(x_5); -x_377 = l_Array_mkArray1___rarg(x_376); -x_378 = l_Array_append___rarg(x_314, x_377); -lean_dec(x_377); -lean_inc(x_331); -x_379 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_379, 0, x_331); -lean_ctor_set(x_379, 1, x_313); -lean_ctor_set(x_379, 2, x_378); -x_380 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_331); -x_381 = l_Lean_Syntax_node4(x_331, x_380, x_33, x_326, x_360, x_379); -x_382 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_331); -x_383 = l_Lean_Syntax_node4(x_331, x_382, x_351, x_31, x_357, x_381); -x_384 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_385 = l_Lean_Syntax_node2(x_331, x_384, x_365, x_383); -lean_inc(x_385); -x_386 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_386, 0, x_385); -x_387 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_385, x_386, x_12, x_13, x_336); -return x_387; -} -} -} -} -else -{ -lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; size_t x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_391 = lean_ctor_get(x_54, 1); -lean_inc(x_391); +x_381 = l_Array_mkArray1___rarg(x_380); +x_382 = l_Array_append___rarg(x_316, x_381); +lean_dec(x_381); +lean_inc(x_335); +x_383 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_383, 0, x_335); +lean_ctor_set(x_383, 1, x_315); +lean_ctor_set(x_383, 2, x_382); +x_384 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_335); +x_385 = l_Lean_Syntax_node4(x_335, x_384, x_33, x_330, x_364, x_383); +x_386 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_335); +x_387 = l_Lean_Syntax_node4(x_335, x_386, x_355, x_31, x_361, x_385); +x_388 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_389 = l_Lean_Syntax_node2(x_335, x_388, x_369, x_387); +lean_inc(x_389); +x_390 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_390, 0, x_389); +x_391 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_389, x_390, x_12, x_13, x_340); +return x_391; +} +} +} +} +else +{ +lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; size_t x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; +x_395 = lean_ctor_get(x_54, 1); +lean_inc(x_395); lean_dec(x_54); -x_392 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_391); -x_393 = lean_ctor_get(x_392, 1); -lean_inc(x_393); -if (lean_is_exclusive(x_392)) { - lean_ctor_release(x_392, 0); - lean_ctor_release(x_392, 1); - x_394 = x_392; +x_396 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_395); +x_397 = lean_ctor_get(x_396, 1); +lean_inc(x_397); +if (lean_is_exclusive(x_396)) { + lean_ctor_release(x_396, 0); + lean_ctor_release(x_396, 1); + x_398 = x_396; } else { - lean_dec_ref(x_392); - x_394 = lean_box(0); + lean_dec_ref(x_396); + x_398 = lean_box(0); } -x_395 = l_Lake_DSL_structVal___closed__3; +x_399 = l_Lake_DSL_structVal___closed__3; lean_inc(x_53); -if (lean_is_scalar(x_394)) { - x_396 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_398)) { + x_400 = lean_alloc_ctor(2, 2, 0); } else { - x_396 = x_394; - lean_ctor_set_tag(x_396, 2); + x_400 = x_398; + lean_ctor_set_tag(x_400, 2); } -lean_ctor_set(x_396, 0, x_53); -lean_ctor_set(x_396, 1, x_395); -x_397 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_398 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_53); -x_399 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_399, 0, x_53); -lean_ctor_set(x_399, 1, x_397); -lean_ctor_set(x_399, 2, x_398); -x_400 = lean_array_size(x_23); -x_401 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_400, x_17, x_23); -x_402 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_403 = l_Lean_mkSepArray(x_401, x_402); -lean_dec(x_401); -x_404 = l_Array_append___rarg(x_398, x_403); -lean_dec(x_403); +lean_ctor_set(x_400, 0, x_53); +lean_ctor_set(x_400, 1, x_399); +x_401 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_402 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_53); -x_405 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_405, 0, x_53); -lean_ctor_set(x_405, 1, x_397); -lean_ctor_set(x_405, 2, x_404); +x_403 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_403, 0, x_53); +lean_ctor_set(x_403, 1, x_401); +lean_ctor_set(x_403, 2, x_402); +x_404 = lean_array_size(x_23); +x_405 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_404, x_17, x_23); x_406 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_399); -lean_inc(x_53); -x_407 = l_Lean_Syntax_node1(x_53, x_406, x_399); -x_408 = l_Lake_DSL_structVal___closed__18; +x_407 = l_Lean_mkSepArray(x_405, x_406); +lean_dec(x_405); +x_408 = l_Array_append___rarg(x_402, x_407); +lean_dec(x_407); lean_inc(x_53); -x_409 = lean_alloc_ctor(2, 2, 0); +x_409 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_409, 0, x_53); -lean_ctor_set(x_409, 1, x_408); -x_410 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_399); -x_411 = l_Lean_Syntax_node6(x_53, x_410, x_396, x_399, x_405, x_407, x_399, x_409); -x_412 = l_Lean_Elab_Command_getRef(x_12, x_13, x_393); -x_413 = lean_ctor_get(x_412, 0); -lean_inc(x_413); -x_414 = lean_ctor_get(x_412, 1); -lean_inc(x_414); -if (lean_is_exclusive(x_412)) { - lean_ctor_release(x_412, 0); - lean_ctor_release(x_412, 1); - x_415 = x_412; +lean_ctor_set(x_409, 1, x_401); +lean_ctor_set(x_409, 2, x_408); +x_410 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_53); +x_411 = l_Lean_Syntax_node1(x_53, x_410, x_409); +x_412 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_403); +lean_inc(x_53); +x_413 = l_Lean_Syntax_node1(x_53, x_412, x_403); +x_414 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_53); +x_415 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_415, 0, x_53); +lean_ctor_set(x_415, 1, x_414); +x_416 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_403); +x_417 = l_Lean_Syntax_node6(x_53, x_416, x_400, x_403, x_411, x_413, x_403, x_415); +x_418 = l_Lean_Elab_Command_getRef(x_12, x_13, x_397); +x_419 = lean_ctor_get(x_418, 0); +lean_inc(x_419); +x_420 = lean_ctor_get(x_418, 1); +lean_inc(x_420); +if (lean_is_exclusive(x_418)) { + lean_ctor_release(x_418, 0); + lean_ctor_release(x_418, 1); + x_421 = x_418; } else { - lean_dec_ref(x_412); - x_415 = lean_box(0); + lean_dec_ref(x_418); + x_421 = lean_box(0); } -x_416 = l_Lean_SourceInfo_fromRef(x_413, x_29); -lean_dec(x_413); -x_417 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_414); -x_418 = lean_ctor_get(x_417, 1); -lean_inc(x_418); -if (lean_is_exclusive(x_417)) { - lean_ctor_release(x_417, 0); - lean_ctor_release(x_417, 1); - x_419 = x_417; -} else { - lean_dec_ref(x_417); - x_419 = lean_box(0); -} -x_420 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_418); -x_421 = lean_ctor_get(x_420, 1); -lean_inc(x_421); -if (lean_is_exclusive(x_420)) { - lean_ctor_release(x_420, 0); - lean_ctor_release(x_420, 1); - x_422 = x_420; -} else { - lean_dec_ref(x_420); - x_422 = lean_box(0); -} -x_423 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_416); -if (lean_is_scalar(x_422)) { - x_424 = lean_alloc_ctor(2, 2, 0); -} else { - x_424 = x_422; - lean_ctor_set_tag(x_424, 2); -} -lean_ctor_set(x_424, 0, x_416); -lean_ctor_set(x_424, 1, x_423); -x_425 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_426 = l_Lean_Syntax_SepArray_ofElems(x_425, x_4); -x_427 = l_Array_append___rarg(x_398, x_426); -lean_dec(x_426); -lean_inc(x_416); -x_428 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_428, 0, x_416); -lean_ctor_set(x_428, 1, x_397); -lean_ctor_set(x_428, 2, x_427); -x_429 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_416); -if (lean_is_scalar(x_419)) { +x_422 = l_Lean_SourceInfo_fromRef(x_419, x_29); +lean_dec(x_419); +x_423 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_420); +x_424 = lean_ctor_get(x_423, 1); +lean_inc(x_424); +if (lean_is_exclusive(x_423)) { + lean_ctor_release(x_423, 0); + lean_ctor_release(x_423, 1); + x_425 = x_423; +} else { + lean_dec_ref(x_423); + x_425 = lean_box(0); +} +x_426 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_424); +x_427 = lean_ctor_get(x_426, 1); +lean_inc(x_427); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + x_428 = x_426; +} else { + lean_dec_ref(x_426); + x_428 = lean_box(0); +} +x_429 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_422); +if (lean_is_scalar(x_428)) { x_430 = lean_alloc_ctor(2, 2, 0); } else { - x_430 = x_419; + x_430 = x_428; lean_ctor_set_tag(x_430, 2); } -lean_ctor_set(x_430, 0, x_416); +lean_ctor_set(x_430, 0, x_422); lean_ctor_set(x_430, 1, x_429); -x_431 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_416); -x_432 = l_Lean_Syntax_node3(x_416, x_431, x_424, x_428, x_430); -lean_inc(x_416); -x_433 = l_Lean_Syntax_node1(x_416, x_397, x_432); -lean_inc(x_416); +x_431 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_432 = l_Lean_Syntax_SepArray_ofElems(x_431, x_4); +x_433 = l_Array_append___rarg(x_402, x_432); +lean_dec(x_432); +lean_inc(x_422); x_434 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_434, 0, x_416); -lean_ctor_set(x_434, 1, x_397); -lean_ctor_set(x_434, 2, x_398); -x_435 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_416); -if (lean_is_scalar(x_415)) { +lean_ctor_set(x_434, 0, x_422); +lean_ctor_set(x_434, 1, x_401); +lean_ctor_set(x_434, 2, x_433); +x_435 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_422); +if (lean_is_scalar(x_425)) { x_436 = lean_alloc_ctor(2, 2, 0); } else { - x_436 = x_415; + x_436 = x_425; lean_ctor_set_tag(x_436, 2); } -lean_ctor_set(x_436, 0, x_416); +lean_ctor_set(x_436, 0, x_422); lean_ctor_set(x_436, 1, x_435); -x_437 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_416); +x_437 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_422); +x_438 = l_Lean_Syntax_node3(x_422, x_437, x_430, x_434, x_436); +lean_inc(x_422); +x_439 = l_Lean_Syntax_node1(x_422, x_401, x_438); +lean_inc(x_422); +x_440 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_440, 0, x_422); +lean_ctor_set(x_440, 1, x_401); +lean_ctor_set(x_440, 2, x_402); +x_441 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_422); +if (lean_is_scalar(x_421)) { + x_442 = lean_alloc_ctor(2, 2, 0); +} else { + x_442 = x_421; + lean_ctor_set_tag(x_442, 2); +} +lean_ctor_set(x_442, 0, x_422); +lean_ctor_set(x_442, 1, x_441); +x_443 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_422); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_437); -lean_ctor_set(x_49, 0, x_416); -x_438 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_416); -x_439 = l_Lean_Syntax_node2(x_416, x_438, x_49, x_30); -lean_inc(x_416); -x_440 = l_Lean_Syntax_node1(x_416, x_397, x_439); -x_441 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_434); -lean_inc(x_416); -x_442 = l_Lean_Syntax_node2(x_416, x_441, x_434, x_440); -x_443 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_416); +lean_ctor_set(x_49, 1, x_443); +lean_ctor_set(x_49, 0, x_422); +x_444 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_422); +x_445 = l_Lean_Syntax_node2(x_422, x_444, x_49, x_30); +lean_inc(x_422); +x_446 = l_Lean_Syntax_node1(x_422, x_401, x_445); +x_447 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_440); +lean_inc(x_422); +x_448 = l_Lean_Syntax_node2(x_422, x_447, x_440, x_446); +x_449 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_422); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_443); -lean_ctor_set(x_33, 0, x_416); -x_444 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_434, 2); -lean_inc(x_416); -x_445 = l_Lean_Syntax_node2(x_416, x_444, x_434, x_434); +lean_ctor_set(x_33, 1, x_449); +lean_ctor_set(x_33, 0, x_422); +x_450 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_440, 2); +lean_inc(x_422); +x_451 = l_Lean_Syntax_node2(x_422, x_450, x_440, x_440); if (lean_obj_tag(x_7) == 0) { -x_446 = x_398; -goto block_473; +x_452 = x_402; +goto block_479; } else { -lean_object* x_474; lean_object* x_475; -x_474 = lean_ctor_get(x_7, 0); -lean_inc(x_474); +lean_object* x_480; lean_object* x_481; +x_480 = lean_ctor_get(x_7, 0); +lean_inc(x_480); lean_dec(x_7); -x_475 = l_Array_mkArray1___rarg(x_474); -x_446 = x_475; -goto block_473; -} -block_473: -{ -lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; -x_447 = l_Array_append___rarg(x_398, x_446); -lean_dec(x_446); -lean_inc(x_416); -x_448 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_448, 0, x_416); -lean_ctor_set(x_448, 1, x_397); -lean_ctor_set(x_448, 2, x_447); -x_449 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_434, 3); -lean_inc(x_416); -x_450 = l_Lean_Syntax_node6(x_416, x_449, x_448, x_433, x_434, x_434, x_434, x_434); +x_481 = l_Array_mkArray1___rarg(x_480); +x_452 = x_481; +goto block_479; +} +block_479: +{ +lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; +x_453 = l_Array_append___rarg(x_402, x_452); +lean_dec(x_452); +lean_inc(x_422); +x_454 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_454, 0, x_422); +lean_ctor_set(x_454, 1, x_401); +lean_ctor_set(x_454, 2, x_453); +x_455 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_440, 3); +lean_inc(x_422); +x_456 = l_Lean_Syntax_node6(x_422, x_455, x_454, x_439, x_440, x_440, x_440, x_440); if (lean_obj_tag(x_5) == 0) { -lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; -x_451 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_416); -x_452 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_452, 0, x_416); -lean_ctor_set(x_452, 1, x_397); -lean_ctor_set(x_452, 2, x_451); -x_453 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_416); -x_454 = l_Lean_Syntax_node4(x_416, x_453, x_33, x_411, x_445, x_452); -x_455 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_416); -x_456 = l_Lean_Syntax_node4(x_416, x_455, x_436, x_31, x_442, x_454); -x_457 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_458 = l_Lean_Syntax_node2(x_416, x_457, x_450, x_456); -lean_inc(x_458); -x_459 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_459, 0, x_458); -x_460 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_458, x_459, x_12, x_13, x_421); -return x_460; -} -else -{ -lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; -x_461 = lean_ctor_get(x_5, 0); -lean_inc(x_461); +lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; +x_457 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_422); +x_458 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_458, 0, x_422); +lean_ctor_set(x_458, 1, x_401); +lean_ctor_set(x_458, 2, x_457); +x_459 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_422); +x_460 = l_Lean_Syntax_node4(x_422, x_459, x_33, x_417, x_451, x_458); +x_461 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_422); +x_462 = l_Lean_Syntax_node4(x_422, x_461, x_442, x_31, x_448, x_460); +x_463 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_464 = l_Lean_Syntax_node2(x_422, x_463, x_456, x_462); +lean_inc(x_464); +x_465 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_465, 0, x_464); +x_466 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_464, x_465, x_12, x_13, x_427); +return x_466; +} +else +{ +lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; +x_467 = lean_ctor_get(x_5, 0); +lean_inc(x_467); lean_dec(x_5); -x_462 = l_Array_mkArray1___rarg(x_461); -x_463 = l_Array_append___rarg(x_398, x_462); -lean_dec(x_462); -lean_inc(x_416); -x_464 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_464, 0, x_416); -lean_ctor_set(x_464, 1, x_397); -lean_ctor_set(x_464, 2, x_463); -x_465 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_416); -x_466 = l_Lean_Syntax_node4(x_416, x_465, x_33, x_411, x_445, x_464); -x_467 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_416); -x_468 = l_Lean_Syntax_node4(x_416, x_467, x_436, x_31, x_442, x_466); -x_469 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_470 = l_Lean_Syntax_node2(x_416, x_469, x_450, x_468); -lean_inc(x_470); -x_471 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_471, 0, x_470); -x_472 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_470, x_471, x_12, x_13, x_421); -return x_472; -} -} -} -} -else -{ -lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; size_t x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; -x_476 = lean_ctor_get(x_49, 0); -x_477 = lean_ctor_get(x_49, 1); -lean_inc(x_477); +x_468 = l_Array_mkArray1___rarg(x_467); +x_469 = l_Array_append___rarg(x_402, x_468); +lean_dec(x_468); +lean_inc(x_422); +x_470 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_470, 0, x_422); +lean_ctor_set(x_470, 1, x_401); +lean_ctor_set(x_470, 2, x_469); +x_471 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_422); +x_472 = l_Lean_Syntax_node4(x_422, x_471, x_33, x_417, x_451, x_470); +x_473 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_422); +x_474 = l_Lean_Syntax_node4(x_422, x_473, x_442, x_31, x_448, x_472); +x_475 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_476 = l_Lean_Syntax_node2(x_422, x_475, x_456, x_474); lean_inc(x_476); -lean_dec(x_49); -x_478 = l_Lean_SourceInfo_fromRef(x_476, x_29); -lean_dec(x_476); -x_479 = l_Lean_Elab_Command_getCurrMacroScope(x_48, x_13, x_477); -lean_dec(x_48); -x_480 = lean_ctor_get(x_479, 1); -lean_inc(x_480); -if (lean_is_exclusive(x_479)) { - lean_ctor_release(x_479, 0); - lean_ctor_release(x_479, 1); - x_481 = x_479; -} else { - lean_dec_ref(x_479); - x_481 = lean_box(0); +x_477 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_477, 0, x_476); +x_478 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_476, x_477, x_12, x_13, x_427); +return x_478; } -x_482 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_480); -x_483 = lean_ctor_get(x_482, 1); +} +} +} +else +{ +lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; size_t x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; +x_482 = lean_ctor_get(x_49, 0); +x_483 = lean_ctor_get(x_49, 1); lean_inc(x_483); -if (lean_is_exclusive(x_482)) { - lean_ctor_release(x_482, 0); - lean_ctor_release(x_482, 1); - x_484 = x_482; +lean_inc(x_482); +lean_dec(x_49); +x_484 = l_Lean_SourceInfo_fromRef(x_482, x_29); +lean_dec(x_482); +x_485 = l_Lean_Elab_Command_getCurrMacroScope(x_48, x_13, x_483); +lean_dec(x_48); +x_486 = lean_ctor_get(x_485, 1); +lean_inc(x_486); +if (lean_is_exclusive(x_485)) { + lean_ctor_release(x_485, 0); + lean_ctor_release(x_485, 1); + x_487 = x_485; +} else { + lean_dec_ref(x_485); + x_487 = lean_box(0); +} +x_488 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_486); +x_489 = lean_ctor_get(x_488, 1); +lean_inc(x_489); +if (lean_is_exclusive(x_488)) { + lean_ctor_release(x_488, 0); + lean_ctor_release(x_488, 1); + x_490 = x_488; } else { - lean_dec_ref(x_482); - x_484 = lean_box(0); + lean_dec_ref(x_488); + x_490 = lean_box(0); } -x_485 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_478); -if (lean_is_scalar(x_484)) { - x_486 = lean_alloc_ctor(2, 2, 0); +x_491 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_484); +if (lean_is_scalar(x_490)) { + x_492 = lean_alloc_ctor(2, 2, 0); } else { - x_486 = x_484; - lean_ctor_set_tag(x_486, 2); + x_492 = x_490; + lean_ctor_set_tag(x_492, 2); } -lean_ctor_set(x_486, 0, x_478); -lean_ctor_set(x_486, 1, x_485); -x_487 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_488 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_478); -x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_478); -lean_ctor_set(x_489, 1, x_487); -lean_ctor_set(x_489, 2, x_488); -x_490 = lean_array_size(x_23); -x_491 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_490, x_17, x_23); -x_492 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_493 = l_Lean_mkSepArray(x_491, x_492); -lean_dec(x_491); -x_494 = l_Array_append___rarg(x_488, x_493); -lean_dec(x_493); -lean_inc(x_478); +lean_ctor_set(x_492, 0, x_484); +lean_ctor_set(x_492, 1, x_491); +x_493 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_494 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_484); x_495 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_495, 0, x_478); -lean_ctor_set(x_495, 1, x_487); +lean_ctor_set(x_495, 0, x_484); +lean_ctor_set(x_495, 1, x_493); lean_ctor_set(x_495, 2, x_494); -x_496 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_489); -lean_inc(x_478); -x_497 = l_Lean_Syntax_node1(x_478, x_496, x_489); -x_498 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_478); -if (lean_is_scalar(x_481)) { - x_499 = lean_alloc_ctor(2, 2, 0); -} else { - x_499 = x_481; - lean_ctor_set_tag(x_499, 2); -} -lean_ctor_set(x_499, 0, x_478); -lean_ctor_set(x_499, 1, x_498); -x_500 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_489); -x_501 = l_Lean_Syntax_node6(x_478, x_500, x_486, x_489, x_495, x_497, x_489, x_499); -x_502 = l_Lean_Elab_Command_getRef(x_12, x_13, x_483); -x_503 = lean_ctor_get(x_502, 0); -lean_inc(x_503); -x_504 = lean_ctor_get(x_502, 1); -lean_inc(x_504); -if (lean_is_exclusive(x_502)) { - lean_ctor_release(x_502, 0); - lean_ctor_release(x_502, 1); - x_505 = x_502; -} else { - lean_dec_ref(x_502); - x_505 = lean_box(0); -} -x_506 = l_Lean_SourceInfo_fromRef(x_503, x_29); -lean_dec(x_503); -x_507 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_504); -x_508 = lean_ctor_get(x_507, 1); -lean_inc(x_508); -if (lean_is_exclusive(x_507)) { - lean_ctor_release(x_507, 0); - lean_ctor_release(x_507, 1); - x_509 = x_507; -} else { - lean_dec_ref(x_507); - x_509 = lean_box(0); -} -x_510 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_508); -x_511 = lean_ctor_get(x_510, 1); +x_496 = lean_array_size(x_23); +x_497 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_496, x_17, x_23); +x_498 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_499 = l_Lean_mkSepArray(x_497, x_498); +lean_dec(x_497); +x_500 = l_Array_append___rarg(x_494, x_499); +lean_dec(x_499); +lean_inc(x_484); +x_501 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_501, 0, x_484); +lean_ctor_set(x_501, 1, x_493); +lean_ctor_set(x_501, 2, x_500); +x_502 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_484); +x_503 = l_Lean_Syntax_node1(x_484, x_502, x_501); +x_504 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_495); +lean_inc(x_484); +x_505 = l_Lean_Syntax_node1(x_484, x_504, x_495); +x_506 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_484); +if (lean_is_scalar(x_487)) { + x_507 = lean_alloc_ctor(2, 2, 0); +} else { + x_507 = x_487; + lean_ctor_set_tag(x_507, 2); +} +lean_ctor_set(x_507, 0, x_484); +lean_ctor_set(x_507, 1, x_506); +x_508 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_495); +x_509 = l_Lean_Syntax_node6(x_484, x_508, x_492, x_495, x_503, x_505, x_495, x_507); +x_510 = l_Lean_Elab_Command_getRef(x_12, x_13, x_489); +x_511 = lean_ctor_get(x_510, 0); lean_inc(x_511); +x_512 = lean_ctor_get(x_510, 1); +lean_inc(x_512); if (lean_is_exclusive(x_510)) { lean_ctor_release(x_510, 0); lean_ctor_release(x_510, 1); - x_512 = x_510; + x_513 = x_510; } else { lean_dec_ref(x_510); - x_512 = lean_box(0); + x_513 = lean_box(0); } -x_513 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_506); -if (lean_is_scalar(x_512)) { - x_514 = lean_alloc_ctor(2, 2, 0); -} else { - x_514 = x_512; - lean_ctor_set_tag(x_514, 2); -} -lean_ctor_set(x_514, 0, x_506); -lean_ctor_set(x_514, 1, x_513); -x_515 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_516 = l_Lean_Syntax_SepArray_ofElems(x_515, x_4); -x_517 = l_Array_append___rarg(x_488, x_516); -lean_dec(x_516); -lean_inc(x_506); -x_518 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_518, 0, x_506); -lean_ctor_set(x_518, 1, x_487); -lean_ctor_set(x_518, 2, x_517); -x_519 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_506); -if (lean_is_scalar(x_509)) { - x_520 = lean_alloc_ctor(2, 2, 0); +x_514 = l_Lean_SourceInfo_fromRef(x_511, x_29); +lean_dec(x_511); +x_515 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_512); +x_516 = lean_ctor_get(x_515, 1); +lean_inc(x_516); +if (lean_is_exclusive(x_515)) { + lean_ctor_release(x_515, 0); + lean_ctor_release(x_515, 1); + x_517 = x_515; } else { - x_520 = x_509; - lean_ctor_set_tag(x_520, 2); + lean_dec_ref(x_515); + x_517 = lean_box(0); } -lean_ctor_set(x_520, 0, x_506); -lean_ctor_set(x_520, 1, x_519); -x_521 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_506); -x_522 = l_Lean_Syntax_node3(x_506, x_521, x_514, x_518, x_520); -lean_inc(x_506); -x_523 = l_Lean_Syntax_node1(x_506, x_487, x_522); -lean_inc(x_506); -x_524 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_524, 0, x_506); -lean_ctor_set(x_524, 1, x_487); -lean_ctor_set(x_524, 2, x_488); -x_525 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_506); -if (lean_is_scalar(x_505)) { - x_526 = lean_alloc_ctor(2, 2, 0); +x_518 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_516); +x_519 = lean_ctor_get(x_518, 1); +lean_inc(x_519); +if (lean_is_exclusive(x_518)) { + lean_ctor_release(x_518, 0); + lean_ctor_release(x_518, 1); + x_520 = x_518; } else { - x_526 = x_505; - lean_ctor_set_tag(x_526, 2); + lean_dec_ref(x_518); + x_520 = lean_box(0); } -lean_ctor_set(x_526, 0, x_506); -lean_ctor_set(x_526, 1, x_525); -x_527 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_506); -x_528 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_528, 0, x_506); +x_521 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_514); +if (lean_is_scalar(x_520)) { + x_522 = lean_alloc_ctor(2, 2, 0); +} else { + x_522 = x_520; + lean_ctor_set_tag(x_522, 2); +} +lean_ctor_set(x_522, 0, x_514); +lean_ctor_set(x_522, 1, x_521); +x_523 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_524 = l_Lean_Syntax_SepArray_ofElems(x_523, x_4); +x_525 = l_Array_append___rarg(x_494, x_524); +lean_dec(x_524); +lean_inc(x_514); +x_526 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_526, 0, x_514); +lean_ctor_set(x_526, 1, x_493); +lean_ctor_set(x_526, 2, x_525); +x_527 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_514); +if (lean_is_scalar(x_517)) { + x_528 = lean_alloc_ctor(2, 2, 0); +} else { + x_528 = x_517; + lean_ctor_set_tag(x_528, 2); +} +lean_ctor_set(x_528, 0, x_514); lean_ctor_set(x_528, 1, x_527); -x_529 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_506); -x_530 = l_Lean_Syntax_node2(x_506, x_529, x_528, x_30); -lean_inc(x_506); -x_531 = l_Lean_Syntax_node1(x_506, x_487, x_530); -x_532 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_524); -lean_inc(x_506); -x_533 = l_Lean_Syntax_node2(x_506, x_532, x_524, x_531); -x_534 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_506); +x_529 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_514); +x_530 = l_Lean_Syntax_node3(x_514, x_529, x_522, x_526, x_528); +lean_inc(x_514); +x_531 = l_Lean_Syntax_node1(x_514, x_493, x_530); +lean_inc(x_514); +x_532 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_532, 0, x_514); +lean_ctor_set(x_532, 1, x_493); +lean_ctor_set(x_532, 2, x_494); +x_533 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_514); +if (lean_is_scalar(x_513)) { + x_534 = lean_alloc_ctor(2, 2, 0); +} else { + x_534 = x_513; + lean_ctor_set_tag(x_534, 2); +} +lean_ctor_set(x_534, 0, x_514); +lean_ctor_set(x_534, 1, x_533); +x_535 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_514); +x_536 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_536, 0, x_514); +lean_ctor_set(x_536, 1, x_535); +x_537 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_514); +x_538 = l_Lean_Syntax_node2(x_514, x_537, x_536, x_30); +lean_inc(x_514); +x_539 = l_Lean_Syntax_node1(x_514, x_493, x_538); +x_540 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_532); +lean_inc(x_514); +x_541 = l_Lean_Syntax_node2(x_514, x_540, x_532, x_539); +x_542 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_514); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_534); -lean_ctor_set(x_33, 0, x_506); -x_535 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_524, 2); -lean_inc(x_506); -x_536 = l_Lean_Syntax_node2(x_506, x_535, x_524, x_524); +lean_ctor_set(x_33, 1, x_542); +lean_ctor_set(x_33, 0, x_514); +x_543 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_532, 2); +lean_inc(x_514); +x_544 = l_Lean_Syntax_node2(x_514, x_543, x_532, x_532); if (lean_obj_tag(x_7) == 0) { -x_537 = x_488; -goto block_564; +x_545 = x_494; +goto block_572; } else { -lean_object* x_565; lean_object* x_566; -x_565 = lean_ctor_get(x_7, 0); -lean_inc(x_565); +lean_object* x_573; lean_object* x_574; +x_573 = lean_ctor_get(x_7, 0); +lean_inc(x_573); lean_dec(x_7); -x_566 = l_Array_mkArray1___rarg(x_565); -x_537 = x_566; -goto block_564; +x_574 = l_Array_mkArray1___rarg(x_573); +x_545 = x_574; +goto block_572; } -block_564: +block_572: { -lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; -x_538 = l_Array_append___rarg(x_488, x_537); -lean_dec(x_537); -lean_inc(x_506); -x_539 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_539, 0, x_506); -lean_ctor_set(x_539, 1, x_487); -lean_ctor_set(x_539, 2, x_538); -x_540 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_524, 3); -lean_inc(x_506); -x_541 = l_Lean_Syntax_node6(x_506, x_540, x_539, x_523, x_524, x_524, x_524, x_524); +lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; +x_546 = l_Array_append___rarg(x_494, x_545); +lean_dec(x_545); +lean_inc(x_514); +x_547 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_547, 0, x_514); +lean_ctor_set(x_547, 1, x_493); +lean_ctor_set(x_547, 2, x_546); +x_548 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_532, 3); +lean_inc(x_514); +x_549 = l_Lean_Syntax_node6(x_514, x_548, x_547, x_531, x_532, x_532, x_532, x_532); if (lean_obj_tag(x_5) == 0) { -lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; -x_542 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_506); -x_543 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_543, 0, x_506); -lean_ctor_set(x_543, 1, x_487); -lean_ctor_set(x_543, 2, x_542); -x_544 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_506); -x_545 = l_Lean_Syntax_node4(x_506, x_544, x_33, x_501, x_536, x_543); -x_546 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_506); -x_547 = l_Lean_Syntax_node4(x_506, x_546, x_526, x_31, x_533, x_545); -x_548 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_549 = l_Lean_Syntax_node2(x_506, x_548, x_541, x_547); -lean_inc(x_549); -x_550 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_550, 0, x_549); -x_551 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_549, x_550, x_12, x_13, x_511); -return x_551; +lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; +x_550 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_514); +x_551 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_551, 0, x_514); +lean_ctor_set(x_551, 1, x_493); +lean_ctor_set(x_551, 2, x_550); +x_552 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_514); +x_553 = l_Lean_Syntax_node4(x_514, x_552, x_33, x_509, x_544, x_551); +x_554 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_514); +x_555 = l_Lean_Syntax_node4(x_514, x_554, x_534, x_31, x_541, x_553); +x_556 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_557 = l_Lean_Syntax_node2(x_514, x_556, x_549, x_555); +lean_inc(x_557); +x_558 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_558, 0, x_557); +x_559 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_557, x_558, x_12, x_13, x_519); +return x_559; } else { -lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; -x_552 = lean_ctor_get(x_5, 0); -lean_inc(x_552); +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; +x_560 = lean_ctor_get(x_5, 0); +lean_inc(x_560); lean_dec(x_5); -x_553 = l_Array_mkArray1___rarg(x_552); -x_554 = l_Array_append___rarg(x_488, x_553); -lean_dec(x_553); -lean_inc(x_506); -x_555 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_555, 0, x_506); -lean_ctor_set(x_555, 1, x_487); -lean_ctor_set(x_555, 2, x_554); -x_556 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_506); -x_557 = l_Lean_Syntax_node4(x_506, x_556, x_33, x_501, x_536, x_555); -x_558 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_506); -x_559 = l_Lean_Syntax_node4(x_506, x_558, x_526, x_31, x_533, x_557); -x_560 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_561 = l_Lean_Syntax_node2(x_506, x_560, x_541, x_559); -lean_inc(x_561); -x_562 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_562, 0, x_561); -x_563 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_561, x_562, x_12, x_13, x_511); -return x_563; +x_561 = l_Array_mkArray1___rarg(x_560); +x_562 = l_Array_append___rarg(x_494, x_561); +lean_dec(x_561); +lean_inc(x_514); +x_563 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_563, 0, x_514); +lean_ctor_set(x_563, 1, x_493); +lean_ctor_set(x_563, 2, x_562); +x_564 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_514); +x_565 = l_Lean_Syntax_node4(x_514, x_564, x_33, x_509, x_544, x_563); +x_566 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_514); +x_567 = l_Lean_Syntax_node4(x_514, x_566, x_534, x_31, x_541, x_565); +x_568 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_569 = l_Lean_Syntax_node2(x_514, x_568, x_549, x_567); +lean_inc(x_569); +x_570 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_570, 0, x_569); +x_571 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_569, x_570, x_12, x_13, x_519); +return x_571; } } } } else { -lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; uint8_t x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; size_t x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; -x_567 = lean_ctor_get(x_33, 0); -x_568 = lean_ctor_get(x_33, 1); -lean_inc(x_568); -lean_inc(x_567); -lean_dec(x_33); -x_569 = l_Lean_replaceRef(x_3, x_567); -lean_dec(x_567); -x_570 = lean_ctor_get(x_12, 0); -lean_inc(x_570); -x_571 = lean_ctor_get(x_12, 1); -lean_inc(x_571); -x_572 = lean_ctor_get(x_12, 2); -lean_inc(x_572); -x_573 = lean_ctor_get(x_12, 3); -lean_inc(x_573); -x_574 = lean_ctor_get(x_12, 4); -lean_inc(x_574); -x_575 = lean_ctor_get(x_12, 5); -lean_inc(x_575); -x_576 = lean_ctor_get(x_12, 7); +lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; uint8_t x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; size_t x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; +x_575 = lean_ctor_get(x_33, 0); +x_576 = lean_ctor_get(x_33, 1); lean_inc(x_576); -x_577 = lean_ctor_get(x_12, 8); -lean_inc(x_577); -x_578 = lean_ctor_get(x_12, 9); +lean_inc(x_575); +lean_dec(x_33); +x_577 = l_Lean_replaceRef(x_3, x_575); +lean_dec(x_575); +x_578 = lean_ctor_get(x_12, 0); lean_inc(x_578); -x_579 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); -x_580 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_580, 0, x_570); -lean_ctor_set(x_580, 1, x_571); -lean_ctor_set(x_580, 2, x_572); -lean_ctor_set(x_580, 3, x_573); -lean_ctor_set(x_580, 4, x_574); -lean_ctor_set(x_580, 5, x_575); -lean_ctor_set(x_580, 6, x_569); -lean_ctor_set(x_580, 7, x_576); -lean_ctor_set(x_580, 8, x_577); -lean_ctor_set(x_580, 9, x_578); -lean_ctor_set_uint8(x_580, sizeof(void*)*10, x_579); -x_581 = l_Lean_Elab_Command_getRef(x_580, x_13, x_568); -x_582 = lean_ctor_get(x_581, 0); +x_579 = lean_ctor_get(x_12, 1); +lean_inc(x_579); +x_580 = lean_ctor_get(x_12, 2); +lean_inc(x_580); +x_581 = lean_ctor_get(x_12, 3); +lean_inc(x_581); +x_582 = lean_ctor_get(x_12, 4); lean_inc(x_582); -x_583 = lean_ctor_get(x_581, 1); +x_583 = lean_ctor_get(x_12, 5); lean_inc(x_583); -if (lean_is_exclusive(x_581)) { - lean_ctor_release(x_581, 0); - lean_ctor_release(x_581, 1); - x_584 = x_581; -} else { - lean_dec_ref(x_581); - x_584 = lean_box(0); -} -x_585 = l_Lean_SourceInfo_fromRef(x_582, x_29); -lean_dec(x_582); -x_586 = l_Lean_Elab_Command_getCurrMacroScope(x_580, x_13, x_583); -lean_dec(x_580); -x_587 = lean_ctor_get(x_586, 1); -lean_inc(x_587); -if (lean_is_exclusive(x_586)) { - lean_ctor_release(x_586, 0); - lean_ctor_release(x_586, 1); - x_588 = x_586; -} else { - lean_dec_ref(x_586); - x_588 = lean_box(0); -} -x_589 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_587); -x_590 = lean_ctor_get(x_589, 1); +x_584 = lean_ctor_get(x_12, 7); +lean_inc(x_584); +x_585 = lean_ctor_get(x_12, 8); +lean_inc(x_585); +x_586 = lean_ctor_get(x_12, 9); +lean_inc(x_586); +x_587 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); +x_588 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_588, 0, x_578); +lean_ctor_set(x_588, 1, x_579); +lean_ctor_set(x_588, 2, x_580); +lean_ctor_set(x_588, 3, x_581); +lean_ctor_set(x_588, 4, x_582); +lean_ctor_set(x_588, 5, x_583); +lean_ctor_set(x_588, 6, x_577); +lean_ctor_set(x_588, 7, x_584); +lean_ctor_set(x_588, 8, x_585); +lean_ctor_set(x_588, 9, x_586); +lean_ctor_set_uint8(x_588, sizeof(void*)*10, x_587); +x_589 = l_Lean_Elab_Command_getRef(x_588, x_13, x_576); +x_590 = lean_ctor_get(x_589, 0); lean_inc(x_590); +x_591 = lean_ctor_get(x_589, 1); +lean_inc(x_591); if (lean_is_exclusive(x_589)) { lean_ctor_release(x_589, 0); lean_ctor_release(x_589, 1); - x_591 = x_589; + x_592 = x_589; } else { lean_dec_ref(x_589); - x_591 = lean_box(0); -} -x_592 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_585); -if (lean_is_scalar(x_591)) { - x_593 = lean_alloc_ctor(2, 2, 0); -} else { - x_593 = x_591; - lean_ctor_set_tag(x_593, 2); + x_592 = lean_box(0); } -lean_ctor_set(x_593, 0, x_585); -lean_ctor_set(x_593, 1, x_592); -x_594 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_595 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_585); -x_596 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_596, 0, x_585); -lean_ctor_set(x_596, 1, x_594); -lean_ctor_set(x_596, 2, x_595); -x_597 = lean_array_size(x_23); -x_598 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_597, x_17, x_23); -x_599 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_600 = l_Lean_mkSepArray(x_598, x_599); -lean_dec(x_598); -x_601 = l_Array_append___rarg(x_595, x_600); -lean_dec(x_600); -lean_inc(x_585); -x_602 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_602, 0, x_585); -lean_ctor_set(x_602, 1, x_594); -lean_ctor_set(x_602, 2, x_601); -x_603 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_596); -lean_inc(x_585); -x_604 = l_Lean_Syntax_node1(x_585, x_603, x_596); -x_605 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_585); -if (lean_is_scalar(x_588)) { - x_606 = lean_alloc_ctor(2, 2, 0); -} else { - x_606 = x_588; - lean_ctor_set_tag(x_606, 2); -} -lean_ctor_set(x_606, 0, x_585); -lean_ctor_set(x_606, 1, x_605); -x_607 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_596); -x_608 = l_Lean_Syntax_node6(x_585, x_607, x_593, x_596, x_602, x_604, x_596, x_606); -x_609 = l_Lean_Elab_Command_getRef(x_12, x_13, x_590); -x_610 = lean_ctor_get(x_609, 0); -lean_inc(x_610); -x_611 = lean_ctor_get(x_609, 1); -lean_inc(x_611); -if (lean_is_exclusive(x_609)) { - lean_ctor_release(x_609, 0); - lean_ctor_release(x_609, 1); - x_612 = x_609; +x_593 = l_Lean_SourceInfo_fromRef(x_590, x_29); +lean_dec(x_590); +x_594 = l_Lean_Elab_Command_getCurrMacroScope(x_588, x_13, x_591); +lean_dec(x_588); +x_595 = lean_ctor_get(x_594, 1); +lean_inc(x_595); +if (lean_is_exclusive(x_594)) { + lean_ctor_release(x_594, 0); + lean_ctor_release(x_594, 1); + x_596 = x_594; +} else { + lean_dec_ref(x_594); + x_596 = lean_box(0); +} +x_597 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_595); +x_598 = lean_ctor_get(x_597, 1); +lean_inc(x_598); +if (lean_is_exclusive(x_597)) { + lean_ctor_release(x_597, 0); + lean_ctor_release(x_597, 1); + x_599 = x_597; +} else { + lean_dec_ref(x_597); + x_599 = lean_box(0); +} +x_600 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_593); +if (lean_is_scalar(x_599)) { + x_601 = lean_alloc_ctor(2, 2, 0); +} else { + x_601 = x_599; + lean_ctor_set_tag(x_601, 2); +} +lean_ctor_set(x_601, 0, x_593); +lean_ctor_set(x_601, 1, x_600); +x_602 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_603 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_593); +x_604 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_604, 0, x_593); +lean_ctor_set(x_604, 1, x_602); +lean_ctor_set(x_604, 2, x_603); +x_605 = lean_array_size(x_23); +x_606 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_605, x_17, x_23); +x_607 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_608 = l_Lean_mkSepArray(x_606, x_607); +lean_dec(x_606); +x_609 = l_Array_append___rarg(x_603, x_608); +lean_dec(x_608); +lean_inc(x_593); +x_610 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_610, 0, x_593); +lean_ctor_set(x_610, 1, x_602); +lean_ctor_set(x_610, 2, x_609); +x_611 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_593); +x_612 = l_Lean_Syntax_node1(x_593, x_611, x_610); +x_613 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_604); +lean_inc(x_593); +x_614 = l_Lean_Syntax_node1(x_593, x_613, x_604); +x_615 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_593); +if (lean_is_scalar(x_596)) { + x_616 = lean_alloc_ctor(2, 2, 0); +} else { + x_616 = x_596; + lean_ctor_set_tag(x_616, 2); +} +lean_ctor_set(x_616, 0, x_593); +lean_ctor_set(x_616, 1, x_615); +x_617 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_604); +x_618 = l_Lean_Syntax_node6(x_593, x_617, x_601, x_604, x_612, x_614, x_604, x_616); +x_619 = l_Lean_Elab_Command_getRef(x_12, x_13, x_598); +x_620 = lean_ctor_get(x_619, 0); +lean_inc(x_620); +x_621 = lean_ctor_get(x_619, 1); +lean_inc(x_621); +if (lean_is_exclusive(x_619)) { + lean_ctor_release(x_619, 0); + lean_ctor_release(x_619, 1); + x_622 = x_619; +} else { + lean_dec_ref(x_619); + x_622 = lean_box(0); +} +x_623 = l_Lean_SourceInfo_fromRef(x_620, x_29); +lean_dec(x_620); +x_624 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_621); +x_625 = lean_ctor_get(x_624, 1); +lean_inc(x_625); +if (lean_is_exclusive(x_624)) { + lean_ctor_release(x_624, 0); + lean_ctor_release(x_624, 1); + x_626 = x_624; +} else { + lean_dec_ref(x_624); + x_626 = lean_box(0); +} +x_627 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_625); +x_628 = lean_ctor_get(x_627, 1); +lean_inc(x_628); +if (lean_is_exclusive(x_627)) { + lean_ctor_release(x_627, 0); + lean_ctor_release(x_627, 1); + x_629 = x_627; +} else { + lean_dec_ref(x_627); + x_629 = lean_box(0); +} +x_630 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_623); +if (lean_is_scalar(x_629)) { + x_631 = lean_alloc_ctor(2, 2, 0); } else { - lean_dec_ref(x_609); - x_612 = lean_box(0); + x_631 = x_629; + lean_ctor_set_tag(x_631, 2); } -x_613 = l_Lean_SourceInfo_fromRef(x_610, x_29); -lean_dec(x_610); -x_614 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_611); -x_615 = lean_ctor_get(x_614, 1); -lean_inc(x_615); -if (lean_is_exclusive(x_614)) { - lean_ctor_release(x_614, 0); - lean_ctor_release(x_614, 1); - x_616 = x_614; -} else { - lean_dec_ref(x_614); - x_616 = lean_box(0); -} -x_617 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_615); -x_618 = lean_ctor_get(x_617, 1); -lean_inc(x_618); -if (lean_is_exclusive(x_617)) { - lean_ctor_release(x_617, 0); - lean_ctor_release(x_617, 1); - x_619 = x_617; -} else { - lean_dec_ref(x_617); - x_619 = lean_box(0); -} -x_620 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_613); -if (lean_is_scalar(x_619)) { - x_621 = lean_alloc_ctor(2, 2, 0); -} else { - x_621 = x_619; - lean_ctor_set_tag(x_621, 2); -} -lean_ctor_set(x_621, 0, x_613); -lean_ctor_set(x_621, 1, x_620); -x_622 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_623 = l_Lean_Syntax_SepArray_ofElems(x_622, x_4); -x_624 = l_Array_append___rarg(x_595, x_623); -lean_dec(x_623); -lean_inc(x_613); -x_625 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_625, 0, x_613); -lean_ctor_set(x_625, 1, x_594); -lean_ctor_set(x_625, 2, x_624); -x_626 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_613); -if (lean_is_scalar(x_616)) { - x_627 = lean_alloc_ctor(2, 2, 0); -} else { - x_627 = x_616; - lean_ctor_set_tag(x_627, 2); -} -lean_ctor_set(x_627, 0, x_613); -lean_ctor_set(x_627, 1, x_626); -x_628 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_613); -x_629 = l_Lean_Syntax_node3(x_613, x_628, x_621, x_625, x_627); -lean_inc(x_613); -x_630 = l_Lean_Syntax_node1(x_613, x_594, x_629); -lean_inc(x_613); -x_631 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_631, 0, x_613); -lean_ctor_set(x_631, 1, x_594); -lean_ctor_set(x_631, 2, x_595); -x_632 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_613); -if (lean_is_scalar(x_612)) { - x_633 = lean_alloc_ctor(2, 2, 0); -} else { - x_633 = x_612; - lean_ctor_set_tag(x_633, 2); -} -lean_ctor_set(x_633, 0, x_613); -lean_ctor_set(x_633, 1, x_632); -x_634 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_613); -if (lean_is_scalar(x_584)) { - x_635 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_631, 0, x_623); +lean_ctor_set(x_631, 1, x_630); +x_632 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_633 = l_Lean_Syntax_SepArray_ofElems(x_632, x_4); +x_634 = l_Array_append___rarg(x_603, x_633); +lean_dec(x_633); +lean_inc(x_623); +x_635 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_635, 0, x_623); +lean_ctor_set(x_635, 1, x_602); +lean_ctor_set(x_635, 2, x_634); +x_636 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_623); +if (lean_is_scalar(x_626)) { + x_637 = lean_alloc_ctor(2, 2, 0); } else { - x_635 = x_584; - lean_ctor_set_tag(x_635, 2); + x_637 = x_626; + lean_ctor_set_tag(x_637, 2); } -lean_ctor_set(x_635, 0, x_613); -lean_ctor_set(x_635, 1, x_634); -x_636 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_613); -x_637 = l_Lean_Syntax_node2(x_613, x_636, x_635, x_30); -lean_inc(x_613); -x_638 = l_Lean_Syntax_node1(x_613, x_594, x_637); -x_639 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_631); -lean_inc(x_613); -x_640 = l_Lean_Syntax_node2(x_613, x_639, x_631, x_638); -x_641 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_613); -x_642 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_642, 0, x_613); -lean_ctor_set(x_642, 1, x_641); -x_643 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_631, 2); -lean_inc(x_613); -x_644 = l_Lean_Syntax_node2(x_613, x_643, x_631, x_631); +lean_ctor_set(x_637, 0, x_623); +lean_ctor_set(x_637, 1, x_636); +x_638 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_623); +x_639 = l_Lean_Syntax_node3(x_623, x_638, x_631, x_635, x_637); +lean_inc(x_623); +x_640 = l_Lean_Syntax_node1(x_623, x_602, x_639); +lean_inc(x_623); +x_641 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_641, 0, x_623); +lean_ctor_set(x_641, 1, x_602); +lean_ctor_set(x_641, 2, x_603); +x_642 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_623); +if (lean_is_scalar(x_622)) { + x_643 = lean_alloc_ctor(2, 2, 0); +} else { + x_643 = x_622; + lean_ctor_set_tag(x_643, 2); +} +lean_ctor_set(x_643, 0, x_623); +lean_ctor_set(x_643, 1, x_642); +x_644 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_623); +if (lean_is_scalar(x_592)) { + x_645 = lean_alloc_ctor(2, 2, 0); +} else { + x_645 = x_592; + lean_ctor_set_tag(x_645, 2); +} +lean_ctor_set(x_645, 0, x_623); +lean_ctor_set(x_645, 1, x_644); +x_646 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_623); +x_647 = l_Lean_Syntax_node2(x_623, x_646, x_645, x_30); +lean_inc(x_623); +x_648 = l_Lean_Syntax_node1(x_623, x_602, x_647); +x_649 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_641); +lean_inc(x_623); +x_650 = l_Lean_Syntax_node2(x_623, x_649, x_641, x_648); +x_651 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_623); +x_652 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_652, 0, x_623); +lean_ctor_set(x_652, 1, x_651); +x_653 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_641, 2); +lean_inc(x_623); +x_654 = l_Lean_Syntax_node2(x_623, x_653, x_641, x_641); if (lean_obj_tag(x_7) == 0) { -x_645 = x_595; -goto block_672; +x_655 = x_603; +goto block_682; } else { -lean_object* x_673; lean_object* x_674; -x_673 = lean_ctor_get(x_7, 0); -lean_inc(x_673); +lean_object* x_683; lean_object* x_684; +x_683 = lean_ctor_get(x_7, 0); +lean_inc(x_683); lean_dec(x_7); -x_674 = l_Array_mkArray1___rarg(x_673); -x_645 = x_674; -goto block_672; -} -block_672: -{ -lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; -x_646 = l_Array_append___rarg(x_595, x_645); -lean_dec(x_645); -lean_inc(x_613); -x_647 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_647, 0, x_613); -lean_ctor_set(x_647, 1, x_594); -lean_ctor_set(x_647, 2, x_646); -x_648 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_631, 3); -lean_inc(x_613); -x_649 = l_Lean_Syntax_node6(x_613, x_648, x_647, x_630, x_631, x_631, x_631, x_631); +x_684 = l_Array_mkArray1___rarg(x_683); +x_655 = x_684; +goto block_682; +} +block_682: +{ +lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; +x_656 = l_Array_append___rarg(x_603, x_655); +lean_dec(x_655); +lean_inc(x_623); +x_657 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_657, 0, x_623); +lean_ctor_set(x_657, 1, x_602); +lean_ctor_set(x_657, 2, x_656); +x_658 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_641, 3); +lean_inc(x_623); +x_659 = l_Lean_Syntax_node6(x_623, x_658, x_657, x_640, x_641, x_641, x_641, x_641); if (lean_obj_tag(x_5) == 0) { -lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; -x_650 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_613); -x_651 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_651, 0, x_613); -lean_ctor_set(x_651, 1, x_594); -lean_ctor_set(x_651, 2, x_650); -x_652 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_613); -x_653 = l_Lean_Syntax_node4(x_613, x_652, x_642, x_608, x_644, x_651); -x_654 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_613); -x_655 = l_Lean_Syntax_node4(x_613, x_654, x_633, x_31, x_640, x_653); -x_656 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_657 = l_Lean_Syntax_node2(x_613, x_656, x_649, x_655); -lean_inc(x_657); -x_658 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_658, 0, x_657); -x_659 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_657, x_658, x_12, x_13, x_618); -return x_659; -} -else -{ -lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; -x_660 = lean_ctor_get(x_5, 0); -lean_inc(x_660); +lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; +x_660 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_623); +x_661 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_661, 0, x_623); +lean_ctor_set(x_661, 1, x_602); +lean_ctor_set(x_661, 2, x_660); +x_662 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_623); +x_663 = l_Lean_Syntax_node4(x_623, x_662, x_652, x_618, x_654, x_661); +x_664 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_623); +x_665 = l_Lean_Syntax_node4(x_623, x_664, x_643, x_31, x_650, x_663); +x_666 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_667 = l_Lean_Syntax_node2(x_623, x_666, x_659, x_665); +lean_inc(x_667); +x_668 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_668, 0, x_667); +x_669 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_667, x_668, x_12, x_13, x_628); +return x_669; +} +else +{ +lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; +x_670 = lean_ctor_get(x_5, 0); +lean_inc(x_670); lean_dec(x_5); -x_661 = l_Array_mkArray1___rarg(x_660); -x_662 = l_Array_append___rarg(x_595, x_661); -lean_dec(x_661); -lean_inc(x_613); -x_663 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_663, 0, x_613); -lean_ctor_set(x_663, 1, x_594); -lean_ctor_set(x_663, 2, x_662); -x_664 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_613); -x_665 = l_Lean_Syntax_node4(x_613, x_664, x_642, x_608, x_644, x_663); -x_666 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_613); -x_667 = l_Lean_Syntax_node4(x_613, x_666, x_633, x_31, x_640, x_665); -x_668 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_669 = l_Lean_Syntax_node2(x_613, x_668, x_649, x_667); -lean_inc(x_669); -x_670 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_670, 0, x_669); -x_671 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_669, x_670, x_12, x_13, x_618); -return x_671; -} -} -} -} -} -else -{ -lean_object* x_738; lean_object* x_739; uint8_t x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; -x_738 = lean_ctor_get(x_25, 0); -x_739 = lean_ctor_get(x_25, 1); -lean_inc(x_739); -lean_inc(x_738); +x_671 = l_Array_mkArray1___rarg(x_670); +x_672 = l_Array_append___rarg(x_603, x_671); +lean_dec(x_671); +lean_inc(x_623); +x_673 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_673, 0, x_623); +lean_ctor_set(x_673, 1, x_602); +lean_ctor_set(x_673, 2, x_672); +x_674 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_623); +x_675 = l_Lean_Syntax_node4(x_623, x_674, x_652, x_618, x_654, x_673); +x_676 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_623); +x_677 = l_Lean_Syntax_node4(x_623, x_676, x_643, x_31, x_650, x_675); +x_678 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_679 = l_Lean_Syntax_node2(x_623, x_678, x_659, x_677); +lean_inc(x_679); +x_680 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_680, 0, x_679); +x_681 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_679, x_680, x_12, x_13, x_628); +return x_681; +} +} +} +} +} +else +{ +lean_object* x_748; lean_object* x_749; uint8_t x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; +x_748 = lean_ctor_get(x_25, 0); +x_749 = lean_ctor_get(x_25, 1); +lean_inc(x_749); +lean_inc(x_748); lean_dec(x_25); -x_740 = 0; -x_741 = l_Lean_mkCIdentFrom(x_738, x_2, x_740); -lean_dec(x_738); +x_750 = 0; +x_751 = l_Lean_mkCIdentFrom(x_748, x_2, x_750); +lean_dec(x_748); if (lean_obj_tag(x_8) == 0) { if (lean_obj_tag(x_9) == 0) { -lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; -x_855 = l_Lean_Elab_Command_getRef(x_12, x_13, x_739); -x_856 = lean_ctor_get(x_855, 0); -lean_inc(x_856); -x_857 = lean_ctor_get(x_855, 1); -lean_inc(x_857); -if (lean_is_exclusive(x_855)) { - lean_ctor_release(x_855, 0); - lean_ctor_release(x_855, 1); - x_858 = x_855; +lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; +x_867 = l_Lean_Elab_Command_getRef(x_12, x_13, x_749); +x_868 = lean_ctor_get(x_867, 0); +lean_inc(x_868); +x_869 = lean_ctor_get(x_867, 1); +lean_inc(x_869); +if (lean_is_exclusive(x_867)) { + lean_ctor_release(x_867, 0); + lean_ctor_release(x_867, 1); + x_870 = x_867; } else { - lean_dec_ref(x_855); - x_858 = lean_box(0); + lean_dec_ref(x_867); + x_870 = lean_box(0); } lean_inc(x_13); lean_inc(x_12); -x_859 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_856, x_740, x_12, x_13, x_857); -lean_dec(x_856); -if (lean_obj_tag(x_859) == 0) +x_871 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_868, x_750, x_12, x_13, x_869); +lean_dec(x_868); +if (lean_obj_tag(x_871) == 0) { -lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; -x_860 = lean_ctor_get(x_859, 0); -lean_inc(x_860); -x_861 = lean_ctor_get(x_859, 1); -lean_inc(x_861); -lean_dec(x_859); -x_862 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -if (lean_is_scalar(x_858)) { - x_863 = lean_alloc_ctor(1, 2, 0); -} else { - x_863 = x_858; - lean_ctor_set_tag(x_863, 1); -} -lean_ctor_set(x_863, 0, x_860); -lean_ctor_set(x_863, 1, x_862); -x_864 = lean_array_mk(x_863); -x_865 = lean_box(2); -x_866 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_867 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_867, 0, x_865); -lean_ctor_set(x_867, 1, x_866); -lean_ctor_set(x_867, 2, x_864); -x_742 = x_867; -x_743 = x_861; -goto block_854; -} -else -{ -lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; -lean_dec(x_858); -lean_dec(x_741); +lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; +x_872 = lean_ctor_get(x_871, 0); +lean_inc(x_872); +x_873 = lean_ctor_get(x_871, 1); +lean_inc(x_873); +lean_dec(x_871); +x_874 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +if (lean_is_scalar(x_870)) { + x_875 = lean_alloc_ctor(1, 2, 0); +} else { + x_875 = x_870; + lean_ctor_set_tag(x_875, 1); +} +lean_ctor_set(x_875, 0, x_872); +lean_ctor_set(x_875, 1, x_874); +x_876 = lean_array_mk(x_875); +x_877 = lean_box(2); +x_878 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_879 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_879, 0, x_877); +lean_ctor_set(x_879, 1, x_878); +lean_ctor_set(x_879, 2, x_876); +x_752 = x_879; +x_753 = x_873; +goto block_866; +} +else +{ +lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_dec(x_870); +lean_dec(x_751); lean_dec(x_23); lean_dec(x_13); lean_dec(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_868 = lean_ctor_get(x_859, 0); -lean_inc(x_868); -x_869 = lean_ctor_get(x_859, 1); -lean_inc(x_869); -if (lean_is_exclusive(x_859)) { - lean_ctor_release(x_859, 0); - lean_ctor_release(x_859, 1); - x_870 = x_859; +x_880 = lean_ctor_get(x_871, 0); +lean_inc(x_880); +x_881 = lean_ctor_get(x_871, 1); +lean_inc(x_881); +if (lean_is_exclusive(x_871)) { + lean_ctor_release(x_871, 0); + lean_ctor_release(x_871, 1); + x_882 = x_871; } else { - lean_dec_ref(x_859); - x_870 = lean_box(0); + lean_dec_ref(x_871); + x_882 = lean_box(0); } -if (lean_is_scalar(x_870)) { - x_871 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_882)) { + x_883 = lean_alloc_ctor(1, 2, 0); } else { - x_871 = x_870; + x_883 = x_882; } -lean_ctor_set(x_871, 0, x_868); -lean_ctor_set(x_871, 1, x_869); -return x_871; +lean_ctor_set(x_883, 0, x_880); +lean_ctor_set(x_883, 1, x_881); +return x_883; } } else { -lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; -x_872 = lean_ctor_get(x_9, 0); -lean_inc(x_872); +lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; +x_884 = lean_ctor_get(x_9, 0); +lean_inc(x_884); lean_dec(x_9); -x_873 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_872, x_740, x_12, x_13, x_739); -x_874 = lean_ctor_get(x_873, 0); -lean_inc(x_874); -x_875 = lean_ctor_get(x_873, 1); -lean_inc(x_875); -if (lean_is_exclusive(x_873)) { - lean_ctor_release(x_873, 0); - lean_ctor_release(x_873, 1); - x_876 = x_873; -} else { - lean_dec_ref(x_873); - x_876 = lean_box(0); -} -x_877 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -if (lean_is_scalar(x_876)) { - x_878 = lean_alloc_ctor(1, 2, 0); -} else { - x_878 = x_876; - lean_ctor_set_tag(x_878, 1); -} -lean_ctor_set(x_878, 0, x_874); -lean_ctor_set(x_878, 1, x_877); -x_879 = lean_array_mk(x_878); -x_880 = lean_box(2); -x_881 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_882 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_882, 0, x_880); -lean_ctor_set(x_882, 1, x_881); -lean_ctor_set(x_882, 2, x_879); -x_742 = x_882; -x_743 = x_875; -goto block_854; +x_885 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_884, x_750, x_12, x_13, x_749); +x_886 = lean_ctor_get(x_885, 0); +lean_inc(x_886); +x_887 = lean_ctor_get(x_885, 1); +lean_inc(x_887); +if (lean_is_exclusive(x_885)) { + lean_ctor_release(x_885, 0); + lean_ctor_release(x_885, 1); + x_888 = x_885; +} else { + lean_dec_ref(x_885); + x_888 = lean_box(0); +} +x_889 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +if (lean_is_scalar(x_888)) { + x_890 = lean_alloc_ctor(1, 2, 0); +} else { + x_890 = x_888; + lean_ctor_set_tag(x_890, 1); +} +lean_ctor_set(x_890, 0, x_886); +lean_ctor_set(x_890, 1, x_889); +x_891 = lean_array_mk(x_890); +x_892 = lean_box(2); +x_893 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_894 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_894, 0, x_892); +lean_ctor_set(x_894, 1, x_893); +lean_ctor_set(x_894, 2, x_891); +x_752 = x_894; +x_753 = x_887; +goto block_866; } } else { if (lean_obj_tag(x_9) == 0) { -lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; -x_883 = lean_ctor_get(x_8, 0); -x_884 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -lean_inc(x_883); -x_885 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_885, 0, x_883); -lean_ctor_set(x_885, 1, x_884); -x_886 = lean_array_mk(x_885); -x_887 = lean_box(2); -x_888 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_889 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_889, 0, x_887); -lean_ctor_set(x_889, 1, x_888); -lean_ctor_set(x_889, 2, x_886); -x_742 = x_889; -x_743 = x_739; -goto block_854; +lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; +x_895 = lean_ctor_get(x_8, 0); +x_896 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +lean_inc(x_895); +x_897 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_897, 0, x_895); +lean_ctor_set(x_897, 1, x_896); +x_898 = lean_array_mk(x_897); +x_899 = lean_box(2); +x_900 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_901 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_901, 0, x_899); +lean_ctor_set(x_901, 1, x_900); +lean_ctor_set(x_901, 2, x_898); +x_752 = x_901; +x_753 = x_749; +goto block_866; } else { -lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; -x_890 = lean_ctor_get(x_8, 0); -x_891 = lean_ctor_get(x_9, 0); -lean_inc(x_891); +lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; +x_902 = lean_ctor_get(x_8, 0); +x_903 = lean_ctor_get(x_9, 0); +lean_inc(x_903); lean_dec(x_9); -x_892 = l_Lean_mkIdentFrom(x_890, x_891, x_740); -x_893 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; -x_894 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_894, 0, x_892); -lean_ctor_set(x_894, 1, x_893); -x_895 = lean_array_mk(x_894); -x_896 = lean_box(2); -x_897 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_898 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_898, 0, x_896); -lean_ctor_set(x_898, 1, x_897); -lean_ctor_set(x_898, 2, x_895); -x_742 = x_898; -x_743 = x_739; -goto block_854; -} -} -block_854: -{ -lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; size_t x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; -x_744 = l_Lean_Elab_Command_getRef(x_12, x_13, x_743); -x_745 = lean_ctor_get(x_744, 0); -lean_inc(x_745); -x_746 = lean_ctor_get(x_744, 1); -lean_inc(x_746); -if (lean_is_exclusive(x_744)) { - lean_ctor_release(x_744, 0); - lean_ctor_release(x_744, 1); - x_747 = x_744; -} else { - lean_dec_ref(x_744); - x_747 = lean_box(0); -} -x_748 = l_Lean_replaceRef(x_3, x_745); -lean_dec(x_745); -x_749 = lean_ctor_get(x_12, 0); -lean_inc(x_749); -x_750 = lean_ctor_get(x_12, 1); -lean_inc(x_750); -x_751 = lean_ctor_get(x_12, 2); -lean_inc(x_751); -x_752 = lean_ctor_get(x_12, 3); -lean_inc(x_752); -x_753 = lean_ctor_get(x_12, 4); -lean_inc(x_753); -x_754 = lean_ctor_get(x_12, 5); -lean_inc(x_754); -x_755 = lean_ctor_get(x_12, 7); +x_904 = l_Lean_mkIdentFrom(x_902, x_903, x_750); +x_905 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__26; +x_906 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_906, 0, x_904); +lean_ctor_set(x_906, 1, x_905); +x_907 = lean_array_mk(x_906); +x_908 = lean_box(2); +x_909 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_910 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_910, 0, x_908); +lean_ctor_set(x_910, 1, x_909); +lean_ctor_set(x_910, 2, x_907); +x_752 = x_910; +x_753 = x_749; +goto block_866; +} +} +block_866: +{ +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; uint8_t x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; size_t x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; +x_754 = l_Lean_Elab_Command_getRef(x_12, x_13, x_753); +x_755 = lean_ctor_get(x_754, 0); lean_inc(x_755); -x_756 = lean_ctor_get(x_12, 8); +x_756 = lean_ctor_get(x_754, 1); lean_inc(x_756); -x_757 = lean_ctor_get(x_12, 9); -lean_inc(x_757); -x_758 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); -x_759 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_759, 0, x_749); -lean_ctor_set(x_759, 1, x_750); -lean_ctor_set(x_759, 2, x_751); -lean_ctor_set(x_759, 3, x_752); -lean_ctor_set(x_759, 4, x_753); -lean_ctor_set(x_759, 5, x_754); -lean_ctor_set(x_759, 6, x_748); -lean_ctor_set(x_759, 7, x_755); -lean_ctor_set(x_759, 8, x_756); -lean_ctor_set(x_759, 9, x_757); -lean_ctor_set_uint8(x_759, sizeof(void*)*10, x_758); -x_760 = l_Lean_Elab_Command_getRef(x_759, x_13, x_746); -x_761 = lean_ctor_get(x_760, 0); +if (lean_is_exclusive(x_754)) { + lean_ctor_release(x_754, 0); + lean_ctor_release(x_754, 1); + x_757 = x_754; +} else { + lean_dec_ref(x_754); + x_757 = lean_box(0); +} +x_758 = l_Lean_replaceRef(x_3, x_755); +lean_dec(x_755); +x_759 = lean_ctor_get(x_12, 0); +lean_inc(x_759); +x_760 = lean_ctor_get(x_12, 1); +lean_inc(x_760); +x_761 = lean_ctor_get(x_12, 2); lean_inc(x_761); -x_762 = lean_ctor_get(x_760, 1); +x_762 = lean_ctor_get(x_12, 3); lean_inc(x_762); -if (lean_is_exclusive(x_760)) { - lean_ctor_release(x_760, 0); - lean_ctor_release(x_760, 1); - x_763 = x_760; -} else { - lean_dec_ref(x_760); - x_763 = lean_box(0); -} -x_764 = l_Lean_SourceInfo_fromRef(x_761, x_740); -lean_dec(x_761); -x_765 = l_Lean_Elab_Command_getCurrMacroScope(x_759, x_13, x_762); -lean_dec(x_759); -x_766 = lean_ctor_get(x_765, 1); +x_763 = lean_ctor_get(x_12, 4); +lean_inc(x_763); +x_764 = lean_ctor_get(x_12, 5); +lean_inc(x_764); +x_765 = lean_ctor_get(x_12, 7); +lean_inc(x_765); +x_766 = lean_ctor_get(x_12, 8); lean_inc(x_766); -if (lean_is_exclusive(x_765)) { - lean_ctor_release(x_765, 0); - lean_ctor_release(x_765, 1); - x_767 = x_765; -} else { - lean_dec_ref(x_765); - x_767 = lean_box(0); -} -x_768 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_766); -x_769 = lean_ctor_get(x_768, 1); -lean_inc(x_769); -if (lean_is_exclusive(x_768)) { - lean_ctor_release(x_768, 0); - lean_ctor_release(x_768, 1); - x_770 = x_768; +x_767 = lean_ctor_get(x_12, 9); +lean_inc(x_767); +x_768 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); +x_769 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_769, 0, x_759); +lean_ctor_set(x_769, 1, x_760); +lean_ctor_set(x_769, 2, x_761); +lean_ctor_set(x_769, 3, x_762); +lean_ctor_set(x_769, 4, x_763); +lean_ctor_set(x_769, 5, x_764); +lean_ctor_set(x_769, 6, x_758); +lean_ctor_set(x_769, 7, x_765); +lean_ctor_set(x_769, 8, x_766); +lean_ctor_set(x_769, 9, x_767); +lean_ctor_set_uint8(x_769, sizeof(void*)*10, x_768); +x_770 = l_Lean_Elab_Command_getRef(x_769, x_13, x_756); +x_771 = lean_ctor_get(x_770, 0); +lean_inc(x_771); +x_772 = lean_ctor_get(x_770, 1); +lean_inc(x_772); +if (lean_is_exclusive(x_770)) { + lean_ctor_release(x_770, 0); + lean_ctor_release(x_770, 1); + x_773 = x_770; +} else { + lean_dec_ref(x_770); + x_773 = lean_box(0); +} +x_774 = l_Lean_SourceInfo_fromRef(x_771, x_750); +lean_dec(x_771); +x_775 = l_Lean_Elab_Command_getCurrMacroScope(x_769, x_13, x_772); +lean_dec(x_769); +x_776 = lean_ctor_get(x_775, 1); +lean_inc(x_776); +if (lean_is_exclusive(x_775)) { + lean_ctor_release(x_775, 0); + lean_ctor_release(x_775, 1); + x_777 = x_775; +} else { + lean_dec_ref(x_775); + x_777 = lean_box(0); +} +x_778 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_776); +x_779 = lean_ctor_get(x_778, 1); +lean_inc(x_779); +if (lean_is_exclusive(x_778)) { + lean_ctor_release(x_778, 0); + lean_ctor_release(x_778, 1); + x_780 = x_778; } else { - lean_dec_ref(x_768); - x_770 = lean_box(0); + lean_dec_ref(x_778); + x_780 = lean_box(0); } -x_771 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_764); -if (lean_is_scalar(x_770)) { - x_772 = lean_alloc_ctor(2, 2, 0); +x_781 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_774); +if (lean_is_scalar(x_780)) { + x_782 = lean_alloc_ctor(2, 2, 0); } else { - x_772 = x_770; - lean_ctor_set_tag(x_772, 2); + x_782 = x_780; + lean_ctor_set_tag(x_782, 2); } -lean_ctor_set(x_772, 0, x_764); -lean_ctor_set(x_772, 1, x_771); -x_773 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_774 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_764); -x_775 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_775, 0, x_764); -lean_ctor_set(x_775, 1, x_773); -lean_ctor_set(x_775, 2, x_774); -x_776 = lean_array_size(x_23); -x_777 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_776, x_17, x_23); -x_778 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_779 = l_Lean_mkSepArray(x_777, x_778); -lean_dec(x_777); -x_780 = l_Array_append___rarg(x_774, x_779); -lean_dec(x_779); -lean_inc(x_764); -x_781 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_781, 0, x_764); -lean_ctor_set(x_781, 1, x_773); -lean_ctor_set(x_781, 2, x_780); -x_782 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_775); -lean_inc(x_764); -x_783 = l_Lean_Syntax_node1(x_764, x_782, x_775); -x_784 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_764); -if (lean_is_scalar(x_767)) { - x_785 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_782, 0, x_774); +lean_ctor_set(x_782, 1, x_781); +x_783 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_784 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_774); +x_785 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_785, 0, x_774); +lean_ctor_set(x_785, 1, x_783); +lean_ctor_set(x_785, 2, x_784); +x_786 = lean_array_size(x_23); +x_787 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_786, x_17, x_23); +x_788 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_789 = l_Lean_mkSepArray(x_787, x_788); +lean_dec(x_787); +x_790 = l_Array_append___rarg(x_784, x_789); +lean_dec(x_789); +lean_inc(x_774); +x_791 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_791, 0, x_774); +lean_ctor_set(x_791, 1, x_783); +lean_ctor_set(x_791, 2, x_790); +x_792 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_774); +x_793 = l_Lean_Syntax_node1(x_774, x_792, x_791); +x_794 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_785); +lean_inc(x_774); +x_795 = l_Lean_Syntax_node1(x_774, x_794, x_785); +x_796 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_774); +if (lean_is_scalar(x_777)) { + x_797 = lean_alloc_ctor(2, 2, 0); +} else { + x_797 = x_777; + lean_ctor_set_tag(x_797, 2); +} +lean_ctor_set(x_797, 0, x_774); +lean_ctor_set(x_797, 1, x_796); +x_798 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_785); +x_799 = l_Lean_Syntax_node6(x_774, x_798, x_782, x_785, x_793, x_795, x_785, x_797); +x_800 = l_Lean_Elab_Command_getRef(x_12, x_13, x_779); +x_801 = lean_ctor_get(x_800, 0); +lean_inc(x_801); +x_802 = lean_ctor_get(x_800, 1); +lean_inc(x_802); +if (lean_is_exclusive(x_800)) { + lean_ctor_release(x_800, 0); + lean_ctor_release(x_800, 1); + x_803 = x_800; +} else { + lean_dec_ref(x_800); + x_803 = lean_box(0); +} +x_804 = l_Lean_SourceInfo_fromRef(x_801, x_750); +lean_dec(x_801); +x_805 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_802); +x_806 = lean_ctor_get(x_805, 1); +lean_inc(x_806); +if (lean_is_exclusive(x_805)) { + lean_ctor_release(x_805, 0); + lean_ctor_release(x_805, 1); + x_807 = x_805; } else { - x_785 = x_767; - lean_ctor_set_tag(x_785, 2); + lean_dec_ref(x_805); + x_807 = lean_box(0); } -lean_ctor_set(x_785, 0, x_764); -lean_ctor_set(x_785, 1, x_784); -x_786 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_775); -x_787 = l_Lean_Syntax_node6(x_764, x_786, x_772, x_775, x_781, x_783, x_775, x_785); -x_788 = l_Lean_Elab_Command_getRef(x_12, x_13, x_769); -x_789 = lean_ctor_get(x_788, 0); -lean_inc(x_789); -x_790 = lean_ctor_get(x_788, 1); -lean_inc(x_790); -if (lean_is_exclusive(x_788)) { - lean_ctor_release(x_788, 0); - lean_ctor_release(x_788, 1); - x_791 = x_788; +x_808 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_806); +x_809 = lean_ctor_get(x_808, 1); +lean_inc(x_809); +if (lean_is_exclusive(x_808)) { + lean_ctor_release(x_808, 0); + lean_ctor_release(x_808, 1); + x_810 = x_808; } else { - lean_dec_ref(x_788); - x_791 = lean_box(0); + lean_dec_ref(x_808); + x_810 = lean_box(0); } -x_792 = l_Lean_SourceInfo_fromRef(x_789, x_740); -lean_dec(x_789); -x_793 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_790); -x_794 = lean_ctor_get(x_793, 1); -lean_inc(x_794); -if (lean_is_exclusive(x_793)) { - lean_ctor_release(x_793, 0); - lean_ctor_release(x_793, 1); - x_795 = x_793; -} else { - lean_dec_ref(x_793); - x_795 = lean_box(0); -} -x_796 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_794); -x_797 = lean_ctor_get(x_796, 1); -lean_inc(x_797); -if (lean_is_exclusive(x_796)) { - lean_ctor_release(x_796, 0); - lean_ctor_release(x_796, 1); - x_798 = x_796; -} else { - lean_dec_ref(x_796); - x_798 = lean_box(0); -} -x_799 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_792); -if (lean_is_scalar(x_798)) { - x_800 = lean_alloc_ctor(2, 2, 0); -} else { - x_800 = x_798; - lean_ctor_set_tag(x_800, 2); -} -lean_ctor_set(x_800, 0, x_792); -lean_ctor_set(x_800, 1, x_799); -x_801 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_802 = l_Lean_Syntax_SepArray_ofElems(x_801, x_4); -x_803 = l_Array_append___rarg(x_774, x_802); -lean_dec(x_802); -lean_inc(x_792); -x_804 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_804, 0, x_792); -lean_ctor_set(x_804, 1, x_773); -lean_ctor_set(x_804, 2, x_803); -x_805 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_792); -if (lean_is_scalar(x_795)) { - x_806 = lean_alloc_ctor(2, 2, 0); -} else { - x_806 = x_795; - lean_ctor_set_tag(x_806, 2); -} -lean_ctor_set(x_806, 0, x_792); -lean_ctor_set(x_806, 1, x_805); -x_807 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_792); -x_808 = l_Lean_Syntax_node3(x_792, x_807, x_800, x_804, x_806); -lean_inc(x_792); -x_809 = l_Lean_Syntax_node1(x_792, x_773, x_808); -lean_inc(x_792); -x_810 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_810, 0, x_792); -lean_ctor_set(x_810, 1, x_773); -lean_ctor_set(x_810, 2, x_774); x_811 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_792); -if (lean_is_scalar(x_791)) { +lean_inc(x_804); +if (lean_is_scalar(x_810)) { x_812 = lean_alloc_ctor(2, 2, 0); } else { - x_812 = x_791; + x_812 = x_810; lean_ctor_set_tag(x_812, 2); } -lean_ctor_set(x_812, 0, x_792); +lean_ctor_set(x_812, 0, x_804); lean_ctor_set(x_812, 1, x_811); -x_813 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_792); -if (lean_is_scalar(x_763)) { - x_814 = lean_alloc_ctor(2, 2, 0); -} else { - x_814 = x_763; - lean_ctor_set_tag(x_814, 2); -} -lean_ctor_set(x_814, 0, x_792); -lean_ctor_set(x_814, 1, x_813); -x_815 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_792); -x_816 = l_Lean_Syntax_node2(x_792, x_815, x_814, x_741); -lean_inc(x_792); -x_817 = l_Lean_Syntax_node1(x_792, x_773, x_816); -x_818 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_810); -lean_inc(x_792); -x_819 = l_Lean_Syntax_node2(x_792, x_818, x_810, x_817); -x_820 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_792); -if (lean_is_scalar(x_747)) { - x_821 = lean_alloc_ctor(2, 2, 0); -} else { - x_821 = x_747; - lean_ctor_set_tag(x_821, 2); -} -lean_ctor_set(x_821, 0, x_792); -lean_ctor_set(x_821, 1, x_820); -x_822 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_810, 2); -lean_inc(x_792); -x_823 = l_Lean_Syntax_node2(x_792, x_822, x_810, x_810); +x_813 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_814 = l_Lean_Syntax_SepArray_ofElems(x_813, x_4); +x_815 = l_Array_append___rarg(x_784, x_814); +lean_dec(x_814); +lean_inc(x_804); +x_816 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_816, 0, x_804); +lean_ctor_set(x_816, 1, x_783); +lean_ctor_set(x_816, 2, x_815); +x_817 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_804); +if (lean_is_scalar(x_807)) { + x_818 = lean_alloc_ctor(2, 2, 0); +} else { + x_818 = x_807; + lean_ctor_set_tag(x_818, 2); +} +lean_ctor_set(x_818, 0, x_804); +lean_ctor_set(x_818, 1, x_817); +x_819 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_804); +x_820 = l_Lean_Syntax_node3(x_804, x_819, x_812, x_816, x_818); +lean_inc(x_804); +x_821 = l_Lean_Syntax_node1(x_804, x_783, x_820); +lean_inc(x_804); +x_822 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_822, 0, x_804); +lean_ctor_set(x_822, 1, x_783); +lean_ctor_set(x_822, 2, x_784); +x_823 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_804); +if (lean_is_scalar(x_803)) { + x_824 = lean_alloc_ctor(2, 2, 0); +} else { + x_824 = x_803; + lean_ctor_set_tag(x_824, 2); +} +lean_ctor_set(x_824, 0, x_804); +lean_ctor_set(x_824, 1, x_823); +x_825 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_804); +if (lean_is_scalar(x_773)) { + x_826 = lean_alloc_ctor(2, 2, 0); +} else { + x_826 = x_773; + lean_ctor_set_tag(x_826, 2); +} +lean_ctor_set(x_826, 0, x_804); +lean_ctor_set(x_826, 1, x_825); +x_827 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_804); +x_828 = l_Lean_Syntax_node2(x_804, x_827, x_826, x_751); +lean_inc(x_804); +x_829 = l_Lean_Syntax_node1(x_804, x_783, x_828); +x_830 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_822); +lean_inc(x_804); +x_831 = l_Lean_Syntax_node2(x_804, x_830, x_822, x_829); +x_832 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_804); +if (lean_is_scalar(x_757)) { + x_833 = lean_alloc_ctor(2, 2, 0); +} else { + x_833 = x_757; + lean_ctor_set_tag(x_833, 2); +} +lean_ctor_set(x_833, 0, x_804); +lean_ctor_set(x_833, 1, x_832); +x_834 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_822, 2); +lean_inc(x_804); +x_835 = l_Lean_Syntax_node2(x_804, x_834, x_822, x_822); if (lean_obj_tag(x_7) == 0) { -x_824 = x_774; -goto block_851; +x_836 = x_784; +goto block_863; } else { -lean_object* x_852; lean_object* x_853; -x_852 = lean_ctor_get(x_7, 0); -lean_inc(x_852); +lean_object* x_864; lean_object* x_865; +x_864 = lean_ctor_get(x_7, 0); +lean_inc(x_864); lean_dec(x_7); -x_853 = l_Array_mkArray1___rarg(x_852); -x_824 = x_853; -goto block_851; -} -block_851: -{ -lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; -x_825 = l_Array_append___rarg(x_774, x_824); -lean_dec(x_824); -lean_inc(x_792); -x_826 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_826, 0, x_792); -lean_ctor_set(x_826, 1, x_773); -lean_ctor_set(x_826, 2, x_825); -x_827 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_810, 3); -lean_inc(x_792); -x_828 = l_Lean_Syntax_node6(x_792, x_827, x_826, x_809, x_810, x_810, x_810, x_810); +x_865 = l_Array_mkArray1___rarg(x_864); +x_836 = x_865; +goto block_863; +} +block_863: +{ +lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; +x_837 = l_Array_append___rarg(x_784, x_836); +lean_dec(x_836); +lean_inc(x_804); +x_838 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_838, 0, x_804); +lean_ctor_set(x_838, 1, x_783); +lean_ctor_set(x_838, 2, x_837); +x_839 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_822, 3); +lean_inc(x_804); +x_840 = l_Lean_Syntax_node6(x_804, x_839, x_838, x_821, x_822, x_822, x_822, x_822); if (lean_obj_tag(x_5) == 0) { -lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; -x_829 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_792); -x_830 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_830, 0, x_792); -lean_ctor_set(x_830, 1, x_773); -lean_ctor_set(x_830, 2, x_829); -x_831 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_792); -x_832 = l_Lean_Syntax_node4(x_792, x_831, x_821, x_787, x_823, x_830); -x_833 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_792); -x_834 = l_Lean_Syntax_node4(x_792, x_833, x_812, x_742, x_819, x_832); -x_835 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_836 = l_Lean_Syntax_node2(x_792, x_835, x_828, x_834); -lean_inc(x_836); -x_837 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_837, 0, x_836); -x_838 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_836, x_837, x_12, x_13, x_797); -return x_838; -} -else -{ -lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; -x_839 = lean_ctor_get(x_5, 0); -lean_inc(x_839); -lean_dec(x_5); -x_840 = l_Array_mkArray1___rarg(x_839); -x_841 = l_Array_append___rarg(x_774, x_840); -lean_dec(x_840); -lean_inc(x_792); +lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; +x_841 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_804); x_842 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_842, 0, x_792); -lean_ctor_set(x_842, 1, x_773); +lean_ctor_set(x_842, 0, x_804); +lean_ctor_set(x_842, 1, x_783); lean_ctor_set(x_842, 2, x_841); x_843 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_792); -x_844 = l_Lean_Syntax_node4(x_792, x_843, x_821, x_787, x_823, x_842); -x_845 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_792); -x_846 = l_Lean_Syntax_node4(x_792, x_845, x_812, x_742, x_819, x_844); -x_847 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_848 = l_Lean_Syntax_node2(x_792, x_847, x_828, x_846); +lean_inc(x_804); +x_844 = l_Lean_Syntax_node4(x_804, x_843, x_833, x_799, x_835, x_842); +x_845 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_804); +x_846 = l_Lean_Syntax_node4(x_804, x_845, x_824, x_752, x_831, x_844); +x_847 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_848 = l_Lean_Syntax_node2(x_804, x_847, x_840, x_846); lean_inc(x_848); x_849 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); lean_closure_set(x_849, 0, x_848); -x_850 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_848, x_849, x_12, x_13, x_797); +x_850 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_848, x_849, x_12, x_13, x_809); return x_850; } +else +{ +lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; +x_851 = lean_ctor_get(x_5, 0); +lean_inc(x_851); +lean_dec(x_5); +x_852 = l_Array_mkArray1___rarg(x_851); +x_853 = l_Array_append___rarg(x_784, x_852); +lean_dec(x_852); +lean_inc(x_804); +x_854 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_854, 0, x_804); +lean_ctor_set(x_854, 1, x_783); +lean_ctor_set(x_854, 2, x_853); +x_855 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_804); +x_856 = l_Lean_Syntax_node4(x_804, x_855, x_833, x_799, x_835, x_854); +x_857 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_804); +x_858 = l_Lean_Syntax_node4(x_804, x_857, x_824, x_752, x_831, x_856); +x_859 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_860 = l_Lean_Syntax_node2(x_804, x_859, x_840, x_858); +lean_inc(x_860); +x_861 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_861, 0, x_860); +x_862 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_6, x_860, x_861, x_12, x_13, x_809); +return x_862; +} } } } } else { -uint8_t x_899; +uint8_t x_911; lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); @@ -7401,23 +7447,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_899 = !lean_is_exclusive(x_18); -if (x_899 == 0) +x_911 = !lean_is_exclusive(x_18); +if (x_911 == 0) { return x_18; } else { -lean_object* x_900; lean_object* x_901; lean_object* x_902; -x_900 = lean_ctor_get(x_18, 0); -x_901 = lean_ctor_get(x_18, 1); -lean_inc(x_901); -lean_inc(x_900); +lean_object* x_912; lean_object* x_913; lean_object* x_914; +x_912 = lean_ctor_get(x_18, 0); +x_913 = lean_ctor_get(x_18, 1); +lean_inc(x_913); +lean_inc(x_912); lean_dec(x_18); -x_902 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_902, 0, x_900); -lean_ctor_set(x_902, 1, x_901); -return x_902; +x_914 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_914, 0, x_912); +lean_ctor_set(x_914, 1, x_913); +return x_914; } } } @@ -7550,52 +7596,52 @@ if (lean_obj_tag(x_11) == 0) lean_free_object(x_25); if (lean_obj_tag(x_12) == 0) { -lean_object* x_680; uint8_t x_681; -x_680 = l_Lean_Elab_Command_getRef(x_15, x_16, x_32); -x_681 = !lean_is_exclusive(x_680); -if (x_681 == 0) +lean_object* x_690; uint8_t x_691; +x_690 = l_Lean_Elab_Command_getRef(x_15, x_16, x_32); +x_691 = !lean_is_exclusive(x_690); +if (x_691 == 0) { -lean_object* x_682; lean_object* x_683; lean_object* x_684; -x_682 = lean_ctor_get(x_680, 0); -x_683 = lean_ctor_get(x_680, 1); +lean_object* x_692; lean_object* x_693; lean_object* x_694; +x_692 = lean_ctor_get(x_690, 0); +x_693 = lean_ctor_get(x_690, 1); lean_inc(x_16); lean_inc(x_15); -x_684 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_682, x_33, x_15, x_16, x_683); -lean_dec(x_682); -if (lean_obj_tag(x_684) == 0) -{ -lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; -x_685 = lean_ctor_get(x_684, 0); -lean_inc(x_685); -x_686 = lean_ctor_get(x_684, 1); -lean_inc(x_686); -lean_dec(x_684); -x_687 = lean_box(2); -x_688 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_689 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_689, 0, x_687); -lean_ctor_set(x_689, 1, x_688); -lean_ctor_set(x_689, 2, x_24); -lean_ctor_set_tag(x_680, 1); -lean_ctor_set(x_680, 1, x_4); -lean_ctor_set(x_680, 0, x_689); +x_694 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_692, x_33, x_15, x_16, x_693); +lean_dec(x_692); +if (lean_obj_tag(x_694) == 0) +{ +lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; +x_695 = lean_ctor_get(x_694, 0); +lean_inc(x_695); +x_696 = lean_ctor_get(x_694, 1); +lean_inc(x_696); +lean_dec(x_694); +x_697 = lean_box(2); +x_698 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_699 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_699, 0, x_697); +lean_ctor_set(x_699, 1, x_698); +lean_ctor_set(x_699, 2, x_24); +lean_ctor_set_tag(x_690, 1); +lean_ctor_set(x_690, 1, x_4); +lean_ctor_set(x_690, 0, x_699); lean_ctor_set_tag(x_29, 1); -lean_ctor_set(x_29, 1, x_680); -lean_ctor_set(x_29, 0, x_685); -x_690 = lean_array_mk(x_29); -x_691 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_692 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_692, 0, x_687); -lean_ctor_set(x_692, 1, x_691); -lean_ctor_set(x_692, 2, x_690); -x_35 = x_692; -x_36 = x_686; -goto block_679; -} -else -{ -uint8_t x_693; -lean_free_object(x_680); +lean_ctor_set(x_29, 1, x_690); +lean_ctor_set(x_29, 0, x_695); +x_700 = lean_array_mk(x_29); +x_701 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_702 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_702, 0, x_697); +lean_ctor_set(x_702, 1, x_701); +lean_ctor_set(x_702, 2, x_700); +x_35 = x_702; +x_36 = x_696; +goto block_689; +} +else +{ +uint8_t x_703; +lean_free_object(x_690); lean_dec(x_34); lean_free_object(x_29); lean_dec(x_27); @@ -7606,71 +7652,71 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_693 = !lean_is_exclusive(x_684); -if (x_693 == 0) +x_703 = !lean_is_exclusive(x_694); +if (x_703 == 0) { -return x_684; +return x_694; } else { -lean_object* x_694; lean_object* x_695; lean_object* x_696; -x_694 = lean_ctor_get(x_684, 0); -x_695 = lean_ctor_get(x_684, 1); -lean_inc(x_695); -lean_inc(x_694); -lean_dec(x_684); -x_696 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_696, 0, x_694); -lean_ctor_set(x_696, 1, x_695); -return x_696; +lean_object* x_704; lean_object* x_705; lean_object* x_706; +x_704 = lean_ctor_get(x_694, 0); +x_705 = lean_ctor_get(x_694, 1); +lean_inc(x_705); +lean_inc(x_704); +lean_dec(x_694); +x_706 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_706, 0, x_704); +lean_ctor_set(x_706, 1, x_705); +return x_706; } } } else { -lean_object* x_697; lean_object* x_698; lean_object* x_699; -x_697 = lean_ctor_get(x_680, 0); -x_698 = lean_ctor_get(x_680, 1); -lean_inc(x_698); -lean_inc(x_697); -lean_dec(x_680); +lean_object* x_707; lean_object* x_708; lean_object* x_709; +x_707 = lean_ctor_get(x_690, 0); +x_708 = lean_ctor_get(x_690, 1); +lean_inc(x_708); +lean_inc(x_707); +lean_dec(x_690); lean_inc(x_16); lean_inc(x_15); -x_699 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_697, x_33, x_15, x_16, x_698); -lean_dec(x_697); -if (lean_obj_tag(x_699) == 0) +x_709 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_707, x_33, x_15, x_16, x_708); +lean_dec(x_707); +if (lean_obj_tag(x_709) == 0) { -lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; -x_700 = lean_ctor_get(x_699, 0); -lean_inc(x_700); -x_701 = lean_ctor_get(x_699, 1); -lean_inc(x_701); -lean_dec(x_699); -x_702 = lean_box(2); -x_703 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_704 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_704, 0, x_702); -lean_ctor_set(x_704, 1, x_703); -lean_ctor_set(x_704, 2, x_24); -x_705 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_705, 0, x_704); -lean_ctor_set(x_705, 1, x_4); +lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; +x_710 = lean_ctor_get(x_709, 0); +lean_inc(x_710); +x_711 = lean_ctor_get(x_709, 1); +lean_inc(x_711); +lean_dec(x_709); +x_712 = lean_box(2); +x_713 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_714 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_714, 0, x_712); +lean_ctor_set(x_714, 1, x_713); +lean_ctor_set(x_714, 2, x_24); +x_715 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_715, 0, x_714); +lean_ctor_set(x_715, 1, x_4); lean_ctor_set_tag(x_29, 1); -lean_ctor_set(x_29, 1, x_705); -lean_ctor_set(x_29, 0, x_700); -x_706 = lean_array_mk(x_29); -x_707 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_708 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_708, 0, x_702); -lean_ctor_set(x_708, 1, x_707); -lean_ctor_set(x_708, 2, x_706); -x_35 = x_708; -x_36 = x_701; -goto block_679; +lean_ctor_set(x_29, 1, x_715); +lean_ctor_set(x_29, 0, x_710); +x_716 = lean_array_mk(x_29); +x_717 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_718 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_718, 0, x_712); +lean_ctor_set(x_718, 1, x_717); +lean_ctor_set(x_718, 2, x_716); +x_35 = x_718; +x_36 = x_711; +goto block_689; } else { -lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; +lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_dec(x_34); lean_free_object(x_29); lean_dec(x_27); @@ -7681,93 +7727,93 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_709 = lean_ctor_get(x_699, 0); -lean_inc(x_709); -x_710 = lean_ctor_get(x_699, 1); -lean_inc(x_710); -if (lean_is_exclusive(x_699)) { - lean_ctor_release(x_699, 0); - lean_ctor_release(x_699, 1); - x_711 = x_699; +x_719 = lean_ctor_get(x_709, 0); +lean_inc(x_719); +x_720 = lean_ctor_get(x_709, 1); +lean_inc(x_720); +if (lean_is_exclusive(x_709)) { + lean_ctor_release(x_709, 0); + lean_ctor_release(x_709, 1); + x_721 = x_709; } else { - lean_dec_ref(x_699); - x_711 = lean_box(0); + lean_dec_ref(x_709); + x_721 = lean_box(0); } -if (lean_is_scalar(x_711)) { - x_712 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_721)) { + x_722 = lean_alloc_ctor(1, 2, 0); } else { - x_712 = x_711; + x_722 = x_721; } -lean_ctor_set(x_712, 0, x_709); -lean_ctor_set(x_712, 1, x_710); -return x_712; +lean_ctor_set(x_722, 0, x_719); +lean_ctor_set(x_722, 1, x_720); +return x_722; } } } else { -lean_object* x_713; lean_object* x_714; uint8_t x_715; -x_713 = lean_ctor_get(x_12, 0); -lean_inc(x_713); +lean_object* x_723; lean_object* x_724; uint8_t x_725; +x_723 = lean_ctor_get(x_12, 0); +lean_inc(x_723); lean_dec(x_12); -x_714 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_713, x_33, x_15, x_16, x_32); -x_715 = !lean_is_exclusive(x_714); -if (x_715 == 0) -{ -lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; -x_716 = lean_ctor_get(x_714, 0); -x_717 = lean_ctor_get(x_714, 1); -x_718 = lean_box(2); -x_719 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_720 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_720, 0, x_718); -lean_ctor_set(x_720, 1, x_719); -lean_ctor_set(x_720, 2, x_24); -lean_ctor_set_tag(x_714, 1); -lean_ctor_set(x_714, 1, x_4); -lean_ctor_set(x_714, 0, x_720); +x_724 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_723, x_33, x_15, x_16, x_32); +x_725 = !lean_is_exclusive(x_724); +if (x_725 == 0) +{ +lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; +x_726 = lean_ctor_get(x_724, 0); +x_727 = lean_ctor_get(x_724, 1); +x_728 = lean_box(2); +x_729 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_730 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_730, 0, x_728); +lean_ctor_set(x_730, 1, x_729); +lean_ctor_set(x_730, 2, x_24); +lean_ctor_set_tag(x_724, 1); +lean_ctor_set(x_724, 1, x_4); +lean_ctor_set(x_724, 0, x_730); lean_ctor_set_tag(x_29, 1); -lean_ctor_set(x_29, 1, x_714); -lean_ctor_set(x_29, 0, x_716); -x_721 = lean_array_mk(x_29); -x_722 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_723 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_723, 0, x_718); -lean_ctor_set(x_723, 1, x_722); -lean_ctor_set(x_723, 2, x_721); -x_35 = x_723; -x_36 = x_717; -goto block_679; -} -else -{ -lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; -x_724 = lean_ctor_get(x_714, 0); -x_725 = lean_ctor_get(x_714, 1); -lean_inc(x_725); -lean_inc(x_724); -lean_dec(x_714); -x_726 = lean_box(2); -x_727 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_728 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_728, 0, x_726); -lean_ctor_set(x_728, 1, x_727); -lean_ctor_set(x_728, 2, x_24); -x_729 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_729, 0, x_728); -lean_ctor_set(x_729, 1, x_4); +lean_ctor_set(x_29, 1, x_724); +lean_ctor_set(x_29, 0, x_726); +x_731 = lean_array_mk(x_29); +x_732 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_733 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_733, 0, x_728); +lean_ctor_set(x_733, 1, x_732); +lean_ctor_set(x_733, 2, x_731); +x_35 = x_733; +x_36 = x_727; +goto block_689; +} +else +{ +lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; +x_734 = lean_ctor_get(x_724, 0); +x_735 = lean_ctor_get(x_724, 1); +lean_inc(x_735); +lean_inc(x_734); +lean_dec(x_724); +x_736 = lean_box(2); +x_737 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_738 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_738, 0, x_736); +lean_ctor_set(x_738, 1, x_737); +lean_ctor_set(x_738, 2, x_24); +x_739 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_739, 0, x_738); +lean_ctor_set(x_739, 1, x_4); lean_ctor_set_tag(x_29, 1); -lean_ctor_set(x_29, 1, x_729); -lean_ctor_set(x_29, 0, x_724); -x_730 = lean_array_mk(x_29); -x_731 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_732 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_732, 0, x_726); -lean_ctor_set(x_732, 1, x_731); -lean_ctor_set(x_732, 2, x_730); -x_35 = x_732; -x_36 = x_725; -goto block_679; +lean_ctor_set(x_29, 1, x_739); +lean_ctor_set(x_29, 0, x_734); +x_740 = lean_array_mk(x_29); +x_741 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_742 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_742, 0, x_736); +lean_ctor_set(x_742, 1, x_741); +lean_ctor_set(x_742, 2, x_740); +x_35 = x_742; +x_36 = x_735; +goto block_689; } } } @@ -7775,63 +7821,63 @@ else { if (lean_obj_tag(x_12) == 0) { -lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; -x_733 = lean_ctor_get(x_11, 0); -x_734 = lean_box(2); -x_735 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_736 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_736, 0, x_734); -lean_ctor_set(x_736, 1, x_735); -lean_ctor_set(x_736, 2, x_24); +lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; +x_743 = lean_ctor_get(x_11, 0); +x_744 = lean_box(2); +x_745 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_746 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_746, 0, x_744); +lean_ctor_set(x_746, 1, x_745); +lean_ctor_set(x_746, 2, x_24); lean_ctor_set_tag(x_29, 1); lean_ctor_set(x_29, 1, x_4); -lean_ctor_set(x_29, 0, x_736); -lean_inc(x_733); +lean_ctor_set(x_29, 0, x_746); +lean_inc(x_743); lean_ctor_set_tag(x_25, 1); lean_ctor_set(x_25, 1, x_29); -lean_ctor_set(x_25, 0, x_733); -x_737 = lean_array_mk(x_25); -x_738 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_739 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_739, 0, x_734); -lean_ctor_set(x_739, 1, x_738); -lean_ctor_set(x_739, 2, x_737); -x_35 = x_739; +lean_ctor_set(x_25, 0, x_743); +x_747 = lean_array_mk(x_25); +x_748 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_749 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_749, 0, x_744); +lean_ctor_set(x_749, 1, x_748); +lean_ctor_set(x_749, 2, x_747); +x_35 = x_749; x_36 = x_32; -goto block_679; +goto block_689; } else { -lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; -x_740 = lean_ctor_get(x_11, 0); -x_741 = lean_ctor_get(x_12, 0); -lean_inc(x_741); +lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; +x_750 = lean_ctor_get(x_11, 0); +x_751 = lean_ctor_get(x_12, 0); +lean_inc(x_751); lean_dec(x_12); -x_742 = l_Lean_mkIdentFrom(x_740, x_741, x_33); -x_743 = lean_box(2); -x_744 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_745 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_745, 0, x_743); -lean_ctor_set(x_745, 1, x_744); -lean_ctor_set(x_745, 2, x_24); +x_752 = l_Lean_mkIdentFrom(x_750, x_751, x_33); +x_753 = lean_box(2); +x_754 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_755 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_755, 0, x_753); +lean_ctor_set(x_755, 1, x_754); +lean_ctor_set(x_755, 2, x_24); lean_ctor_set_tag(x_29, 1); lean_ctor_set(x_29, 1, x_4); -lean_ctor_set(x_29, 0, x_745); +lean_ctor_set(x_29, 0, x_755); lean_ctor_set_tag(x_25, 1); lean_ctor_set(x_25, 1, x_29); -lean_ctor_set(x_25, 0, x_742); -x_746 = lean_array_mk(x_25); -x_747 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_748 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_748, 0, x_743); -lean_ctor_set(x_748, 1, x_747); -lean_ctor_set(x_748, 2, x_746); -x_35 = x_748; +lean_ctor_set(x_25, 0, x_752); +x_756 = lean_array_mk(x_25); +x_757 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_758 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_758, 0, x_753); +lean_ctor_set(x_758, 1, x_757); +lean_ctor_set(x_758, 2, x_756); +x_35 = x_758; x_36 = x_32; -goto block_679; +goto block_689; } } -block_679: +block_689: { lean_object* x_37; uint8_t x_38; x_37 = l_Lean_Elab_Command_getRef(x_15, x_16, x_36); @@ -7896,7 +7942,7 @@ x_62 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_60); x_63 = !lean_is_exclusive(x_62); if (x_63 == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; size_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; size_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; x_64 = lean_ctor_get(x_62, 1); x_65 = lean_ctor_get(x_62, 0); lean_dec(x_65); @@ -7914,7 +7960,7 @@ lean_ctor_set(x_69, 1, x_67); lean_ctor_set(x_69, 2, x_68); x_70 = lean_array_size(x_27); x_71 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_70, x_20, x_27); -x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; x_73 = l_Lean_mkSepArray(x_71, x_72); lean_dec(x_71); x_74 = l_Array_append___rarg(x_68, x_73); @@ -7924,1745 +7970,1760 @@ x_75 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_75, 0, x_57); lean_ctor_set(x_75, 1, x_67); lean_ctor_set(x_75, 2, x_74); -x_76 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_76 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_57); +x_77 = l_Lean_Syntax_node1(x_57, x_76, x_75); +x_78 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; lean_inc(x_69); lean_inc(x_57); -x_77 = l_Lean_Syntax_node1(x_57, x_76, x_69); -x_78 = l_Lake_DSL_structVal___closed__18; +x_79 = l_Lean_Syntax_node1(x_57, x_78, x_69); +x_80 = l_Lake_DSL_structVal___closed__18; lean_inc(x_57); lean_ctor_set_tag(x_58, 2); -lean_ctor_set(x_58, 1, x_78); +lean_ctor_set(x_58, 1, x_80); lean_ctor_set(x_58, 0, x_57); -x_79 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +x_81 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; lean_inc(x_69); -x_80 = l_Lean_Syntax_node6(x_57, x_79, x_62, x_69, x_75, x_77, x_69, x_58); -x_81 = l_Lean_Elab_Command_getRef(x_15, x_16, x_64); -x_82 = !lean_is_exclusive(x_81); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_83 = lean_ctor_get(x_81, 0); -x_84 = lean_ctor_get(x_81, 1); -x_85 = l_Lean_SourceInfo_fromRef(x_83, x_33); -lean_dec(x_83); -x_86 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_84); -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) +x_82 = l_Lean_Syntax_node6(x_57, x_81, x_62, x_69, x_77, x_79, x_69, x_58); +x_83 = l_Lean_Elab_Command_getRef(x_15, x_16, x_64); +x_84 = !lean_is_exclusive(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_85 = lean_ctor_get(x_83, 0); +x_86 = lean_ctor_get(x_83, 1); +x_87 = l_Lean_SourceInfo_fromRef(x_85, x_33); +lean_dec(x_85); +x_88 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_86); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_88 = lean_ctor_get(x_86, 1); -x_89 = lean_ctor_get(x_86, 0); -lean_dec(x_89); -x_90 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_88); -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) -{ -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_92 = lean_ctor_get(x_90, 1); -x_93 = lean_ctor_get(x_90, 0); -lean_dec(x_93); -x_94 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_85); -lean_ctor_set_tag(x_90, 2); -lean_ctor_set(x_90, 1, x_94); -lean_ctor_set(x_90, 0, x_85); -x_95 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_96 = l_Lean_Syntax_SepArray_ofElems(x_95, x_7); -x_97 = l_Array_append___rarg(x_68, x_96); -lean_dec(x_96); -lean_inc(x_85); -x_98 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_98, 0, x_85); -lean_ctor_set(x_98, 1, x_67); -lean_ctor_set(x_98, 2, x_97); -x_99 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_85); -lean_ctor_set_tag(x_86, 2); -lean_ctor_set(x_86, 1, x_99); -lean_ctor_set(x_86, 0, x_85); -x_100 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_85); -x_101 = l_Lean_Syntax_node3(x_85, x_100, x_90, x_98, x_86); -lean_inc(x_85); -x_102 = l_Lean_Syntax_node1(x_85, x_67, x_101); -lean_inc(x_85); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_85); -lean_ctor_set(x_103, 1, x_67); -lean_ctor_set(x_103, 2, x_68); -x_104 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_85); -lean_ctor_set_tag(x_81, 2); -lean_ctor_set(x_81, 1, x_104); -lean_ctor_set(x_81, 0, x_85); -x_105 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_85); +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_90); +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_94 = lean_ctor_get(x_92, 1); +x_95 = lean_ctor_get(x_92, 0); +lean_dec(x_95); +x_96 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_87); +lean_ctor_set_tag(x_92, 2); +lean_ctor_set(x_92, 1, x_96); +lean_ctor_set(x_92, 0, x_87); +x_97 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_98 = l_Lean_Syntax_SepArray_ofElems(x_97, x_7); +x_99 = l_Array_append___rarg(x_68, x_98); +lean_dec(x_98); +lean_inc(x_87); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_87); +lean_ctor_set(x_100, 1, x_67); +lean_ctor_set(x_100, 2, x_99); +x_101 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_87); +lean_ctor_set_tag(x_88, 2); +lean_ctor_set(x_88, 1, x_101); +lean_ctor_set(x_88, 0, x_87); +x_102 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_87); +x_103 = l_Lean_Syntax_node3(x_87, x_102, x_92, x_100, x_88); +lean_inc(x_87); +x_104 = l_Lean_Syntax_node1(x_87, x_67, x_103); +lean_inc(x_87); +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_87); +lean_ctor_set(x_105, 1, x_67); +lean_ctor_set(x_105, 2, x_68); +x_106 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_87); +lean_ctor_set_tag(x_83, 2); +lean_ctor_set(x_83, 1, x_106); +lean_ctor_set(x_83, 0, x_87); +x_107 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_105); -lean_ctor_set(x_53, 0, x_85); -x_106 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_85); -x_107 = l_Lean_Syntax_node2(x_85, x_106, x_53, x_34); -lean_inc(x_85); -x_108 = l_Lean_Syntax_node1(x_85, x_67, x_107); -x_109 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_103); -lean_inc(x_85); -x_110 = l_Lean_Syntax_node2(x_85, x_109, x_103, x_108); -x_111 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_85); +lean_ctor_set(x_53, 1, x_107); +lean_ctor_set(x_53, 0, x_87); +x_108 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_87); +x_109 = l_Lean_Syntax_node2(x_87, x_108, x_53, x_34); +lean_inc(x_87); +x_110 = l_Lean_Syntax_node1(x_87, x_67, x_109); +x_111 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_105); +lean_inc(x_87); +x_112 = l_Lean_Syntax_node2(x_87, x_111, x_105, x_110); +x_113 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_111); -lean_ctor_set(x_37, 0, x_85); -x_112 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_103, 2); -lean_inc(x_85); -x_113 = l_Lean_Syntax_node2(x_85, x_112, x_103, x_103); +lean_ctor_set(x_37, 1, x_113); +lean_ctor_set(x_37, 0, x_87); +x_114 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_105, 2); +lean_inc(x_87); +x_115 = l_Lean_Syntax_node2(x_87, x_114, x_105, x_105); if (lean_obj_tag(x_10) == 0) { -x_114 = x_68; -goto block_141; +x_116 = x_68; +goto block_143; } else { -lean_object* x_142; lean_object* x_143; -x_142 = lean_ctor_get(x_10, 0); -lean_inc(x_142); +lean_object* x_144; lean_object* x_145; +x_144 = lean_ctor_get(x_10, 0); +lean_inc(x_144); lean_dec(x_10); -x_143 = l_Array_mkArray1___rarg(x_142); -x_114 = x_143; -goto block_141; +x_145 = l_Array_mkArray1___rarg(x_144); +x_116 = x_145; +goto block_143; } -block_141: +block_143: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = l_Array_append___rarg(x_68, x_114); -lean_dec(x_114); -lean_inc(x_85); -x_116 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_116, 0, x_85); -lean_ctor_set(x_116, 1, x_67); -lean_ctor_set(x_116, 2, x_115); -x_117 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_103, 3); -lean_inc(x_85); -x_118 = l_Lean_Syntax_node6(x_85, x_117, x_116, x_102, x_103, x_103, x_103, x_103); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_117 = l_Array_append___rarg(x_68, x_116); +lean_dec(x_116); +lean_inc(x_87); +x_118 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_118, 0, x_87); +lean_ctor_set(x_118, 1, x_67); +lean_ctor_set(x_118, 2, x_117); +x_119 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_105, 3); +lean_inc(x_87); +x_120 = l_Lean_Syntax_node6(x_87, x_119, x_118, x_104, x_105, x_105, x_105, x_105); if (lean_obj_tag(x_8) == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_119 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_85); -x_120 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_120, 0, x_85); -lean_ctor_set(x_120, 1, x_67); -lean_ctor_set(x_120, 2, x_119); -x_121 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_122 = l_Lean_Syntax_node4(x_85, x_121, x_37, x_80, x_113, x_120); -x_123 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_124 = l_Lean_Syntax_node4(x_85, x_123, x_81, x_35, x_110, x_122); -x_125 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_126 = l_Lean_Syntax_node2(x_85, x_125, x_118, x_124); -lean_inc(x_126); -x_127 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_127, 0, x_126); -x_128 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_126, x_127, x_15, x_16, x_92); -return x_128; +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_121 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_87); +x_122 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_122, 0, x_87); +lean_ctor_set(x_122, 1, x_67); +lean_ctor_set(x_122, 2, x_121); +x_123 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_124 = l_Lean_Syntax_node4(x_87, x_123, x_37, x_82, x_115, x_122); +x_125 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_126 = l_Lean_Syntax_node4(x_87, x_125, x_83, x_35, x_112, x_124); +x_127 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_128 = l_Lean_Syntax_node2(x_87, x_127, x_120, x_126); +lean_inc(x_128); +x_129 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_129, 0, x_128); +x_130 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_128, x_129, x_15, x_16, x_94); +return x_130; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_129 = lean_ctor_get(x_8, 0); -lean_inc(x_129); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_131 = lean_ctor_get(x_8, 0); +lean_inc(x_131); lean_dec(x_8); -x_130 = l_Array_mkArray1___rarg(x_129); -x_131 = l_Array_append___rarg(x_68, x_130); -lean_dec(x_130); -lean_inc(x_85); -x_132 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_132, 0, x_85); -lean_ctor_set(x_132, 1, x_67); -lean_ctor_set(x_132, 2, x_131); -x_133 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_134 = l_Lean_Syntax_node4(x_85, x_133, x_37, x_80, x_113, x_132); -x_135 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_136 = l_Lean_Syntax_node4(x_85, x_135, x_81, x_35, x_110, x_134); -x_137 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_138 = l_Lean_Syntax_node2(x_85, x_137, x_118, x_136); -lean_inc(x_138); -x_139 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_139, 0, x_138); -x_140 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_138, x_139, x_15, x_16, x_92); -return x_140; +x_132 = l_Array_mkArray1___rarg(x_131); +x_133 = l_Array_append___rarg(x_68, x_132); +lean_dec(x_132); +lean_inc(x_87); +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_87); +lean_ctor_set(x_134, 1, x_67); +lean_ctor_set(x_134, 2, x_133); +x_135 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_136 = l_Lean_Syntax_node4(x_87, x_135, x_37, x_82, x_115, x_134); +x_137 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_138 = l_Lean_Syntax_node4(x_87, x_137, x_83, x_35, x_112, x_136); +x_139 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_140 = l_Lean_Syntax_node2(x_87, x_139, x_120, x_138); +lean_inc(x_140); +x_141 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_141, 0, x_140); +x_142 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_140, x_141, x_15, x_16, x_94); +return x_142; } } } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_144 = lean_ctor_get(x_90, 1); -lean_inc(x_144); -lean_dec(x_90); -x_145 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_85); -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_85); -lean_ctor_set(x_146, 1, x_145); -x_147 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_148 = l_Lean_Syntax_SepArray_ofElems(x_147, x_7); -x_149 = l_Array_append___rarg(x_68, x_148); -lean_dec(x_148); -lean_inc(x_85); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_85); -lean_ctor_set(x_150, 1, x_67); -lean_ctor_set(x_150, 2, x_149); -x_151 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_85); -lean_ctor_set_tag(x_86, 2); -lean_ctor_set(x_86, 1, x_151); -lean_ctor_set(x_86, 0, x_85); -x_152 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_85); -x_153 = l_Lean_Syntax_node3(x_85, x_152, x_146, x_150, x_86); -lean_inc(x_85); -x_154 = l_Lean_Syntax_node1(x_85, x_67, x_153); -lean_inc(x_85); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_85); -lean_ctor_set(x_155, 1, x_67); -lean_ctor_set(x_155, 2, x_68); -x_156 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_85); -lean_ctor_set_tag(x_81, 2); -lean_ctor_set(x_81, 1, x_156); -lean_ctor_set(x_81, 0, x_85); -x_157 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_85); +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_146 = lean_ctor_get(x_92, 1); +lean_inc(x_146); +lean_dec(x_92); +x_147 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_87); +x_148 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_148, 0, x_87); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_150 = l_Lean_Syntax_SepArray_ofElems(x_149, x_7); +x_151 = l_Array_append___rarg(x_68, x_150); +lean_dec(x_150); +lean_inc(x_87); +x_152 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_152, 0, x_87); +lean_ctor_set(x_152, 1, x_67); +lean_ctor_set(x_152, 2, x_151); +x_153 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_87); +lean_ctor_set_tag(x_88, 2); +lean_ctor_set(x_88, 1, x_153); +lean_ctor_set(x_88, 0, x_87); +x_154 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_87); +x_155 = l_Lean_Syntax_node3(x_87, x_154, x_148, x_152, x_88); +lean_inc(x_87); +x_156 = l_Lean_Syntax_node1(x_87, x_67, x_155); +lean_inc(x_87); +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_87); +lean_ctor_set(x_157, 1, x_67); +lean_ctor_set(x_157, 2, x_68); +x_158 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_87); +lean_ctor_set_tag(x_83, 2); +lean_ctor_set(x_83, 1, x_158); +lean_ctor_set(x_83, 0, x_87); +x_159 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_157); -lean_ctor_set(x_53, 0, x_85); -x_158 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_85); -x_159 = l_Lean_Syntax_node2(x_85, x_158, x_53, x_34); -lean_inc(x_85); -x_160 = l_Lean_Syntax_node1(x_85, x_67, x_159); -x_161 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_155); -lean_inc(x_85); -x_162 = l_Lean_Syntax_node2(x_85, x_161, x_155, x_160); -x_163 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_85); +lean_ctor_set(x_53, 1, x_159); +lean_ctor_set(x_53, 0, x_87); +x_160 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_87); +x_161 = l_Lean_Syntax_node2(x_87, x_160, x_53, x_34); +lean_inc(x_87); +x_162 = l_Lean_Syntax_node1(x_87, x_67, x_161); +x_163 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_157); +lean_inc(x_87); +x_164 = l_Lean_Syntax_node2(x_87, x_163, x_157, x_162); +x_165 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_163); -lean_ctor_set(x_37, 0, x_85); -x_164 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_155, 2); -lean_inc(x_85); -x_165 = l_Lean_Syntax_node2(x_85, x_164, x_155, x_155); +lean_ctor_set(x_37, 1, x_165); +lean_ctor_set(x_37, 0, x_87); +x_166 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_157, 2); +lean_inc(x_87); +x_167 = l_Lean_Syntax_node2(x_87, x_166, x_157, x_157); if (lean_obj_tag(x_10) == 0) { -x_166 = x_68; -goto block_193; +x_168 = x_68; +goto block_195; } else { -lean_object* x_194; lean_object* x_195; -x_194 = lean_ctor_get(x_10, 0); -lean_inc(x_194); +lean_object* x_196; lean_object* x_197; +x_196 = lean_ctor_get(x_10, 0); +lean_inc(x_196); lean_dec(x_10); -x_195 = l_Array_mkArray1___rarg(x_194); -x_166 = x_195; -goto block_193; +x_197 = l_Array_mkArray1___rarg(x_196); +x_168 = x_197; +goto block_195; } -block_193: +block_195: { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_167 = l_Array_append___rarg(x_68, x_166); -lean_dec(x_166); -lean_inc(x_85); -x_168 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_168, 0, x_85); -lean_ctor_set(x_168, 1, x_67); -lean_ctor_set(x_168, 2, x_167); -x_169 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_155, 3); -lean_inc(x_85); -x_170 = l_Lean_Syntax_node6(x_85, x_169, x_168, x_154, x_155, x_155, x_155, x_155); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_169 = l_Array_append___rarg(x_68, x_168); +lean_dec(x_168); +lean_inc(x_87); +x_170 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_170, 0, x_87); +lean_ctor_set(x_170, 1, x_67); +lean_ctor_set(x_170, 2, x_169); +x_171 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_157, 3); +lean_inc(x_87); +x_172 = l_Lean_Syntax_node6(x_87, x_171, x_170, x_156, x_157, x_157, x_157, x_157); if (lean_obj_tag(x_8) == 0) { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_171 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_85); -x_172 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_172, 0, x_85); -lean_ctor_set(x_172, 1, x_67); -lean_ctor_set(x_172, 2, x_171); -x_173 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_174 = l_Lean_Syntax_node4(x_85, x_173, x_37, x_80, x_165, x_172); -x_175 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_176 = l_Lean_Syntax_node4(x_85, x_175, x_81, x_35, x_162, x_174); -x_177 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_178 = l_Lean_Syntax_node2(x_85, x_177, x_170, x_176); -lean_inc(x_178); -x_179 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_179, 0, x_178); -x_180 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_178, x_179, x_15, x_16, x_144); -return x_180; +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_173 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_87); +x_174 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_174, 0, x_87); +lean_ctor_set(x_174, 1, x_67); +lean_ctor_set(x_174, 2, x_173); +x_175 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_176 = l_Lean_Syntax_node4(x_87, x_175, x_37, x_82, x_167, x_174); +x_177 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_178 = l_Lean_Syntax_node4(x_87, x_177, x_83, x_35, x_164, x_176); +x_179 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_180 = l_Lean_Syntax_node2(x_87, x_179, x_172, x_178); +lean_inc(x_180); +x_181 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_181, 0, x_180); +x_182 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_180, x_181, x_15, x_16, x_146); +return x_182; } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_181 = lean_ctor_get(x_8, 0); -lean_inc(x_181); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_183 = lean_ctor_get(x_8, 0); +lean_inc(x_183); lean_dec(x_8); -x_182 = l_Array_mkArray1___rarg(x_181); -x_183 = l_Array_append___rarg(x_68, x_182); -lean_dec(x_182); -lean_inc(x_85); -x_184 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_184, 0, x_85); -lean_ctor_set(x_184, 1, x_67); -lean_ctor_set(x_184, 2, x_183); -x_185 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_186 = l_Lean_Syntax_node4(x_85, x_185, x_37, x_80, x_165, x_184); -x_187 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_188 = l_Lean_Syntax_node4(x_85, x_187, x_81, x_35, x_162, x_186); -x_189 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_190 = l_Lean_Syntax_node2(x_85, x_189, x_170, x_188); -lean_inc(x_190); -x_191 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_191, 0, x_190); -x_192 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_190, x_191, x_15, x_16, x_144); -return x_192; +x_184 = l_Array_mkArray1___rarg(x_183); +x_185 = l_Array_append___rarg(x_68, x_184); +lean_dec(x_184); +lean_inc(x_87); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_87); +lean_ctor_set(x_186, 1, x_67); +lean_ctor_set(x_186, 2, x_185); +x_187 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_188 = l_Lean_Syntax_node4(x_87, x_187, x_37, x_82, x_167, x_186); +x_189 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_190 = l_Lean_Syntax_node4(x_87, x_189, x_83, x_35, x_164, x_188); +x_191 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_192 = l_Lean_Syntax_node2(x_87, x_191, x_172, x_190); +lean_inc(x_192); +x_193 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_193, 0, x_192); +x_194 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_192, x_193, x_15, x_16, x_146); +return x_194; } } } } else { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_196 = lean_ctor_get(x_86, 1); -lean_inc(x_196); -lean_dec(x_86); -x_197 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_196); -x_198 = lean_ctor_get(x_197, 1); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_198 = lean_ctor_get(x_88, 1); lean_inc(x_198); -if (lean_is_exclusive(x_197)) { - lean_ctor_release(x_197, 0); - lean_ctor_release(x_197, 1); - x_199 = x_197; +lean_dec(x_88); +x_199 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_198); +x_200 = lean_ctor_get(x_199, 1); +lean_inc(x_200); +if (lean_is_exclusive(x_199)) { + lean_ctor_release(x_199, 0); + lean_ctor_release(x_199, 1); + x_201 = x_199; } else { - lean_dec_ref(x_197); - x_199 = lean_box(0); + lean_dec_ref(x_199); + x_201 = lean_box(0); } -x_200 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_85); -if (lean_is_scalar(x_199)) { - x_201 = lean_alloc_ctor(2, 2, 0); +x_202 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_87); +if (lean_is_scalar(x_201)) { + x_203 = lean_alloc_ctor(2, 2, 0); } else { - x_201 = x_199; - lean_ctor_set_tag(x_201, 2); -} -lean_ctor_set(x_201, 0, x_85); -lean_ctor_set(x_201, 1, x_200); -x_202 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_203 = l_Lean_Syntax_SepArray_ofElems(x_202, x_7); -x_204 = l_Array_append___rarg(x_68, x_203); -lean_dec(x_203); -lean_inc(x_85); -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_85); -lean_ctor_set(x_205, 1, x_67); -lean_ctor_set(x_205, 2, x_204); -x_206 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_85); -x_207 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_207, 0, x_85); -lean_ctor_set(x_207, 1, x_206); -x_208 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_85); -x_209 = l_Lean_Syntax_node3(x_85, x_208, x_201, x_205, x_207); -lean_inc(x_85); -x_210 = l_Lean_Syntax_node1(x_85, x_67, x_209); -lean_inc(x_85); -x_211 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_211, 0, x_85); -lean_ctor_set(x_211, 1, x_67); -lean_ctor_set(x_211, 2, x_68); -x_212 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_85); -lean_ctor_set_tag(x_81, 2); -lean_ctor_set(x_81, 1, x_212); -lean_ctor_set(x_81, 0, x_85); -x_213 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_85); + x_203 = x_201; + lean_ctor_set_tag(x_203, 2); +} +lean_ctor_set(x_203, 0, x_87); +lean_ctor_set(x_203, 1, x_202); +x_204 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_205 = l_Lean_Syntax_SepArray_ofElems(x_204, x_7); +x_206 = l_Array_append___rarg(x_68, x_205); +lean_dec(x_205); +lean_inc(x_87); +x_207 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_207, 0, x_87); +lean_ctor_set(x_207, 1, x_67); +lean_ctor_set(x_207, 2, x_206); +x_208 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_87); +x_209 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_209, 0, x_87); +lean_ctor_set(x_209, 1, x_208); +x_210 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_87); +x_211 = l_Lean_Syntax_node3(x_87, x_210, x_203, x_207, x_209); +lean_inc(x_87); +x_212 = l_Lean_Syntax_node1(x_87, x_67, x_211); +lean_inc(x_87); +x_213 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_213, 0, x_87); +lean_ctor_set(x_213, 1, x_67); +lean_ctor_set(x_213, 2, x_68); +x_214 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_87); +lean_ctor_set_tag(x_83, 2); +lean_ctor_set(x_83, 1, x_214); +lean_ctor_set(x_83, 0, x_87); +x_215 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_213); -lean_ctor_set(x_53, 0, x_85); -x_214 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_85); -x_215 = l_Lean_Syntax_node2(x_85, x_214, x_53, x_34); -lean_inc(x_85); -x_216 = l_Lean_Syntax_node1(x_85, x_67, x_215); -x_217 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_211); -lean_inc(x_85); -x_218 = l_Lean_Syntax_node2(x_85, x_217, x_211, x_216); -x_219 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_85); +lean_ctor_set(x_53, 1, x_215); +lean_ctor_set(x_53, 0, x_87); +x_216 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_87); +x_217 = l_Lean_Syntax_node2(x_87, x_216, x_53, x_34); +lean_inc(x_87); +x_218 = l_Lean_Syntax_node1(x_87, x_67, x_217); +x_219 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_213); +lean_inc(x_87); +x_220 = l_Lean_Syntax_node2(x_87, x_219, x_213, x_218); +x_221 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_87); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_219); -lean_ctor_set(x_37, 0, x_85); -x_220 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_211, 2); -lean_inc(x_85); -x_221 = l_Lean_Syntax_node2(x_85, x_220, x_211, x_211); +lean_ctor_set(x_37, 1, x_221); +lean_ctor_set(x_37, 0, x_87); +x_222 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_213, 2); +lean_inc(x_87); +x_223 = l_Lean_Syntax_node2(x_87, x_222, x_213, x_213); if (lean_obj_tag(x_10) == 0) { -x_222 = x_68; -goto block_249; +x_224 = x_68; +goto block_251; } else { -lean_object* x_250; lean_object* x_251; -x_250 = lean_ctor_get(x_10, 0); -lean_inc(x_250); +lean_object* x_252; lean_object* x_253; +x_252 = lean_ctor_get(x_10, 0); +lean_inc(x_252); lean_dec(x_10); -x_251 = l_Array_mkArray1___rarg(x_250); -x_222 = x_251; -goto block_249; +x_253 = l_Array_mkArray1___rarg(x_252); +x_224 = x_253; +goto block_251; } -block_249: +block_251: { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_223 = l_Array_append___rarg(x_68, x_222); -lean_dec(x_222); -lean_inc(x_85); -x_224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_224, 0, x_85); -lean_ctor_set(x_224, 1, x_67); -lean_ctor_set(x_224, 2, x_223); -x_225 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_211, 3); -lean_inc(x_85); -x_226 = l_Lean_Syntax_node6(x_85, x_225, x_224, x_210, x_211, x_211, x_211, x_211); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_225 = l_Array_append___rarg(x_68, x_224); +lean_dec(x_224); +lean_inc(x_87); +x_226 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_226, 0, x_87); +lean_ctor_set(x_226, 1, x_67); +lean_ctor_set(x_226, 2, x_225); +x_227 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_213, 3); +lean_inc(x_87); +x_228 = l_Lean_Syntax_node6(x_87, x_227, x_226, x_212, x_213, x_213, x_213, x_213); if (lean_obj_tag(x_8) == 0) { -lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_227 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_85); -x_228 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_228, 0, x_85); -lean_ctor_set(x_228, 1, x_67); -lean_ctor_set(x_228, 2, x_227); -x_229 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_230 = l_Lean_Syntax_node4(x_85, x_229, x_37, x_80, x_221, x_228); -x_231 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_232 = l_Lean_Syntax_node4(x_85, x_231, x_81, x_35, x_218, x_230); -x_233 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_234 = l_Lean_Syntax_node2(x_85, x_233, x_226, x_232); -lean_inc(x_234); -x_235 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_235, 0, x_234); -x_236 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_234, x_235, x_15, x_16, x_198); -return x_236; +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_229 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_87); +x_230 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_230, 0, x_87); +lean_ctor_set(x_230, 1, x_67); +lean_ctor_set(x_230, 2, x_229); +x_231 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_232 = l_Lean_Syntax_node4(x_87, x_231, x_37, x_82, x_223, x_230); +x_233 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_234 = l_Lean_Syntax_node4(x_87, x_233, x_83, x_35, x_220, x_232); +x_235 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_236 = l_Lean_Syntax_node2(x_87, x_235, x_228, x_234); +lean_inc(x_236); +x_237 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_237, 0, x_236); +x_238 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_236, x_237, x_15, x_16, x_200); +return x_238; } else { -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; -x_237 = lean_ctor_get(x_8, 0); -lean_inc(x_237); +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_239 = lean_ctor_get(x_8, 0); +lean_inc(x_239); lean_dec(x_8); -x_238 = l_Array_mkArray1___rarg(x_237); -x_239 = l_Array_append___rarg(x_68, x_238); -lean_dec(x_238); -lean_inc(x_85); -x_240 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_240, 0, x_85); -lean_ctor_set(x_240, 1, x_67); -lean_ctor_set(x_240, 2, x_239); -x_241 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_85); -x_242 = l_Lean_Syntax_node4(x_85, x_241, x_37, x_80, x_221, x_240); -x_243 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_85); -x_244 = l_Lean_Syntax_node4(x_85, x_243, x_81, x_35, x_218, x_242); -x_245 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_246 = l_Lean_Syntax_node2(x_85, x_245, x_226, x_244); -lean_inc(x_246); -x_247 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_247, 0, x_246); -x_248 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_246, x_247, x_15, x_16, x_198); -return x_248; +x_240 = l_Array_mkArray1___rarg(x_239); +x_241 = l_Array_append___rarg(x_68, x_240); +lean_dec(x_240); +lean_inc(x_87); +x_242 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_242, 0, x_87); +lean_ctor_set(x_242, 1, x_67); +lean_ctor_set(x_242, 2, x_241); +x_243 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_87); +x_244 = l_Lean_Syntax_node4(x_87, x_243, x_37, x_82, x_223, x_242); +x_245 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_87); +x_246 = l_Lean_Syntax_node4(x_87, x_245, x_83, x_35, x_220, x_244); +x_247 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_248 = l_Lean_Syntax_node2(x_87, x_247, x_228, x_246); +lean_inc(x_248); +x_249 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_249, 0, x_248); +x_250 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_248, x_249, x_15, x_16, x_200); +return x_250; } } } } else { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_252 = lean_ctor_get(x_81, 0); -x_253 = lean_ctor_get(x_81, 1); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_81); -x_254 = l_Lean_SourceInfo_fromRef(x_252, x_33); -lean_dec(x_252); -x_255 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_253); -x_256 = lean_ctor_get(x_255, 1); -lean_inc(x_256); -if (lean_is_exclusive(x_255)) { - lean_ctor_release(x_255, 0); - lean_ctor_release(x_255, 1); - x_257 = x_255; +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_254 = lean_ctor_get(x_83, 0); +x_255 = lean_ctor_get(x_83, 1); +lean_inc(x_255); +lean_inc(x_254); +lean_dec(x_83); +x_256 = l_Lean_SourceInfo_fromRef(x_254, x_33); +lean_dec(x_254); +x_257 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_255); +x_258 = lean_ctor_get(x_257, 1); +lean_inc(x_258); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_259 = x_257; +} else { + lean_dec_ref(x_257); + x_259 = lean_box(0); +} +x_260 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_258); +x_261 = lean_ctor_get(x_260, 1); +lean_inc(x_261); +if (lean_is_exclusive(x_260)) { + lean_ctor_release(x_260, 0); + lean_ctor_release(x_260, 1); + x_262 = x_260; } else { - lean_dec_ref(x_255); - x_257 = lean_box(0); + lean_dec_ref(x_260); + x_262 = lean_box(0); } -x_258 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_256); -x_259 = lean_ctor_get(x_258, 1); -lean_inc(x_259); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_260 = x_258; +x_263 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_256); +if (lean_is_scalar(x_262)) { + x_264 = lean_alloc_ctor(2, 2, 0); } else { - lean_dec_ref(x_258); - x_260 = lean_box(0); + x_264 = x_262; + lean_ctor_set_tag(x_264, 2); } -x_261 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_254); -if (lean_is_scalar(x_260)) { - x_262 = lean_alloc_ctor(2, 2, 0); -} else { - x_262 = x_260; - lean_ctor_set_tag(x_262, 2); -} -lean_ctor_set(x_262, 0, x_254); -lean_ctor_set(x_262, 1, x_261); -x_263 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_264 = l_Lean_Syntax_SepArray_ofElems(x_263, x_7); -x_265 = l_Array_append___rarg(x_68, x_264); -lean_dec(x_264); -lean_inc(x_254); -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_254); -lean_ctor_set(x_266, 1, x_67); -lean_ctor_set(x_266, 2, x_265); -x_267 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_254); -if (lean_is_scalar(x_257)) { - x_268 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_256); +lean_ctor_set(x_264, 1, x_263); +x_265 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_266 = l_Lean_Syntax_SepArray_ofElems(x_265, x_7); +x_267 = l_Array_append___rarg(x_68, x_266); +lean_dec(x_266); +lean_inc(x_256); +x_268 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_268, 0, x_256); +lean_ctor_set(x_268, 1, x_67); +lean_ctor_set(x_268, 2, x_267); +x_269 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_256); +if (lean_is_scalar(x_259)) { + x_270 = lean_alloc_ctor(2, 2, 0); } else { - x_268 = x_257; - lean_ctor_set_tag(x_268, 2); + x_270 = x_259; + lean_ctor_set_tag(x_270, 2); } -lean_ctor_set(x_268, 0, x_254); -lean_ctor_set(x_268, 1, x_267); -x_269 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_254); -x_270 = l_Lean_Syntax_node3(x_254, x_269, x_262, x_266, x_268); -lean_inc(x_254); -x_271 = l_Lean_Syntax_node1(x_254, x_67, x_270); -lean_inc(x_254); -x_272 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_272, 0, x_254); -lean_ctor_set(x_272, 1, x_67); -lean_ctor_set(x_272, 2, x_68); -x_273 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_254); -x_274 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_274, 0, x_254); -lean_ctor_set(x_274, 1, x_273); -x_275 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_254); +lean_ctor_set(x_270, 0, x_256); +lean_ctor_set(x_270, 1, x_269); +x_271 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_256); +x_272 = l_Lean_Syntax_node3(x_256, x_271, x_264, x_268, x_270); +lean_inc(x_256); +x_273 = l_Lean_Syntax_node1(x_256, x_67, x_272); +lean_inc(x_256); +x_274 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_274, 0, x_256); +lean_ctor_set(x_274, 1, x_67); +lean_ctor_set(x_274, 2, x_68); +x_275 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_256); +x_276 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_276, 0, x_256); +lean_ctor_set(x_276, 1, x_275); +x_277 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_256); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_275); -lean_ctor_set(x_53, 0, x_254); -x_276 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_254); -x_277 = l_Lean_Syntax_node2(x_254, x_276, x_53, x_34); -lean_inc(x_254); -x_278 = l_Lean_Syntax_node1(x_254, x_67, x_277); -x_279 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_272); -lean_inc(x_254); -x_280 = l_Lean_Syntax_node2(x_254, x_279, x_272, x_278); -x_281 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_254); +lean_ctor_set(x_53, 1, x_277); +lean_ctor_set(x_53, 0, x_256); +x_278 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_256); +x_279 = l_Lean_Syntax_node2(x_256, x_278, x_53, x_34); +lean_inc(x_256); +x_280 = l_Lean_Syntax_node1(x_256, x_67, x_279); +x_281 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_274); +lean_inc(x_256); +x_282 = l_Lean_Syntax_node2(x_256, x_281, x_274, x_280); +x_283 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_256); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_281); -lean_ctor_set(x_37, 0, x_254); -x_282 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_272, 2); -lean_inc(x_254); -x_283 = l_Lean_Syntax_node2(x_254, x_282, x_272, x_272); +lean_ctor_set(x_37, 1, x_283); +lean_ctor_set(x_37, 0, x_256); +x_284 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_274, 2); +lean_inc(x_256); +x_285 = l_Lean_Syntax_node2(x_256, x_284, x_274, x_274); if (lean_obj_tag(x_10) == 0) { -x_284 = x_68; -goto block_311; +x_286 = x_68; +goto block_313; } else { -lean_object* x_312; lean_object* x_313; -x_312 = lean_ctor_get(x_10, 0); -lean_inc(x_312); +lean_object* x_314; lean_object* x_315; +x_314 = lean_ctor_get(x_10, 0); +lean_inc(x_314); lean_dec(x_10); -x_313 = l_Array_mkArray1___rarg(x_312); -x_284 = x_313; -goto block_311; +x_315 = l_Array_mkArray1___rarg(x_314); +x_286 = x_315; +goto block_313; } -block_311: +block_313: { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; -x_285 = l_Array_append___rarg(x_68, x_284); -lean_dec(x_284); -lean_inc(x_254); -x_286 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_286, 0, x_254); -lean_ctor_set(x_286, 1, x_67); -lean_ctor_set(x_286, 2, x_285); -x_287 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_272, 3); -lean_inc(x_254); -x_288 = l_Lean_Syntax_node6(x_254, x_287, x_286, x_271, x_272, x_272, x_272, x_272); +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; +x_287 = l_Array_append___rarg(x_68, x_286); +lean_dec(x_286); +lean_inc(x_256); +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_256); +lean_ctor_set(x_288, 1, x_67); +lean_ctor_set(x_288, 2, x_287); +x_289 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_274, 3); +lean_inc(x_256); +x_290 = l_Lean_Syntax_node6(x_256, x_289, x_288, x_273, x_274, x_274, x_274, x_274); if (lean_obj_tag(x_8) == 0) { -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_289 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_254); -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_254); -lean_ctor_set(x_290, 1, x_67); -lean_ctor_set(x_290, 2, x_289); -x_291 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_254); -x_292 = l_Lean_Syntax_node4(x_254, x_291, x_37, x_80, x_283, x_290); -x_293 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_254); -x_294 = l_Lean_Syntax_node4(x_254, x_293, x_274, x_35, x_280, x_292); -x_295 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_296 = l_Lean_Syntax_node2(x_254, x_295, x_288, x_294); -lean_inc(x_296); -x_297 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_297, 0, x_296); -x_298 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_296, x_297, x_15, x_16, x_259); -return x_298; +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_291 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_256); +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_256); +lean_ctor_set(x_292, 1, x_67); +lean_ctor_set(x_292, 2, x_291); +x_293 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_256); +x_294 = l_Lean_Syntax_node4(x_256, x_293, x_37, x_82, x_285, x_292); +x_295 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_256); +x_296 = l_Lean_Syntax_node4(x_256, x_295, x_276, x_35, x_282, x_294); +x_297 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_298 = l_Lean_Syntax_node2(x_256, x_297, x_290, x_296); +lean_inc(x_298); +x_299 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_299, 0, x_298); +x_300 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_298, x_299, x_15, x_16, x_261); +return x_300; } else { -lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -x_299 = lean_ctor_get(x_8, 0); -lean_inc(x_299); +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_301 = lean_ctor_get(x_8, 0); +lean_inc(x_301); lean_dec(x_8); -x_300 = l_Array_mkArray1___rarg(x_299); -x_301 = l_Array_append___rarg(x_68, x_300); -lean_dec(x_300); -lean_inc(x_254); -x_302 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_302, 0, x_254); -lean_ctor_set(x_302, 1, x_67); -lean_ctor_set(x_302, 2, x_301); -x_303 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_254); -x_304 = l_Lean_Syntax_node4(x_254, x_303, x_37, x_80, x_283, x_302); -x_305 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_254); -x_306 = l_Lean_Syntax_node4(x_254, x_305, x_274, x_35, x_280, x_304); -x_307 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_308 = l_Lean_Syntax_node2(x_254, x_307, x_288, x_306); -lean_inc(x_308); -x_309 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_309, 0, x_308); -x_310 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_308, x_309, x_15, x_16, x_259); -return x_310; +x_302 = l_Array_mkArray1___rarg(x_301); +x_303 = l_Array_append___rarg(x_68, x_302); +lean_dec(x_302); +lean_inc(x_256); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_256); +lean_ctor_set(x_304, 1, x_67); +lean_ctor_set(x_304, 2, x_303); +x_305 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_256); +x_306 = l_Lean_Syntax_node4(x_256, x_305, x_37, x_82, x_285, x_304); +x_307 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_256); +x_308 = l_Lean_Syntax_node4(x_256, x_307, x_276, x_35, x_282, x_306); +x_309 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_310 = l_Lean_Syntax_node2(x_256, x_309, x_290, x_308); +lean_inc(x_310); +x_311 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_311, 0, x_310); +x_312 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_310, x_311, x_15, x_16, x_261); +return x_312; } } } } else { -lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; size_t x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; -x_314 = lean_ctor_get(x_62, 1); -lean_inc(x_314); +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; size_t x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +x_316 = lean_ctor_get(x_62, 1); +lean_inc(x_316); lean_dec(x_62); -x_315 = l_Lake_DSL_structVal___closed__3; +x_317 = l_Lake_DSL_structVal___closed__3; lean_inc(x_57); -x_316 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_316, 0, x_57); -lean_ctor_set(x_316, 1, x_315); -x_317 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_318 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +x_318 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_318, 0, x_57); +lean_ctor_set(x_318, 1, x_317); +x_319 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_320 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_57); -x_319 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_319, 0, x_57); -lean_ctor_set(x_319, 1, x_317); -lean_ctor_set(x_319, 2, x_318); -x_320 = lean_array_size(x_27); -x_321 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_320, x_20, x_27); -x_322 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_323 = l_Lean_mkSepArray(x_321, x_322); -lean_dec(x_321); -x_324 = l_Array_append___rarg(x_318, x_323); +x_321 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_321, 0, x_57); +lean_ctor_set(x_321, 1, x_319); +lean_ctor_set(x_321, 2, x_320); +x_322 = lean_array_size(x_27); +x_323 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_322, x_20, x_27); +x_324 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_325 = l_Lean_mkSepArray(x_323, x_324); lean_dec(x_323); +x_326 = l_Array_append___rarg(x_320, x_325); +lean_dec(x_325); lean_inc(x_57); -x_325 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_325, 0, x_57); -lean_ctor_set(x_325, 1, x_317); -lean_ctor_set(x_325, 2, x_324); -x_326 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_319); +x_327 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_327, 0, x_57); +lean_ctor_set(x_327, 1, x_319); +lean_ctor_set(x_327, 2, x_326); +x_328 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; lean_inc(x_57); -x_327 = l_Lean_Syntax_node1(x_57, x_326, x_319); -x_328 = l_Lake_DSL_structVal___closed__18; +x_329 = l_Lean_Syntax_node1(x_57, x_328, x_327); +x_330 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_321); +lean_inc(x_57); +x_331 = l_Lean_Syntax_node1(x_57, x_330, x_321); +x_332 = l_Lake_DSL_structVal___closed__18; lean_inc(x_57); lean_ctor_set_tag(x_58, 2); -lean_ctor_set(x_58, 1, x_328); +lean_ctor_set(x_58, 1, x_332); lean_ctor_set(x_58, 0, x_57); -x_329 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_319); -x_330 = l_Lean_Syntax_node6(x_57, x_329, x_316, x_319, x_325, x_327, x_319, x_58); -x_331 = l_Lean_Elab_Command_getRef(x_15, x_16, x_314); -x_332 = lean_ctor_get(x_331, 0); -lean_inc(x_332); -x_333 = lean_ctor_get(x_331, 1); -lean_inc(x_333); -if (lean_is_exclusive(x_331)) { - lean_ctor_release(x_331, 0); - lean_ctor_release(x_331, 1); - x_334 = x_331; -} else { - lean_dec_ref(x_331); - x_334 = lean_box(0); -} -x_335 = l_Lean_SourceInfo_fromRef(x_332, x_33); -lean_dec(x_332); -x_336 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_333); -x_337 = lean_ctor_get(x_336, 1); -lean_inc(x_337); -if (lean_is_exclusive(x_336)) { - lean_ctor_release(x_336, 0); - lean_ctor_release(x_336, 1); - x_338 = x_336; +x_333 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_321); +x_334 = l_Lean_Syntax_node6(x_57, x_333, x_318, x_321, x_329, x_331, x_321, x_58); +x_335 = l_Lean_Elab_Command_getRef(x_15, x_16, x_316); +x_336 = lean_ctor_get(x_335, 0); +lean_inc(x_336); +x_337 = lean_ctor_get(x_335, 1); +lean_inc(x_337); +if (lean_is_exclusive(x_335)) { + lean_ctor_release(x_335, 0); + lean_ctor_release(x_335, 1); + x_338 = x_335; } else { - lean_dec_ref(x_336); + lean_dec_ref(x_335); x_338 = lean_box(0); } -x_339 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_337); -x_340 = lean_ctor_get(x_339, 1); -lean_inc(x_340); -if (lean_is_exclusive(x_339)) { - lean_ctor_release(x_339, 0); - lean_ctor_release(x_339, 1); - x_341 = x_339; -} else { - lean_dec_ref(x_339); - x_341 = lean_box(0); -} -x_342 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_335); -if (lean_is_scalar(x_341)) { - x_343 = lean_alloc_ctor(2, 2, 0); -} else { - x_343 = x_341; - lean_ctor_set_tag(x_343, 2); -} -lean_ctor_set(x_343, 0, x_335); -lean_ctor_set(x_343, 1, x_342); -x_344 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_345 = l_Lean_Syntax_SepArray_ofElems(x_344, x_7); -x_346 = l_Array_append___rarg(x_318, x_345); -lean_dec(x_345); -lean_inc(x_335); -x_347 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_347, 0, x_335); -lean_ctor_set(x_347, 1, x_317); -lean_ctor_set(x_347, 2, x_346); -x_348 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_335); +x_339 = l_Lean_SourceInfo_fromRef(x_336, x_33); +lean_dec(x_336); +x_340 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_337); +x_341 = lean_ctor_get(x_340, 1); +lean_inc(x_341); +if (lean_is_exclusive(x_340)) { + lean_ctor_release(x_340, 0); + lean_ctor_release(x_340, 1); + x_342 = x_340; +} else { + lean_dec_ref(x_340); + x_342 = lean_box(0); +} +x_343 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_341); +x_344 = lean_ctor_get(x_343, 1); +lean_inc(x_344); +if (lean_is_exclusive(x_343)) { + lean_ctor_release(x_343, 0); + lean_ctor_release(x_343, 1); + x_345 = x_343; +} else { + lean_dec_ref(x_343); + x_345 = lean_box(0); +} +x_346 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_339); +if (lean_is_scalar(x_345)) { + x_347 = lean_alloc_ctor(2, 2, 0); +} else { + x_347 = x_345; + lean_ctor_set_tag(x_347, 2); +} +lean_ctor_set(x_347, 0, x_339); +lean_ctor_set(x_347, 1, x_346); +x_348 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_349 = l_Lean_Syntax_SepArray_ofElems(x_348, x_7); +x_350 = l_Array_append___rarg(x_320, x_349); +lean_dec(x_349); +lean_inc(x_339); +x_351 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_351, 0, x_339); +lean_ctor_set(x_351, 1, x_319); +lean_ctor_set(x_351, 2, x_350); +x_352 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_339); +if (lean_is_scalar(x_342)) { + x_353 = lean_alloc_ctor(2, 2, 0); +} else { + x_353 = x_342; + lean_ctor_set_tag(x_353, 2); +} +lean_ctor_set(x_353, 0, x_339); +lean_ctor_set(x_353, 1, x_352); +x_354 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_339); +x_355 = l_Lean_Syntax_node3(x_339, x_354, x_347, x_351, x_353); +lean_inc(x_339); +x_356 = l_Lean_Syntax_node1(x_339, x_319, x_355); +lean_inc(x_339); +x_357 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_357, 0, x_339); +lean_ctor_set(x_357, 1, x_319); +lean_ctor_set(x_357, 2, x_320); +x_358 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_339); if (lean_is_scalar(x_338)) { - x_349 = lean_alloc_ctor(2, 2, 0); -} else { - x_349 = x_338; - lean_ctor_set_tag(x_349, 2); -} -lean_ctor_set(x_349, 0, x_335); -lean_ctor_set(x_349, 1, x_348); -x_350 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_335); -x_351 = l_Lean_Syntax_node3(x_335, x_350, x_343, x_347, x_349); -lean_inc(x_335); -x_352 = l_Lean_Syntax_node1(x_335, x_317, x_351); -lean_inc(x_335); -x_353 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_353, 0, x_335); -lean_ctor_set(x_353, 1, x_317); -lean_ctor_set(x_353, 2, x_318); -x_354 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_335); -if (lean_is_scalar(x_334)) { - x_355 = lean_alloc_ctor(2, 2, 0); + x_359 = lean_alloc_ctor(2, 2, 0); } else { - x_355 = x_334; - lean_ctor_set_tag(x_355, 2); + x_359 = x_338; + lean_ctor_set_tag(x_359, 2); } -lean_ctor_set(x_355, 0, x_335); -lean_ctor_set(x_355, 1, x_354); -x_356 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_335); +lean_ctor_set(x_359, 0, x_339); +lean_ctor_set(x_359, 1, x_358); +x_360 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_339); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_356); -lean_ctor_set(x_53, 0, x_335); -x_357 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_335); -x_358 = l_Lean_Syntax_node2(x_335, x_357, x_53, x_34); -lean_inc(x_335); -x_359 = l_Lean_Syntax_node1(x_335, x_317, x_358); -x_360 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_353); -lean_inc(x_335); -x_361 = l_Lean_Syntax_node2(x_335, x_360, x_353, x_359); -x_362 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_335); +lean_ctor_set(x_53, 1, x_360); +lean_ctor_set(x_53, 0, x_339); +x_361 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_339); +x_362 = l_Lean_Syntax_node2(x_339, x_361, x_53, x_34); +lean_inc(x_339); +x_363 = l_Lean_Syntax_node1(x_339, x_319, x_362); +x_364 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_357); +lean_inc(x_339); +x_365 = l_Lean_Syntax_node2(x_339, x_364, x_357, x_363); +x_366 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_339); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_362); -lean_ctor_set(x_37, 0, x_335); -x_363 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_353, 2); -lean_inc(x_335); -x_364 = l_Lean_Syntax_node2(x_335, x_363, x_353, x_353); +lean_ctor_set(x_37, 1, x_366); +lean_ctor_set(x_37, 0, x_339); +x_367 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_357, 2); +lean_inc(x_339); +x_368 = l_Lean_Syntax_node2(x_339, x_367, x_357, x_357); if (lean_obj_tag(x_10) == 0) { -x_365 = x_318; -goto block_392; +x_369 = x_320; +goto block_396; } else { -lean_object* x_393; lean_object* x_394; -x_393 = lean_ctor_get(x_10, 0); -lean_inc(x_393); +lean_object* x_397; lean_object* x_398; +x_397 = lean_ctor_get(x_10, 0); +lean_inc(x_397); lean_dec(x_10); -x_394 = l_Array_mkArray1___rarg(x_393); -x_365 = x_394; -goto block_392; +x_398 = l_Array_mkArray1___rarg(x_397); +x_369 = x_398; +goto block_396; } -block_392: -{ -lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; -x_366 = l_Array_append___rarg(x_318, x_365); -lean_dec(x_365); -lean_inc(x_335); -x_367 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_367, 0, x_335); -lean_ctor_set(x_367, 1, x_317); -lean_ctor_set(x_367, 2, x_366); -x_368 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_353, 3); -lean_inc(x_335); -x_369 = l_Lean_Syntax_node6(x_335, x_368, x_367, x_352, x_353, x_353, x_353, x_353); -if (lean_obj_tag(x_8) == 0) +block_396: { -lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -x_370 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_335); +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +x_370 = l_Array_append___rarg(x_320, x_369); +lean_dec(x_369); +lean_inc(x_339); x_371 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_371, 0, x_335); -lean_ctor_set(x_371, 1, x_317); +lean_ctor_set(x_371, 0, x_339); +lean_ctor_set(x_371, 1, x_319); lean_ctor_set(x_371, 2, x_370); -x_372 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_335); -x_373 = l_Lean_Syntax_node4(x_335, x_372, x_37, x_330, x_364, x_371); -x_374 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_335); -x_375 = l_Lean_Syntax_node4(x_335, x_374, x_355, x_35, x_361, x_373); -x_376 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_377 = l_Lean_Syntax_node2(x_335, x_376, x_369, x_375); -lean_inc(x_377); -x_378 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_378, 0, x_377); -x_379 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_377, x_378, x_15, x_16, x_340); -return x_379; +x_372 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_357, 3); +lean_inc(x_339); +x_373 = l_Lean_Syntax_node6(x_339, x_372, x_371, x_356, x_357, x_357, x_357, x_357); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_374 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_339); +x_375 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_375, 0, x_339); +lean_ctor_set(x_375, 1, x_319); +lean_ctor_set(x_375, 2, x_374); +x_376 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_339); +x_377 = l_Lean_Syntax_node4(x_339, x_376, x_37, x_334, x_368, x_375); +x_378 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_339); +x_379 = l_Lean_Syntax_node4(x_339, x_378, x_359, x_35, x_365, x_377); +x_380 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_381 = l_Lean_Syntax_node2(x_339, x_380, x_373, x_379); +lean_inc(x_381); +x_382 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_382, 0, x_381); +x_383 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_381, x_382, x_15, x_16, x_344); +return x_383; } else { -lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; -x_380 = lean_ctor_get(x_8, 0); -lean_inc(x_380); +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; +x_384 = lean_ctor_get(x_8, 0); +lean_inc(x_384); lean_dec(x_8); -x_381 = l_Array_mkArray1___rarg(x_380); -x_382 = l_Array_append___rarg(x_318, x_381); -lean_dec(x_381); -lean_inc(x_335); -x_383 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_383, 0, x_335); -lean_ctor_set(x_383, 1, x_317); -lean_ctor_set(x_383, 2, x_382); -x_384 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_335); -x_385 = l_Lean_Syntax_node4(x_335, x_384, x_37, x_330, x_364, x_383); -x_386 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_335); -x_387 = l_Lean_Syntax_node4(x_335, x_386, x_355, x_35, x_361, x_385); -x_388 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_389 = l_Lean_Syntax_node2(x_335, x_388, x_369, x_387); -lean_inc(x_389); -x_390 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_390, 0, x_389); -x_391 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_389, x_390, x_15, x_16, x_340); -return x_391; +x_385 = l_Array_mkArray1___rarg(x_384); +x_386 = l_Array_append___rarg(x_320, x_385); +lean_dec(x_385); +lean_inc(x_339); +x_387 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_387, 0, x_339); +lean_ctor_set(x_387, 1, x_319); +lean_ctor_set(x_387, 2, x_386); +x_388 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_339); +x_389 = l_Lean_Syntax_node4(x_339, x_388, x_37, x_334, x_368, x_387); +x_390 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_339); +x_391 = l_Lean_Syntax_node4(x_339, x_390, x_359, x_35, x_365, x_389); +x_392 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_393 = l_Lean_Syntax_node2(x_339, x_392, x_373, x_391); +lean_inc(x_393); +x_394 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_394, 0, x_393); +x_395 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_393, x_394, x_15, x_16, x_344); +return x_395; } } } } else { -lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; size_t x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; -x_395 = lean_ctor_get(x_58, 1); -lean_inc(x_395); +lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; size_t x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; +x_399 = lean_ctor_get(x_58, 1); +lean_inc(x_399); lean_dec(x_58); -x_396 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_395); -x_397 = lean_ctor_get(x_396, 1); -lean_inc(x_397); -if (lean_is_exclusive(x_396)) { - lean_ctor_release(x_396, 0); - lean_ctor_release(x_396, 1); - x_398 = x_396; -} else { - lean_dec_ref(x_396); - x_398 = lean_box(0); -} -x_399 = l_Lake_DSL_structVal___closed__3; +x_400 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_399); +x_401 = lean_ctor_get(x_400, 1); +lean_inc(x_401); +if (lean_is_exclusive(x_400)) { + lean_ctor_release(x_400, 0); + lean_ctor_release(x_400, 1); + x_402 = x_400; +} else { + lean_dec_ref(x_400); + x_402 = lean_box(0); +} +x_403 = l_Lake_DSL_structVal___closed__3; lean_inc(x_57); -if (lean_is_scalar(x_398)) { - x_400 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_402)) { + x_404 = lean_alloc_ctor(2, 2, 0); } else { - x_400 = x_398; - lean_ctor_set_tag(x_400, 2); + x_404 = x_402; + lean_ctor_set_tag(x_404, 2); } -lean_ctor_set(x_400, 0, x_57); -lean_ctor_set(x_400, 1, x_399); -x_401 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_402 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_57); -x_403 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_403, 0, x_57); -lean_ctor_set(x_403, 1, x_401); -lean_ctor_set(x_403, 2, x_402); -x_404 = lean_array_size(x_27); -x_405 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_404, x_20, x_27); -x_406 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_407 = l_Lean_mkSepArray(x_405, x_406); -lean_dec(x_405); -x_408 = l_Array_append___rarg(x_402, x_407); -lean_dec(x_407); +lean_ctor_set(x_404, 0, x_57); +lean_ctor_set(x_404, 1, x_403); +x_405 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_406 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_57); -x_409 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_409, 0, x_57); -lean_ctor_set(x_409, 1, x_401); -lean_ctor_set(x_409, 2, x_408); +x_407 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_407, 0, x_57); +lean_ctor_set(x_407, 1, x_405); +lean_ctor_set(x_407, 2, x_406); +x_408 = lean_array_size(x_27); +x_409 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_408, x_20, x_27); x_410 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_403); -lean_inc(x_57); -x_411 = l_Lean_Syntax_node1(x_57, x_410, x_403); -x_412 = l_Lake_DSL_structVal___closed__18; +x_411 = l_Lean_mkSepArray(x_409, x_410); +lean_dec(x_409); +x_412 = l_Array_append___rarg(x_406, x_411); +lean_dec(x_411); lean_inc(x_57); -x_413 = lean_alloc_ctor(2, 2, 0); +x_413 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_413, 0, x_57); -lean_ctor_set(x_413, 1, x_412); -x_414 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_403); -x_415 = l_Lean_Syntax_node6(x_57, x_414, x_400, x_403, x_409, x_411, x_403, x_413); -x_416 = l_Lean_Elab_Command_getRef(x_15, x_16, x_397); -x_417 = lean_ctor_get(x_416, 0); -lean_inc(x_417); -x_418 = lean_ctor_get(x_416, 1); -lean_inc(x_418); -if (lean_is_exclusive(x_416)) { - lean_ctor_release(x_416, 0); - lean_ctor_release(x_416, 1); - x_419 = x_416; -} else { - lean_dec_ref(x_416); - x_419 = lean_box(0); -} -x_420 = l_Lean_SourceInfo_fromRef(x_417, x_33); -lean_dec(x_417); -x_421 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_418); -x_422 = lean_ctor_get(x_421, 1); -lean_inc(x_422); -if (lean_is_exclusive(x_421)) { - lean_ctor_release(x_421, 0); - lean_ctor_release(x_421, 1); - x_423 = x_421; -} else { - lean_dec_ref(x_421); - x_423 = lean_box(0); -} -x_424 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_422); -x_425 = lean_ctor_get(x_424, 1); -lean_inc(x_425); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - x_426 = x_424; -} else { - lean_dec_ref(x_424); - x_426 = lean_box(0); -} -x_427 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_420); -if (lean_is_scalar(x_426)) { - x_428 = lean_alloc_ctor(2, 2, 0); -} else { - x_428 = x_426; - lean_ctor_set_tag(x_428, 2); -} -lean_ctor_set(x_428, 0, x_420); -lean_ctor_set(x_428, 1, x_427); -x_429 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_430 = l_Lean_Syntax_SepArray_ofElems(x_429, x_7); -x_431 = l_Array_append___rarg(x_402, x_430); -lean_dec(x_430); -lean_inc(x_420); -x_432 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_432, 0, x_420); -lean_ctor_set(x_432, 1, x_401); -lean_ctor_set(x_432, 2, x_431); -x_433 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_420); -if (lean_is_scalar(x_423)) { +lean_ctor_set(x_413, 1, x_405); +lean_ctor_set(x_413, 2, x_412); +x_414 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_57); +x_415 = l_Lean_Syntax_node1(x_57, x_414, x_413); +x_416 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_407); +lean_inc(x_57); +x_417 = l_Lean_Syntax_node1(x_57, x_416, x_407); +x_418 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_57); +x_419 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_419, 0, x_57); +lean_ctor_set(x_419, 1, x_418); +x_420 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_407); +x_421 = l_Lean_Syntax_node6(x_57, x_420, x_404, x_407, x_415, x_417, x_407, x_419); +x_422 = l_Lean_Elab_Command_getRef(x_15, x_16, x_401); +x_423 = lean_ctor_get(x_422, 0); +lean_inc(x_423); +x_424 = lean_ctor_get(x_422, 1); +lean_inc(x_424); +if (lean_is_exclusive(x_422)) { + lean_ctor_release(x_422, 0); + lean_ctor_release(x_422, 1); + x_425 = x_422; +} else { + lean_dec_ref(x_422); + x_425 = lean_box(0); +} +x_426 = l_Lean_SourceInfo_fromRef(x_423, x_33); +lean_dec(x_423); +x_427 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_424); +x_428 = lean_ctor_get(x_427, 1); +lean_inc(x_428); +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + x_429 = x_427; +} else { + lean_dec_ref(x_427); + x_429 = lean_box(0); +} +x_430 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_428); +x_431 = lean_ctor_get(x_430, 1); +lean_inc(x_431); +if (lean_is_exclusive(x_430)) { + lean_ctor_release(x_430, 0); + lean_ctor_release(x_430, 1); + x_432 = x_430; +} else { + lean_dec_ref(x_430); + x_432 = lean_box(0); +} +x_433 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_426); +if (lean_is_scalar(x_432)) { x_434 = lean_alloc_ctor(2, 2, 0); } else { - x_434 = x_423; + x_434 = x_432; lean_ctor_set_tag(x_434, 2); } -lean_ctor_set(x_434, 0, x_420); +lean_ctor_set(x_434, 0, x_426); lean_ctor_set(x_434, 1, x_433); -x_435 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_420); -x_436 = l_Lean_Syntax_node3(x_420, x_435, x_428, x_432, x_434); -lean_inc(x_420); -x_437 = l_Lean_Syntax_node1(x_420, x_401, x_436); -lean_inc(x_420); +x_435 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_436 = l_Lean_Syntax_SepArray_ofElems(x_435, x_7); +x_437 = l_Array_append___rarg(x_406, x_436); +lean_dec(x_436); +lean_inc(x_426); x_438 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_438, 0, x_420); -lean_ctor_set(x_438, 1, x_401); -lean_ctor_set(x_438, 2, x_402); -x_439 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_420); -if (lean_is_scalar(x_419)) { +lean_ctor_set(x_438, 0, x_426); +lean_ctor_set(x_438, 1, x_405); +lean_ctor_set(x_438, 2, x_437); +x_439 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_426); +if (lean_is_scalar(x_429)) { x_440 = lean_alloc_ctor(2, 2, 0); } else { - x_440 = x_419; + x_440 = x_429; lean_ctor_set_tag(x_440, 2); } -lean_ctor_set(x_440, 0, x_420); +lean_ctor_set(x_440, 0, x_426); lean_ctor_set(x_440, 1, x_439); -x_441 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_420); +x_441 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_426); +x_442 = l_Lean_Syntax_node3(x_426, x_441, x_434, x_438, x_440); +lean_inc(x_426); +x_443 = l_Lean_Syntax_node1(x_426, x_405, x_442); +lean_inc(x_426); +x_444 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_444, 0, x_426); +lean_ctor_set(x_444, 1, x_405); +lean_ctor_set(x_444, 2, x_406); +x_445 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_426); +if (lean_is_scalar(x_425)) { + x_446 = lean_alloc_ctor(2, 2, 0); +} else { + x_446 = x_425; + lean_ctor_set_tag(x_446, 2); +} +lean_ctor_set(x_446, 0, x_426); +lean_ctor_set(x_446, 1, x_445); +x_447 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_426); lean_ctor_set_tag(x_53, 2); -lean_ctor_set(x_53, 1, x_441); -lean_ctor_set(x_53, 0, x_420); -x_442 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_420); -x_443 = l_Lean_Syntax_node2(x_420, x_442, x_53, x_34); -lean_inc(x_420); -x_444 = l_Lean_Syntax_node1(x_420, x_401, x_443); -x_445 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_438); -lean_inc(x_420); -x_446 = l_Lean_Syntax_node2(x_420, x_445, x_438, x_444); -x_447 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_420); +lean_ctor_set(x_53, 1, x_447); +lean_ctor_set(x_53, 0, x_426); +x_448 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_426); +x_449 = l_Lean_Syntax_node2(x_426, x_448, x_53, x_34); +lean_inc(x_426); +x_450 = l_Lean_Syntax_node1(x_426, x_405, x_449); +x_451 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_444); +lean_inc(x_426); +x_452 = l_Lean_Syntax_node2(x_426, x_451, x_444, x_450); +x_453 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_426); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_447); -lean_ctor_set(x_37, 0, x_420); -x_448 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_438, 2); -lean_inc(x_420); -x_449 = l_Lean_Syntax_node2(x_420, x_448, x_438, x_438); +lean_ctor_set(x_37, 1, x_453); +lean_ctor_set(x_37, 0, x_426); +x_454 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_444, 2); +lean_inc(x_426); +x_455 = l_Lean_Syntax_node2(x_426, x_454, x_444, x_444); if (lean_obj_tag(x_10) == 0) { -x_450 = x_402; -goto block_477; +x_456 = x_406; +goto block_483; } else { -lean_object* x_478; lean_object* x_479; -x_478 = lean_ctor_get(x_10, 0); -lean_inc(x_478); +lean_object* x_484; lean_object* x_485; +x_484 = lean_ctor_get(x_10, 0); +lean_inc(x_484); lean_dec(x_10); -x_479 = l_Array_mkArray1___rarg(x_478); -x_450 = x_479; -goto block_477; -} -block_477: -{ -lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; -x_451 = l_Array_append___rarg(x_402, x_450); -lean_dec(x_450); -lean_inc(x_420); -x_452 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_452, 0, x_420); -lean_ctor_set(x_452, 1, x_401); -lean_ctor_set(x_452, 2, x_451); -x_453 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_438, 3); -lean_inc(x_420); -x_454 = l_Lean_Syntax_node6(x_420, x_453, x_452, x_437, x_438, x_438, x_438, x_438); +x_485 = l_Array_mkArray1___rarg(x_484); +x_456 = x_485; +goto block_483; +} +block_483: +{ +lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; +x_457 = l_Array_append___rarg(x_406, x_456); +lean_dec(x_456); +lean_inc(x_426); +x_458 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_458, 0, x_426); +lean_ctor_set(x_458, 1, x_405); +lean_ctor_set(x_458, 2, x_457); +x_459 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_444, 3); +lean_inc(x_426); +x_460 = l_Lean_Syntax_node6(x_426, x_459, x_458, x_443, x_444, x_444, x_444, x_444); if (lean_obj_tag(x_8) == 0) { -lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; -x_455 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_420); -x_456 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_456, 0, x_420); -lean_ctor_set(x_456, 1, x_401); -lean_ctor_set(x_456, 2, x_455); -x_457 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_420); -x_458 = l_Lean_Syntax_node4(x_420, x_457, x_37, x_415, x_449, x_456); -x_459 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_420); -x_460 = l_Lean_Syntax_node4(x_420, x_459, x_440, x_35, x_446, x_458); -x_461 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_462 = l_Lean_Syntax_node2(x_420, x_461, x_454, x_460); -lean_inc(x_462); -x_463 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_463, 0, x_462); -x_464 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_462, x_463, x_15, x_16, x_425); -return x_464; -} -else -{ -lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; -x_465 = lean_ctor_get(x_8, 0); -lean_inc(x_465); +lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; +x_461 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_426); +x_462 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_462, 0, x_426); +lean_ctor_set(x_462, 1, x_405); +lean_ctor_set(x_462, 2, x_461); +x_463 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_426); +x_464 = l_Lean_Syntax_node4(x_426, x_463, x_37, x_421, x_455, x_462); +x_465 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_426); +x_466 = l_Lean_Syntax_node4(x_426, x_465, x_446, x_35, x_452, x_464); +x_467 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_468 = l_Lean_Syntax_node2(x_426, x_467, x_460, x_466); +lean_inc(x_468); +x_469 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_469, 0, x_468); +x_470 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_468, x_469, x_15, x_16, x_431); +return x_470; +} +else +{ +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; +x_471 = lean_ctor_get(x_8, 0); +lean_inc(x_471); lean_dec(x_8); -x_466 = l_Array_mkArray1___rarg(x_465); -x_467 = l_Array_append___rarg(x_402, x_466); -lean_dec(x_466); -lean_inc(x_420); -x_468 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_468, 0, x_420); -lean_ctor_set(x_468, 1, x_401); -lean_ctor_set(x_468, 2, x_467); -x_469 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_420); -x_470 = l_Lean_Syntax_node4(x_420, x_469, x_37, x_415, x_449, x_468); -x_471 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_420); -x_472 = l_Lean_Syntax_node4(x_420, x_471, x_440, x_35, x_446, x_470); -x_473 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_474 = l_Lean_Syntax_node2(x_420, x_473, x_454, x_472); -lean_inc(x_474); -x_475 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_475, 0, x_474); -x_476 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_474, x_475, x_15, x_16, x_425); -return x_476; +x_472 = l_Array_mkArray1___rarg(x_471); +x_473 = l_Array_append___rarg(x_406, x_472); +lean_dec(x_472); +lean_inc(x_426); +x_474 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_474, 0, x_426); +lean_ctor_set(x_474, 1, x_405); +lean_ctor_set(x_474, 2, x_473); +x_475 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_426); +x_476 = l_Lean_Syntax_node4(x_426, x_475, x_37, x_421, x_455, x_474); +x_477 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_426); +x_478 = l_Lean_Syntax_node4(x_426, x_477, x_446, x_35, x_452, x_476); +x_479 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_480 = l_Lean_Syntax_node2(x_426, x_479, x_460, x_478); +lean_inc(x_480); +x_481 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_481, 0, x_480); +x_482 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_480, x_481, x_15, x_16, x_431); +return x_482; } } } } else { -lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; size_t x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; -x_480 = lean_ctor_get(x_53, 0); -x_481 = lean_ctor_get(x_53, 1); -lean_inc(x_481); -lean_inc(x_480); +lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; size_t x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; +x_486 = lean_ctor_get(x_53, 0); +x_487 = lean_ctor_get(x_53, 1); +lean_inc(x_487); +lean_inc(x_486); lean_dec(x_53); -x_482 = l_Lean_SourceInfo_fromRef(x_480, x_33); -lean_dec(x_480); -x_483 = l_Lean_Elab_Command_getCurrMacroScope(x_52, x_16, x_481); +x_488 = l_Lean_SourceInfo_fromRef(x_486, x_33); +lean_dec(x_486); +x_489 = l_Lean_Elab_Command_getCurrMacroScope(x_52, x_16, x_487); lean_dec(x_52); -x_484 = lean_ctor_get(x_483, 1); -lean_inc(x_484); -if (lean_is_exclusive(x_483)) { - lean_ctor_release(x_483, 0); - lean_ctor_release(x_483, 1); - x_485 = x_483; +x_490 = lean_ctor_get(x_489, 1); +lean_inc(x_490); +if (lean_is_exclusive(x_489)) { + lean_ctor_release(x_489, 0); + lean_ctor_release(x_489, 1); + x_491 = x_489; } else { - lean_dec_ref(x_483); - x_485 = lean_box(0); + lean_dec_ref(x_489); + x_491 = lean_box(0); } -x_486 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_484); -x_487 = lean_ctor_get(x_486, 1); -lean_inc(x_487); -if (lean_is_exclusive(x_486)) { - lean_ctor_release(x_486, 0); - lean_ctor_release(x_486, 1); - x_488 = x_486; +x_492 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_490); +x_493 = lean_ctor_get(x_492, 1); +lean_inc(x_493); +if (lean_is_exclusive(x_492)) { + lean_ctor_release(x_492, 0); + lean_ctor_release(x_492, 1); + x_494 = x_492; +} else { + lean_dec_ref(x_492); + x_494 = lean_box(0); +} +x_495 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_488); +if (lean_is_scalar(x_494)) { + x_496 = lean_alloc_ctor(2, 2, 0); +} else { + x_496 = x_494; + lean_ctor_set_tag(x_496, 2); +} +lean_ctor_set(x_496, 0, x_488); +lean_ctor_set(x_496, 1, x_495); +x_497 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_498 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_488); +x_499 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_499, 0, x_488); +lean_ctor_set(x_499, 1, x_497); +lean_ctor_set(x_499, 2, x_498); +x_500 = lean_array_size(x_27); +x_501 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_500, x_20, x_27); +x_502 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_503 = l_Lean_mkSepArray(x_501, x_502); +lean_dec(x_501); +x_504 = l_Array_append___rarg(x_498, x_503); +lean_dec(x_503); +lean_inc(x_488); +x_505 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_505, 0, x_488); +lean_ctor_set(x_505, 1, x_497); +lean_ctor_set(x_505, 2, x_504); +x_506 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_488); +x_507 = l_Lean_Syntax_node1(x_488, x_506, x_505); +x_508 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_499); +lean_inc(x_488); +x_509 = l_Lean_Syntax_node1(x_488, x_508, x_499); +x_510 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_488); +if (lean_is_scalar(x_491)) { + x_511 = lean_alloc_ctor(2, 2, 0); +} else { + x_511 = x_491; + lean_ctor_set_tag(x_511, 2); +} +lean_ctor_set(x_511, 0, x_488); +lean_ctor_set(x_511, 1, x_510); +x_512 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_499); +x_513 = l_Lean_Syntax_node6(x_488, x_512, x_496, x_499, x_507, x_509, x_499, x_511); +x_514 = l_Lean_Elab_Command_getRef(x_15, x_16, x_493); +x_515 = lean_ctor_get(x_514, 0); +lean_inc(x_515); +x_516 = lean_ctor_get(x_514, 1); +lean_inc(x_516); +if (lean_is_exclusive(x_514)) { + lean_ctor_release(x_514, 0); + lean_ctor_release(x_514, 1); + x_517 = x_514; } else { - lean_dec_ref(x_486); - x_488 = lean_box(0); + lean_dec_ref(x_514); + x_517 = lean_box(0); } -x_489 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_482); -if (lean_is_scalar(x_488)) { - x_490 = lean_alloc_ctor(2, 2, 0); +x_518 = l_Lean_SourceInfo_fromRef(x_515, x_33); +lean_dec(x_515); +x_519 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_516); +x_520 = lean_ctor_get(x_519, 1); +lean_inc(x_520); +if (lean_is_exclusive(x_519)) { + lean_ctor_release(x_519, 0); + lean_ctor_release(x_519, 1); + x_521 = x_519; } else { - x_490 = x_488; - lean_ctor_set_tag(x_490, 2); + lean_dec_ref(x_519); + x_521 = lean_box(0); } -lean_ctor_set(x_490, 0, x_482); -lean_ctor_set(x_490, 1, x_489); -x_491 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_492 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_482); -x_493 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_493, 0, x_482); -lean_ctor_set(x_493, 1, x_491); -lean_ctor_set(x_493, 2, x_492); -x_494 = lean_array_size(x_27); -x_495 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_494, x_20, x_27); -x_496 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_497 = l_Lean_mkSepArray(x_495, x_496); -lean_dec(x_495); -x_498 = l_Array_append___rarg(x_492, x_497); -lean_dec(x_497); -lean_inc(x_482); -x_499 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_499, 0, x_482); -lean_ctor_set(x_499, 1, x_491); -lean_ctor_set(x_499, 2, x_498); -x_500 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_493); -lean_inc(x_482); -x_501 = l_Lean_Syntax_node1(x_482, x_500, x_493); -x_502 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_482); -if (lean_is_scalar(x_485)) { - x_503 = lean_alloc_ctor(2, 2, 0); +x_522 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_520); +x_523 = lean_ctor_get(x_522, 1); +lean_inc(x_523); +if (lean_is_exclusive(x_522)) { + lean_ctor_release(x_522, 0); + lean_ctor_release(x_522, 1); + x_524 = x_522; } else { - x_503 = x_485; - lean_ctor_set_tag(x_503, 2); + lean_dec_ref(x_522); + x_524 = lean_box(0); } -lean_ctor_set(x_503, 0, x_482); -lean_ctor_set(x_503, 1, x_502); -x_504 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_493); -x_505 = l_Lean_Syntax_node6(x_482, x_504, x_490, x_493, x_499, x_501, x_493, x_503); -x_506 = l_Lean_Elab_Command_getRef(x_15, x_16, x_487); -x_507 = lean_ctor_get(x_506, 0); -lean_inc(x_507); -x_508 = lean_ctor_get(x_506, 1); -lean_inc(x_508); -if (lean_is_exclusive(x_506)) { - lean_ctor_release(x_506, 0); - lean_ctor_release(x_506, 1); - x_509 = x_506; -} else { - lean_dec_ref(x_506); - x_509 = lean_box(0); -} -x_510 = l_Lean_SourceInfo_fromRef(x_507, x_33); -lean_dec(x_507); -x_511 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_508); -x_512 = lean_ctor_get(x_511, 1); -lean_inc(x_512); -if (lean_is_exclusive(x_511)) { - lean_ctor_release(x_511, 0); - lean_ctor_release(x_511, 1); - x_513 = x_511; +x_525 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_518); +if (lean_is_scalar(x_524)) { + x_526 = lean_alloc_ctor(2, 2, 0); } else { - lean_dec_ref(x_511); - x_513 = lean_box(0); + x_526 = x_524; + lean_ctor_set_tag(x_526, 2); } -x_514 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_512); -x_515 = lean_ctor_get(x_514, 1); -lean_inc(x_515); -if (lean_is_exclusive(x_514)) { - lean_ctor_release(x_514, 0); - lean_ctor_release(x_514, 1); - x_516 = x_514; +lean_ctor_set(x_526, 0, x_518); +lean_ctor_set(x_526, 1, x_525); +x_527 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_528 = l_Lean_Syntax_SepArray_ofElems(x_527, x_7); +x_529 = l_Array_append___rarg(x_498, x_528); +lean_dec(x_528); +lean_inc(x_518); +x_530 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_530, 0, x_518); +lean_ctor_set(x_530, 1, x_497); +lean_ctor_set(x_530, 2, x_529); +x_531 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_518); +if (lean_is_scalar(x_521)) { + x_532 = lean_alloc_ctor(2, 2, 0); } else { - lean_dec_ref(x_514); - x_516 = lean_box(0); -} -x_517 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_510); -if (lean_is_scalar(x_516)) { - x_518 = lean_alloc_ctor(2, 2, 0); -} else { - x_518 = x_516; - lean_ctor_set_tag(x_518, 2); -} -lean_ctor_set(x_518, 0, x_510); -lean_ctor_set(x_518, 1, x_517); -x_519 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_520 = l_Lean_Syntax_SepArray_ofElems(x_519, x_7); -x_521 = l_Array_append___rarg(x_492, x_520); -lean_dec(x_520); -lean_inc(x_510); -x_522 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_522, 0, x_510); -lean_ctor_set(x_522, 1, x_491); -lean_ctor_set(x_522, 2, x_521); -x_523 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_510); -if (lean_is_scalar(x_513)) { - x_524 = lean_alloc_ctor(2, 2, 0); -} else { - x_524 = x_513; - lean_ctor_set_tag(x_524, 2); -} -lean_ctor_set(x_524, 0, x_510); -lean_ctor_set(x_524, 1, x_523); -x_525 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_510); -x_526 = l_Lean_Syntax_node3(x_510, x_525, x_518, x_522, x_524); -lean_inc(x_510); -x_527 = l_Lean_Syntax_node1(x_510, x_491, x_526); -lean_inc(x_510); -x_528 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_528, 0, x_510); -lean_ctor_set(x_528, 1, x_491); -lean_ctor_set(x_528, 2, x_492); -x_529 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_510); -if (lean_is_scalar(x_509)) { - x_530 = lean_alloc_ctor(2, 2, 0); -} else { - x_530 = x_509; - lean_ctor_set_tag(x_530, 2); -} -lean_ctor_set(x_530, 0, x_510); -lean_ctor_set(x_530, 1, x_529); -x_531 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_510); -x_532 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_532, 0, x_510); + x_532 = x_521; + lean_ctor_set_tag(x_532, 2); +} +lean_ctor_set(x_532, 0, x_518); lean_ctor_set(x_532, 1, x_531); -x_533 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_510); -x_534 = l_Lean_Syntax_node2(x_510, x_533, x_532, x_34); -lean_inc(x_510); -x_535 = l_Lean_Syntax_node1(x_510, x_491, x_534); -x_536 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_528); -lean_inc(x_510); -x_537 = l_Lean_Syntax_node2(x_510, x_536, x_528, x_535); -x_538 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_510); +x_533 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_518); +x_534 = l_Lean_Syntax_node3(x_518, x_533, x_526, x_530, x_532); +lean_inc(x_518); +x_535 = l_Lean_Syntax_node1(x_518, x_497, x_534); +lean_inc(x_518); +x_536 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_536, 0, x_518); +lean_ctor_set(x_536, 1, x_497); +lean_ctor_set(x_536, 2, x_498); +x_537 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_518); +if (lean_is_scalar(x_517)) { + x_538 = lean_alloc_ctor(2, 2, 0); +} else { + x_538 = x_517; + lean_ctor_set_tag(x_538, 2); +} +lean_ctor_set(x_538, 0, x_518); +lean_ctor_set(x_538, 1, x_537); +x_539 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_518); +x_540 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_540, 0, x_518); +lean_ctor_set(x_540, 1, x_539); +x_541 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_518); +x_542 = l_Lean_Syntax_node2(x_518, x_541, x_540, x_34); +lean_inc(x_518); +x_543 = l_Lean_Syntax_node1(x_518, x_497, x_542); +x_544 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_536); +lean_inc(x_518); +x_545 = l_Lean_Syntax_node2(x_518, x_544, x_536, x_543); +x_546 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_518); lean_ctor_set_tag(x_37, 2); -lean_ctor_set(x_37, 1, x_538); -lean_ctor_set(x_37, 0, x_510); -x_539 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_528, 2); -lean_inc(x_510); -x_540 = l_Lean_Syntax_node2(x_510, x_539, x_528, x_528); +lean_ctor_set(x_37, 1, x_546); +lean_ctor_set(x_37, 0, x_518); +x_547 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_536, 2); +lean_inc(x_518); +x_548 = l_Lean_Syntax_node2(x_518, x_547, x_536, x_536); if (lean_obj_tag(x_10) == 0) { -x_541 = x_492; -goto block_568; +x_549 = x_498; +goto block_576; +} +else +{ +lean_object* x_577; lean_object* x_578; +x_577 = lean_ctor_get(x_10, 0); +lean_inc(x_577); +lean_dec(x_10); +x_578 = l_Array_mkArray1___rarg(x_577); +x_549 = x_578; +goto block_576; +} +block_576: +{ +lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; +x_550 = l_Array_append___rarg(x_498, x_549); +lean_dec(x_549); +lean_inc(x_518); +x_551 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_551, 0, x_518); +lean_ctor_set(x_551, 1, x_497); +lean_ctor_set(x_551, 2, x_550); +x_552 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_536, 3); +lean_inc(x_518); +x_553 = l_Lean_Syntax_node6(x_518, x_552, x_551, x_535, x_536, x_536, x_536, x_536); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; +x_554 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_518); +x_555 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_555, 0, x_518); +lean_ctor_set(x_555, 1, x_497); +lean_ctor_set(x_555, 2, x_554); +x_556 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_518); +x_557 = l_Lean_Syntax_node4(x_518, x_556, x_37, x_513, x_548, x_555); +x_558 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_518); +x_559 = l_Lean_Syntax_node4(x_518, x_558, x_538, x_35, x_545, x_557); +x_560 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_561 = l_Lean_Syntax_node2(x_518, x_560, x_553, x_559); +lean_inc(x_561); +x_562 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_562, 0, x_561); +x_563 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_561, x_562, x_15, x_16, x_523); +return x_563; } else { -lean_object* x_569; lean_object* x_570; -x_569 = lean_ctor_get(x_10, 0); -lean_inc(x_569); -lean_dec(x_10); -x_570 = l_Array_mkArray1___rarg(x_569); -x_541 = x_570; -goto block_568; -} -block_568: -{ -lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; -x_542 = l_Array_append___rarg(x_492, x_541); -lean_dec(x_541); -lean_inc(x_510); -x_543 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_543, 0, x_510); -lean_ctor_set(x_543, 1, x_491); -lean_ctor_set(x_543, 2, x_542); -x_544 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_528, 3); -lean_inc(x_510); -x_545 = l_Lean_Syntax_node6(x_510, x_544, x_543, x_527, x_528, x_528, x_528, x_528); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; -x_546 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_510); -x_547 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_547, 0, x_510); -lean_ctor_set(x_547, 1, x_491); -lean_ctor_set(x_547, 2, x_546); -x_548 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_510); -x_549 = l_Lean_Syntax_node4(x_510, x_548, x_37, x_505, x_540, x_547); -x_550 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_510); -x_551 = l_Lean_Syntax_node4(x_510, x_550, x_530, x_35, x_537, x_549); -x_552 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_553 = l_Lean_Syntax_node2(x_510, x_552, x_545, x_551); -lean_inc(x_553); -x_554 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_554, 0, x_553); -x_555 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_553, x_554, x_15, x_16, x_515); -return x_555; -} -else -{ -lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; -x_556 = lean_ctor_get(x_8, 0); -lean_inc(x_556); +lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; +x_564 = lean_ctor_get(x_8, 0); +lean_inc(x_564); lean_dec(x_8); -x_557 = l_Array_mkArray1___rarg(x_556); -x_558 = l_Array_append___rarg(x_492, x_557); -lean_dec(x_557); -lean_inc(x_510); -x_559 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_559, 0, x_510); -lean_ctor_set(x_559, 1, x_491); -lean_ctor_set(x_559, 2, x_558); -x_560 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_510); -x_561 = l_Lean_Syntax_node4(x_510, x_560, x_37, x_505, x_540, x_559); -x_562 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_510); -x_563 = l_Lean_Syntax_node4(x_510, x_562, x_530, x_35, x_537, x_561); -x_564 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_565 = l_Lean_Syntax_node2(x_510, x_564, x_545, x_563); -lean_inc(x_565); -x_566 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_566, 0, x_565); -x_567 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_565, x_566, x_15, x_16, x_515); -return x_567; +x_565 = l_Array_mkArray1___rarg(x_564); +x_566 = l_Array_append___rarg(x_498, x_565); +lean_dec(x_565); +lean_inc(x_518); +x_567 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_567, 0, x_518); +lean_ctor_set(x_567, 1, x_497); +lean_ctor_set(x_567, 2, x_566); +x_568 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_518); +x_569 = l_Lean_Syntax_node4(x_518, x_568, x_37, x_513, x_548, x_567); +x_570 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_518); +x_571 = l_Lean_Syntax_node4(x_518, x_570, x_538, x_35, x_545, x_569); +x_572 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_573 = l_Lean_Syntax_node2(x_518, x_572, x_553, x_571); +lean_inc(x_573); +x_574 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_574, 0, x_573); +x_575 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_573, x_574, x_15, x_16, x_523); +return x_575; } } } } else { -lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; uint8_t x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; size_t x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; -x_571 = lean_ctor_get(x_37, 0); -x_572 = lean_ctor_get(x_37, 1); -lean_inc(x_572); -lean_inc(x_571); -lean_dec(x_37); -x_573 = l_Lean_replaceRef(x_6, x_571); -lean_dec(x_571); -x_574 = lean_ctor_get(x_15, 0); -lean_inc(x_574); -x_575 = lean_ctor_get(x_15, 1); -lean_inc(x_575); -x_576 = lean_ctor_get(x_15, 2); -lean_inc(x_576); -x_577 = lean_ctor_get(x_15, 3); -lean_inc(x_577); -x_578 = lean_ctor_get(x_15, 4); -lean_inc(x_578); -x_579 = lean_ctor_get(x_15, 5); -lean_inc(x_579); -x_580 = lean_ctor_get(x_15, 7); +lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; uint8_t x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; size_t x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; +x_579 = lean_ctor_get(x_37, 0); +x_580 = lean_ctor_get(x_37, 1); lean_inc(x_580); -x_581 = lean_ctor_get(x_15, 8); -lean_inc(x_581); -x_582 = lean_ctor_get(x_15, 9); +lean_inc(x_579); +lean_dec(x_37); +x_581 = l_Lean_replaceRef(x_6, x_579); +lean_dec(x_579); +x_582 = lean_ctor_get(x_15, 0); lean_inc(x_582); -x_583 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); -x_584 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_584, 0, x_574); -lean_ctor_set(x_584, 1, x_575); -lean_ctor_set(x_584, 2, x_576); -lean_ctor_set(x_584, 3, x_577); -lean_ctor_set(x_584, 4, x_578); -lean_ctor_set(x_584, 5, x_579); -lean_ctor_set(x_584, 6, x_573); -lean_ctor_set(x_584, 7, x_580); -lean_ctor_set(x_584, 8, x_581); -lean_ctor_set(x_584, 9, x_582); -lean_ctor_set_uint8(x_584, sizeof(void*)*10, x_583); -x_585 = l_Lean_Elab_Command_getRef(x_584, x_16, x_572); -x_586 = lean_ctor_get(x_585, 0); +x_583 = lean_ctor_get(x_15, 1); +lean_inc(x_583); +x_584 = lean_ctor_get(x_15, 2); +lean_inc(x_584); +x_585 = lean_ctor_get(x_15, 3); +lean_inc(x_585); +x_586 = lean_ctor_get(x_15, 4); lean_inc(x_586); -x_587 = lean_ctor_get(x_585, 1); +x_587 = lean_ctor_get(x_15, 5); lean_inc(x_587); -if (lean_is_exclusive(x_585)) { - lean_ctor_release(x_585, 0); - lean_ctor_release(x_585, 1); - x_588 = x_585; -} else { - lean_dec_ref(x_585); - x_588 = lean_box(0); -} -x_589 = l_Lean_SourceInfo_fromRef(x_586, x_33); -lean_dec(x_586); -x_590 = l_Lean_Elab_Command_getCurrMacroScope(x_584, x_16, x_587); -lean_dec(x_584); -x_591 = lean_ctor_get(x_590, 1); -lean_inc(x_591); -if (lean_is_exclusive(x_590)) { - lean_ctor_release(x_590, 0); - lean_ctor_release(x_590, 1); - x_592 = x_590; -} else { - lean_dec_ref(x_590); - x_592 = lean_box(0); -} -x_593 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_591); -x_594 = lean_ctor_get(x_593, 1); +x_588 = lean_ctor_get(x_15, 7); +lean_inc(x_588); +x_589 = lean_ctor_get(x_15, 8); +lean_inc(x_589); +x_590 = lean_ctor_get(x_15, 9); +lean_inc(x_590); +x_591 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); +x_592 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_592, 0, x_582); +lean_ctor_set(x_592, 1, x_583); +lean_ctor_set(x_592, 2, x_584); +lean_ctor_set(x_592, 3, x_585); +lean_ctor_set(x_592, 4, x_586); +lean_ctor_set(x_592, 5, x_587); +lean_ctor_set(x_592, 6, x_581); +lean_ctor_set(x_592, 7, x_588); +lean_ctor_set(x_592, 8, x_589); +lean_ctor_set(x_592, 9, x_590); +lean_ctor_set_uint8(x_592, sizeof(void*)*10, x_591); +x_593 = l_Lean_Elab_Command_getRef(x_592, x_16, x_580); +x_594 = lean_ctor_get(x_593, 0); lean_inc(x_594); +x_595 = lean_ctor_get(x_593, 1); +lean_inc(x_595); if (lean_is_exclusive(x_593)) { lean_ctor_release(x_593, 0); lean_ctor_release(x_593, 1); - x_595 = x_593; + x_596 = x_593; } else { lean_dec_ref(x_593); - x_595 = lean_box(0); -} -x_596 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_589); -if (lean_is_scalar(x_595)) { - x_597 = lean_alloc_ctor(2, 2, 0); + x_596 = lean_box(0); +} +x_597 = l_Lean_SourceInfo_fromRef(x_594, x_33); +lean_dec(x_594); +x_598 = l_Lean_Elab_Command_getCurrMacroScope(x_592, x_16, x_595); +lean_dec(x_592); +x_599 = lean_ctor_get(x_598, 1); +lean_inc(x_599); +if (lean_is_exclusive(x_598)) { + lean_ctor_release(x_598, 0); + lean_ctor_release(x_598, 1); + x_600 = x_598; +} else { + lean_dec_ref(x_598); + x_600 = lean_box(0); +} +x_601 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_599); +x_602 = lean_ctor_get(x_601, 1); +lean_inc(x_602); +if (lean_is_exclusive(x_601)) { + lean_ctor_release(x_601, 0); + lean_ctor_release(x_601, 1); + x_603 = x_601; +} else { + lean_dec_ref(x_601); + x_603 = lean_box(0); +} +x_604 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_597); +if (lean_is_scalar(x_603)) { + x_605 = lean_alloc_ctor(2, 2, 0); +} else { + x_605 = x_603; + lean_ctor_set_tag(x_605, 2); +} +lean_ctor_set(x_605, 0, x_597); +lean_ctor_set(x_605, 1, x_604); +x_606 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_607 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_597); +x_608 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_608, 0, x_597); +lean_ctor_set(x_608, 1, x_606); +lean_ctor_set(x_608, 2, x_607); +x_609 = lean_array_size(x_27); +x_610 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_609, x_20, x_27); +x_611 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_612 = l_Lean_mkSepArray(x_610, x_611); +lean_dec(x_610); +x_613 = l_Array_append___rarg(x_607, x_612); +lean_dec(x_612); +lean_inc(x_597); +x_614 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_614, 0, x_597); +lean_ctor_set(x_614, 1, x_606); +lean_ctor_set(x_614, 2, x_613); +x_615 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_597); +x_616 = l_Lean_Syntax_node1(x_597, x_615, x_614); +x_617 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_608); +lean_inc(x_597); +x_618 = l_Lean_Syntax_node1(x_597, x_617, x_608); +x_619 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_597); +if (lean_is_scalar(x_600)) { + x_620 = lean_alloc_ctor(2, 2, 0); +} else { + x_620 = x_600; + lean_ctor_set_tag(x_620, 2); +} +lean_ctor_set(x_620, 0, x_597); +lean_ctor_set(x_620, 1, x_619); +x_621 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_608); +x_622 = l_Lean_Syntax_node6(x_597, x_621, x_605, x_608, x_616, x_618, x_608, x_620); +x_623 = l_Lean_Elab_Command_getRef(x_15, x_16, x_602); +x_624 = lean_ctor_get(x_623, 0); +lean_inc(x_624); +x_625 = lean_ctor_get(x_623, 1); +lean_inc(x_625); +if (lean_is_exclusive(x_623)) { + lean_ctor_release(x_623, 0); + lean_ctor_release(x_623, 1); + x_626 = x_623; +} else { + lean_dec_ref(x_623); + x_626 = lean_box(0); +} +x_627 = l_Lean_SourceInfo_fromRef(x_624, x_33); +lean_dec(x_624); +x_628 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_625); +x_629 = lean_ctor_get(x_628, 1); +lean_inc(x_629); +if (lean_is_exclusive(x_628)) { + lean_ctor_release(x_628, 0); + lean_ctor_release(x_628, 1); + x_630 = x_628; +} else { + lean_dec_ref(x_628); + x_630 = lean_box(0); +} +x_631 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_629); +x_632 = lean_ctor_get(x_631, 1); +lean_inc(x_632); +if (lean_is_exclusive(x_631)) { + lean_ctor_release(x_631, 0); + lean_ctor_release(x_631, 1); + x_633 = x_631; +} else { + lean_dec_ref(x_631); + x_633 = lean_box(0); +} +x_634 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_627); +if (lean_is_scalar(x_633)) { + x_635 = lean_alloc_ctor(2, 2, 0); } else { - x_597 = x_595; - lean_ctor_set_tag(x_597, 2); + x_635 = x_633; + lean_ctor_set_tag(x_635, 2); } -lean_ctor_set(x_597, 0, x_589); -lean_ctor_set(x_597, 1, x_596); -x_598 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_599 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_589); -x_600 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_600, 0, x_589); -lean_ctor_set(x_600, 1, x_598); -lean_ctor_set(x_600, 2, x_599); -x_601 = lean_array_size(x_27); -x_602 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_601, x_20, x_27); -x_603 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_604 = l_Lean_mkSepArray(x_602, x_603); -lean_dec(x_602); -x_605 = l_Array_append___rarg(x_599, x_604); -lean_dec(x_604); -lean_inc(x_589); -x_606 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_606, 0, x_589); -lean_ctor_set(x_606, 1, x_598); -lean_ctor_set(x_606, 2, x_605); -x_607 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_600); -lean_inc(x_589); -x_608 = l_Lean_Syntax_node1(x_589, x_607, x_600); -x_609 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_589); -if (lean_is_scalar(x_592)) { - x_610 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_635, 0, x_627); +lean_ctor_set(x_635, 1, x_634); +x_636 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_637 = l_Lean_Syntax_SepArray_ofElems(x_636, x_7); +x_638 = l_Array_append___rarg(x_607, x_637); +lean_dec(x_637); +lean_inc(x_627); +x_639 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_639, 0, x_627); +lean_ctor_set(x_639, 1, x_606); +lean_ctor_set(x_639, 2, x_638); +x_640 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_627); +if (lean_is_scalar(x_630)) { + x_641 = lean_alloc_ctor(2, 2, 0); } else { - x_610 = x_592; - lean_ctor_set_tag(x_610, 2); + x_641 = x_630; + lean_ctor_set_tag(x_641, 2); } -lean_ctor_set(x_610, 0, x_589); -lean_ctor_set(x_610, 1, x_609); -x_611 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_600); -x_612 = l_Lean_Syntax_node6(x_589, x_611, x_597, x_600, x_606, x_608, x_600, x_610); -x_613 = l_Lean_Elab_Command_getRef(x_15, x_16, x_594); -x_614 = lean_ctor_get(x_613, 0); -lean_inc(x_614); -x_615 = lean_ctor_get(x_613, 1); -lean_inc(x_615); -if (lean_is_exclusive(x_613)) { - lean_ctor_release(x_613, 0); - lean_ctor_release(x_613, 1); - x_616 = x_613; -} else { - lean_dec_ref(x_613); - x_616 = lean_box(0); -} -x_617 = l_Lean_SourceInfo_fromRef(x_614, x_33); -lean_dec(x_614); -x_618 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_615); -x_619 = lean_ctor_get(x_618, 1); -lean_inc(x_619); -if (lean_is_exclusive(x_618)) { - lean_ctor_release(x_618, 0); - lean_ctor_release(x_618, 1); - x_620 = x_618; -} else { - lean_dec_ref(x_618); - x_620 = lean_box(0); -} -x_621 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_619); -x_622 = lean_ctor_get(x_621, 1); -lean_inc(x_622); -if (lean_is_exclusive(x_621)) { - lean_ctor_release(x_621, 0); - lean_ctor_release(x_621, 1); - x_623 = x_621; -} else { - lean_dec_ref(x_621); - x_623 = lean_box(0); -} -x_624 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_617); -if (lean_is_scalar(x_623)) { - x_625 = lean_alloc_ctor(2, 2, 0); -} else { - x_625 = x_623; - lean_ctor_set_tag(x_625, 2); -} -lean_ctor_set(x_625, 0, x_617); -lean_ctor_set(x_625, 1, x_624); -x_626 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_627 = l_Lean_Syntax_SepArray_ofElems(x_626, x_7); -x_628 = l_Array_append___rarg(x_599, x_627); -lean_dec(x_627); -lean_inc(x_617); -x_629 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_629, 0, x_617); -lean_ctor_set(x_629, 1, x_598); -lean_ctor_set(x_629, 2, x_628); -x_630 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_617); -if (lean_is_scalar(x_620)) { - x_631 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_641, 0, x_627); +lean_ctor_set(x_641, 1, x_640); +x_642 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_627); +x_643 = l_Lean_Syntax_node3(x_627, x_642, x_635, x_639, x_641); +lean_inc(x_627); +x_644 = l_Lean_Syntax_node1(x_627, x_606, x_643); +lean_inc(x_627); +x_645 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_645, 0, x_627); +lean_ctor_set(x_645, 1, x_606); +lean_ctor_set(x_645, 2, x_607); +x_646 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_627); +if (lean_is_scalar(x_626)) { + x_647 = lean_alloc_ctor(2, 2, 0); } else { - x_631 = x_620; - lean_ctor_set_tag(x_631, 2); + x_647 = x_626; + lean_ctor_set_tag(x_647, 2); } -lean_ctor_set(x_631, 0, x_617); -lean_ctor_set(x_631, 1, x_630); -x_632 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_617); -x_633 = l_Lean_Syntax_node3(x_617, x_632, x_625, x_629, x_631); -lean_inc(x_617); -x_634 = l_Lean_Syntax_node1(x_617, x_598, x_633); -lean_inc(x_617); -x_635 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_635, 0, x_617); -lean_ctor_set(x_635, 1, x_598); -lean_ctor_set(x_635, 2, x_599); -x_636 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_617); -if (lean_is_scalar(x_616)) { - x_637 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_647, 0, x_627); +lean_ctor_set(x_647, 1, x_646); +x_648 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_627); +if (lean_is_scalar(x_596)) { + x_649 = lean_alloc_ctor(2, 2, 0); } else { - x_637 = x_616; - lean_ctor_set_tag(x_637, 2); + x_649 = x_596; + lean_ctor_set_tag(x_649, 2); } -lean_ctor_set(x_637, 0, x_617); -lean_ctor_set(x_637, 1, x_636); -x_638 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_617); -if (lean_is_scalar(x_588)) { - x_639 = lean_alloc_ctor(2, 2, 0); -} else { - x_639 = x_588; - lean_ctor_set_tag(x_639, 2); -} -lean_ctor_set(x_639, 0, x_617); -lean_ctor_set(x_639, 1, x_638); -x_640 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_617); -x_641 = l_Lean_Syntax_node2(x_617, x_640, x_639, x_34); -lean_inc(x_617); -x_642 = l_Lean_Syntax_node1(x_617, x_598, x_641); -x_643 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_635); -lean_inc(x_617); -x_644 = l_Lean_Syntax_node2(x_617, x_643, x_635, x_642); -x_645 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_617); -x_646 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_646, 0, x_617); -lean_ctor_set(x_646, 1, x_645); -x_647 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_635, 2); -lean_inc(x_617); -x_648 = l_Lean_Syntax_node2(x_617, x_647, x_635, x_635); +lean_ctor_set(x_649, 0, x_627); +lean_ctor_set(x_649, 1, x_648); +x_650 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_627); +x_651 = l_Lean_Syntax_node2(x_627, x_650, x_649, x_34); +lean_inc(x_627); +x_652 = l_Lean_Syntax_node1(x_627, x_606, x_651); +x_653 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_645); +lean_inc(x_627); +x_654 = l_Lean_Syntax_node2(x_627, x_653, x_645, x_652); +x_655 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_627); +x_656 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_656, 0, x_627); +lean_ctor_set(x_656, 1, x_655); +x_657 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_645, 2); +lean_inc(x_627); +x_658 = l_Lean_Syntax_node2(x_627, x_657, x_645, x_645); if (lean_obj_tag(x_10) == 0) { -x_649 = x_599; -goto block_676; +x_659 = x_607; +goto block_686; } else { -lean_object* x_677; lean_object* x_678; -x_677 = lean_ctor_get(x_10, 0); -lean_inc(x_677); +lean_object* x_687; lean_object* x_688; +x_687 = lean_ctor_get(x_10, 0); +lean_inc(x_687); lean_dec(x_10); -x_678 = l_Array_mkArray1___rarg(x_677); -x_649 = x_678; -goto block_676; -} -block_676: -{ -lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; -x_650 = l_Array_append___rarg(x_599, x_649); -lean_dec(x_649); -lean_inc(x_617); -x_651 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_651, 0, x_617); -lean_ctor_set(x_651, 1, x_598); -lean_ctor_set(x_651, 2, x_650); -x_652 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_635, 3); -lean_inc(x_617); -x_653 = l_Lean_Syntax_node6(x_617, x_652, x_651, x_634, x_635, x_635, x_635, x_635); +x_688 = l_Array_mkArray1___rarg(x_687); +x_659 = x_688; +goto block_686; +} +block_686: +{ +lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; +x_660 = l_Array_append___rarg(x_607, x_659); +lean_dec(x_659); +lean_inc(x_627); +x_661 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_661, 0, x_627); +lean_ctor_set(x_661, 1, x_606); +lean_ctor_set(x_661, 2, x_660); +x_662 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_645, 3); +lean_inc(x_627); +x_663 = l_Lean_Syntax_node6(x_627, x_662, x_661, x_644, x_645, x_645, x_645, x_645); if (lean_obj_tag(x_8) == 0) { -lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; -x_654 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_617); -x_655 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_655, 0, x_617); -lean_ctor_set(x_655, 1, x_598); -lean_ctor_set(x_655, 2, x_654); -x_656 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_617); -x_657 = l_Lean_Syntax_node4(x_617, x_656, x_646, x_612, x_648, x_655); -x_658 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_617); -x_659 = l_Lean_Syntax_node4(x_617, x_658, x_637, x_35, x_644, x_657); -x_660 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_661 = l_Lean_Syntax_node2(x_617, x_660, x_653, x_659); -lean_inc(x_661); -x_662 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_662, 0, x_661); -x_663 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_661, x_662, x_15, x_16, x_622); -return x_663; -} -else -{ -lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; -x_664 = lean_ctor_get(x_8, 0); -lean_inc(x_664); +lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; +x_664 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_627); +x_665 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_665, 0, x_627); +lean_ctor_set(x_665, 1, x_606); +lean_ctor_set(x_665, 2, x_664); +x_666 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_627); +x_667 = l_Lean_Syntax_node4(x_627, x_666, x_656, x_622, x_658, x_665); +x_668 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_627); +x_669 = l_Lean_Syntax_node4(x_627, x_668, x_647, x_35, x_654, x_667); +x_670 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_671 = l_Lean_Syntax_node2(x_627, x_670, x_663, x_669); +lean_inc(x_671); +x_672 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_672, 0, x_671); +x_673 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_671, x_672, x_15, x_16, x_632); +return x_673; +} +else +{ +lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; +x_674 = lean_ctor_get(x_8, 0); +lean_inc(x_674); lean_dec(x_8); -x_665 = l_Array_mkArray1___rarg(x_664); -x_666 = l_Array_append___rarg(x_599, x_665); -lean_dec(x_665); -lean_inc(x_617); -x_667 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_667, 0, x_617); -lean_ctor_set(x_667, 1, x_598); -lean_ctor_set(x_667, 2, x_666); -x_668 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_617); -x_669 = l_Lean_Syntax_node4(x_617, x_668, x_646, x_612, x_648, x_667); -x_670 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_617); -x_671 = l_Lean_Syntax_node4(x_617, x_670, x_637, x_35, x_644, x_669); -x_672 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_673 = l_Lean_Syntax_node2(x_617, x_672, x_653, x_671); -lean_inc(x_673); -x_674 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_674, 0, x_673); -x_675 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_673, x_674, x_15, x_16, x_622); -return x_675; +x_675 = l_Array_mkArray1___rarg(x_674); +x_676 = l_Array_append___rarg(x_607, x_675); +lean_dec(x_675); +lean_inc(x_627); +x_677 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_677, 0, x_627); +lean_ctor_set(x_677, 1, x_606); +lean_ctor_set(x_677, 2, x_676); +x_678 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_627); +x_679 = l_Lean_Syntax_node4(x_627, x_678, x_656, x_622, x_658, x_677); +x_680 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_627); +x_681 = l_Lean_Syntax_node4(x_627, x_680, x_647, x_35, x_654, x_679); +x_682 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_683 = l_Lean_Syntax_node2(x_627, x_682, x_663, x_681); +lean_inc(x_683); +x_684 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_684, 0, x_683); +x_685 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_683, x_684, x_15, x_16, x_632); +return x_685; } } } @@ -9670,78 +9731,78 @@ return x_675; } else { -lean_object* x_749; lean_object* x_750; uint8_t x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; -x_749 = lean_ctor_get(x_29, 0); -x_750 = lean_ctor_get(x_29, 1); -lean_inc(x_750); -lean_inc(x_749); +lean_object* x_759; lean_object* x_760; uint8_t x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; +x_759 = lean_ctor_get(x_29, 0); +x_760 = lean_ctor_get(x_29, 1); +lean_inc(x_760); +lean_inc(x_759); lean_dec(x_29); -x_751 = 0; -x_752 = l_Lean_mkCIdentFrom(x_749, x_2, x_751); -lean_dec(x_749); +x_761 = 0; +x_762 = l_Lean_mkCIdentFrom(x_759, x_2, x_761); +lean_dec(x_759); if (lean_obj_tag(x_11) == 0) { lean_free_object(x_25); if (lean_obj_tag(x_12) == 0) { -lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; -x_866 = l_Lean_Elab_Command_getRef(x_15, x_16, x_750); -x_867 = lean_ctor_get(x_866, 0); -lean_inc(x_867); -x_868 = lean_ctor_get(x_866, 1); -lean_inc(x_868); -if (lean_is_exclusive(x_866)) { - lean_ctor_release(x_866, 0); - lean_ctor_release(x_866, 1); - x_869 = x_866; +lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; +x_878 = l_Lean_Elab_Command_getRef(x_15, x_16, x_760); +x_879 = lean_ctor_get(x_878, 0); +lean_inc(x_879); +x_880 = lean_ctor_get(x_878, 1); +lean_inc(x_880); +if (lean_is_exclusive(x_878)) { + lean_ctor_release(x_878, 0); + lean_ctor_release(x_878, 1); + x_881 = x_878; } else { - lean_dec_ref(x_866); - x_869 = lean_box(0); + lean_dec_ref(x_878); + x_881 = lean_box(0); } lean_inc(x_16); lean_inc(x_15); -x_870 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_867, x_751, x_15, x_16, x_868); -lean_dec(x_867); -if (lean_obj_tag(x_870) == 0) +x_882 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_879, x_761, x_15, x_16, x_880); +lean_dec(x_879); +if (lean_obj_tag(x_882) == 0) { -lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; -x_871 = lean_ctor_get(x_870, 0); -lean_inc(x_871); -x_872 = lean_ctor_get(x_870, 1); -lean_inc(x_872); -lean_dec(x_870); -x_873 = lean_box(2); -x_874 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_875 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_875, 0, x_873); -lean_ctor_set(x_875, 1, x_874); -lean_ctor_set(x_875, 2, x_24); -if (lean_is_scalar(x_869)) { - x_876 = lean_alloc_ctor(1, 2, 0); -} else { - x_876 = x_869; - lean_ctor_set_tag(x_876, 1); -} -lean_ctor_set(x_876, 0, x_875); -lean_ctor_set(x_876, 1, x_4); -x_877 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_877, 0, x_871); -lean_ctor_set(x_877, 1, x_876); -x_878 = lean_array_mk(x_877); -x_879 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_880 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_880, 0, x_873); -lean_ctor_set(x_880, 1, x_879); -lean_ctor_set(x_880, 2, x_878); -x_753 = x_880; -x_754 = x_872; -goto block_865; +lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; +x_883 = lean_ctor_get(x_882, 0); +lean_inc(x_883); +x_884 = lean_ctor_get(x_882, 1); +lean_inc(x_884); +lean_dec(x_882); +x_885 = lean_box(2); +x_886 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_887 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_887, 0, x_885); +lean_ctor_set(x_887, 1, x_886); +lean_ctor_set(x_887, 2, x_24); +if (lean_is_scalar(x_881)) { + x_888 = lean_alloc_ctor(1, 2, 0); +} else { + x_888 = x_881; + lean_ctor_set_tag(x_888, 1); +} +lean_ctor_set(x_888, 0, x_887); +lean_ctor_set(x_888, 1, x_4); +x_889 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_889, 0, x_883); +lean_ctor_set(x_889, 1, x_888); +x_890 = lean_array_mk(x_889); +x_891 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_892 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_892, 0, x_885); +lean_ctor_set(x_892, 1, x_891); +lean_ctor_set(x_892, 2, x_890); +x_763 = x_892; +x_764 = x_884; +goto block_877; } else { -lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; -lean_dec(x_869); -lean_dec(x_752); +lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; +lean_dec(x_881); +lean_dec(x_762); lean_dec(x_27); lean_dec(x_24); lean_dec(x_16); @@ -9750,564 +9811,567 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_881 = lean_ctor_get(x_870, 0); -lean_inc(x_881); -x_882 = lean_ctor_get(x_870, 1); -lean_inc(x_882); -if (lean_is_exclusive(x_870)) { - lean_ctor_release(x_870, 0); - lean_ctor_release(x_870, 1); - x_883 = x_870; +x_893 = lean_ctor_get(x_882, 0); +lean_inc(x_893); +x_894 = lean_ctor_get(x_882, 1); +lean_inc(x_894); +if (lean_is_exclusive(x_882)) { + lean_ctor_release(x_882, 0); + lean_ctor_release(x_882, 1); + x_895 = x_882; } else { - lean_dec_ref(x_870); - x_883 = lean_box(0); + lean_dec_ref(x_882); + x_895 = lean_box(0); } -if (lean_is_scalar(x_883)) { - x_884 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_895)) { + x_896 = lean_alloc_ctor(1, 2, 0); } else { - x_884 = x_883; + x_896 = x_895; } -lean_ctor_set(x_884, 0, x_881); -lean_ctor_set(x_884, 1, x_882); -return x_884; +lean_ctor_set(x_896, 0, x_893); +lean_ctor_set(x_896, 1, x_894); +return x_896; } } else { -lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; -x_885 = lean_ctor_get(x_12, 0); -lean_inc(x_885); +lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; +x_897 = lean_ctor_get(x_12, 0); +lean_inc(x_897); lean_dec(x_12); -x_886 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_885, x_751, x_15, x_16, x_750); -x_887 = lean_ctor_get(x_886, 0); -lean_inc(x_887); -x_888 = lean_ctor_get(x_886, 1); -lean_inc(x_888); -if (lean_is_exclusive(x_886)) { - lean_ctor_release(x_886, 0); - lean_ctor_release(x_886, 1); - x_889 = x_886; -} else { - lean_dec_ref(x_886); - x_889 = lean_box(0); -} -x_890 = lean_box(2); -x_891 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_892 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_892, 0, x_890); -lean_ctor_set(x_892, 1, x_891); -lean_ctor_set(x_892, 2, x_24); -if (lean_is_scalar(x_889)) { - x_893 = lean_alloc_ctor(1, 2, 0); -} else { - x_893 = x_889; - lean_ctor_set_tag(x_893, 1); -} -lean_ctor_set(x_893, 0, x_892); -lean_ctor_set(x_893, 1, x_4); -x_894 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_894, 0, x_887); -lean_ctor_set(x_894, 1, x_893); -x_895 = lean_array_mk(x_894); -x_896 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_897 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_897, 0, x_890); -lean_ctor_set(x_897, 1, x_896); -lean_ctor_set(x_897, 2, x_895); -x_753 = x_897; -x_754 = x_888; -goto block_865; +x_898 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_897, x_761, x_15, x_16, x_760); +x_899 = lean_ctor_get(x_898, 0); +lean_inc(x_899); +x_900 = lean_ctor_get(x_898, 1); +lean_inc(x_900); +if (lean_is_exclusive(x_898)) { + lean_ctor_release(x_898, 0); + lean_ctor_release(x_898, 1); + x_901 = x_898; +} else { + lean_dec_ref(x_898); + x_901 = lean_box(0); +} +x_902 = lean_box(2); +x_903 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_904 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_904, 0, x_902); +lean_ctor_set(x_904, 1, x_903); +lean_ctor_set(x_904, 2, x_24); +if (lean_is_scalar(x_901)) { + x_905 = lean_alloc_ctor(1, 2, 0); +} else { + x_905 = x_901; + lean_ctor_set_tag(x_905, 1); +} +lean_ctor_set(x_905, 0, x_904); +lean_ctor_set(x_905, 1, x_4); +x_906 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_906, 0, x_899); +lean_ctor_set(x_906, 1, x_905); +x_907 = lean_array_mk(x_906); +x_908 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_909 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_909, 0, x_902); +lean_ctor_set(x_909, 1, x_908); +lean_ctor_set(x_909, 2, x_907); +x_763 = x_909; +x_764 = x_900; +goto block_877; } } else { if (lean_obj_tag(x_12) == 0) { -lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; -x_898 = lean_ctor_get(x_11, 0); -x_899 = lean_box(2); -x_900 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_901 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_901, 0, x_899); -lean_ctor_set(x_901, 1, x_900); -lean_ctor_set(x_901, 2, x_24); -x_902 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_902, 0, x_901); -lean_ctor_set(x_902, 1, x_4); -lean_inc(x_898); +lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; +x_910 = lean_ctor_get(x_11, 0); +x_911 = lean_box(2); +x_912 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_913 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_913, 0, x_911); +lean_ctor_set(x_913, 1, x_912); +lean_ctor_set(x_913, 2, x_24); +x_914 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_914, 0, x_913); +lean_ctor_set(x_914, 1, x_4); +lean_inc(x_910); lean_ctor_set_tag(x_25, 1); -lean_ctor_set(x_25, 1, x_902); -lean_ctor_set(x_25, 0, x_898); -x_903 = lean_array_mk(x_25); -x_904 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_905 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_905, 0, x_899); -lean_ctor_set(x_905, 1, x_904); -lean_ctor_set(x_905, 2, x_903); -x_753 = x_905; -x_754 = x_750; -goto block_865; -} -else -{ -lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; -x_906 = lean_ctor_get(x_11, 0); -x_907 = lean_ctor_get(x_12, 0); -lean_inc(x_907); +lean_ctor_set(x_25, 1, x_914); +lean_ctor_set(x_25, 0, x_910); +x_915 = lean_array_mk(x_25); +x_916 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_917 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_917, 0, x_911); +lean_ctor_set(x_917, 1, x_916); +lean_ctor_set(x_917, 2, x_915); +x_763 = x_917; +x_764 = x_760; +goto block_877; +} +else +{ +lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; +x_918 = lean_ctor_get(x_11, 0); +x_919 = lean_ctor_get(x_12, 0); +lean_inc(x_919); lean_dec(x_12); -x_908 = l_Lean_mkIdentFrom(x_906, x_907, x_751); -x_909 = lean_box(2); -x_910 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_911 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_911, 0, x_909); -lean_ctor_set(x_911, 1, x_910); -lean_ctor_set(x_911, 2, x_24); -x_912 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_912, 0, x_911); -lean_ctor_set(x_912, 1, x_4); +x_920 = l_Lean_mkIdentFrom(x_918, x_919, x_761); +x_921 = lean_box(2); +x_922 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_923 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_923, 0, x_921); +lean_ctor_set(x_923, 1, x_922); +lean_ctor_set(x_923, 2, x_24); +x_924 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_924, 0, x_923); +lean_ctor_set(x_924, 1, x_4); lean_ctor_set_tag(x_25, 1); -lean_ctor_set(x_25, 1, x_912); -lean_ctor_set(x_25, 0, x_908); -x_913 = lean_array_mk(x_25); -x_914 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_915 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_915, 0, x_909); -lean_ctor_set(x_915, 1, x_914); -lean_ctor_set(x_915, 2, x_913); -x_753 = x_915; -x_754 = x_750; -goto block_865; -} -} -block_865: -{ -lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; uint8_t x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; size_t x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; -x_755 = l_Lean_Elab_Command_getRef(x_15, x_16, x_754); -x_756 = lean_ctor_get(x_755, 0); -lean_inc(x_756); -x_757 = lean_ctor_get(x_755, 1); -lean_inc(x_757); -if (lean_is_exclusive(x_755)) { - lean_ctor_release(x_755, 0); - lean_ctor_release(x_755, 1); - x_758 = x_755; -} else { - lean_dec_ref(x_755); - x_758 = lean_box(0); -} -x_759 = l_Lean_replaceRef(x_6, x_756); -lean_dec(x_756); -x_760 = lean_ctor_get(x_15, 0); -lean_inc(x_760); -x_761 = lean_ctor_get(x_15, 1); -lean_inc(x_761); -x_762 = lean_ctor_get(x_15, 2); -lean_inc(x_762); -x_763 = lean_ctor_get(x_15, 3); -lean_inc(x_763); -x_764 = lean_ctor_get(x_15, 4); -lean_inc(x_764); -x_765 = lean_ctor_get(x_15, 5); -lean_inc(x_765); -x_766 = lean_ctor_get(x_15, 7); +lean_ctor_set(x_25, 1, x_924); +lean_ctor_set(x_25, 0, x_920); +x_925 = lean_array_mk(x_25); +x_926 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_927 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_927, 0, x_921); +lean_ctor_set(x_927, 1, x_926); +lean_ctor_set(x_927, 2, x_925); +x_763 = x_927; +x_764 = x_760; +goto block_877; +} +} +block_877: +{ +lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; uint8_t x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; size_t x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; +x_765 = l_Lean_Elab_Command_getRef(x_15, x_16, x_764); +x_766 = lean_ctor_get(x_765, 0); lean_inc(x_766); -x_767 = lean_ctor_get(x_15, 8); +x_767 = lean_ctor_get(x_765, 1); lean_inc(x_767); -x_768 = lean_ctor_get(x_15, 9); -lean_inc(x_768); -x_769 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); -x_770 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_770, 0, x_760); -lean_ctor_set(x_770, 1, x_761); -lean_ctor_set(x_770, 2, x_762); -lean_ctor_set(x_770, 3, x_763); -lean_ctor_set(x_770, 4, x_764); -lean_ctor_set(x_770, 5, x_765); -lean_ctor_set(x_770, 6, x_759); -lean_ctor_set(x_770, 7, x_766); -lean_ctor_set(x_770, 8, x_767); -lean_ctor_set(x_770, 9, x_768); -lean_ctor_set_uint8(x_770, sizeof(void*)*10, x_769); -x_771 = l_Lean_Elab_Command_getRef(x_770, x_16, x_757); -x_772 = lean_ctor_get(x_771, 0); -lean_inc(x_772); -x_773 = lean_ctor_get(x_771, 1); -lean_inc(x_773); -if (lean_is_exclusive(x_771)) { - lean_ctor_release(x_771, 0); - lean_ctor_release(x_771, 1); - x_774 = x_771; -} else { - lean_dec_ref(x_771); - x_774 = lean_box(0); -} -x_775 = l_Lean_SourceInfo_fromRef(x_772, x_751); -lean_dec(x_772); -x_776 = l_Lean_Elab_Command_getCurrMacroScope(x_770, x_16, x_773); -lean_dec(x_770); -x_777 = lean_ctor_get(x_776, 1); -lean_inc(x_777); -if (lean_is_exclusive(x_776)) { - lean_ctor_release(x_776, 0); - lean_ctor_release(x_776, 1); - x_778 = x_776; -} else { - lean_dec_ref(x_776); - x_778 = lean_box(0); -} -x_779 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_777); -x_780 = lean_ctor_get(x_779, 1); -lean_inc(x_780); -if (lean_is_exclusive(x_779)) { - lean_ctor_release(x_779, 0); - lean_ctor_release(x_779, 1); - x_781 = x_779; -} else { - lean_dec_ref(x_779); - x_781 = lean_box(0); -} -x_782 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_775); -if (lean_is_scalar(x_781)) { - x_783 = lean_alloc_ctor(2, 2, 0); +if (lean_is_exclusive(x_765)) { + lean_ctor_release(x_765, 0); + lean_ctor_release(x_765, 1); + x_768 = x_765; } else { - x_783 = x_781; - lean_ctor_set_tag(x_783, 2); + lean_dec_ref(x_765); + x_768 = lean_box(0); } -lean_ctor_set(x_783, 0, x_775); -lean_ctor_set(x_783, 1, x_782); -x_784 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_785 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_775); -x_786 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_786, 0, x_775); -lean_ctor_set(x_786, 1, x_784); -lean_ctor_set(x_786, 2, x_785); -x_787 = lean_array_size(x_27); -x_788 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_787, x_20, x_27); -x_789 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_790 = l_Lean_mkSepArray(x_788, x_789); -lean_dec(x_788); -x_791 = l_Array_append___rarg(x_785, x_790); -lean_dec(x_790); -lean_inc(x_775); -x_792 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_792, 0, x_775); -lean_ctor_set(x_792, 1, x_784); -lean_ctor_set(x_792, 2, x_791); -x_793 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_786); -lean_inc(x_775); -x_794 = l_Lean_Syntax_node1(x_775, x_793, x_786); -x_795 = l_Lake_DSL_structVal___closed__18; +x_769 = l_Lean_replaceRef(x_6, x_766); +lean_dec(x_766); +x_770 = lean_ctor_get(x_15, 0); +lean_inc(x_770); +x_771 = lean_ctor_get(x_15, 1); +lean_inc(x_771); +x_772 = lean_ctor_get(x_15, 2); +lean_inc(x_772); +x_773 = lean_ctor_get(x_15, 3); +lean_inc(x_773); +x_774 = lean_ctor_get(x_15, 4); +lean_inc(x_774); +x_775 = lean_ctor_get(x_15, 5); lean_inc(x_775); -if (lean_is_scalar(x_778)) { - x_796 = lean_alloc_ctor(2, 2, 0); +x_776 = lean_ctor_get(x_15, 7); +lean_inc(x_776); +x_777 = lean_ctor_get(x_15, 8); +lean_inc(x_777); +x_778 = lean_ctor_get(x_15, 9); +lean_inc(x_778); +x_779 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); +x_780 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_780, 0, x_770); +lean_ctor_set(x_780, 1, x_771); +lean_ctor_set(x_780, 2, x_772); +lean_ctor_set(x_780, 3, x_773); +lean_ctor_set(x_780, 4, x_774); +lean_ctor_set(x_780, 5, x_775); +lean_ctor_set(x_780, 6, x_769); +lean_ctor_set(x_780, 7, x_776); +lean_ctor_set(x_780, 8, x_777); +lean_ctor_set(x_780, 9, x_778); +lean_ctor_set_uint8(x_780, sizeof(void*)*10, x_779); +x_781 = l_Lean_Elab_Command_getRef(x_780, x_16, x_767); +x_782 = lean_ctor_get(x_781, 0); +lean_inc(x_782); +x_783 = lean_ctor_get(x_781, 1); +lean_inc(x_783); +if (lean_is_exclusive(x_781)) { + lean_ctor_release(x_781, 0); + lean_ctor_release(x_781, 1); + x_784 = x_781; } else { - x_796 = x_778; - lean_ctor_set_tag(x_796, 2); + lean_dec_ref(x_781); + x_784 = lean_box(0); } -lean_ctor_set(x_796, 0, x_775); -lean_ctor_set(x_796, 1, x_795); -x_797 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_786); -x_798 = l_Lean_Syntax_node6(x_775, x_797, x_783, x_786, x_792, x_794, x_786, x_796); -x_799 = l_Lean_Elab_Command_getRef(x_15, x_16, x_780); -x_800 = lean_ctor_get(x_799, 0); -lean_inc(x_800); -x_801 = lean_ctor_get(x_799, 1); -lean_inc(x_801); -if (lean_is_exclusive(x_799)) { - lean_ctor_release(x_799, 0); - lean_ctor_release(x_799, 1); - x_802 = x_799; +x_785 = l_Lean_SourceInfo_fromRef(x_782, x_761); +lean_dec(x_782); +x_786 = l_Lean_Elab_Command_getCurrMacroScope(x_780, x_16, x_783); +lean_dec(x_780); +x_787 = lean_ctor_get(x_786, 1); +lean_inc(x_787); +if (lean_is_exclusive(x_786)) { + lean_ctor_release(x_786, 0); + lean_ctor_release(x_786, 1); + x_788 = x_786; +} else { + lean_dec_ref(x_786); + x_788 = lean_box(0); +} +x_789 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_787); +x_790 = lean_ctor_get(x_789, 1); +lean_inc(x_790); +if (lean_is_exclusive(x_789)) { + lean_ctor_release(x_789, 0); + lean_ctor_release(x_789, 1); + x_791 = x_789; } else { - lean_dec_ref(x_799); - x_802 = lean_box(0); + lean_dec_ref(x_789); + x_791 = lean_box(0); } -x_803 = l_Lean_SourceInfo_fromRef(x_800, x_751); +x_792 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_785); +if (lean_is_scalar(x_791)) { + x_793 = lean_alloc_ctor(2, 2, 0); +} else { + x_793 = x_791; + lean_ctor_set_tag(x_793, 2); +} +lean_ctor_set(x_793, 0, x_785); +lean_ctor_set(x_793, 1, x_792); +x_794 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_795 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_785); +x_796 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_796, 0, x_785); +lean_ctor_set(x_796, 1, x_794); +lean_ctor_set(x_796, 2, x_795); +x_797 = lean_array_size(x_27); +x_798 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_797, x_20, x_27); +x_799 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_800 = l_Lean_mkSepArray(x_798, x_799); +lean_dec(x_798); +x_801 = l_Array_append___rarg(x_795, x_800); lean_dec(x_800); -x_804 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_801); -x_805 = lean_ctor_get(x_804, 1); -lean_inc(x_805); -if (lean_is_exclusive(x_804)) { - lean_ctor_release(x_804, 0); - lean_ctor_release(x_804, 1); - x_806 = x_804; -} else { - lean_dec_ref(x_804); - x_806 = lean_box(0); -} -x_807 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_805); -x_808 = lean_ctor_get(x_807, 1); -lean_inc(x_808); -if (lean_is_exclusive(x_807)) { - lean_ctor_release(x_807, 0); - lean_ctor_release(x_807, 1); - x_809 = x_807; -} else { - lean_dec_ref(x_807); - x_809 = lean_box(0); -} -x_810 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_803); -if (lean_is_scalar(x_809)) { - x_811 = lean_alloc_ctor(2, 2, 0); -} else { - x_811 = x_809; - lean_ctor_set_tag(x_811, 2); -} -lean_ctor_set(x_811, 0, x_803); -lean_ctor_set(x_811, 1, x_810); -x_812 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_813 = l_Lean_Syntax_SepArray_ofElems(x_812, x_7); -x_814 = l_Array_append___rarg(x_785, x_813); -lean_dec(x_813); -lean_inc(x_803); -x_815 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_815, 0, x_803); -lean_ctor_set(x_815, 1, x_784); -lean_ctor_set(x_815, 2, x_814); -x_816 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_803); -if (lean_is_scalar(x_806)) { - x_817 = lean_alloc_ctor(2, 2, 0); -} else { - x_817 = x_806; - lean_ctor_set_tag(x_817, 2); -} -lean_ctor_set(x_817, 0, x_803); -lean_ctor_set(x_817, 1, x_816); -x_818 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_803); -x_819 = l_Lean_Syntax_node3(x_803, x_818, x_811, x_815, x_817); -lean_inc(x_803); -x_820 = l_Lean_Syntax_node1(x_803, x_784, x_819); -lean_inc(x_803); -x_821 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_821, 0, x_803); -lean_ctor_set(x_821, 1, x_784); -lean_ctor_set(x_821, 2, x_785); +lean_inc(x_785); +x_802 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_802, 0, x_785); +lean_ctor_set(x_802, 1, x_794); +lean_ctor_set(x_802, 2, x_801); +x_803 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_785); +x_804 = l_Lean_Syntax_node1(x_785, x_803, x_802); +x_805 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_796); +lean_inc(x_785); +x_806 = l_Lean_Syntax_node1(x_785, x_805, x_796); +x_807 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_785); +if (lean_is_scalar(x_788)) { + x_808 = lean_alloc_ctor(2, 2, 0); +} else { + x_808 = x_788; + lean_ctor_set_tag(x_808, 2); +} +lean_ctor_set(x_808, 0, x_785); +lean_ctor_set(x_808, 1, x_807); +x_809 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_796); +x_810 = l_Lean_Syntax_node6(x_785, x_809, x_793, x_796, x_804, x_806, x_796, x_808); +x_811 = l_Lean_Elab_Command_getRef(x_15, x_16, x_790); +x_812 = lean_ctor_get(x_811, 0); +lean_inc(x_812); +x_813 = lean_ctor_get(x_811, 1); +lean_inc(x_813); +if (lean_is_exclusive(x_811)) { + lean_ctor_release(x_811, 0); + lean_ctor_release(x_811, 1); + x_814 = x_811; +} else { + lean_dec_ref(x_811); + x_814 = lean_box(0); +} +x_815 = l_Lean_SourceInfo_fromRef(x_812, x_761); +lean_dec(x_812); +x_816 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_813); +x_817 = lean_ctor_get(x_816, 1); +lean_inc(x_817); +if (lean_is_exclusive(x_816)) { + lean_ctor_release(x_816, 0); + lean_ctor_release(x_816, 1); + x_818 = x_816; +} else { + lean_dec_ref(x_816); + x_818 = lean_box(0); +} +x_819 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_817); +x_820 = lean_ctor_get(x_819, 1); +lean_inc(x_820); +if (lean_is_exclusive(x_819)) { + lean_ctor_release(x_819, 0); + lean_ctor_release(x_819, 1); + x_821 = x_819; +} else { + lean_dec_ref(x_819); + x_821 = lean_box(0); +} x_822 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_803); -if (lean_is_scalar(x_802)) { +lean_inc(x_815); +if (lean_is_scalar(x_821)) { x_823 = lean_alloc_ctor(2, 2, 0); } else { - x_823 = x_802; + x_823 = x_821; lean_ctor_set_tag(x_823, 2); } -lean_ctor_set(x_823, 0, x_803); +lean_ctor_set(x_823, 0, x_815); lean_ctor_set(x_823, 1, x_822); -x_824 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_803); -if (lean_is_scalar(x_774)) { - x_825 = lean_alloc_ctor(2, 2, 0); -} else { - x_825 = x_774; - lean_ctor_set_tag(x_825, 2); -} -lean_ctor_set(x_825, 0, x_803); -lean_ctor_set(x_825, 1, x_824); -x_826 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_803); -x_827 = l_Lean_Syntax_node2(x_803, x_826, x_825, x_752); -lean_inc(x_803); -x_828 = l_Lean_Syntax_node1(x_803, x_784, x_827); -x_829 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_821); -lean_inc(x_803); -x_830 = l_Lean_Syntax_node2(x_803, x_829, x_821, x_828); -x_831 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_803); -if (lean_is_scalar(x_758)) { - x_832 = lean_alloc_ctor(2, 2, 0); -} else { - x_832 = x_758; - lean_ctor_set_tag(x_832, 2); -} -lean_ctor_set(x_832, 0, x_803); -lean_ctor_set(x_832, 1, x_831); -x_833 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_821, 2); -lean_inc(x_803); -x_834 = l_Lean_Syntax_node2(x_803, x_833, x_821, x_821); +x_824 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_825 = l_Lean_Syntax_SepArray_ofElems(x_824, x_7); +x_826 = l_Array_append___rarg(x_795, x_825); +lean_dec(x_825); +lean_inc(x_815); +x_827 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_827, 0, x_815); +lean_ctor_set(x_827, 1, x_794); +lean_ctor_set(x_827, 2, x_826); +x_828 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_815); +if (lean_is_scalar(x_818)) { + x_829 = lean_alloc_ctor(2, 2, 0); +} else { + x_829 = x_818; + lean_ctor_set_tag(x_829, 2); +} +lean_ctor_set(x_829, 0, x_815); +lean_ctor_set(x_829, 1, x_828); +x_830 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_815); +x_831 = l_Lean_Syntax_node3(x_815, x_830, x_823, x_827, x_829); +lean_inc(x_815); +x_832 = l_Lean_Syntax_node1(x_815, x_794, x_831); +lean_inc(x_815); +x_833 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_833, 0, x_815); +lean_ctor_set(x_833, 1, x_794); +lean_ctor_set(x_833, 2, x_795); +x_834 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_815); +if (lean_is_scalar(x_814)) { + x_835 = lean_alloc_ctor(2, 2, 0); +} else { + x_835 = x_814; + lean_ctor_set_tag(x_835, 2); +} +lean_ctor_set(x_835, 0, x_815); +lean_ctor_set(x_835, 1, x_834); +x_836 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_815); +if (lean_is_scalar(x_784)) { + x_837 = lean_alloc_ctor(2, 2, 0); +} else { + x_837 = x_784; + lean_ctor_set_tag(x_837, 2); +} +lean_ctor_set(x_837, 0, x_815); +lean_ctor_set(x_837, 1, x_836); +x_838 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_815); +x_839 = l_Lean_Syntax_node2(x_815, x_838, x_837, x_762); +lean_inc(x_815); +x_840 = l_Lean_Syntax_node1(x_815, x_794, x_839); +x_841 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_833); +lean_inc(x_815); +x_842 = l_Lean_Syntax_node2(x_815, x_841, x_833, x_840); +x_843 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_815); +if (lean_is_scalar(x_768)) { + x_844 = lean_alloc_ctor(2, 2, 0); +} else { + x_844 = x_768; + lean_ctor_set_tag(x_844, 2); +} +lean_ctor_set(x_844, 0, x_815); +lean_ctor_set(x_844, 1, x_843); +x_845 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_833, 2); +lean_inc(x_815); +x_846 = l_Lean_Syntax_node2(x_815, x_845, x_833, x_833); if (lean_obj_tag(x_10) == 0) { -x_835 = x_785; -goto block_862; +x_847 = x_795; +goto block_874; } else { -lean_object* x_863; lean_object* x_864; -x_863 = lean_ctor_get(x_10, 0); -lean_inc(x_863); +lean_object* x_875; lean_object* x_876; +x_875 = lean_ctor_get(x_10, 0); +lean_inc(x_875); lean_dec(x_10); -x_864 = l_Array_mkArray1___rarg(x_863); -x_835 = x_864; -goto block_862; -} -block_862: -{ -lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; -x_836 = l_Array_append___rarg(x_785, x_835); -lean_dec(x_835); -lean_inc(x_803); -x_837 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_837, 0, x_803); -lean_ctor_set(x_837, 1, x_784); -lean_ctor_set(x_837, 2, x_836); -x_838 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_821, 3); -lean_inc(x_803); -x_839 = l_Lean_Syntax_node6(x_803, x_838, x_837, x_820, x_821, x_821, x_821, x_821); +x_876 = l_Array_mkArray1___rarg(x_875); +x_847 = x_876; +goto block_874; +} +block_874: +{ +lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; +x_848 = l_Array_append___rarg(x_795, x_847); +lean_dec(x_847); +lean_inc(x_815); +x_849 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_849, 0, x_815); +lean_ctor_set(x_849, 1, x_794); +lean_ctor_set(x_849, 2, x_848); +x_850 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_833, 3); +lean_inc(x_815); +x_851 = l_Lean_Syntax_node6(x_815, x_850, x_849, x_832, x_833, x_833, x_833, x_833); if (lean_obj_tag(x_8) == 0) { -lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; -x_840 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_803); -x_841 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_841, 0, x_803); -lean_ctor_set(x_841, 1, x_784); -lean_ctor_set(x_841, 2, x_840); -x_842 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_803); -x_843 = l_Lean_Syntax_node4(x_803, x_842, x_832, x_798, x_834, x_841); -x_844 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_803); -x_845 = l_Lean_Syntax_node4(x_803, x_844, x_823, x_753, x_830, x_843); -x_846 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_847 = l_Lean_Syntax_node2(x_803, x_846, x_839, x_845); -lean_inc(x_847); -x_848 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_848, 0, x_847); -x_849 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_847, x_848, x_15, x_16, x_808); -return x_849; -} -else -{ -lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; -x_850 = lean_ctor_get(x_8, 0); -lean_inc(x_850); -lean_dec(x_8); -x_851 = l_Array_mkArray1___rarg(x_850); -x_852 = l_Array_append___rarg(x_785, x_851); -lean_dec(x_851); -lean_inc(x_803); +lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; +x_852 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_815); x_853 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_853, 0, x_803); -lean_ctor_set(x_853, 1, x_784); +lean_ctor_set(x_853, 0, x_815); +lean_ctor_set(x_853, 1, x_794); lean_ctor_set(x_853, 2, x_852); x_854 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_803); -x_855 = l_Lean_Syntax_node4(x_803, x_854, x_832, x_798, x_834, x_853); -x_856 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_803); -x_857 = l_Lean_Syntax_node4(x_803, x_856, x_823, x_753, x_830, x_855); -x_858 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_859 = l_Lean_Syntax_node2(x_803, x_858, x_839, x_857); +lean_inc(x_815); +x_855 = l_Lean_Syntax_node4(x_815, x_854, x_844, x_810, x_846, x_853); +x_856 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_815); +x_857 = l_Lean_Syntax_node4(x_815, x_856, x_835, x_763, x_842, x_855); +x_858 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_859 = l_Lean_Syntax_node2(x_815, x_858, x_851, x_857); lean_inc(x_859); x_860 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); lean_closure_set(x_860, 0, x_859); -x_861 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_859, x_860, x_15, x_16, x_808); +x_861 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_859, x_860, x_15, x_16, x_820); return x_861; } +else +{ +lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; +x_862 = lean_ctor_get(x_8, 0); +lean_inc(x_862); +lean_dec(x_8); +x_863 = l_Array_mkArray1___rarg(x_862); +x_864 = l_Array_append___rarg(x_795, x_863); +lean_dec(x_863); +lean_inc(x_815); +x_865 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_865, 0, x_815); +lean_ctor_set(x_865, 1, x_794); +lean_ctor_set(x_865, 2, x_864); +x_866 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_815); +x_867 = l_Lean_Syntax_node4(x_815, x_866, x_844, x_810, x_846, x_865); +x_868 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_815); +x_869 = l_Lean_Syntax_node4(x_815, x_868, x_835, x_763, x_842, x_867); +x_870 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_871 = l_Lean_Syntax_node2(x_815, x_870, x_851, x_869); +lean_inc(x_871); +x_872 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_872, 0, x_871); +x_873 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_871, x_872, x_15, x_16, x_820); +return x_873; +} } } } } else { -lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; uint8_t x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; -x_916 = lean_ctor_get(x_25, 0); -x_917 = lean_ctor_get(x_25, 1); -lean_inc(x_917); -lean_inc(x_916); +lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; uint8_t x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; +x_928 = lean_ctor_get(x_25, 0); +x_929 = lean_ctor_get(x_25, 1); +lean_inc(x_929); +lean_inc(x_928); lean_dec(x_25); -x_918 = l_Lean_Elab_Command_getRef(x_15, x_16, x_917); -x_919 = lean_ctor_get(x_918, 0); -lean_inc(x_919); -x_920 = lean_ctor_get(x_918, 1); -lean_inc(x_920); -if (lean_is_exclusive(x_918)) { - lean_ctor_release(x_918, 0); - lean_ctor_release(x_918, 1); - x_921 = x_918; -} else { - lean_dec_ref(x_918); - x_921 = lean_box(0); -} -x_922 = 0; -x_923 = l_Lean_mkCIdentFrom(x_919, x_2, x_922); -lean_dec(x_919); +x_930 = l_Lean_Elab_Command_getRef(x_15, x_16, x_929); +x_931 = lean_ctor_get(x_930, 0); +lean_inc(x_931); +x_932 = lean_ctor_get(x_930, 1); +lean_inc(x_932); +if (lean_is_exclusive(x_930)) { + lean_ctor_release(x_930, 0); + lean_ctor_release(x_930, 1); + x_933 = x_930; +} else { + lean_dec_ref(x_930); + x_933 = lean_box(0); +} +x_934 = 0; +x_935 = l_Lean_mkCIdentFrom(x_931, x_2, x_934); +lean_dec(x_931); if (lean_obj_tag(x_11) == 0) { if (lean_obj_tag(x_12) == 0) { -lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; -x_1037 = l_Lean_Elab_Command_getRef(x_15, x_16, x_920); -x_1038 = lean_ctor_get(x_1037, 0); -lean_inc(x_1038); -x_1039 = lean_ctor_get(x_1037, 1); -lean_inc(x_1039); -if (lean_is_exclusive(x_1037)) { - lean_ctor_release(x_1037, 0); - lean_ctor_release(x_1037, 1); - x_1040 = x_1037; +lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; +x_1051 = l_Lean_Elab_Command_getRef(x_15, x_16, x_932); +x_1052 = lean_ctor_get(x_1051, 0); +lean_inc(x_1052); +x_1053 = lean_ctor_get(x_1051, 1); +lean_inc(x_1053); +if (lean_is_exclusive(x_1051)) { + lean_ctor_release(x_1051, 0); + lean_ctor_release(x_1051, 1); + x_1054 = x_1051; } else { - lean_dec_ref(x_1037); - x_1040 = lean_box(0); + lean_dec_ref(x_1051); + x_1054 = lean_box(0); } lean_inc(x_16); lean_inc(x_15); -x_1041 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_1038, x_922, x_15, x_16, x_1039); -lean_dec(x_1038); -if (lean_obj_tag(x_1041) == 0) -{ -lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; -x_1042 = lean_ctor_get(x_1041, 0); -lean_inc(x_1042); -x_1043 = lean_ctor_get(x_1041, 1); -lean_inc(x_1043); -lean_dec(x_1041); -x_1044 = lean_box(2); -x_1045 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_1046 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1046, 0, x_1044); -lean_ctor_set(x_1046, 1, x_1045); -lean_ctor_set(x_1046, 2, x_24); -if (lean_is_scalar(x_1040)) { - x_1047 = lean_alloc_ctor(1, 2, 0); -} else { - x_1047 = x_1040; - lean_ctor_set_tag(x_1047, 1); -} -lean_ctor_set(x_1047, 0, x_1046); -lean_ctor_set(x_1047, 1, x_4); -if (lean_is_scalar(x_921)) { - x_1048 = lean_alloc_ctor(1, 2, 0); -} else { - x_1048 = x_921; - lean_ctor_set_tag(x_1048, 1); -} -lean_ctor_set(x_1048, 0, x_1042); -lean_ctor_set(x_1048, 1, x_1047); -x_1049 = lean_array_mk(x_1048); -x_1050 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_1051 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1051, 0, x_1044); -lean_ctor_set(x_1051, 1, x_1050); -lean_ctor_set(x_1051, 2, x_1049); -x_924 = x_1051; -x_925 = x_1043; -goto block_1036; -} -else -{ -lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; -lean_dec(x_1040); -lean_dec(x_923); -lean_dec(x_921); -lean_dec(x_916); +x_1055 = l_Lean_Elab_Term_mkFreshIdent___at_Lake_DSL_elabConfigDecl___spec__6(x_1052, x_934, x_15, x_16, x_1053); +lean_dec(x_1052); +if (lean_obj_tag(x_1055) == 0) +{ +lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; +x_1056 = lean_ctor_get(x_1055, 0); +lean_inc(x_1056); +x_1057 = lean_ctor_get(x_1055, 1); +lean_inc(x_1057); +lean_dec(x_1055); +x_1058 = lean_box(2); +x_1059 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_1060 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1060, 0, x_1058); +lean_ctor_set(x_1060, 1, x_1059); +lean_ctor_set(x_1060, 2, x_24); +if (lean_is_scalar(x_1054)) { + x_1061 = lean_alloc_ctor(1, 2, 0); +} else { + x_1061 = x_1054; + lean_ctor_set_tag(x_1061, 1); +} +lean_ctor_set(x_1061, 0, x_1060); +lean_ctor_set(x_1061, 1, x_4); +if (lean_is_scalar(x_933)) { + x_1062 = lean_alloc_ctor(1, 2, 0); +} else { + x_1062 = x_933; + lean_ctor_set_tag(x_1062, 1); +} +lean_ctor_set(x_1062, 0, x_1056); +lean_ctor_set(x_1062, 1, x_1061); +x_1063 = lean_array_mk(x_1062); +x_1064 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_1065 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1065, 0, x_1058); +lean_ctor_set(x_1065, 1, x_1064); +lean_ctor_set(x_1065, 2, x_1063); +x_936 = x_1065; +x_937 = x_1057; +goto block_1050; +} +else +{ +lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; +lean_dec(x_1054); +lean_dec(x_935); +lean_dec(x_933); +lean_dec(x_928); lean_dec(x_24); lean_dec(x_16); lean_dec(x_15); @@ -10315,481 +10379,484 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_1052 = lean_ctor_get(x_1041, 0); -lean_inc(x_1052); -x_1053 = lean_ctor_get(x_1041, 1); -lean_inc(x_1053); -if (lean_is_exclusive(x_1041)) { - lean_ctor_release(x_1041, 0); - lean_ctor_release(x_1041, 1); - x_1054 = x_1041; +x_1066 = lean_ctor_get(x_1055, 0); +lean_inc(x_1066); +x_1067 = lean_ctor_get(x_1055, 1); +lean_inc(x_1067); +if (lean_is_exclusive(x_1055)) { + lean_ctor_release(x_1055, 0); + lean_ctor_release(x_1055, 1); + x_1068 = x_1055; } else { - lean_dec_ref(x_1041); - x_1054 = lean_box(0); + lean_dec_ref(x_1055); + x_1068 = lean_box(0); } -if (lean_is_scalar(x_1054)) { - x_1055 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1068)) { + x_1069 = lean_alloc_ctor(1, 2, 0); } else { - x_1055 = x_1054; + x_1069 = x_1068; } -lean_ctor_set(x_1055, 0, x_1052); -lean_ctor_set(x_1055, 1, x_1053); -return x_1055; +lean_ctor_set(x_1069, 0, x_1066); +lean_ctor_set(x_1069, 1, x_1067); +return x_1069; } } else { -lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; -x_1056 = lean_ctor_get(x_12, 0); -lean_inc(x_1056); +lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; +x_1070 = lean_ctor_get(x_12, 0); +lean_inc(x_1070); lean_dec(x_12); -x_1057 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_1056, x_922, x_15, x_16, x_920); -x_1058 = lean_ctor_get(x_1057, 0); -lean_inc(x_1058); -x_1059 = lean_ctor_get(x_1057, 1); -lean_inc(x_1059); -if (lean_is_exclusive(x_1057)) { - lean_ctor_release(x_1057, 0); - lean_ctor_release(x_1057, 1); - x_1060 = x_1057; -} else { - lean_dec_ref(x_1057); - x_1060 = lean_box(0); -} -x_1061 = lean_box(2); -x_1062 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_1063 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1063, 0, x_1061); -lean_ctor_set(x_1063, 1, x_1062); -lean_ctor_set(x_1063, 2, x_24); -if (lean_is_scalar(x_1060)) { - x_1064 = lean_alloc_ctor(1, 2, 0); -} else { - x_1064 = x_1060; - lean_ctor_set_tag(x_1064, 1); -} -lean_ctor_set(x_1064, 0, x_1063); -lean_ctor_set(x_1064, 1, x_4); -if (lean_is_scalar(x_921)) { - x_1065 = lean_alloc_ctor(1, 2, 0); -} else { - x_1065 = x_921; - lean_ctor_set_tag(x_1065, 1); +x_1071 = l_Lean_mkIdentFromRef___at_Lake_DSL_elabConfigDecl___spec__8(x_1070, x_934, x_15, x_16, x_932); +x_1072 = lean_ctor_get(x_1071, 0); +lean_inc(x_1072); +x_1073 = lean_ctor_get(x_1071, 1); +lean_inc(x_1073); +if (lean_is_exclusive(x_1071)) { + lean_ctor_release(x_1071, 0); + lean_ctor_release(x_1071, 1); + x_1074 = x_1071; +} else { + lean_dec_ref(x_1071); + x_1074 = lean_box(0); +} +x_1075 = lean_box(2); +x_1076 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_1077 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1077, 0, x_1075); +lean_ctor_set(x_1077, 1, x_1076); +lean_ctor_set(x_1077, 2, x_24); +if (lean_is_scalar(x_1074)) { + x_1078 = lean_alloc_ctor(1, 2, 0); +} else { + x_1078 = x_1074; + lean_ctor_set_tag(x_1078, 1); } -lean_ctor_set(x_1065, 0, x_1058); -lean_ctor_set(x_1065, 1, x_1064); -x_1066 = lean_array_mk(x_1065); -x_1067 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_1068 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1068, 0, x_1061); -lean_ctor_set(x_1068, 1, x_1067); -lean_ctor_set(x_1068, 2, x_1066); -x_924 = x_1068; -x_925 = x_1059; -goto block_1036; +lean_ctor_set(x_1078, 0, x_1077); +lean_ctor_set(x_1078, 1, x_4); +if (lean_is_scalar(x_933)) { + x_1079 = lean_alloc_ctor(1, 2, 0); +} else { + x_1079 = x_933; + lean_ctor_set_tag(x_1079, 1); +} +lean_ctor_set(x_1079, 0, x_1072); +lean_ctor_set(x_1079, 1, x_1078); +x_1080 = lean_array_mk(x_1079); +x_1081 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_1082 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1082, 0, x_1075); +lean_ctor_set(x_1082, 1, x_1081); +lean_ctor_set(x_1082, 2, x_1080); +x_936 = x_1082; +x_937 = x_1073; +goto block_1050; } } else { if (lean_obj_tag(x_12) == 0) { -lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; -x_1069 = lean_ctor_get(x_11, 0); -x_1070 = lean_box(2); -x_1071 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_1072 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1072, 0, x_1070); -lean_ctor_set(x_1072, 1, x_1071); -lean_ctor_set(x_1072, 2, x_24); -if (lean_is_scalar(x_921)) { - x_1073 = lean_alloc_ctor(1, 2, 0); -} else { - x_1073 = x_921; - lean_ctor_set_tag(x_1073, 1); -} -lean_ctor_set(x_1073, 0, x_1072); -lean_ctor_set(x_1073, 1, x_4); -lean_inc(x_1069); -x_1074 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1074, 0, x_1069); -lean_ctor_set(x_1074, 1, x_1073); -x_1075 = lean_array_mk(x_1074); -x_1076 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_1077 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1077, 0, x_1070); -lean_ctor_set(x_1077, 1, x_1076); -lean_ctor_set(x_1077, 2, x_1075); -x_924 = x_1077; -x_925 = x_920; -goto block_1036; -} -else -{ -lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; -x_1078 = lean_ctor_get(x_11, 0); -x_1079 = lean_ctor_get(x_12, 0); -lean_inc(x_1079); -lean_dec(x_12); -x_1080 = l_Lean_mkIdentFrom(x_1078, x_1079, x_922); -x_1081 = lean_box(2); -x_1082 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_1083 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1083, 0, x_1081); -lean_ctor_set(x_1083, 1, x_1082); -lean_ctor_set(x_1083, 2, x_24); -if (lean_is_scalar(x_921)) { - x_1084 = lean_alloc_ctor(1, 2, 0); -} else { - x_1084 = x_921; - lean_ctor_set_tag(x_1084, 1); -} -lean_ctor_set(x_1084, 0, x_1083); -lean_ctor_set(x_1084, 1, x_4); -x_1085 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1085, 0, x_1080); -lean_ctor_set(x_1085, 1, x_1084); -x_1086 = lean_array_mk(x_1085); -x_1087 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_1088 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1088, 0, x_1081); +lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; +x_1083 = lean_ctor_get(x_11, 0); +x_1084 = lean_box(2); +x_1085 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_1086 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1086, 0, x_1084); +lean_ctor_set(x_1086, 1, x_1085); +lean_ctor_set(x_1086, 2, x_24); +if (lean_is_scalar(x_933)) { + x_1087 = lean_alloc_ctor(1, 2, 0); +} else { + x_1087 = x_933; + lean_ctor_set_tag(x_1087, 1); +} +lean_ctor_set(x_1087, 0, x_1086); +lean_ctor_set(x_1087, 1, x_4); +lean_inc(x_1083); +x_1088 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1088, 0, x_1083); lean_ctor_set(x_1088, 1, x_1087); -lean_ctor_set(x_1088, 2, x_1086); -x_924 = x_1088; -x_925 = x_920; -goto block_1036; -} -} -block_1036: -{ -lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; uint8_t x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; size_t x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; -x_926 = l_Lean_Elab_Command_getRef(x_15, x_16, x_925); -x_927 = lean_ctor_get(x_926, 0); -lean_inc(x_927); -x_928 = lean_ctor_get(x_926, 1); -lean_inc(x_928); -if (lean_is_exclusive(x_926)) { - lean_ctor_release(x_926, 0); - lean_ctor_release(x_926, 1); - x_929 = x_926; -} else { - lean_dec_ref(x_926); - x_929 = lean_box(0); -} -x_930 = l_Lean_replaceRef(x_6, x_927); -lean_dec(x_927); -x_931 = lean_ctor_get(x_15, 0); -lean_inc(x_931); -x_932 = lean_ctor_get(x_15, 1); -lean_inc(x_932); -x_933 = lean_ctor_get(x_15, 2); -lean_inc(x_933); -x_934 = lean_ctor_get(x_15, 3); -lean_inc(x_934); -x_935 = lean_ctor_get(x_15, 4); -lean_inc(x_935); -x_936 = lean_ctor_get(x_15, 5); -lean_inc(x_936); -x_937 = lean_ctor_get(x_15, 7); -lean_inc(x_937); -x_938 = lean_ctor_get(x_15, 8); -lean_inc(x_938); -x_939 = lean_ctor_get(x_15, 9); +x_1089 = lean_array_mk(x_1088); +x_1090 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_1091 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1091, 0, x_1084); +lean_ctor_set(x_1091, 1, x_1090); +lean_ctor_set(x_1091, 2, x_1089); +x_936 = x_1091; +x_937 = x_932; +goto block_1050; +} +else +{ +lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; +x_1092 = lean_ctor_get(x_11, 0); +x_1093 = lean_ctor_get(x_12, 0); +lean_inc(x_1093); +lean_dec(x_12); +x_1094 = l_Lean_mkIdentFrom(x_1092, x_1093, x_934); +x_1095 = lean_box(2); +x_1096 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_1097 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1097, 0, x_1095); +lean_ctor_set(x_1097, 1, x_1096); +lean_ctor_set(x_1097, 2, x_24); +if (lean_is_scalar(x_933)) { + x_1098 = lean_alloc_ctor(1, 2, 0); +} else { + x_1098 = x_933; + lean_ctor_set_tag(x_1098, 1); +} +lean_ctor_set(x_1098, 0, x_1097); +lean_ctor_set(x_1098, 1, x_4); +x_1099 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1099, 0, x_1094); +lean_ctor_set(x_1099, 1, x_1098); +x_1100 = lean_array_mk(x_1099); +x_1101 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_1102 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1102, 0, x_1095); +lean_ctor_set(x_1102, 1, x_1101); +lean_ctor_set(x_1102, 2, x_1100); +x_936 = x_1102; +x_937 = x_932; +goto block_1050; +} +} +block_1050: +{ +lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; uint8_t x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; size_t x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; +x_938 = l_Lean_Elab_Command_getRef(x_15, x_16, x_937); +x_939 = lean_ctor_get(x_938, 0); lean_inc(x_939); -x_940 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); -x_941 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_941, 0, x_931); -lean_ctor_set(x_941, 1, x_932); -lean_ctor_set(x_941, 2, x_933); -lean_ctor_set(x_941, 3, x_934); -lean_ctor_set(x_941, 4, x_935); -lean_ctor_set(x_941, 5, x_936); -lean_ctor_set(x_941, 6, x_930); -lean_ctor_set(x_941, 7, x_937); -lean_ctor_set(x_941, 8, x_938); -lean_ctor_set(x_941, 9, x_939); -lean_ctor_set_uint8(x_941, sizeof(void*)*10, x_940); -x_942 = l_Lean_Elab_Command_getRef(x_941, x_16, x_928); -x_943 = lean_ctor_get(x_942, 0); +x_940 = lean_ctor_get(x_938, 1); +lean_inc(x_940); +if (lean_is_exclusive(x_938)) { + lean_ctor_release(x_938, 0); + lean_ctor_release(x_938, 1); + x_941 = x_938; +} else { + lean_dec_ref(x_938); + x_941 = lean_box(0); +} +x_942 = l_Lean_replaceRef(x_6, x_939); +lean_dec(x_939); +x_943 = lean_ctor_get(x_15, 0); lean_inc(x_943); -x_944 = lean_ctor_get(x_942, 1); +x_944 = lean_ctor_get(x_15, 1); lean_inc(x_944); -if (lean_is_exclusive(x_942)) { - lean_ctor_release(x_942, 0); - lean_ctor_release(x_942, 1); - x_945 = x_942; -} else { - lean_dec_ref(x_942); - x_945 = lean_box(0); -} -x_946 = l_Lean_SourceInfo_fromRef(x_943, x_922); -lean_dec(x_943); -x_947 = l_Lean_Elab_Command_getCurrMacroScope(x_941, x_16, x_944); -lean_dec(x_941); -x_948 = lean_ctor_get(x_947, 1); +x_945 = lean_ctor_get(x_15, 2); +lean_inc(x_945); +x_946 = lean_ctor_get(x_15, 3); +lean_inc(x_946); +x_947 = lean_ctor_get(x_15, 4); +lean_inc(x_947); +x_948 = lean_ctor_get(x_15, 5); lean_inc(x_948); -if (lean_is_exclusive(x_947)) { - lean_ctor_release(x_947, 0); - lean_ctor_release(x_947, 1); - x_949 = x_947; -} else { - lean_dec_ref(x_947); - x_949 = lean_box(0); -} -x_950 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_948); -x_951 = lean_ctor_get(x_950, 1); +x_949 = lean_ctor_get(x_15, 7); +lean_inc(x_949); +x_950 = lean_ctor_get(x_15, 8); +lean_inc(x_950); +x_951 = lean_ctor_get(x_15, 9); lean_inc(x_951); -if (lean_is_exclusive(x_950)) { - lean_ctor_release(x_950, 0); - lean_ctor_release(x_950, 1); - x_952 = x_950; -} else { - lean_dec_ref(x_950); - x_952 = lean_box(0); -} -x_953 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_946); -if (lean_is_scalar(x_952)) { - x_954 = lean_alloc_ctor(2, 2, 0); -} else { - x_954 = x_952; - lean_ctor_set_tag(x_954, 2); -} -lean_ctor_set(x_954, 0, x_946); -lean_ctor_set(x_954, 1, x_953); -x_955 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_956 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_946); -x_957 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_957, 0, x_946); -lean_ctor_set(x_957, 1, x_955); -lean_ctor_set(x_957, 2, x_956); -x_958 = lean_array_size(x_916); -x_959 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_958, x_20, x_916); -x_960 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_961 = l_Lean_mkSepArray(x_959, x_960); -lean_dec(x_959); -x_962 = l_Array_append___rarg(x_956, x_961); -lean_dec(x_961); -lean_inc(x_946); -x_963 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_963, 0, x_946); -lean_ctor_set(x_963, 1, x_955); -lean_ctor_set(x_963, 2, x_962); -x_964 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_957); -lean_inc(x_946); -x_965 = l_Lean_Syntax_node1(x_946, x_964, x_957); -x_966 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_946); -if (lean_is_scalar(x_949)) { - x_967 = lean_alloc_ctor(2, 2, 0); -} else { - x_967 = x_949; - lean_ctor_set_tag(x_967, 2); -} -lean_ctor_set(x_967, 0, x_946); -lean_ctor_set(x_967, 1, x_966); -x_968 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_957); -x_969 = l_Lean_Syntax_node6(x_946, x_968, x_954, x_957, x_963, x_965, x_957, x_967); -x_970 = l_Lean_Elab_Command_getRef(x_15, x_16, x_951); -x_971 = lean_ctor_get(x_970, 0); -lean_inc(x_971); -x_972 = lean_ctor_get(x_970, 1); -lean_inc(x_972); -if (lean_is_exclusive(x_970)) { - lean_ctor_release(x_970, 0); - lean_ctor_release(x_970, 1); - x_973 = x_970; -} else { - lean_dec_ref(x_970); - x_973 = lean_box(0); -} -x_974 = l_Lean_SourceInfo_fromRef(x_971, x_922); +x_952 = lean_ctor_get_uint8(x_15, sizeof(void*)*10); +x_953 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_953, 0, x_943); +lean_ctor_set(x_953, 1, x_944); +lean_ctor_set(x_953, 2, x_945); +lean_ctor_set(x_953, 3, x_946); +lean_ctor_set(x_953, 4, x_947); +lean_ctor_set(x_953, 5, x_948); +lean_ctor_set(x_953, 6, x_942); +lean_ctor_set(x_953, 7, x_949); +lean_ctor_set(x_953, 8, x_950); +lean_ctor_set(x_953, 9, x_951); +lean_ctor_set_uint8(x_953, sizeof(void*)*10, x_952); +x_954 = l_Lean_Elab_Command_getRef(x_953, x_16, x_940); +x_955 = lean_ctor_get(x_954, 0); +lean_inc(x_955); +x_956 = lean_ctor_get(x_954, 1); +lean_inc(x_956); +if (lean_is_exclusive(x_954)) { + lean_ctor_release(x_954, 0); + lean_ctor_release(x_954, 1); + x_957 = x_954; +} else { + lean_dec_ref(x_954); + x_957 = lean_box(0); +} +x_958 = l_Lean_SourceInfo_fromRef(x_955, x_934); +lean_dec(x_955); +x_959 = l_Lean_Elab_Command_getCurrMacroScope(x_953, x_16, x_956); +lean_dec(x_953); +x_960 = lean_ctor_get(x_959, 1); +lean_inc(x_960); +if (lean_is_exclusive(x_959)) { + lean_ctor_release(x_959, 0); + lean_ctor_release(x_959, 1); + x_961 = x_959; +} else { + lean_dec_ref(x_959); + x_961 = lean_box(0); +} +x_962 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_960); +x_963 = lean_ctor_get(x_962, 1); +lean_inc(x_963); +if (lean_is_exclusive(x_962)) { + lean_ctor_release(x_962, 0); + lean_ctor_release(x_962, 1); + x_964 = x_962; +} else { + lean_dec_ref(x_962); + x_964 = lean_box(0); +} +x_965 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_958); +if (lean_is_scalar(x_964)) { + x_966 = lean_alloc_ctor(2, 2, 0); +} else { + x_966 = x_964; + lean_ctor_set_tag(x_966, 2); +} +lean_ctor_set(x_966, 0, x_958); +lean_ctor_set(x_966, 1, x_965); +x_967 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_968 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_958); +x_969 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_969, 0, x_958); +lean_ctor_set(x_969, 1, x_967); +lean_ctor_set(x_969, 2, x_968); +x_970 = lean_array_size(x_928); +x_971 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_970, x_20, x_928); +x_972 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_973 = l_Lean_mkSepArray(x_971, x_972); lean_dec(x_971); -x_975 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_972); -x_976 = lean_ctor_get(x_975, 1); -lean_inc(x_976); -if (lean_is_exclusive(x_975)) { - lean_ctor_release(x_975, 0); - lean_ctor_release(x_975, 1); - x_977 = x_975; -} else { - lean_dec_ref(x_975); - x_977 = lean_box(0); -} -x_978 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_976); -x_979 = lean_ctor_get(x_978, 1); -lean_inc(x_979); -if (lean_is_exclusive(x_978)) { - lean_ctor_release(x_978, 0); - lean_ctor_release(x_978, 1); - x_980 = x_978; -} else { - lean_dec_ref(x_978); - x_980 = lean_box(0); -} -x_981 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_974); -if (lean_is_scalar(x_980)) { - x_982 = lean_alloc_ctor(2, 2, 0); -} else { - x_982 = x_980; - lean_ctor_set_tag(x_982, 2); -} -lean_ctor_set(x_982, 0, x_974); -lean_ctor_set(x_982, 1, x_981); -x_983 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_984 = l_Lean_Syntax_SepArray_ofElems(x_983, x_7); -x_985 = l_Array_append___rarg(x_956, x_984); -lean_dec(x_984); -lean_inc(x_974); -x_986 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_986, 0, x_974); -lean_ctor_set(x_986, 1, x_955); -lean_ctor_set(x_986, 2, x_985); -x_987 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_974); -if (lean_is_scalar(x_977)) { - x_988 = lean_alloc_ctor(2, 2, 0); -} else { - x_988 = x_977; - lean_ctor_set_tag(x_988, 2); -} -lean_ctor_set(x_988, 0, x_974); -lean_ctor_set(x_988, 1, x_987); -x_989 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_974); -x_990 = l_Lean_Syntax_node3(x_974, x_989, x_982, x_986, x_988); -lean_inc(x_974); -x_991 = l_Lean_Syntax_node1(x_974, x_955, x_990); -lean_inc(x_974); -x_992 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_992, 0, x_974); -lean_ctor_set(x_992, 1, x_955); -lean_ctor_set(x_992, 2, x_956); -x_993 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_974); -if (lean_is_scalar(x_973)) { - x_994 = lean_alloc_ctor(2, 2, 0); -} else { - x_994 = x_973; - lean_ctor_set_tag(x_994, 2); -} -lean_ctor_set(x_994, 0, x_974); -lean_ctor_set(x_994, 1, x_993); -x_995 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_974); -if (lean_is_scalar(x_945)) { +x_974 = l_Array_append___rarg(x_968, x_973); +lean_dec(x_973); +lean_inc(x_958); +x_975 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_975, 0, x_958); +lean_ctor_set(x_975, 1, x_967); +lean_ctor_set(x_975, 2, x_974); +x_976 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_958); +x_977 = l_Lean_Syntax_node1(x_958, x_976, x_975); +x_978 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_969); +lean_inc(x_958); +x_979 = l_Lean_Syntax_node1(x_958, x_978, x_969); +x_980 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_958); +if (lean_is_scalar(x_961)) { + x_981 = lean_alloc_ctor(2, 2, 0); +} else { + x_981 = x_961; + lean_ctor_set_tag(x_981, 2); +} +lean_ctor_set(x_981, 0, x_958); +lean_ctor_set(x_981, 1, x_980); +x_982 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_969); +x_983 = l_Lean_Syntax_node6(x_958, x_982, x_966, x_969, x_977, x_979, x_969, x_981); +x_984 = l_Lean_Elab_Command_getRef(x_15, x_16, x_963); +x_985 = lean_ctor_get(x_984, 0); +lean_inc(x_985); +x_986 = lean_ctor_get(x_984, 1); +lean_inc(x_986); +if (lean_is_exclusive(x_984)) { + lean_ctor_release(x_984, 0); + lean_ctor_release(x_984, 1); + x_987 = x_984; +} else { + lean_dec_ref(x_984); + x_987 = lean_box(0); +} +x_988 = l_Lean_SourceInfo_fromRef(x_985, x_934); +lean_dec(x_985); +x_989 = l_Lean_Elab_Command_getCurrMacroScope(x_15, x_16, x_986); +x_990 = lean_ctor_get(x_989, 1); +lean_inc(x_990); +if (lean_is_exclusive(x_989)) { + lean_ctor_release(x_989, 0); + lean_ctor_release(x_989, 1); + x_991 = x_989; +} else { + lean_dec_ref(x_989); + x_991 = lean_box(0); +} +x_992 = l_Lean_Elab_Command_getMainModule___rarg(x_16, x_990); +x_993 = lean_ctor_get(x_992, 1); +lean_inc(x_993); +if (lean_is_exclusive(x_992)) { + lean_ctor_release(x_992, 0); + lean_ctor_release(x_992, 1); + x_994 = x_992; +} else { + lean_dec_ref(x_992); + x_994 = lean_box(0); +} +x_995 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_988); +if (lean_is_scalar(x_994)) { x_996 = lean_alloc_ctor(2, 2, 0); } else { - x_996 = x_945; + x_996 = x_994; lean_ctor_set_tag(x_996, 2); } -lean_ctor_set(x_996, 0, x_974); +lean_ctor_set(x_996, 0, x_988); lean_ctor_set(x_996, 1, x_995); -x_997 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_974); -x_998 = l_Lean_Syntax_node2(x_974, x_997, x_996, x_923); -lean_inc(x_974); -x_999 = l_Lean_Syntax_node1(x_974, x_955, x_998); -x_1000 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_992); -lean_inc(x_974); -x_1001 = l_Lean_Syntax_node2(x_974, x_1000, x_992, x_999); -x_1002 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_974); -if (lean_is_scalar(x_929)) { - x_1003 = lean_alloc_ctor(2, 2, 0); -} else { - x_1003 = x_929; - lean_ctor_set_tag(x_1003, 2); -} -lean_ctor_set(x_1003, 0, x_974); -lean_ctor_set(x_1003, 1, x_1002); -x_1004 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_992, 2); -lean_inc(x_974); -x_1005 = l_Lean_Syntax_node2(x_974, x_1004, x_992, x_992); +x_997 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_998 = l_Lean_Syntax_SepArray_ofElems(x_997, x_7); +x_999 = l_Array_append___rarg(x_968, x_998); +lean_dec(x_998); +lean_inc(x_988); +x_1000 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1000, 0, x_988); +lean_ctor_set(x_1000, 1, x_967); +lean_ctor_set(x_1000, 2, x_999); +x_1001 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_988); +if (lean_is_scalar(x_991)) { + x_1002 = lean_alloc_ctor(2, 2, 0); +} else { + x_1002 = x_991; + lean_ctor_set_tag(x_1002, 2); +} +lean_ctor_set(x_1002, 0, x_988); +lean_ctor_set(x_1002, 1, x_1001); +x_1003 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_988); +x_1004 = l_Lean_Syntax_node3(x_988, x_1003, x_996, x_1000, x_1002); +lean_inc(x_988); +x_1005 = l_Lean_Syntax_node1(x_988, x_967, x_1004); +lean_inc(x_988); +x_1006 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1006, 0, x_988); +lean_ctor_set(x_1006, 1, x_967); +lean_ctor_set(x_1006, 2, x_968); +x_1007 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_988); +if (lean_is_scalar(x_987)) { + x_1008 = lean_alloc_ctor(2, 2, 0); +} else { + x_1008 = x_987; + lean_ctor_set_tag(x_1008, 2); +} +lean_ctor_set(x_1008, 0, x_988); +lean_ctor_set(x_1008, 1, x_1007); +x_1009 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_988); +if (lean_is_scalar(x_957)) { + x_1010 = lean_alloc_ctor(2, 2, 0); +} else { + x_1010 = x_957; + lean_ctor_set_tag(x_1010, 2); +} +lean_ctor_set(x_1010, 0, x_988); +lean_ctor_set(x_1010, 1, x_1009); +x_1011 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_988); +x_1012 = l_Lean_Syntax_node2(x_988, x_1011, x_1010, x_935); +lean_inc(x_988); +x_1013 = l_Lean_Syntax_node1(x_988, x_967, x_1012); +x_1014 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_1006); +lean_inc(x_988); +x_1015 = l_Lean_Syntax_node2(x_988, x_1014, x_1006, x_1013); +x_1016 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_988); +if (lean_is_scalar(x_941)) { + x_1017 = lean_alloc_ctor(2, 2, 0); +} else { + x_1017 = x_941; + lean_ctor_set_tag(x_1017, 2); +} +lean_ctor_set(x_1017, 0, x_988); +lean_ctor_set(x_1017, 1, x_1016); +x_1018 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_1006, 2); +lean_inc(x_988); +x_1019 = l_Lean_Syntax_node2(x_988, x_1018, x_1006, x_1006); if (lean_obj_tag(x_10) == 0) { -x_1006 = x_956; -goto block_1033; +x_1020 = x_968; +goto block_1047; } else { -lean_object* x_1034; lean_object* x_1035; -x_1034 = lean_ctor_get(x_10, 0); -lean_inc(x_1034); +lean_object* x_1048; lean_object* x_1049; +x_1048 = lean_ctor_get(x_10, 0); +lean_inc(x_1048); lean_dec(x_10); -x_1035 = l_Array_mkArray1___rarg(x_1034); -x_1006 = x_1035; -goto block_1033; -} -block_1033: -{ -lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; -x_1007 = l_Array_append___rarg(x_956, x_1006); -lean_dec(x_1006); -lean_inc(x_974); -x_1008 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1008, 0, x_974); -lean_ctor_set(x_1008, 1, x_955); -lean_ctor_set(x_1008, 2, x_1007); -x_1009 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_992, 3); -lean_inc(x_974); -x_1010 = l_Lean_Syntax_node6(x_974, x_1009, x_1008, x_991, x_992, x_992, x_992, x_992); +x_1049 = l_Array_mkArray1___rarg(x_1048); +x_1020 = x_1049; +goto block_1047; +} +block_1047: +{ +lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; +x_1021 = l_Array_append___rarg(x_968, x_1020); +lean_dec(x_1020); +lean_inc(x_988); +x_1022 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1022, 0, x_988); +lean_ctor_set(x_1022, 1, x_967); +lean_ctor_set(x_1022, 2, x_1021); +x_1023 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_1006, 3); +lean_inc(x_988); +x_1024 = l_Lean_Syntax_node6(x_988, x_1023, x_1022, x_1005, x_1006, x_1006, x_1006, x_1006); if (lean_obj_tag(x_8) == 0) { -lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; -x_1011 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_974); -x_1012 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1012, 0, x_974); -lean_ctor_set(x_1012, 1, x_955); -lean_ctor_set(x_1012, 2, x_1011); -x_1013 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_974); -x_1014 = l_Lean_Syntax_node4(x_974, x_1013, x_1003, x_969, x_1005, x_1012); -x_1015 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_974); -x_1016 = l_Lean_Syntax_node4(x_974, x_1015, x_994, x_924, x_1001, x_1014); -x_1017 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_1018 = l_Lean_Syntax_node2(x_974, x_1017, x_1010, x_1016); -lean_inc(x_1018); -x_1019 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_1019, 0, x_1018); -x_1020 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_1018, x_1019, x_15, x_16, x_979); -return x_1020; -} -else -{ -lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; -x_1021 = lean_ctor_get(x_8, 0); -lean_inc(x_1021); +lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; +x_1025 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_988); +x_1026 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1026, 0, x_988); +lean_ctor_set(x_1026, 1, x_967); +lean_ctor_set(x_1026, 2, x_1025); +x_1027 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_988); +x_1028 = l_Lean_Syntax_node4(x_988, x_1027, x_1017, x_983, x_1019, x_1026); +x_1029 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_988); +x_1030 = l_Lean_Syntax_node4(x_988, x_1029, x_1008, x_936, x_1015, x_1028); +x_1031 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_1032 = l_Lean_Syntax_node2(x_988, x_1031, x_1024, x_1030); +lean_inc(x_1032); +x_1033 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_1033, 0, x_1032); +x_1034 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_1032, x_1033, x_15, x_16, x_993); +return x_1034; +} +else +{ +lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; +x_1035 = lean_ctor_get(x_8, 0); +lean_inc(x_1035); lean_dec(x_8); -x_1022 = l_Array_mkArray1___rarg(x_1021); -x_1023 = l_Array_append___rarg(x_956, x_1022); -lean_dec(x_1022); -lean_inc(x_974); -x_1024 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1024, 0, x_974); -lean_ctor_set(x_1024, 1, x_955); -lean_ctor_set(x_1024, 2, x_1023); -x_1025 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_974); -x_1026 = l_Lean_Syntax_node4(x_974, x_1025, x_1003, x_969, x_1005, x_1024); -x_1027 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_974); -x_1028 = l_Lean_Syntax_node4(x_974, x_1027, x_994, x_924, x_1001, x_1026); -x_1029 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_1030 = l_Lean_Syntax_node2(x_974, x_1029, x_1010, x_1028); -lean_inc(x_1030); -x_1031 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_1031, 0, x_1030); -x_1032 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_1030, x_1031, x_15, x_16, x_979); -return x_1032; +x_1036 = l_Array_mkArray1___rarg(x_1035); +x_1037 = l_Array_append___rarg(x_968, x_1036); +lean_dec(x_1036); +lean_inc(x_988); +x_1038 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1038, 0, x_988); +lean_ctor_set(x_1038, 1, x_967); +lean_ctor_set(x_1038, 2, x_1037); +x_1039 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_988); +x_1040 = l_Lean_Syntax_node4(x_988, x_1039, x_1017, x_983, x_1019, x_1038); +x_1041 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_988); +x_1042 = l_Lean_Syntax_node4(x_988, x_1041, x_1008, x_936, x_1015, x_1040); +x_1043 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_1044 = l_Lean_Syntax_node2(x_988, x_1043, x_1024, x_1042); +lean_inc(x_1044); +x_1045 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_1045, 0, x_1044); +x_1046 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_9, x_1044, x_1045, x_15, x_16, x_993); +return x_1046; } } } @@ -10797,7 +10864,7 @@ return x_1032; } else { -uint8_t x_1089; +uint8_t x_1103; lean_dec(x_16); lean_dec(x_15); lean_dec(x_12); @@ -10806,23 +10873,23 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); lean_dec(x_2); -x_1089 = !lean_is_exclusive(x_21); -if (x_1089 == 0) +x_1103 = !lean_is_exclusive(x_21); +if (x_1103 == 0) { return x_21; } else { -lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; -x_1090 = lean_ctor_get(x_21, 0); -x_1091 = lean_ctor_get(x_21, 1); -lean_inc(x_1091); -lean_inc(x_1090); +lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; +x_1104 = lean_ctor_get(x_21, 0); +x_1105 = lean_ctor_get(x_21, 1); +lean_inc(x_1105); +lean_inc(x_1104); lean_dec(x_21); -x_1092 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1092, 0, x_1090); -lean_ctor_set(x_1092, 1, x_1091); -return x_1092; +x_1106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1106, 0, x_1104); +lean_ctor_set(x_1106, 1, x_1105); +return x_1106; } } } @@ -11885,60 +11952,60 @@ x_30 = l_Lean_mkCIdentFrom(x_27, x_2, x_29); lean_dec(x_27); if (lean_obj_tag(x_8) == 0) { -lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; -x_612 = lean_box(2); -x_613 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_614 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_614, 0, x_612); -lean_ctor_set(x_614, 1, x_613); -lean_ctor_set(x_614, 2, x_1); +lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; +x_622 = lean_box(2); +x_623 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_624 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_624, 0, x_622); +lean_ctor_set(x_624, 1, x_623); +lean_ctor_set(x_624, 2, x_1); lean_ctor_set_tag(x_25, 1); lean_ctor_set(x_25, 1, x_3); -lean_ctor_set(x_25, 0, x_614); +lean_ctor_set(x_25, 0, x_624); lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 1, x_25); lean_ctor_set(x_21, 0, x_9); -x_615 = lean_array_mk(x_21); -x_616 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_617 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_617, 0, x_612); -lean_ctor_set(x_617, 1, x_616); -lean_ctor_set(x_617, 2, x_615); -x_31 = x_617; +x_625 = lean_array_mk(x_21); +x_626 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_627 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_627, 0, x_622); +lean_ctor_set(x_627, 1, x_626); +lean_ctor_set(x_627, 2, x_625); +x_31 = x_627; x_32 = x_28; -goto block_611; +goto block_621; } else { -lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; -x_618 = lean_ctor_get(x_8, 0); -lean_inc(x_618); +lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; +x_628 = lean_ctor_get(x_8, 0); +lean_inc(x_628); lean_dec(x_8); -x_619 = l_Lean_mkIdentFrom(x_9, x_618, x_29); +x_629 = l_Lean_mkIdentFrom(x_9, x_628, x_29); lean_dec(x_9); -x_620 = lean_box(2); -x_621 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_622 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_622, 0, x_620); -lean_ctor_set(x_622, 1, x_621); -lean_ctor_set(x_622, 2, x_1); +x_630 = lean_box(2); +x_631 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_632 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_632, 0, x_630); +lean_ctor_set(x_632, 1, x_631); +lean_ctor_set(x_632, 2, x_1); lean_ctor_set_tag(x_25, 1); lean_ctor_set(x_25, 1, x_3); -lean_ctor_set(x_25, 0, x_622); +lean_ctor_set(x_25, 0, x_632); lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 1, x_25); -lean_ctor_set(x_21, 0, x_619); -x_623 = lean_array_mk(x_21); -x_624 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_625 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_625, 0, x_620); -lean_ctor_set(x_625, 1, x_624); -lean_ctor_set(x_625, 2, x_623); -x_31 = x_625; +lean_ctor_set(x_21, 0, x_629); +x_633 = lean_array_mk(x_21); +x_634 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_635 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_635, 0, x_630); +lean_ctor_set(x_635, 1, x_634); +lean_ctor_set(x_635, 2, x_633); +x_31 = x_635; x_32 = x_28; -goto block_611; +goto block_621; } -block_611: +block_621: { lean_object* x_33; uint8_t x_34; x_33 = l_Lean_Elab_Command_getRef(x_12, x_13, x_32); @@ -12003,7 +12070,7 @@ x_58 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_56); x_59 = !lean_is_exclusive(x_58); if (x_59 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; x_60 = lean_ctor_get(x_58, 1); x_61 = lean_ctor_get(x_58, 0); lean_dec(x_61); @@ -12021,7 +12088,7 @@ lean_ctor_set(x_65, 1, x_63); lean_ctor_set(x_65, 2, x_64); x_66 = lean_array_size(x_23); x_67 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_66, x_17, x_23); -x_68 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +x_68 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; x_69 = l_Lean_mkSepArray(x_67, x_68); lean_dec(x_67); x_70 = l_Array_append___rarg(x_64, x_69); @@ -12031,2320 +12098,2341 @@ x_71 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_71, 0, x_53); lean_ctor_set(x_71, 1, x_63); lean_ctor_set(x_71, 2, x_70); -x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_72 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_53); +x_73 = l_Lean_Syntax_node1(x_53, x_72, x_71); +x_74 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; lean_inc(x_65); lean_inc(x_53); -x_73 = l_Lean_Syntax_node1(x_53, x_72, x_65); -x_74 = l_Lake_DSL_structVal___closed__18; +x_75 = l_Lean_Syntax_node1(x_53, x_74, x_65); +x_76 = l_Lake_DSL_structVal___closed__18; lean_inc(x_53); lean_ctor_set_tag(x_54, 2); -lean_ctor_set(x_54, 1, x_74); +lean_ctor_set(x_54, 1, x_76); lean_ctor_set(x_54, 0, x_53); -x_75 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +x_77 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; lean_inc(x_65); -x_76 = l_Lean_Syntax_node6(x_53, x_75, x_58, x_65, x_71, x_73, x_65, x_54); -x_77 = l_Lean_Elab_Command_getRef(x_12, x_13, x_60); -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_79 = lean_ctor_get(x_77, 0); -x_80 = lean_ctor_get(x_77, 1); -x_81 = l_Lean_SourceInfo_fromRef(x_79, x_29); -lean_dec(x_79); -x_82 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_80); -x_83 = !lean_is_exclusive(x_82); -if (x_83 == 0) +x_78 = l_Lean_Syntax_node6(x_53, x_77, x_58, x_65, x_73, x_75, x_65, x_54); +x_79 = l_Lean_Elab_Command_getRef(x_12, x_13, x_60); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_84 = lean_ctor_get(x_82, 1); -x_85 = lean_ctor_get(x_82, 0); -lean_dec(x_85); -x_86 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_84); -x_87 = !lean_is_exclusive(x_86); -if (x_87 == 0) +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_81 = lean_ctor_get(x_79, 0); +x_82 = lean_ctor_get(x_79, 1); +x_83 = l_Lean_SourceInfo_fromRef(x_81, x_29); +lean_dec(x_81); +x_84 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_82); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_88 = lean_ctor_get(x_86, 1); -x_89 = lean_ctor_get(x_86, 0); -lean_dec(x_89); -x_90 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -lean_ctor_set_tag(x_86, 2); -lean_ctor_set(x_86, 1, x_90); -lean_ctor_set(x_86, 0, x_81); -x_91 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_92 = l_Lean_Syntax_SepArray_ofElems(x_91, x_5); -x_93 = l_Array_append___rarg(x_64, x_92); -lean_dec(x_92); -lean_inc(x_81); -x_94 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_94, 0, x_81); -lean_ctor_set(x_94, 1, x_63); -lean_ctor_set(x_94, 2, x_93); -x_95 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -lean_ctor_set_tag(x_82, 2); -lean_ctor_set(x_82, 1, x_95); -lean_ctor_set(x_82, 0, x_81); -x_96 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_97 = l_Lean_Syntax_node3(x_81, x_96, x_86, x_94, x_82); -lean_inc(x_81); -x_98 = l_Lean_Syntax_node1(x_81, x_63, x_97); -lean_inc(x_81); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_81); -lean_ctor_set(x_99, 1, x_63); -lean_ctor_set(x_99, 2, x_64); -x_100 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_100); -lean_ctor_set(x_77, 0, x_81); -x_101 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_86 = lean_ctor_get(x_84, 1); +x_87 = lean_ctor_get(x_84, 0); +lean_dec(x_87); +x_88 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_86); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_90 = lean_ctor_get(x_88, 1); +x_91 = lean_ctor_get(x_88, 0); +lean_dec(x_91); +x_92 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +lean_ctor_set_tag(x_88, 2); +lean_ctor_set(x_88, 1, x_92); +lean_ctor_set(x_88, 0, x_83); +x_93 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_94 = l_Lean_Syntax_SepArray_ofElems(x_93, x_5); +x_95 = l_Array_append___rarg(x_64, x_94); +lean_dec(x_94); +lean_inc(x_83); +x_96 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_96, 0, x_83); +lean_ctor_set(x_96, 1, x_63); +lean_ctor_set(x_96, 2, x_95); +x_97 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +lean_ctor_set_tag(x_84, 2); +lean_ctor_set(x_84, 1, x_97); +lean_ctor_set(x_84, 0, x_83); +x_98 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_99 = l_Lean_Syntax_node3(x_83, x_98, x_88, x_96, x_84); +lean_inc(x_83); +x_100 = l_Lean_Syntax_node1(x_83, x_63, x_99); +lean_inc(x_83); +x_101 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_101, 0, x_83); +lean_ctor_set(x_101, 1, x_63); +lean_ctor_set(x_101, 2, x_64); +x_102 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_102); +lean_ctor_set(x_79, 0, x_83); +x_103 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_101); -lean_ctor_set(x_49, 0, x_81); -x_102 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_103 = l_Lean_Syntax_node2(x_81, x_102, x_49, x_30); -lean_inc(x_81); -x_104 = l_Lean_Syntax_node1(x_81, x_63, x_103); -x_105 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_99); -lean_inc(x_81); -x_106 = l_Lean_Syntax_node2(x_81, x_105, x_99, x_104); -x_107 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_103); +lean_ctor_set(x_49, 0, x_83); +x_104 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_105 = l_Lean_Syntax_node2(x_83, x_104, x_49, x_30); +lean_inc(x_83); +x_106 = l_Lean_Syntax_node1(x_83, x_63, x_105); +x_107 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_101); +lean_inc(x_83); +x_108 = l_Lean_Syntax_node2(x_83, x_107, x_101, x_106); +x_109 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_107); -lean_ctor_set(x_33, 0, x_81); -x_108 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_99, 2); -lean_inc(x_81); -x_109 = l_Lean_Syntax_node2(x_81, x_108, x_99, x_99); -x_110 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_81); -lean_ctor_set(x_111, 1, x_63); -lean_ctor_set(x_111, 2, x_110); -x_112 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_111); -lean_inc(x_81); -x_113 = l_Lean_Syntax_node4(x_81, x_112, x_33, x_76, x_109, x_111); -x_114 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_115 = l_Lean_Syntax_node4(x_81, x_114, x_77, x_31, x_106, x_113); +lean_ctor_set(x_33, 1, x_109); +lean_ctor_set(x_33, 0, x_83); +x_110 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_101, 2); +lean_inc(x_83); +x_111 = l_Lean_Syntax_node2(x_83, x_110, x_101, x_101); +x_112 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_113 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_113, 0, x_83); +lean_ctor_set(x_113, 1, x_63); +lean_ctor_set(x_113, 2, x_112); +x_114 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_113); +lean_inc(x_83); +x_115 = l_Lean_Syntax_node4(x_83, x_114, x_33, x_78, x_111, x_113); +x_116 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_117 = l_Lean_Syntax_node4(x_83, x_116, x_79, x_31, x_108, x_115); if (lean_obj_tag(x_6) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_116 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_99, 3); -lean_inc(x_81); -x_117 = l_Lean_Syntax_node6(x_81, x_116, x_111, x_98, x_99, x_99, x_99, x_99); -x_118 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_119 = l_Lean_Syntax_node2(x_81, x_118, x_117, x_115); -lean_inc(x_119); -x_120 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_120, 0, x_119); -x_121 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_119, x_120, x_12, x_13, x_88); -return x_121; -} -else -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -lean_dec(x_111); -x_122 = lean_ctor_get(x_6, 0); -lean_inc(x_122); +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_118 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_101, 3); +lean_inc(x_83); +x_119 = l_Lean_Syntax_node6(x_83, x_118, x_113, x_100, x_101, x_101, x_101, x_101); +x_120 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_121 = l_Lean_Syntax_node2(x_83, x_120, x_119, x_117); +lean_inc(x_121); +x_122 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_122, 0, x_121); +x_123 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_121, x_122, x_12, x_13, x_90); +return x_123; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +lean_dec(x_113); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_124); lean_dec(x_6); -x_123 = l_Array_mkArray1___rarg(x_122); -x_124 = l_Array_append___rarg(x_64, x_123); -lean_dec(x_123); -lean_inc(x_81); -x_125 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_125, 0, x_81); -lean_ctor_set(x_125, 1, x_63); -lean_ctor_set(x_125, 2, x_124); -x_126 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_99, 3); -lean_inc(x_81); -x_127 = l_Lean_Syntax_node6(x_81, x_126, x_125, x_98, x_99, x_99, x_99, x_99); -x_128 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_129 = l_Lean_Syntax_node2(x_81, x_128, x_127, x_115); -lean_inc(x_129); -x_130 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_130, 0, x_129); -x_131 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_129, x_130, x_12, x_13, x_88); -return x_131; -} -} -else -{ -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_132 = lean_ctor_get(x_86, 1); -lean_inc(x_132); -lean_dec(x_86); -x_133 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -x_134 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_134, 0, x_81); -lean_ctor_set(x_134, 1, x_133); -x_135 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_136 = l_Lean_Syntax_SepArray_ofElems(x_135, x_5); -x_137 = l_Array_append___rarg(x_64, x_136); -lean_dec(x_136); -lean_inc(x_81); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_81); -lean_ctor_set(x_138, 1, x_63); -lean_ctor_set(x_138, 2, x_137); -x_139 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -lean_ctor_set_tag(x_82, 2); -lean_ctor_set(x_82, 1, x_139); -lean_ctor_set(x_82, 0, x_81); -x_140 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_141 = l_Lean_Syntax_node3(x_81, x_140, x_134, x_138, x_82); -lean_inc(x_81); -x_142 = l_Lean_Syntax_node1(x_81, x_63, x_141); -lean_inc(x_81); -x_143 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_143, 0, x_81); -lean_ctor_set(x_143, 1, x_63); -lean_ctor_set(x_143, 2, x_64); -x_144 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_144); -lean_ctor_set(x_77, 0, x_81); -x_145 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); +x_125 = l_Array_mkArray1___rarg(x_124); +x_126 = l_Array_append___rarg(x_64, x_125); +lean_dec(x_125); +lean_inc(x_83); +x_127 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_127, 0, x_83); +lean_ctor_set(x_127, 1, x_63); +lean_ctor_set(x_127, 2, x_126); +x_128 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_101, 3); +lean_inc(x_83); +x_129 = l_Lean_Syntax_node6(x_83, x_128, x_127, x_100, x_101, x_101, x_101, x_101); +x_130 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_131 = l_Lean_Syntax_node2(x_83, x_130, x_129, x_117); +lean_inc(x_131); +x_132 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_132, 0, x_131); +x_133 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_131, x_132, x_12, x_13, x_90); +return x_133; +} +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_134 = lean_ctor_get(x_88, 1); +lean_inc(x_134); +lean_dec(x_88); +x_135 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +x_136 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_136, 0, x_83); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_138 = l_Lean_Syntax_SepArray_ofElems(x_137, x_5); +x_139 = l_Array_append___rarg(x_64, x_138); +lean_dec(x_138); +lean_inc(x_83); +x_140 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_140, 0, x_83); +lean_ctor_set(x_140, 1, x_63); +lean_ctor_set(x_140, 2, x_139); +x_141 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +lean_ctor_set_tag(x_84, 2); +lean_ctor_set(x_84, 1, x_141); +lean_ctor_set(x_84, 0, x_83); +x_142 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_143 = l_Lean_Syntax_node3(x_83, x_142, x_136, x_140, x_84); +lean_inc(x_83); +x_144 = l_Lean_Syntax_node1(x_83, x_63, x_143); +lean_inc(x_83); +x_145 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_145, 0, x_83); +lean_ctor_set(x_145, 1, x_63); +lean_ctor_set(x_145, 2, x_64); +x_146 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_146); +lean_ctor_set(x_79, 0, x_83); +x_147 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_145); -lean_ctor_set(x_49, 0, x_81); -x_146 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_147 = l_Lean_Syntax_node2(x_81, x_146, x_49, x_30); -lean_inc(x_81); -x_148 = l_Lean_Syntax_node1(x_81, x_63, x_147); -x_149 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_143); -lean_inc(x_81); -x_150 = l_Lean_Syntax_node2(x_81, x_149, x_143, x_148); -x_151 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_147); +lean_ctor_set(x_49, 0, x_83); +x_148 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_149 = l_Lean_Syntax_node2(x_83, x_148, x_49, x_30); +lean_inc(x_83); +x_150 = l_Lean_Syntax_node1(x_83, x_63, x_149); +x_151 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_145); +lean_inc(x_83); +x_152 = l_Lean_Syntax_node2(x_83, x_151, x_145, x_150); +x_153 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_151); -lean_ctor_set(x_33, 0, x_81); -x_152 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_143, 2); -lean_inc(x_81); -x_153 = l_Lean_Syntax_node2(x_81, x_152, x_143, x_143); -x_154 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_155 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_155, 0, x_81); -lean_ctor_set(x_155, 1, x_63); -lean_ctor_set(x_155, 2, x_154); -x_156 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_155); -lean_inc(x_81); -x_157 = l_Lean_Syntax_node4(x_81, x_156, x_33, x_76, x_153, x_155); -x_158 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_159 = l_Lean_Syntax_node4(x_81, x_158, x_77, x_31, x_150, x_157); +lean_ctor_set(x_33, 1, x_153); +lean_ctor_set(x_33, 0, x_83); +x_154 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_145, 2); +lean_inc(x_83); +x_155 = l_Lean_Syntax_node2(x_83, x_154, x_145, x_145); +x_156 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_83); +lean_ctor_set(x_157, 1, x_63); +lean_ctor_set(x_157, 2, x_156); +x_158 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_157); +lean_inc(x_83); +x_159 = l_Lean_Syntax_node4(x_83, x_158, x_33, x_78, x_155, x_157); +x_160 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_161 = l_Lean_Syntax_node4(x_83, x_160, x_79, x_31, x_152, x_159); if (lean_obj_tag(x_6) == 0) { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_160 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_143, 3); -lean_inc(x_81); -x_161 = l_Lean_Syntax_node6(x_81, x_160, x_155, x_142, x_143, x_143, x_143, x_143); -x_162 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_163 = l_Lean_Syntax_node2(x_81, x_162, x_161, x_159); -lean_inc(x_163); -x_164 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_164, 0, x_163); -x_165 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_163, x_164, x_12, x_13, x_132); -return x_165; -} -else -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -lean_dec(x_155); -x_166 = lean_ctor_get(x_6, 0); -lean_inc(x_166); +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_162 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_145, 3); +lean_inc(x_83); +x_163 = l_Lean_Syntax_node6(x_83, x_162, x_157, x_144, x_145, x_145, x_145, x_145); +x_164 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_165 = l_Lean_Syntax_node2(x_83, x_164, x_163, x_161); +lean_inc(x_165); +x_166 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_166, 0, x_165); +x_167 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_165, x_166, x_12, x_13, x_134); +return x_167; +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_157); +x_168 = lean_ctor_get(x_6, 0); +lean_inc(x_168); lean_dec(x_6); -x_167 = l_Array_mkArray1___rarg(x_166); -x_168 = l_Array_append___rarg(x_64, x_167); -lean_dec(x_167); -lean_inc(x_81); -x_169 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_169, 0, x_81); -lean_ctor_set(x_169, 1, x_63); -lean_ctor_set(x_169, 2, x_168); -x_170 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_143, 3); -lean_inc(x_81); -x_171 = l_Lean_Syntax_node6(x_81, x_170, x_169, x_142, x_143, x_143, x_143, x_143); -x_172 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_173 = l_Lean_Syntax_node2(x_81, x_172, x_171, x_159); -lean_inc(x_173); -x_174 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_174, 0, x_173); -x_175 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_173, x_174, x_12, x_13, x_132); -return x_175; -} -} -} -else -{ -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_176 = lean_ctor_get(x_82, 1); -lean_inc(x_176); -lean_dec(x_82); -x_177 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_176); -x_178 = lean_ctor_get(x_177, 1); +x_169 = l_Array_mkArray1___rarg(x_168); +x_170 = l_Array_append___rarg(x_64, x_169); +lean_dec(x_169); +lean_inc(x_83); +x_171 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_171, 0, x_83); +lean_ctor_set(x_171, 1, x_63); +lean_ctor_set(x_171, 2, x_170); +x_172 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_145, 3); +lean_inc(x_83); +x_173 = l_Lean_Syntax_node6(x_83, x_172, x_171, x_144, x_145, x_145, x_145, x_145); +x_174 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_175 = l_Lean_Syntax_node2(x_83, x_174, x_173, x_161); +lean_inc(x_175); +x_176 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_176, 0, x_175); +x_177 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_175, x_176, x_12, x_13, x_134); +return x_177; +} +} +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_178 = lean_ctor_get(x_84, 1); lean_inc(x_178); -if (lean_is_exclusive(x_177)) { - lean_ctor_release(x_177, 0); - lean_ctor_release(x_177, 1); - x_179 = x_177; -} else { - lean_dec_ref(x_177); - x_179 = lean_box(0); -} -x_180 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_81); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(2, 2, 0); -} else { +lean_dec(x_84); +x_179 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_178); +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); x_181 = x_179; - lean_ctor_set_tag(x_181, 2); -} -lean_ctor_set(x_181, 0, x_81); -lean_ctor_set(x_181, 1, x_180); -x_182 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_183 = l_Lean_Syntax_SepArray_ofElems(x_182, x_5); -x_184 = l_Array_append___rarg(x_64, x_183); -lean_dec(x_183); -lean_inc(x_81); -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_81); -lean_ctor_set(x_185, 1, x_63); -lean_ctor_set(x_185, 2, x_184); -x_186 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_81); -x_187 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_187, 0, x_81); -lean_ctor_set(x_187, 1, x_186); -x_188 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_81); -x_189 = l_Lean_Syntax_node3(x_81, x_188, x_181, x_185, x_187); -lean_inc(x_81); -x_190 = l_Lean_Syntax_node1(x_81, x_63, x_189); -lean_inc(x_81); -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_81); -lean_ctor_set(x_191, 1, x_63); -lean_ctor_set(x_191, 2, x_64); -x_192 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_81); -lean_ctor_set_tag(x_77, 2); -lean_ctor_set(x_77, 1, x_192); -lean_ctor_set(x_77, 0, x_81); -x_193 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_81); +} else { + lean_dec_ref(x_179); + x_181 = lean_box(0); +} +x_182 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_83); +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(2, 2, 0); +} else { + x_183 = x_181; + lean_ctor_set_tag(x_183, 2); +} +lean_ctor_set(x_183, 0, x_83); +lean_ctor_set(x_183, 1, x_182); +x_184 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_185 = l_Lean_Syntax_SepArray_ofElems(x_184, x_5); +x_186 = l_Array_append___rarg(x_64, x_185); +lean_dec(x_185); +lean_inc(x_83); +x_187 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_187, 0, x_83); +lean_ctor_set(x_187, 1, x_63); +lean_ctor_set(x_187, 2, x_186); +x_188 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_83); +x_189 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_189, 0, x_83); +lean_ctor_set(x_189, 1, x_188); +x_190 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_83); +x_191 = l_Lean_Syntax_node3(x_83, x_190, x_183, x_187, x_189); +lean_inc(x_83); +x_192 = l_Lean_Syntax_node1(x_83, x_63, x_191); +lean_inc(x_83); +x_193 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_193, 0, x_83); +lean_ctor_set(x_193, 1, x_63); +lean_ctor_set(x_193, 2, x_64); +x_194 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_83); +lean_ctor_set_tag(x_79, 2); +lean_ctor_set(x_79, 1, x_194); +lean_ctor_set(x_79, 0, x_83); +x_195 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_193); -lean_ctor_set(x_49, 0, x_81); -x_194 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_81); -x_195 = l_Lean_Syntax_node2(x_81, x_194, x_49, x_30); -lean_inc(x_81); -x_196 = l_Lean_Syntax_node1(x_81, x_63, x_195); -x_197 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_191); -lean_inc(x_81); -x_198 = l_Lean_Syntax_node2(x_81, x_197, x_191, x_196); -x_199 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_81); +lean_ctor_set(x_49, 1, x_195); +lean_ctor_set(x_49, 0, x_83); +x_196 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_83); +x_197 = l_Lean_Syntax_node2(x_83, x_196, x_49, x_30); +lean_inc(x_83); +x_198 = l_Lean_Syntax_node1(x_83, x_63, x_197); +x_199 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_193); +lean_inc(x_83); +x_200 = l_Lean_Syntax_node2(x_83, x_199, x_193, x_198); +x_201 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_83); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_199); -lean_ctor_set(x_33, 0, x_81); -x_200 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_191, 2); -lean_inc(x_81); -x_201 = l_Lean_Syntax_node2(x_81, x_200, x_191, x_191); -x_202 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_81); -x_203 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_203, 0, x_81); -lean_ctor_set(x_203, 1, x_63); -lean_ctor_set(x_203, 2, x_202); -x_204 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_203); -lean_inc(x_81); -x_205 = l_Lean_Syntax_node4(x_81, x_204, x_33, x_76, x_201, x_203); -x_206 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_81); -x_207 = l_Lean_Syntax_node4(x_81, x_206, x_77, x_31, x_198, x_205); +lean_ctor_set(x_33, 1, x_201); +lean_ctor_set(x_33, 0, x_83); +x_202 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_193, 2); +lean_inc(x_83); +x_203 = l_Lean_Syntax_node2(x_83, x_202, x_193, x_193); +x_204 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_83); +x_205 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_205, 0, x_83); +lean_ctor_set(x_205, 1, x_63); +lean_ctor_set(x_205, 2, x_204); +x_206 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_205); +lean_inc(x_83); +x_207 = l_Lean_Syntax_node4(x_83, x_206, x_33, x_78, x_203, x_205); +x_208 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_83); +x_209 = l_Lean_Syntax_node4(x_83, x_208, x_79, x_31, x_200, x_207); if (lean_obj_tag(x_6) == 0) { -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_208 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_191, 3); -lean_inc(x_81); -x_209 = l_Lean_Syntax_node6(x_81, x_208, x_203, x_190, x_191, x_191, x_191, x_191); -x_210 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_211 = l_Lean_Syntax_node2(x_81, x_210, x_209, x_207); -lean_inc(x_211); -x_212 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_212, 0, x_211); -x_213 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_211, x_212, x_12, x_13, x_178); -return x_213; -} -else -{ -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -lean_dec(x_203); -x_214 = lean_ctor_get(x_6, 0); -lean_inc(x_214); +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_210 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_193, 3); +lean_inc(x_83); +x_211 = l_Lean_Syntax_node6(x_83, x_210, x_205, x_192, x_193, x_193, x_193, x_193); +x_212 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_213 = l_Lean_Syntax_node2(x_83, x_212, x_211, x_209); +lean_inc(x_213); +x_214 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_214, 0, x_213); +x_215 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_213, x_214, x_12, x_13, x_180); +return x_215; +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec(x_205); +x_216 = lean_ctor_get(x_6, 0); +lean_inc(x_216); lean_dec(x_6); -x_215 = l_Array_mkArray1___rarg(x_214); -x_216 = l_Array_append___rarg(x_64, x_215); -lean_dec(x_215); -lean_inc(x_81); -x_217 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_217, 0, x_81); -lean_ctor_set(x_217, 1, x_63); -lean_ctor_set(x_217, 2, x_216); -x_218 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_191, 3); -lean_inc(x_81); -x_219 = l_Lean_Syntax_node6(x_81, x_218, x_217, x_190, x_191, x_191, x_191, x_191); -x_220 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_221 = l_Lean_Syntax_node2(x_81, x_220, x_219, x_207); -lean_inc(x_221); -x_222 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_222, 0, x_221); -x_223 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_221, x_222, x_12, x_13, x_178); -return x_223; -} +x_217 = l_Array_mkArray1___rarg(x_216); +x_218 = l_Array_append___rarg(x_64, x_217); +lean_dec(x_217); +lean_inc(x_83); +x_219 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_219, 0, x_83); +lean_ctor_set(x_219, 1, x_63); +lean_ctor_set(x_219, 2, x_218); +x_220 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_193, 3); +lean_inc(x_83); +x_221 = l_Lean_Syntax_node6(x_83, x_220, x_219, x_192, x_193, x_193, x_193, x_193); +x_222 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_223 = l_Lean_Syntax_node2(x_83, x_222, x_221, x_209); +lean_inc(x_223); +x_224 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_224, 0, x_223); +x_225 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_223, x_224, x_12, x_13, x_180); +return x_225; +} +} +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_226 = lean_ctor_get(x_79, 0); +x_227 = lean_ctor_get(x_79, 1); +lean_inc(x_227); +lean_inc(x_226); +lean_dec(x_79); +x_228 = l_Lean_SourceInfo_fromRef(x_226, x_29); +lean_dec(x_226); +x_229 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_227); +x_230 = lean_ctor_get(x_229, 1); +lean_inc(x_230); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + lean_ctor_release(x_229, 1); + x_231 = x_229; +} else { + lean_dec_ref(x_229); + x_231 = lean_box(0); } +x_232 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_230); +x_233 = lean_ctor_get(x_232, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_232)) { + lean_ctor_release(x_232, 0); + lean_ctor_release(x_232, 1); + x_234 = x_232; +} else { + lean_dec_ref(x_232); + x_234 = lean_box(0); } -else -{ -lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_224 = lean_ctor_get(x_77, 0); -x_225 = lean_ctor_get(x_77, 1); -lean_inc(x_225); -lean_inc(x_224); -lean_dec(x_77); -x_226 = l_Lean_SourceInfo_fromRef(x_224, x_29); -lean_dec(x_224); -x_227 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_225); -x_228 = lean_ctor_get(x_227, 1); +x_235 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; lean_inc(x_228); -if (lean_is_exclusive(x_227)) { - lean_ctor_release(x_227, 0); - lean_ctor_release(x_227, 1); - x_229 = x_227; -} else { - lean_dec_ref(x_227); - x_229 = lean_box(0); -} -x_230 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_228); -x_231 = lean_ctor_get(x_230, 1); -lean_inc(x_231); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_232 = x_230; -} else { - lean_dec_ref(x_230); - x_232 = lean_box(0); -} -x_233 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_226); -if (lean_is_scalar(x_232)) { - x_234 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_234)) { + x_236 = lean_alloc_ctor(2, 2, 0); } else { - x_234 = x_232; - lean_ctor_set_tag(x_234, 2); + x_236 = x_234; + lean_ctor_set_tag(x_236, 2); } -lean_ctor_set(x_234, 0, x_226); -lean_ctor_set(x_234, 1, x_233); -x_235 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_236 = l_Lean_Syntax_SepArray_ofElems(x_235, x_5); -x_237 = l_Array_append___rarg(x_64, x_236); -lean_dec(x_236); -lean_inc(x_226); -x_238 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_238, 0, x_226); -lean_ctor_set(x_238, 1, x_63); -lean_ctor_set(x_238, 2, x_237); -x_239 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_226); -if (lean_is_scalar(x_229)) { - x_240 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_236, 0, x_228); +lean_ctor_set(x_236, 1, x_235); +x_237 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_238 = l_Lean_Syntax_SepArray_ofElems(x_237, x_5); +x_239 = l_Array_append___rarg(x_64, x_238); +lean_dec(x_238); +lean_inc(x_228); +x_240 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_240, 0, x_228); +lean_ctor_set(x_240, 1, x_63); +lean_ctor_set(x_240, 2, x_239); +x_241 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_228); +if (lean_is_scalar(x_231)) { + x_242 = lean_alloc_ctor(2, 2, 0); } else { - x_240 = x_229; - lean_ctor_set_tag(x_240, 2); + x_242 = x_231; + lean_ctor_set_tag(x_242, 2); } -lean_ctor_set(x_240, 0, x_226); -lean_ctor_set(x_240, 1, x_239); -x_241 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_226); -x_242 = l_Lean_Syntax_node3(x_226, x_241, x_234, x_238, x_240); -lean_inc(x_226); -x_243 = l_Lean_Syntax_node1(x_226, x_63, x_242); -lean_inc(x_226); -x_244 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_244, 0, x_226); -lean_ctor_set(x_244, 1, x_63); -lean_ctor_set(x_244, 2, x_64); -x_245 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_226); -x_246 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_246, 0, x_226); -lean_ctor_set(x_246, 1, x_245); -x_247 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_226); +lean_ctor_set(x_242, 0, x_228); +lean_ctor_set(x_242, 1, x_241); +x_243 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_228); +x_244 = l_Lean_Syntax_node3(x_228, x_243, x_236, x_240, x_242); +lean_inc(x_228); +x_245 = l_Lean_Syntax_node1(x_228, x_63, x_244); +lean_inc(x_228); +x_246 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_246, 0, x_228); +lean_ctor_set(x_246, 1, x_63); +lean_ctor_set(x_246, 2, x_64); +x_247 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_228); +x_248 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_248, 0, x_228); +lean_ctor_set(x_248, 1, x_247); +x_249 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_228); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_247); -lean_ctor_set(x_49, 0, x_226); -x_248 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_226); -x_249 = l_Lean_Syntax_node2(x_226, x_248, x_49, x_30); -lean_inc(x_226); -x_250 = l_Lean_Syntax_node1(x_226, x_63, x_249); -x_251 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_244); -lean_inc(x_226); -x_252 = l_Lean_Syntax_node2(x_226, x_251, x_244, x_250); -x_253 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_226); +lean_ctor_set(x_49, 1, x_249); +lean_ctor_set(x_49, 0, x_228); +x_250 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_228); +x_251 = l_Lean_Syntax_node2(x_228, x_250, x_49, x_30); +lean_inc(x_228); +x_252 = l_Lean_Syntax_node1(x_228, x_63, x_251); +x_253 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_246); +lean_inc(x_228); +x_254 = l_Lean_Syntax_node2(x_228, x_253, x_246, x_252); +x_255 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_228); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_253); -lean_ctor_set(x_33, 0, x_226); -x_254 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_244, 2); -lean_inc(x_226); -x_255 = l_Lean_Syntax_node2(x_226, x_254, x_244, x_244); -x_256 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_226); -x_257 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_257, 0, x_226); -lean_ctor_set(x_257, 1, x_63); -lean_ctor_set(x_257, 2, x_256); -x_258 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_257); -lean_inc(x_226); -x_259 = l_Lean_Syntax_node4(x_226, x_258, x_33, x_76, x_255, x_257); -x_260 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_226); -x_261 = l_Lean_Syntax_node4(x_226, x_260, x_246, x_31, x_252, x_259); +lean_ctor_set(x_33, 1, x_255); +lean_ctor_set(x_33, 0, x_228); +x_256 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_246, 2); +lean_inc(x_228); +x_257 = l_Lean_Syntax_node2(x_228, x_256, x_246, x_246); +x_258 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_228); +x_259 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_259, 0, x_228); +lean_ctor_set(x_259, 1, x_63); +lean_ctor_set(x_259, 2, x_258); +x_260 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_259); +lean_inc(x_228); +x_261 = l_Lean_Syntax_node4(x_228, x_260, x_33, x_78, x_257, x_259); +x_262 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_228); +x_263 = l_Lean_Syntax_node4(x_228, x_262, x_248, x_31, x_254, x_261); if (lean_obj_tag(x_6) == 0) { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; -x_262 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_244, 3); -lean_inc(x_226); -x_263 = l_Lean_Syntax_node6(x_226, x_262, x_257, x_243, x_244, x_244, x_244, x_244); -x_264 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_265 = l_Lean_Syntax_node2(x_226, x_264, x_263, x_261); -lean_inc(x_265); -x_266 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_266, 0, x_265); -x_267 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_265, x_266, x_12, x_13, x_231); -return x_267; +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_264 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_246, 3); +lean_inc(x_228); +x_265 = l_Lean_Syntax_node6(x_228, x_264, x_259, x_245, x_246, x_246, x_246, x_246); +x_266 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_267 = l_Lean_Syntax_node2(x_228, x_266, x_265, x_263); +lean_inc(x_267); +x_268 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_268, 0, x_267); +x_269 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_267, x_268, x_12, x_13, x_233); +return x_269; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -lean_dec(x_257); -x_268 = lean_ctor_get(x_6, 0); -lean_inc(x_268); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +lean_dec(x_259); +x_270 = lean_ctor_get(x_6, 0); +lean_inc(x_270); lean_dec(x_6); -x_269 = l_Array_mkArray1___rarg(x_268); -x_270 = l_Array_append___rarg(x_64, x_269); -lean_dec(x_269); -lean_inc(x_226); -x_271 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_271, 0, x_226); -lean_ctor_set(x_271, 1, x_63); -lean_ctor_set(x_271, 2, x_270); -x_272 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_244, 3); -lean_inc(x_226); -x_273 = l_Lean_Syntax_node6(x_226, x_272, x_271, x_243, x_244, x_244, x_244, x_244); -x_274 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_275 = l_Lean_Syntax_node2(x_226, x_274, x_273, x_261); -lean_inc(x_275); -x_276 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_276, 0, x_275); -x_277 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_275, x_276, x_12, x_13, x_231); -return x_277; +x_271 = l_Array_mkArray1___rarg(x_270); +x_272 = l_Array_append___rarg(x_64, x_271); +lean_dec(x_271); +lean_inc(x_228); +x_273 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_273, 0, x_228); +lean_ctor_set(x_273, 1, x_63); +lean_ctor_set(x_273, 2, x_272); +x_274 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_246, 3); +lean_inc(x_228); +x_275 = l_Lean_Syntax_node6(x_228, x_274, x_273, x_245, x_246, x_246, x_246, x_246); +x_276 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_277 = l_Lean_Syntax_node2(x_228, x_276, x_275, x_263); +lean_inc(x_277); +x_278 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_278, 0, x_277); +x_279 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_277, x_278, x_12, x_13, x_233); +return x_279; } } } else { -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; size_t x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_278 = lean_ctor_get(x_58, 1); -lean_inc(x_278); +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; size_t x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; +x_280 = lean_ctor_get(x_58, 1); +lean_inc(x_280); lean_dec(x_58); -x_279 = l_Lake_DSL_structVal___closed__3; +x_281 = l_Lake_DSL_structVal___closed__3; lean_inc(x_53); -x_280 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_280, 0, x_53); -lean_ctor_set(x_280, 1, x_279); -x_281 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_282 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +x_282 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_282, 0, x_53); +lean_ctor_set(x_282, 1, x_281); +x_283 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_284 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_53); -x_283 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_283, 0, x_53); -lean_ctor_set(x_283, 1, x_281); -lean_ctor_set(x_283, 2, x_282); -x_284 = lean_array_size(x_23); -x_285 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_284, x_17, x_23); -x_286 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_287 = l_Lean_mkSepArray(x_285, x_286); -lean_dec(x_285); -x_288 = l_Array_append___rarg(x_282, x_287); +x_285 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_285, 0, x_53); +lean_ctor_set(x_285, 1, x_283); +lean_ctor_set(x_285, 2, x_284); +x_286 = lean_array_size(x_23); +x_287 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_286, x_17, x_23); +x_288 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_289 = l_Lean_mkSepArray(x_287, x_288); lean_dec(x_287); +x_290 = l_Array_append___rarg(x_284, x_289); +lean_dec(x_289); +lean_inc(x_53); +x_291 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_291, 0, x_53); +lean_ctor_set(x_291, 1, x_283); +lean_ctor_set(x_291, 2, x_290); +x_292 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; lean_inc(x_53); -x_289 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_289, 0, x_53); -lean_ctor_set(x_289, 1, x_281); -lean_ctor_set(x_289, 2, x_288); -x_290 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_283); +x_293 = l_Lean_Syntax_node1(x_53, x_292, x_291); +x_294 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_285); lean_inc(x_53); -x_291 = l_Lean_Syntax_node1(x_53, x_290, x_283); -x_292 = l_Lake_DSL_structVal___closed__18; +x_295 = l_Lean_Syntax_node1(x_53, x_294, x_285); +x_296 = l_Lake_DSL_structVal___closed__18; lean_inc(x_53); lean_ctor_set_tag(x_54, 2); -lean_ctor_set(x_54, 1, x_292); +lean_ctor_set(x_54, 1, x_296); lean_ctor_set(x_54, 0, x_53); -x_293 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_283); -x_294 = l_Lean_Syntax_node6(x_53, x_293, x_280, x_283, x_289, x_291, x_283, x_54); -x_295 = l_Lean_Elab_Command_getRef(x_12, x_13, x_278); -x_296 = lean_ctor_get(x_295, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_295, 1); -lean_inc(x_297); -if (lean_is_exclusive(x_295)) { - lean_ctor_release(x_295, 0); - lean_ctor_release(x_295, 1); - x_298 = x_295; -} else { - lean_dec_ref(x_295); - x_298 = lean_box(0); -} -x_299 = l_Lean_SourceInfo_fromRef(x_296, x_29); -lean_dec(x_296); -x_300 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_297); -x_301 = lean_ctor_get(x_300, 1); +x_297 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_285); +x_298 = l_Lean_Syntax_node6(x_53, x_297, x_282, x_285, x_293, x_295, x_285, x_54); +x_299 = l_Lean_Elab_Command_getRef(x_12, x_13, x_280); +x_300 = lean_ctor_get(x_299, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_299, 1); lean_inc(x_301); -if (lean_is_exclusive(x_300)) { - lean_ctor_release(x_300, 0); - lean_ctor_release(x_300, 1); - x_302 = x_300; +if (lean_is_exclusive(x_299)) { + lean_ctor_release(x_299, 0); + lean_ctor_release(x_299, 1); + x_302 = x_299; } else { - lean_dec_ref(x_300); + lean_dec_ref(x_299); x_302 = lean_box(0); } -x_303 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_301); -x_304 = lean_ctor_get(x_303, 1); -lean_inc(x_304); -if (lean_is_exclusive(x_303)) { - lean_ctor_release(x_303, 0); - lean_ctor_release(x_303, 1); - x_305 = x_303; -} else { - lean_dec_ref(x_303); - x_305 = lean_box(0); -} -x_306 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_299); -if (lean_is_scalar(x_305)) { - x_307 = lean_alloc_ctor(2, 2, 0); -} else { - x_307 = x_305; - lean_ctor_set_tag(x_307, 2); -} -lean_ctor_set(x_307, 0, x_299); -lean_ctor_set(x_307, 1, x_306); -x_308 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_309 = l_Lean_Syntax_SepArray_ofElems(x_308, x_5); -x_310 = l_Array_append___rarg(x_282, x_309); -lean_dec(x_309); -lean_inc(x_299); -x_311 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_311, 0, x_299); -lean_ctor_set(x_311, 1, x_281); -lean_ctor_set(x_311, 2, x_310); -x_312 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_299); +x_303 = l_Lean_SourceInfo_fromRef(x_300, x_29); +lean_dec(x_300); +x_304 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_301); +x_305 = lean_ctor_get(x_304, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_304)) { + lean_ctor_release(x_304, 0); + lean_ctor_release(x_304, 1); + x_306 = x_304; +} else { + lean_dec_ref(x_304); + x_306 = lean_box(0); +} +x_307 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_305); +x_308 = lean_ctor_get(x_307, 1); +lean_inc(x_308); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_309 = x_307; +} else { + lean_dec_ref(x_307); + x_309 = lean_box(0); +} +x_310 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_303); +if (lean_is_scalar(x_309)) { + x_311 = lean_alloc_ctor(2, 2, 0); +} else { + x_311 = x_309; + lean_ctor_set_tag(x_311, 2); +} +lean_ctor_set(x_311, 0, x_303); +lean_ctor_set(x_311, 1, x_310); +x_312 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_313 = l_Lean_Syntax_SepArray_ofElems(x_312, x_5); +x_314 = l_Array_append___rarg(x_284, x_313); +lean_dec(x_313); +lean_inc(x_303); +x_315 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_315, 0, x_303); +lean_ctor_set(x_315, 1, x_283); +lean_ctor_set(x_315, 2, x_314); +x_316 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_303); +if (lean_is_scalar(x_306)) { + x_317 = lean_alloc_ctor(2, 2, 0); +} else { + x_317 = x_306; + lean_ctor_set_tag(x_317, 2); +} +lean_ctor_set(x_317, 0, x_303); +lean_ctor_set(x_317, 1, x_316); +x_318 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_303); +x_319 = l_Lean_Syntax_node3(x_303, x_318, x_311, x_315, x_317); +lean_inc(x_303); +x_320 = l_Lean_Syntax_node1(x_303, x_283, x_319); +lean_inc(x_303); +x_321 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_321, 0, x_303); +lean_ctor_set(x_321, 1, x_283); +lean_ctor_set(x_321, 2, x_284); +x_322 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_303); if (lean_is_scalar(x_302)) { - x_313 = lean_alloc_ctor(2, 2, 0); -} else { - x_313 = x_302; - lean_ctor_set_tag(x_313, 2); -} -lean_ctor_set(x_313, 0, x_299); -lean_ctor_set(x_313, 1, x_312); -x_314 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_299); -x_315 = l_Lean_Syntax_node3(x_299, x_314, x_307, x_311, x_313); -lean_inc(x_299); -x_316 = l_Lean_Syntax_node1(x_299, x_281, x_315); -lean_inc(x_299); -x_317 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_317, 0, x_299); -lean_ctor_set(x_317, 1, x_281); -lean_ctor_set(x_317, 2, x_282); -x_318 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_299); -if (lean_is_scalar(x_298)) { - x_319 = lean_alloc_ctor(2, 2, 0); -} else { - x_319 = x_298; - lean_ctor_set_tag(x_319, 2); -} -lean_ctor_set(x_319, 0, x_299); -lean_ctor_set(x_319, 1, x_318); -x_320 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_299); + x_323 = lean_alloc_ctor(2, 2, 0); +} else { + x_323 = x_302; + lean_ctor_set_tag(x_323, 2); +} +lean_ctor_set(x_323, 0, x_303); +lean_ctor_set(x_323, 1, x_322); +x_324 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_303); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_320); -lean_ctor_set(x_49, 0, x_299); -x_321 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_299); -x_322 = l_Lean_Syntax_node2(x_299, x_321, x_49, x_30); -lean_inc(x_299); -x_323 = l_Lean_Syntax_node1(x_299, x_281, x_322); -x_324 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_317); -lean_inc(x_299); -x_325 = l_Lean_Syntax_node2(x_299, x_324, x_317, x_323); -x_326 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_299); +lean_ctor_set(x_49, 1, x_324); +lean_ctor_set(x_49, 0, x_303); +x_325 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_303); +x_326 = l_Lean_Syntax_node2(x_303, x_325, x_49, x_30); +lean_inc(x_303); +x_327 = l_Lean_Syntax_node1(x_303, x_283, x_326); +x_328 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_321); +lean_inc(x_303); +x_329 = l_Lean_Syntax_node2(x_303, x_328, x_321, x_327); +x_330 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_303); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_326); -lean_ctor_set(x_33, 0, x_299); -x_327 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_317, 2); -lean_inc(x_299); -x_328 = l_Lean_Syntax_node2(x_299, x_327, x_317, x_317); -x_329 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_299); -x_330 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_330, 0, x_299); -lean_ctor_set(x_330, 1, x_281); -lean_ctor_set(x_330, 2, x_329); -x_331 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_330); -lean_inc(x_299); -x_332 = l_Lean_Syntax_node4(x_299, x_331, x_33, x_294, x_328, x_330); -x_333 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_299); -x_334 = l_Lean_Syntax_node4(x_299, x_333, x_319, x_31, x_325, x_332); +lean_ctor_set(x_33, 1, x_330); +lean_ctor_set(x_33, 0, x_303); +x_331 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_321, 2); +lean_inc(x_303); +x_332 = l_Lean_Syntax_node2(x_303, x_331, x_321, x_321); +x_333 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_303); +x_334 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_334, 0, x_303); +lean_ctor_set(x_334, 1, x_283); +lean_ctor_set(x_334, 2, x_333); +x_335 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_334); +lean_inc(x_303); +x_336 = l_Lean_Syntax_node4(x_303, x_335, x_33, x_298, x_332, x_334); +x_337 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_303); +x_338 = l_Lean_Syntax_node4(x_303, x_337, x_323, x_31, x_329, x_336); if (lean_obj_tag(x_6) == 0) { -lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; -x_335 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_317, 3); -lean_inc(x_299); -x_336 = l_Lean_Syntax_node6(x_299, x_335, x_330, x_316, x_317, x_317, x_317, x_317); -x_337 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_338 = l_Lean_Syntax_node2(x_299, x_337, x_336, x_334); -lean_inc(x_338); -x_339 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_339, 0, x_338); -x_340 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_338, x_339, x_12, x_13, x_304); -return x_340; -} -else -{ -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; -lean_dec(x_330); -x_341 = lean_ctor_get(x_6, 0); -lean_inc(x_341); +lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; +x_339 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_321, 3); +lean_inc(x_303); +x_340 = l_Lean_Syntax_node6(x_303, x_339, x_334, x_320, x_321, x_321, x_321, x_321); +x_341 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_342 = l_Lean_Syntax_node2(x_303, x_341, x_340, x_338); +lean_inc(x_342); +x_343 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_343, 0, x_342); +x_344 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_342, x_343, x_12, x_13, x_308); +return x_344; +} +else +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +lean_dec(x_334); +x_345 = lean_ctor_get(x_6, 0); +lean_inc(x_345); lean_dec(x_6); -x_342 = l_Array_mkArray1___rarg(x_341); -x_343 = l_Array_append___rarg(x_282, x_342); -lean_dec(x_342); -lean_inc(x_299); -x_344 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_344, 0, x_299); -lean_ctor_set(x_344, 1, x_281); -lean_ctor_set(x_344, 2, x_343); -x_345 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_317, 3); -lean_inc(x_299); -x_346 = l_Lean_Syntax_node6(x_299, x_345, x_344, x_316, x_317, x_317, x_317, x_317); -x_347 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_348 = l_Lean_Syntax_node2(x_299, x_347, x_346, x_334); -lean_inc(x_348); -x_349 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_349, 0, x_348); -x_350 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_348, x_349, x_12, x_13, x_304); -return x_350; -} -} -} -else -{ -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; size_t x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; -x_351 = lean_ctor_get(x_54, 1); -lean_inc(x_351); +x_346 = l_Array_mkArray1___rarg(x_345); +x_347 = l_Array_append___rarg(x_284, x_346); +lean_dec(x_346); +lean_inc(x_303); +x_348 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_348, 0, x_303); +lean_ctor_set(x_348, 1, x_283); +lean_ctor_set(x_348, 2, x_347); +x_349 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_321, 3); +lean_inc(x_303); +x_350 = l_Lean_Syntax_node6(x_303, x_349, x_348, x_320, x_321, x_321, x_321, x_321); +x_351 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_352 = l_Lean_Syntax_node2(x_303, x_351, x_350, x_338); +lean_inc(x_352); +x_353 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_353, 0, x_352); +x_354 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_352, x_353, x_12, x_13, x_308); +return x_354; +} +} +} +else +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; size_t x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; +x_355 = lean_ctor_get(x_54, 1); +lean_inc(x_355); lean_dec(x_54); -x_352 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_351); -x_353 = lean_ctor_get(x_352, 1); -lean_inc(x_353); -if (lean_is_exclusive(x_352)) { - lean_ctor_release(x_352, 0); - lean_ctor_release(x_352, 1); - x_354 = x_352; -} else { - lean_dec_ref(x_352); - x_354 = lean_box(0); -} -x_355 = l_Lake_DSL_structVal___closed__3; +x_356 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_355); +x_357 = lean_ctor_get(x_356, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_356)) { + lean_ctor_release(x_356, 0); + lean_ctor_release(x_356, 1); + x_358 = x_356; +} else { + lean_dec_ref(x_356); + x_358 = lean_box(0); +} +x_359 = l_Lake_DSL_structVal___closed__3; lean_inc(x_53); -if (lean_is_scalar(x_354)) { - x_356 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_358)) { + x_360 = lean_alloc_ctor(2, 2, 0); } else { - x_356 = x_354; - lean_ctor_set_tag(x_356, 2); + x_360 = x_358; + lean_ctor_set_tag(x_360, 2); } -lean_ctor_set(x_356, 0, x_53); -lean_ctor_set(x_356, 1, x_355); -x_357 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_358 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_53); -x_359 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_359, 0, x_53); -lean_ctor_set(x_359, 1, x_357); -lean_ctor_set(x_359, 2, x_358); -x_360 = lean_array_size(x_23); -x_361 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_360, x_17, x_23); -x_362 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_363 = l_Lean_mkSepArray(x_361, x_362); -lean_dec(x_361); -x_364 = l_Array_append___rarg(x_358, x_363); -lean_dec(x_363); +lean_ctor_set(x_360, 0, x_53); +lean_ctor_set(x_360, 1, x_359); +x_361 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_362 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; lean_inc(x_53); -x_365 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_365, 0, x_53); -lean_ctor_set(x_365, 1, x_357); -lean_ctor_set(x_365, 2, x_364); +x_363 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_363, 0, x_53); +lean_ctor_set(x_363, 1, x_361); +lean_ctor_set(x_363, 2, x_362); +x_364 = lean_array_size(x_23); +x_365 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_364, x_17, x_23); x_366 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_359); -lean_inc(x_53); -x_367 = l_Lean_Syntax_node1(x_53, x_366, x_359); -x_368 = l_Lake_DSL_structVal___closed__18; +x_367 = l_Lean_mkSepArray(x_365, x_366); +lean_dec(x_365); +x_368 = l_Array_append___rarg(x_362, x_367); +lean_dec(x_367); lean_inc(x_53); -x_369 = lean_alloc_ctor(2, 2, 0); +x_369 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_369, 0, x_53); -lean_ctor_set(x_369, 1, x_368); -x_370 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_359); -x_371 = l_Lean_Syntax_node6(x_53, x_370, x_356, x_359, x_365, x_367, x_359, x_369); -x_372 = l_Lean_Elab_Command_getRef(x_12, x_13, x_353); -x_373 = lean_ctor_get(x_372, 0); -lean_inc(x_373); -x_374 = lean_ctor_get(x_372, 1); -lean_inc(x_374); -if (lean_is_exclusive(x_372)) { - lean_ctor_release(x_372, 0); - lean_ctor_release(x_372, 1); - x_375 = x_372; -} else { - lean_dec_ref(x_372); - x_375 = lean_box(0); -} -x_376 = l_Lean_SourceInfo_fromRef(x_373, x_29); -lean_dec(x_373); -x_377 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_374); -x_378 = lean_ctor_get(x_377, 1); -lean_inc(x_378); -if (lean_is_exclusive(x_377)) { - lean_ctor_release(x_377, 0); - lean_ctor_release(x_377, 1); - x_379 = x_377; -} else { - lean_dec_ref(x_377); - x_379 = lean_box(0); -} -x_380 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_378); -x_381 = lean_ctor_get(x_380, 1); -lean_inc(x_381); -if (lean_is_exclusive(x_380)) { - lean_ctor_release(x_380, 0); - lean_ctor_release(x_380, 1); - x_382 = x_380; -} else { - lean_dec_ref(x_380); - x_382 = lean_box(0); -} -x_383 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_376); -if (lean_is_scalar(x_382)) { - x_384 = lean_alloc_ctor(2, 2, 0); -} else { - x_384 = x_382; - lean_ctor_set_tag(x_384, 2); -} -lean_ctor_set(x_384, 0, x_376); -lean_ctor_set(x_384, 1, x_383); -x_385 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_386 = l_Lean_Syntax_SepArray_ofElems(x_385, x_5); -x_387 = l_Array_append___rarg(x_358, x_386); -lean_dec(x_386); -lean_inc(x_376); -x_388 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_388, 0, x_376); -lean_ctor_set(x_388, 1, x_357); -lean_ctor_set(x_388, 2, x_387); -x_389 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_376); -if (lean_is_scalar(x_379)) { +lean_ctor_set(x_369, 1, x_361); +lean_ctor_set(x_369, 2, x_368); +x_370 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_53); +x_371 = l_Lean_Syntax_node1(x_53, x_370, x_369); +x_372 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_363); +lean_inc(x_53); +x_373 = l_Lean_Syntax_node1(x_53, x_372, x_363); +x_374 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_53); +x_375 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_375, 0, x_53); +lean_ctor_set(x_375, 1, x_374); +x_376 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_363); +x_377 = l_Lean_Syntax_node6(x_53, x_376, x_360, x_363, x_371, x_373, x_363, x_375); +x_378 = l_Lean_Elab_Command_getRef(x_12, x_13, x_357); +x_379 = lean_ctor_get(x_378, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_378, 1); +lean_inc(x_380); +if (lean_is_exclusive(x_378)) { + lean_ctor_release(x_378, 0); + lean_ctor_release(x_378, 1); + x_381 = x_378; +} else { + lean_dec_ref(x_378); + x_381 = lean_box(0); +} +x_382 = l_Lean_SourceInfo_fromRef(x_379, x_29); +lean_dec(x_379); +x_383 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_380); +x_384 = lean_ctor_get(x_383, 1); +lean_inc(x_384); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_385 = x_383; +} else { + lean_dec_ref(x_383); + x_385 = lean_box(0); +} +x_386 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_384); +x_387 = lean_ctor_get(x_386, 1); +lean_inc(x_387); +if (lean_is_exclusive(x_386)) { + lean_ctor_release(x_386, 0); + lean_ctor_release(x_386, 1); + x_388 = x_386; +} else { + lean_dec_ref(x_386); + x_388 = lean_box(0); +} +x_389 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_382); +if (lean_is_scalar(x_388)) { x_390 = lean_alloc_ctor(2, 2, 0); } else { - x_390 = x_379; + x_390 = x_388; lean_ctor_set_tag(x_390, 2); } -lean_ctor_set(x_390, 0, x_376); +lean_ctor_set(x_390, 0, x_382); lean_ctor_set(x_390, 1, x_389); -x_391 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_376); -x_392 = l_Lean_Syntax_node3(x_376, x_391, x_384, x_388, x_390); -lean_inc(x_376); -x_393 = l_Lean_Syntax_node1(x_376, x_357, x_392); -lean_inc(x_376); +x_391 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_392 = l_Lean_Syntax_SepArray_ofElems(x_391, x_5); +x_393 = l_Array_append___rarg(x_362, x_392); +lean_dec(x_392); +lean_inc(x_382); x_394 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_394, 0, x_376); -lean_ctor_set(x_394, 1, x_357); -lean_ctor_set(x_394, 2, x_358); -x_395 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_376); -if (lean_is_scalar(x_375)) { +lean_ctor_set(x_394, 0, x_382); +lean_ctor_set(x_394, 1, x_361); +lean_ctor_set(x_394, 2, x_393); +x_395 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_382); +if (lean_is_scalar(x_385)) { x_396 = lean_alloc_ctor(2, 2, 0); } else { - x_396 = x_375; + x_396 = x_385; lean_ctor_set_tag(x_396, 2); } -lean_ctor_set(x_396, 0, x_376); +lean_ctor_set(x_396, 0, x_382); lean_ctor_set(x_396, 1, x_395); -x_397 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_376); +x_397 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_382); +x_398 = l_Lean_Syntax_node3(x_382, x_397, x_390, x_394, x_396); +lean_inc(x_382); +x_399 = l_Lean_Syntax_node1(x_382, x_361, x_398); +lean_inc(x_382); +x_400 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_400, 0, x_382); +lean_ctor_set(x_400, 1, x_361); +lean_ctor_set(x_400, 2, x_362); +x_401 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_382); +if (lean_is_scalar(x_381)) { + x_402 = lean_alloc_ctor(2, 2, 0); +} else { + x_402 = x_381; + lean_ctor_set_tag(x_402, 2); +} +lean_ctor_set(x_402, 0, x_382); +lean_ctor_set(x_402, 1, x_401); +x_403 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_382); lean_ctor_set_tag(x_49, 2); -lean_ctor_set(x_49, 1, x_397); -lean_ctor_set(x_49, 0, x_376); -x_398 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_376); -x_399 = l_Lean_Syntax_node2(x_376, x_398, x_49, x_30); -lean_inc(x_376); -x_400 = l_Lean_Syntax_node1(x_376, x_357, x_399); -x_401 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_394); -lean_inc(x_376); -x_402 = l_Lean_Syntax_node2(x_376, x_401, x_394, x_400); -x_403 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_376); +lean_ctor_set(x_49, 1, x_403); +lean_ctor_set(x_49, 0, x_382); +x_404 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_382); +x_405 = l_Lean_Syntax_node2(x_382, x_404, x_49, x_30); +lean_inc(x_382); +x_406 = l_Lean_Syntax_node1(x_382, x_361, x_405); +x_407 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_400); +lean_inc(x_382); +x_408 = l_Lean_Syntax_node2(x_382, x_407, x_400, x_406); +x_409 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_382); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_403); -lean_ctor_set(x_33, 0, x_376); -x_404 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_394, 2); -lean_inc(x_376); -x_405 = l_Lean_Syntax_node2(x_376, x_404, x_394, x_394); -x_406 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_376); -x_407 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_407, 0, x_376); -lean_ctor_set(x_407, 1, x_357); -lean_ctor_set(x_407, 2, x_406); -x_408 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_407); -lean_inc(x_376); -x_409 = l_Lean_Syntax_node4(x_376, x_408, x_33, x_371, x_405, x_407); -x_410 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_376); -x_411 = l_Lean_Syntax_node4(x_376, x_410, x_396, x_31, x_402, x_409); +lean_ctor_set(x_33, 1, x_409); +lean_ctor_set(x_33, 0, x_382); +x_410 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_400, 2); +lean_inc(x_382); +x_411 = l_Lean_Syntax_node2(x_382, x_410, x_400, x_400); +x_412 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_382); +x_413 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_413, 0, x_382); +lean_ctor_set(x_413, 1, x_361); +lean_ctor_set(x_413, 2, x_412); +x_414 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_413); +lean_inc(x_382); +x_415 = l_Lean_Syntax_node4(x_382, x_414, x_33, x_377, x_411, x_413); +x_416 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_382); +x_417 = l_Lean_Syntax_node4(x_382, x_416, x_402, x_31, x_408, x_415); if (lean_obj_tag(x_6) == 0) { -lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; -x_412 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_394, 3); -lean_inc(x_376); -x_413 = l_Lean_Syntax_node6(x_376, x_412, x_407, x_393, x_394, x_394, x_394, x_394); -x_414 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_415 = l_Lean_Syntax_node2(x_376, x_414, x_413, x_411); -lean_inc(x_415); -x_416 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_416, 0, x_415); -x_417 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_415, x_416, x_12, x_13, x_381); -return x_417; +lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +x_418 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_400, 3); +lean_inc(x_382); +x_419 = l_Lean_Syntax_node6(x_382, x_418, x_413, x_399, x_400, x_400, x_400, x_400); +x_420 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_421 = l_Lean_Syntax_node2(x_382, x_420, x_419, x_417); +lean_inc(x_421); +x_422 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_422, 0, x_421); +x_423 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_421, x_422, x_12, x_13, x_387); +return x_423; } else { -lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; -lean_dec(x_407); -x_418 = lean_ctor_get(x_6, 0); -lean_inc(x_418); +lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; +lean_dec(x_413); +x_424 = lean_ctor_get(x_6, 0); +lean_inc(x_424); lean_dec(x_6); -x_419 = l_Array_mkArray1___rarg(x_418); -x_420 = l_Array_append___rarg(x_358, x_419); -lean_dec(x_419); -lean_inc(x_376); -x_421 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_421, 0, x_376); -lean_ctor_set(x_421, 1, x_357); -lean_ctor_set(x_421, 2, x_420); -x_422 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_394, 3); -lean_inc(x_376); -x_423 = l_Lean_Syntax_node6(x_376, x_422, x_421, x_393, x_394, x_394, x_394, x_394); -x_424 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_425 = l_Lean_Syntax_node2(x_376, x_424, x_423, x_411); -lean_inc(x_425); -x_426 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_426, 0, x_425); -x_427 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_425, x_426, x_12, x_13, x_381); -return x_427; -} -} -} -else -{ -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; size_t x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; -x_428 = lean_ctor_get(x_49, 0); -x_429 = lean_ctor_get(x_49, 1); -lean_inc(x_429); -lean_inc(x_428); +x_425 = l_Array_mkArray1___rarg(x_424); +x_426 = l_Array_append___rarg(x_362, x_425); +lean_dec(x_425); +lean_inc(x_382); +x_427 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_427, 0, x_382); +lean_ctor_set(x_427, 1, x_361); +lean_ctor_set(x_427, 2, x_426); +x_428 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_400, 3); +lean_inc(x_382); +x_429 = l_Lean_Syntax_node6(x_382, x_428, x_427, x_399, x_400, x_400, x_400, x_400); +x_430 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_431 = l_Lean_Syntax_node2(x_382, x_430, x_429, x_417); +lean_inc(x_431); +x_432 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_432, 0, x_431); +x_433 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_431, x_432, x_12, x_13, x_387); +return x_433; +} +} +} +else +{ +lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; size_t x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; +x_434 = lean_ctor_get(x_49, 0); +x_435 = lean_ctor_get(x_49, 1); +lean_inc(x_435); +lean_inc(x_434); lean_dec(x_49); -x_430 = l_Lean_SourceInfo_fromRef(x_428, x_29); -lean_dec(x_428); -x_431 = l_Lean_Elab_Command_getCurrMacroScope(x_48, x_13, x_429); +x_436 = l_Lean_SourceInfo_fromRef(x_434, x_29); +lean_dec(x_434); +x_437 = l_Lean_Elab_Command_getCurrMacroScope(x_48, x_13, x_435); lean_dec(x_48); -x_432 = lean_ctor_get(x_431, 1); -lean_inc(x_432); -if (lean_is_exclusive(x_431)) { - lean_ctor_release(x_431, 0); - lean_ctor_release(x_431, 1); - x_433 = x_431; -} else { - lean_dec_ref(x_431); - x_433 = lean_box(0); -} -x_434 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_432); -x_435 = lean_ctor_get(x_434, 1); -lean_inc(x_435); -if (lean_is_exclusive(x_434)) { - lean_ctor_release(x_434, 0); - lean_ctor_release(x_434, 1); - x_436 = x_434; -} else { - lean_dec_ref(x_434); - x_436 = lean_box(0); -} -x_437 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_430); -if (lean_is_scalar(x_436)) { - x_438 = lean_alloc_ctor(2, 2, 0); -} else { - x_438 = x_436; - lean_ctor_set_tag(x_438, 2); -} -lean_ctor_set(x_438, 0, x_430); -lean_ctor_set(x_438, 1, x_437); -x_439 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_440 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_430); -x_441 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_441, 0, x_430); -lean_ctor_set(x_441, 1, x_439); -lean_ctor_set(x_441, 2, x_440); -x_442 = lean_array_size(x_23); -x_443 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_442, x_17, x_23); -x_444 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_445 = l_Lean_mkSepArray(x_443, x_444); -lean_dec(x_443); -x_446 = l_Array_append___rarg(x_440, x_445); -lean_dec(x_445); -lean_inc(x_430); +x_438 = lean_ctor_get(x_437, 1); +lean_inc(x_438); +if (lean_is_exclusive(x_437)) { + lean_ctor_release(x_437, 0); + lean_ctor_release(x_437, 1); + x_439 = x_437; +} else { + lean_dec_ref(x_437); + x_439 = lean_box(0); +} +x_440 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_438); +x_441 = lean_ctor_get(x_440, 1); +lean_inc(x_441); +if (lean_is_exclusive(x_440)) { + lean_ctor_release(x_440, 0); + lean_ctor_release(x_440, 1); + x_442 = x_440; +} else { + lean_dec_ref(x_440); + x_442 = lean_box(0); +} +x_443 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_436); +if (lean_is_scalar(x_442)) { + x_444 = lean_alloc_ctor(2, 2, 0); +} else { + x_444 = x_442; + lean_ctor_set_tag(x_444, 2); +} +lean_ctor_set(x_444, 0, x_436); +lean_ctor_set(x_444, 1, x_443); +x_445 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_446 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_436); x_447 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_447, 0, x_430); -lean_ctor_set(x_447, 1, x_439); +lean_ctor_set(x_447, 0, x_436); +lean_ctor_set(x_447, 1, x_445); lean_ctor_set(x_447, 2, x_446); -x_448 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_441); -lean_inc(x_430); -x_449 = l_Lean_Syntax_node1(x_430, x_448, x_441); -x_450 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_430); -if (lean_is_scalar(x_433)) { - x_451 = lean_alloc_ctor(2, 2, 0); -} else { - x_451 = x_433; - lean_ctor_set_tag(x_451, 2); -} -lean_ctor_set(x_451, 0, x_430); -lean_ctor_set(x_451, 1, x_450); -x_452 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_441); -x_453 = l_Lean_Syntax_node6(x_430, x_452, x_438, x_441, x_447, x_449, x_441, x_451); -x_454 = l_Lean_Elab_Command_getRef(x_12, x_13, x_435); -x_455 = lean_ctor_get(x_454, 0); -lean_inc(x_455); -x_456 = lean_ctor_get(x_454, 1); -lean_inc(x_456); -if (lean_is_exclusive(x_454)) { - lean_ctor_release(x_454, 0); - lean_ctor_release(x_454, 1); - x_457 = x_454; -} else { - lean_dec_ref(x_454); - x_457 = lean_box(0); -} -x_458 = l_Lean_SourceInfo_fromRef(x_455, x_29); -lean_dec(x_455); -x_459 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_456); -x_460 = lean_ctor_get(x_459, 1); -lean_inc(x_460); -if (lean_is_exclusive(x_459)) { - lean_ctor_release(x_459, 0); - lean_ctor_release(x_459, 1); - x_461 = x_459; -} else { - lean_dec_ref(x_459); - x_461 = lean_box(0); -} -x_462 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_460); -x_463 = lean_ctor_get(x_462, 1); +x_448 = lean_array_size(x_23); +x_449 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_448, x_17, x_23); +x_450 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_451 = l_Lean_mkSepArray(x_449, x_450); +lean_dec(x_449); +x_452 = l_Array_append___rarg(x_446, x_451); +lean_dec(x_451); +lean_inc(x_436); +x_453 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_453, 0, x_436); +lean_ctor_set(x_453, 1, x_445); +lean_ctor_set(x_453, 2, x_452); +x_454 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_436); +x_455 = l_Lean_Syntax_node1(x_436, x_454, x_453); +x_456 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_447); +lean_inc(x_436); +x_457 = l_Lean_Syntax_node1(x_436, x_456, x_447); +x_458 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_436); +if (lean_is_scalar(x_439)) { + x_459 = lean_alloc_ctor(2, 2, 0); +} else { + x_459 = x_439; + lean_ctor_set_tag(x_459, 2); +} +lean_ctor_set(x_459, 0, x_436); +lean_ctor_set(x_459, 1, x_458); +x_460 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_447); +x_461 = l_Lean_Syntax_node6(x_436, x_460, x_444, x_447, x_455, x_457, x_447, x_459); +x_462 = l_Lean_Elab_Command_getRef(x_12, x_13, x_441); +x_463 = lean_ctor_get(x_462, 0); lean_inc(x_463); +x_464 = lean_ctor_get(x_462, 1); +lean_inc(x_464); if (lean_is_exclusive(x_462)) { lean_ctor_release(x_462, 0); lean_ctor_release(x_462, 1); - x_464 = x_462; + x_465 = x_462; } else { lean_dec_ref(x_462); - x_464 = lean_box(0); -} -x_465 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_458); -if (lean_is_scalar(x_464)) { - x_466 = lean_alloc_ctor(2, 2, 0); + x_465 = lean_box(0); +} +x_466 = l_Lean_SourceInfo_fromRef(x_463, x_29); +lean_dec(x_463); +x_467 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_464); +x_468 = lean_ctor_get(x_467, 1); +lean_inc(x_468); +if (lean_is_exclusive(x_467)) { + lean_ctor_release(x_467, 0); + lean_ctor_release(x_467, 1); + x_469 = x_467; +} else { + lean_dec_ref(x_467); + x_469 = lean_box(0); +} +x_470 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_468); +x_471 = lean_ctor_get(x_470, 1); +lean_inc(x_471); +if (lean_is_exclusive(x_470)) { + lean_ctor_release(x_470, 0); + lean_ctor_release(x_470, 1); + x_472 = x_470; +} else { + lean_dec_ref(x_470); + x_472 = lean_box(0); +} +x_473 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_466); +if (lean_is_scalar(x_472)) { + x_474 = lean_alloc_ctor(2, 2, 0); +} else { + x_474 = x_472; + lean_ctor_set_tag(x_474, 2); +} +lean_ctor_set(x_474, 0, x_466); +lean_ctor_set(x_474, 1, x_473); +x_475 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_476 = l_Lean_Syntax_SepArray_ofElems(x_475, x_5); +x_477 = l_Array_append___rarg(x_446, x_476); +lean_dec(x_476); +lean_inc(x_466); +x_478 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_478, 0, x_466); +lean_ctor_set(x_478, 1, x_445); +lean_ctor_set(x_478, 2, x_477); +x_479 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_466); +if (lean_is_scalar(x_469)) { + x_480 = lean_alloc_ctor(2, 2, 0); +} else { + x_480 = x_469; + lean_ctor_set_tag(x_480, 2); +} +lean_ctor_set(x_480, 0, x_466); +lean_ctor_set(x_480, 1, x_479); +x_481 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_466); +x_482 = l_Lean_Syntax_node3(x_466, x_481, x_474, x_478, x_480); +lean_inc(x_466); +x_483 = l_Lean_Syntax_node1(x_466, x_445, x_482); +lean_inc(x_466); +x_484 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_484, 0, x_466); +lean_ctor_set(x_484, 1, x_445); +lean_ctor_set(x_484, 2, x_446); +x_485 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_466); +if (lean_is_scalar(x_465)) { + x_486 = lean_alloc_ctor(2, 2, 0); } else { - x_466 = x_464; - lean_ctor_set_tag(x_466, 2); + x_486 = x_465; + lean_ctor_set_tag(x_486, 2); } -lean_ctor_set(x_466, 0, x_458); -lean_ctor_set(x_466, 1, x_465); -x_467 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_468 = l_Lean_Syntax_SepArray_ofElems(x_467, x_5); -x_469 = l_Array_append___rarg(x_440, x_468); -lean_dec(x_468); -lean_inc(x_458); -x_470 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_470, 0, x_458); -lean_ctor_set(x_470, 1, x_439); -lean_ctor_set(x_470, 2, x_469); -x_471 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_458); -if (lean_is_scalar(x_461)) { - x_472 = lean_alloc_ctor(2, 2, 0); -} else { - x_472 = x_461; - lean_ctor_set_tag(x_472, 2); -} -lean_ctor_set(x_472, 0, x_458); -lean_ctor_set(x_472, 1, x_471); -x_473 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_458); -x_474 = l_Lean_Syntax_node3(x_458, x_473, x_466, x_470, x_472); -lean_inc(x_458); -x_475 = l_Lean_Syntax_node1(x_458, x_439, x_474); -lean_inc(x_458); -x_476 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_476, 0, x_458); -lean_ctor_set(x_476, 1, x_439); -lean_ctor_set(x_476, 2, x_440); -x_477 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_458); -if (lean_is_scalar(x_457)) { - x_478 = lean_alloc_ctor(2, 2, 0); -} else { - x_478 = x_457; - lean_ctor_set_tag(x_478, 2); -} -lean_ctor_set(x_478, 0, x_458); -lean_ctor_set(x_478, 1, x_477); -x_479 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_458); -x_480 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_480, 0, x_458); -lean_ctor_set(x_480, 1, x_479); -x_481 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_458); -x_482 = l_Lean_Syntax_node2(x_458, x_481, x_480, x_30); -lean_inc(x_458); -x_483 = l_Lean_Syntax_node1(x_458, x_439, x_482); -x_484 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_476); -lean_inc(x_458); -x_485 = l_Lean_Syntax_node2(x_458, x_484, x_476, x_483); -x_486 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_458); +lean_ctor_set(x_486, 0, x_466); +lean_ctor_set(x_486, 1, x_485); +x_487 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_466); +x_488 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_488, 0, x_466); +lean_ctor_set(x_488, 1, x_487); +x_489 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_466); +x_490 = l_Lean_Syntax_node2(x_466, x_489, x_488, x_30); +lean_inc(x_466); +x_491 = l_Lean_Syntax_node1(x_466, x_445, x_490); +x_492 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_484); +lean_inc(x_466); +x_493 = l_Lean_Syntax_node2(x_466, x_492, x_484, x_491); +x_494 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_466); lean_ctor_set_tag(x_33, 2); -lean_ctor_set(x_33, 1, x_486); -lean_ctor_set(x_33, 0, x_458); -x_487 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_476, 2); -lean_inc(x_458); -x_488 = l_Lean_Syntax_node2(x_458, x_487, x_476, x_476); -x_489 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_458); -x_490 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_490, 0, x_458); -lean_ctor_set(x_490, 1, x_439); -lean_ctor_set(x_490, 2, x_489); -x_491 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_490); -lean_inc(x_458); -x_492 = l_Lean_Syntax_node4(x_458, x_491, x_33, x_453, x_488, x_490); -x_493 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_458); -x_494 = l_Lean_Syntax_node4(x_458, x_493, x_478, x_31, x_485, x_492); +lean_ctor_set(x_33, 1, x_494); +lean_ctor_set(x_33, 0, x_466); +x_495 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_484, 2); +lean_inc(x_466); +x_496 = l_Lean_Syntax_node2(x_466, x_495, x_484, x_484); +x_497 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_466); +x_498 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_498, 0, x_466); +lean_ctor_set(x_498, 1, x_445); +lean_ctor_set(x_498, 2, x_497); +x_499 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_498); +lean_inc(x_466); +x_500 = l_Lean_Syntax_node4(x_466, x_499, x_33, x_461, x_496, x_498); +x_501 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_466); +x_502 = l_Lean_Syntax_node4(x_466, x_501, x_486, x_31, x_493, x_500); if (lean_obj_tag(x_6) == 0) { -lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; -x_495 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_476, 3); -lean_inc(x_458); -x_496 = l_Lean_Syntax_node6(x_458, x_495, x_490, x_475, x_476, x_476, x_476, x_476); -x_497 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_498 = l_Lean_Syntax_node2(x_458, x_497, x_496, x_494); -lean_inc(x_498); -x_499 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_499, 0, x_498); -x_500 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_498, x_499, x_12, x_13, x_463); -return x_500; +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; +x_503 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_484, 3); +lean_inc(x_466); +x_504 = l_Lean_Syntax_node6(x_466, x_503, x_498, x_483, x_484, x_484, x_484, x_484); +x_505 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_506 = l_Lean_Syntax_node2(x_466, x_505, x_504, x_502); +lean_inc(x_506); +x_507 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_507, 0, x_506); +x_508 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_506, x_507, x_12, x_13, x_471); +return x_508; } else { -lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; -lean_dec(x_490); -x_501 = lean_ctor_get(x_6, 0); -lean_inc(x_501); +lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +lean_dec(x_498); +x_509 = lean_ctor_get(x_6, 0); +lean_inc(x_509); lean_dec(x_6); -x_502 = l_Array_mkArray1___rarg(x_501); -x_503 = l_Array_append___rarg(x_440, x_502); -lean_dec(x_502); -lean_inc(x_458); -x_504 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_504, 0, x_458); -lean_ctor_set(x_504, 1, x_439); -lean_ctor_set(x_504, 2, x_503); -x_505 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_476, 3); -lean_inc(x_458); -x_506 = l_Lean_Syntax_node6(x_458, x_505, x_504, x_475, x_476, x_476, x_476, x_476); -x_507 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_508 = l_Lean_Syntax_node2(x_458, x_507, x_506, x_494); -lean_inc(x_508); -x_509 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_509, 0, x_508); -x_510 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_508, x_509, x_12, x_13, x_463); -return x_510; +x_510 = l_Array_mkArray1___rarg(x_509); +x_511 = l_Array_append___rarg(x_446, x_510); +lean_dec(x_510); +lean_inc(x_466); +x_512 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_512, 0, x_466); +lean_ctor_set(x_512, 1, x_445); +lean_ctor_set(x_512, 2, x_511); +x_513 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_484, 3); +lean_inc(x_466); +x_514 = l_Lean_Syntax_node6(x_466, x_513, x_512, x_483, x_484, x_484, x_484, x_484); +x_515 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_516 = l_Lean_Syntax_node2(x_466, x_515, x_514, x_502); +lean_inc(x_516); +x_517 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_517, 0, x_516); +x_518 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_516, x_517, x_12, x_13, x_471); +return x_518; } } } else { -lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; uint8_t x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; size_t x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; -x_511 = lean_ctor_get(x_33, 0); -x_512 = lean_ctor_get(x_33, 1); -lean_inc(x_512); -lean_inc(x_511); -lean_dec(x_33); -x_513 = l_Lean_replaceRef(x_4, x_511); -lean_dec(x_511); -x_514 = lean_ctor_get(x_12, 0); -lean_inc(x_514); -x_515 = lean_ctor_get(x_12, 1); -lean_inc(x_515); -x_516 = lean_ctor_get(x_12, 2); -lean_inc(x_516); -x_517 = lean_ctor_get(x_12, 3); -lean_inc(x_517); -x_518 = lean_ctor_get(x_12, 4); -lean_inc(x_518); -x_519 = lean_ctor_get(x_12, 5); -lean_inc(x_519); -x_520 = lean_ctor_get(x_12, 7); +lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; uint8_t x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; size_t x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; +x_519 = lean_ctor_get(x_33, 0); +x_520 = lean_ctor_get(x_33, 1); lean_inc(x_520); -x_521 = lean_ctor_get(x_12, 8); -lean_inc(x_521); -x_522 = lean_ctor_get(x_12, 9); +lean_inc(x_519); +lean_dec(x_33); +x_521 = l_Lean_replaceRef(x_4, x_519); +lean_dec(x_519); +x_522 = lean_ctor_get(x_12, 0); lean_inc(x_522); -x_523 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); -x_524 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_524, 0, x_514); -lean_ctor_set(x_524, 1, x_515); -lean_ctor_set(x_524, 2, x_516); -lean_ctor_set(x_524, 3, x_517); -lean_ctor_set(x_524, 4, x_518); -lean_ctor_set(x_524, 5, x_519); -lean_ctor_set(x_524, 6, x_513); -lean_ctor_set(x_524, 7, x_520); -lean_ctor_set(x_524, 8, x_521); -lean_ctor_set(x_524, 9, x_522); -lean_ctor_set_uint8(x_524, sizeof(void*)*10, x_523); -x_525 = l_Lean_Elab_Command_getRef(x_524, x_13, x_512); -x_526 = lean_ctor_get(x_525, 0); +x_523 = lean_ctor_get(x_12, 1); +lean_inc(x_523); +x_524 = lean_ctor_get(x_12, 2); +lean_inc(x_524); +x_525 = lean_ctor_get(x_12, 3); +lean_inc(x_525); +x_526 = lean_ctor_get(x_12, 4); lean_inc(x_526); -x_527 = lean_ctor_get(x_525, 1); +x_527 = lean_ctor_get(x_12, 5); lean_inc(x_527); -if (lean_is_exclusive(x_525)) { - lean_ctor_release(x_525, 0); - lean_ctor_release(x_525, 1); - x_528 = x_525; -} else { - lean_dec_ref(x_525); - x_528 = lean_box(0); -} -x_529 = l_Lean_SourceInfo_fromRef(x_526, x_29); -lean_dec(x_526); -x_530 = l_Lean_Elab_Command_getCurrMacroScope(x_524, x_13, x_527); -lean_dec(x_524); -x_531 = lean_ctor_get(x_530, 1); -lean_inc(x_531); -if (lean_is_exclusive(x_530)) { - lean_ctor_release(x_530, 0); - lean_ctor_release(x_530, 1); - x_532 = x_530; -} else { - lean_dec_ref(x_530); - x_532 = lean_box(0); -} -x_533 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_531); -x_534 = lean_ctor_get(x_533, 1); +x_528 = lean_ctor_get(x_12, 7); +lean_inc(x_528); +x_529 = lean_ctor_get(x_12, 8); +lean_inc(x_529); +x_530 = lean_ctor_get(x_12, 9); +lean_inc(x_530); +x_531 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); +x_532 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_532, 0, x_522); +lean_ctor_set(x_532, 1, x_523); +lean_ctor_set(x_532, 2, x_524); +lean_ctor_set(x_532, 3, x_525); +lean_ctor_set(x_532, 4, x_526); +lean_ctor_set(x_532, 5, x_527); +lean_ctor_set(x_532, 6, x_521); +lean_ctor_set(x_532, 7, x_528); +lean_ctor_set(x_532, 8, x_529); +lean_ctor_set(x_532, 9, x_530); +lean_ctor_set_uint8(x_532, sizeof(void*)*10, x_531); +x_533 = l_Lean_Elab_Command_getRef(x_532, x_13, x_520); +x_534 = lean_ctor_get(x_533, 0); lean_inc(x_534); +x_535 = lean_ctor_get(x_533, 1); +lean_inc(x_535); if (lean_is_exclusive(x_533)) { lean_ctor_release(x_533, 0); lean_ctor_release(x_533, 1); - x_535 = x_533; + x_536 = x_533; } else { lean_dec_ref(x_533); - x_535 = lean_box(0); -} -x_536 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_529); -if (lean_is_scalar(x_535)) { - x_537 = lean_alloc_ctor(2, 2, 0); + x_536 = lean_box(0); +} +x_537 = l_Lean_SourceInfo_fromRef(x_534, x_29); +lean_dec(x_534); +x_538 = l_Lean_Elab_Command_getCurrMacroScope(x_532, x_13, x_535); +lean_dec(x_532); +x_539 = lean_ctor_get(x_538, 1); +lean_inc(x_539); +if (lean_is_exclusive(x_538)) { + lean_ctor_release(x_538, 0); + lean_ctor_release(x_538, 1); + x_540 = x_538; +} else { + lean_dec_ref(x_538); + x_540 = lean_box(0); +} +x_541 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_539); +x_542 = lean_ctor_get(x_541, 1); +lean_inc(x_542); +if (lean_is_exclusive(x_541)) { + lean_ctor_release(x_541, 0); + lean_ctor_release(x_541, 1); + x_543 = x_541; +} else { + lean_dec_ref(x_541); + x_543 = lean_box(0); +} +x_544 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_537); +if (lean_is_scalar(x_543)) { + x_545 = lean_alloc_ctor(2, 2, 0); +} else { + x_545 = x_543; + lean_ctor_set_tag(x_545, 2); +} +lean_ctor_set(x_545, 0, x_537); +lean_ctor_set(x_545, 1, x_544); +x_546 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_547 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_537); +x_548 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_548, 0, x_537); +lean_ctor_set(x_548, 1, x_546); +lean_ctor_set(x_548, 2, x_547); +x_549 = lean_array_size(x_23); +x_550 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_549, x_17, x_23); +x_551 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_552 = l_Lean_mkSepArray(x_550, x_551); +lean_dec(x_550); +x_553 = l_Array_append___rarg(x_547, x_552); +lean_dec(x_552); +lean_inc(x_537); +x_554 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_554, 0, x_537); +lean_ctor_set(x_554, 1, x_546); +lean_ctor_set(x_554, 2, x_553); +x_555 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_537); +x_556 = l_Lean_Syntax_node1(x_537, x_555, x_554); +x_557 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_548); +lean_inc(x_537); +x_558 = l_Lean_Syntax_node1(x_537, x_557, x_548); +x_559 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_537); +if (lean_is_scalar(x_540)) { + x_560 = lean_alloc_ctor(2, 2, 0); +} else { + x_560 = x_540; + lean_ctor_set_tag(x_560, 2); +} +lean_ctor_set(x_560, 0, x_537); +lean_ctor_set(x_560, 1, x_559); +x_561 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_548); +x_562 = l_Lean_Syntax_node6(x_537, x_561, x_545, x_548, x_556, x_558, x_548, x_560); +x_563 = l_Lean_Elab_Command_getRef(x_12, x_13, x_542); +x_564 = lean_ctor_get(x_563, 0); +lean_inc(x_564); +x_565 = lean_ctor_get(x_563, 1); +lean_inc(x_565); +if (lean_is_exclusive(x_563)) { + lean_ctor_release(x_563, 0); + lean_ctor_release(x_563, 1); + x_566 = x_563; +} else { + lean_dec_ref(x_563); + x_566 = lean_box(0); +} +x_567 = l_Lean_SourceInfo_fromRef(x_564, x_29); +lean_dec(x_564); +x_568 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_565); +x_569 = lean_ctor_get(x_568, 1); +lean_inc(x_569); +if (lean_is_exclusive(x_568)) { + lean_ctor_release(x_568, 0); + lean_ctor_release(x_568, 1); + x_570 = x_568; } else { - x_537 = x_535; - lean_ctor_set_tag(x_537, 2); -} -lean_ctor_set(x_537, 0, x_529); -lean_ctor_set(x_537, 1, x_536); -x_538 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_539 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_529); -x_540 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_540, 0, x_529); -lean_ctor_set(x_540, 1, x_538); -lean_ctor_set(x_540, 2, x_539); -x_541 = lean_array_size(x_23); -x_542 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_541, x_17, x_23); -x_543 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_544 = l_Lean_mkSepArray(x_542, x_543); -lean_dec(x_542); -x_545 = l_Array_append___rarg(x_539, x_544); -lean_dec(x_544); -lean_inc(x_529); -x_546 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_546, 0, x_529); -lean_ctor_set(x_546, 1, x_538); -lean_ctor_set(x_546, 2, x_545); -x_547 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_540); -lean_inc(x_529); -x_548 = l_Lean_Syntax_node1(x_529, x_547, x_540); -x_549 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_529); -if (lean_is_scalar(x_532)) { - x_550 = lean_alloc_ctor(2, 2, 0); -} else { - x_550 = x_532; - lean_ctor_set_tag(x_550, 2); -} -lean_ctor_set(x_550, 0, x_529); -lean_ctor_set(x_550, 1, x_549); -x_551 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_540); -x_552 = l_Lean_Syntax_node6(x_529, x_551, x_537, x_540, x_546, x_548, x_540, x_550); -x_553 = l_Lean_Elab_Command_getRef(x_12, x_13, x_534); -x_554 = lean_ctor_get(x_553, 0); -lean_inc(x_554); -x_555 = lean_ctor_get(x_553, 1); -lean_inc(x_555); -if (lean_is_exclusive(x_553)) { - lean_ctor_release(x_553, 0); - lean_ctor_release(x_553, 1); - x_556 = x_553; -} else { - lean_dec_ref(x_553); - x_556 = lean_box(0); -} -x_557 = l_Lean_SourceInfo_fromRef(x_554, x_29); -lean_dec(x_554); -x_558 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_555); -x_559 = lean_ctor_get(x_558, 1); -lean_inc(x_559); -if (lean_is_exclusive(x_558)) { - lean_ctor_release(x_558, 0); - lean_ctor_release(x_558, 1); - x_560 = x_558; -} else { - lean_dec_ref(x_558); - x_560 = lean_box(0); -} -x_561 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_559); -x_562 = lean_ctor_get(x_561, 1); -lean_inc(x_562); -if (lean_is_exclusive(x_561)) { - lean_ctor_release(x_561, 0); - lean_ctor_release(x_561, 1); - x_563 = x_561; -} else { - lean_dec_ref(x_561); - x_563 = lean_box(0); -} -x_564 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_557); -if (lean_is_scalar(x_563)) { - x_565 = lean_alloc_ctor(2, 2, 0); -} else { - x_565 = x_563; - lean_ctor_set_tag(x_565, 2); -} -lean_ctor_set(x_565, 0, x_557); -lean_ctor_set(x_565, 1, x_564); -x_566 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_567 = l_Lean_Syntax_SepArray_ofElems(x_566, x_5); -x_568 = l_Array_append___rarg(x_539, x_567); -lean_dec(x_567); -lean_inc(x_557); -x_569 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_569, 0, x_557); -lean_ctor_set(x_569, 1, x_538); -lean_ctor_set(x_569, 2, x_568); -x_570 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_557); -if (lean_is_scalar(x_560)) { - x_571 = lean_alloc_ctor(2, 2, 0); + lean_dec_ref(x_568); + x_570 = lean_box(0); +} +x_571 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_569); +x_572 = lean_ctor_get(x_571, 1); +lean_inc(x_572); +if (lean_is_exclusive(x_571)) { + lean_ctor_release(x_571, 0); + lean_ctor_release(x_571, 1); + x_573 = x_571; } else { - x_571 = x_560; - lean_ctor_set_tag(x_571, 2); + lean_dec_ref(x_571); + x_573 = lean_box(0); } -lean_ctor_set(x_571, 0, x_557); -lean_ctor_set(x_571, 1, x_570); -x_572 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_557); -x_573 = l_Lean_Syntax_node3(x_557, x_572, x_565, x_569, x_571); -lean_inc(x_557); -x_574 = l_Lean_Syntax_node1(x_557, x_538, x_573); -lean_inc(x_557); -x_575 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_575, 0, x_557); -lean_ctor_set(x_575, 1, x_538); -lean_ctor_set(x_575, 2, x_539); -x_576 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_557); -if (lean_is_scalar(x_556)) { - x_577 = lean_alloc_ctor(2, 2, 0); +x_574 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_567); +if (lean_is_scalar(x_573)) { + x_575 = lean_alloc_ctor(2, 2, 0); +} else { + x_575 = x_573; + lean_ctor_set_tag(x_575, 2); +} +lean_ctor_set(x_575, 0, x_567); +lean_ctor_set(x_575, 1, x_574); +x_576 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_577 = l_Lean_Syntax_SepArray_ofElems(x_576, x_5); +x_578 = l_Array_append___rarg(x_547, x_577); +lean_dec(x_577); +lean_inc(x_567); +x_579 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_579, 0, x_567); +lean_ctor_set(x_579, 1, x_546); +lean_ctor_set(x_579, 2, x_578); +x_580 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_567); +if (lean_is_scalar(x_570)) { + x_581 = lean_alloc_ctor(2, 2, 0); } else { - x_577 = x_556; - lean_ctor_set_tag(x_577, 2); + x_581 = x_570; + lean_ctor_set_tag(x_581, 2); } -lean_ctor_set(x_577, 0, x_557); -lean_ctor_set(x_577, 1, x_576); -x_578 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_557); -if (lean_is_scalar(x_528)) { - x_579 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_581, 0, x_567); +lean_ctor_set(x_581, 1, x_580); +x_582 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_567); +x_583 = l_Lean_Syntax_node3(x_567, x_582, x_575, x_579, x_581); +lean_inc(x_567); +x_584 = l_Lean_Syntax_node1(x_567, x_546, x_583); +lean_inc(x_567); +x_585 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_585, 0, x_567); +lean_ctor_set(x_585, 1, x_546); +lean_ctor_set(x_585, 2, x_547); +x_586 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_567); +if (lean_is_scalar(x_566)) { + x_587 = lean_alloc_ctor(2, 2, 0); } else { - x_579 = x_528; - lean_ctor_set_tag(x_579, 2); + x_587 = x_566; + lean_ctor_set_tag(x_587, 2); } -lean_ctor_set(x_579, 0, x_557); -lean_ctor_set(x_579, 1, x_578); -x_580 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_557); -x_581 = l_Lean_Syntax_node2(x_557, x_580, x_579, x_30); -lean_inc(x_557); -x_582 = l_Lean_Syntax_node1(x_557, x_538, x_581); -x_583 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_575); -lean_inc(x_557); -x_584 = l_Lean_Syntax_node2(x_557, x_583, x_575, x_582); -x_585 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_557); -x_586 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_586, 0, x_557); -lean_ctor_set(x_586, 1, x_585); -x_587 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_575, 2); -lean_inc(x_557); -x_588 = l_Lean_Syntax_node2(x_557, x_587, x_575, x_575); -x_589 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_557); -x_590 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_590, 0, x_557); -lean_ctor_set(x_590, 1, x_538); -lean_ctor_set(x_590, 2, x_589); -x_591 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_590); -lean_inc(x_557); -x_592 = l_Lean_Syntax_node4(x_557, x_591, x_586, x_552, x_588, x_590); -x_593 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_557); -x_594 = l_Lean_Syntax_node4(x_557, x_593, x_577, x_31, x_584, x_592); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; -x_595 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_575, 3); -lean_inc(x_557); -x_596 = l_Lean_Syntax_node6(x_557, x_595, x_590, x_574, x_575, x_575, x_575, x_575); -x_597 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_598 = l_Lean_Syntax_node2(x_557, x_597, x_596, x_594); -lean_inc(x_598); -x_599 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_599, 0, x_598); -x_600 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_598, x_599, x_12, x_13, x_562); -return x_600; +lean_ctor_set(x_587, 0, x_567); +lean_ctor_set(x_587, 1, x_586); +x_588 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_567); +if (lean_is_scalar(x_536)) { + x_589 = lean_alloc_ctor(2, 2, 0); +} else { + x_589 = x_536; + lean_ctor_set_tag(x_589, 2); } -else +lean_ctor_set(x_589, 0, x_567); +lean_ctor_set(x_589, 1, x_588); +x_590 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_567); +x_591 = l_Lean_Syntax_node2(x_567, x_590, x_589, x_30); +lean_inc(x_567); +x_592 = l_Lean_Syntax_node1(x_567, x_546, x_591); +x_593 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_585); +lean_inc(x_567); +x_594 = l_Lean_Syntax_node2(x_567, x_593, x_585, x_592); +x_595 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_567); +x_596 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_596, 0, x_567); +lean_ctor_set(x_596, 1, x_595); +x_597 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_585, 2); +lean_inc(x_567); +x_598 = l_Lean_Syntax_node2(x_567, x_597, x_585, x_585); +x_599 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_567); +x_600 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_600, 0, x_567); +lean_ctor_set(x_600, 1, x_546); +lean_ctor_set(x_600, 2, x_599); +x_601 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_600); +lean_inc(x_567); +x_602 = l_Lean_Syntax_node4(x_567, x_601, x_596, x_562, x_598, x_600); +x_603 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_567); +x_604 = l_Lean_Syntax_node4(x_567, x_603, x_587, x_31, x_594, x_602); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; -lean_dec(x_590); -x_601 = lean_ctor_get(x_6, 0); -lean_inc(x_601); -lean_dec(x_6); -x_602 = l_Array_mkArray1___rarg(x_601); -x_603 = l_Array_append___rarg(x_539, x_602); -lean_dec(x_602); -lean_inc(x_557); -x_604 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_604, 0, x_557); -lean_ctor_set(x_604, 1, x_538); -lean_ctor_set(x_604, 2, x_603); -x_605 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_575, 3); -lean_inc(x_557); -x_606 = l_Lean_Syntax_node6(x_557, x_605, x_604, x_574, x_575, x_575, x_575, x_575); -x_607 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_608 = l_Lean_Syntax_node2(x_557, x_607, x_606, x_594); +lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; +x_605 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_585, 3); +lean_inc(x_567); +x_606 = l_Lean_Syntax_node6(x_567, x_605, x_600, x_584, x_585, x_585, x_585, x_585); +x_607 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_608 = l_Lean_Syntax_node2(x_567, x_607, x_606, x_604); lean_inc(x_608); x_609 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); lean_closure_set(x_609, 0, x_608); -x_610 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_608, x_609, x_12, x_13, x_562); +x_610 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_608, x_609, x_12, x_13, x_572); return x_610; } +else +{ +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; +lean_dec(x_600); +x_611 = lean_ctor_get(x_6, 0); +lean_inc(x_611); +lean_dec(x_6); +x_612 = l_Array_mkArray1___rarg(x_611); +x_613 = l_Array_append___rarg(x_547, x_612); +lean_dec(x_612); +lean_inc(x_567); +x_614 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_614, 0, x_567); +lean_ctor_set(x_614, 1, x_546); +lean_ctor_set(x_614, 2, x_613); +x_615 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_585, 3); +lean_inc(x_567); +x_616 = l_Lean_Syntax_node6(x_567, x_615, x_614, x_584, x_585, x_585, x_585, x_585); +x_617 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_618 = l_Lean_Syntax_node2(x_567, x_617, x_616, x_604); +lean_inc(x_618); +x_619 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_619, 0, x_618); +x_620 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_618, x_619, x_12, x_13, x_572); +return x_620; +} } } } else { -lean_object* x_626; lean_object* x_627; uint8_t x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; -x_626 = lean_ctor_get(x_25, 0); -x_627 = lean_ctor_get(x_25, 1); -lean_inc(x_627); -lean_inc(x_626); +lean_object* x_636; lean_object* x_637; uint8_t x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; +x_636 = lean_ctor_get(x_25, 0); +x_637 = lean_ctor_get(x_25, 1); +lean_inc(x_637); +lean_inc(x_636); lean_dec(x_25); -x_628 = 0; -x_629 = l_Lean_mkCIdentFrom(x_626, x_2, x_628); -lean_dec(x_626); +x_638 = 0; +x_639 = l_Lean_mkCIdentFrom(x_636, x_2, x_638); +lean_dec(x_636); if (lean_obj_tag(x_8) == 0) { -lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; -x_735 = lean_box(2); -x_736 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_737 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_737, 0, x_735); -lean_ctor_set(x_737, 1, x_736); -lean_ctor_set(x_737, 2, x_1); -x_738 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_738, 0, x_737); -lean_ctor_set(x_738, 1, x_3); +lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; +x_747 = lean_box(2); +x_748 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_749 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_749, 0, x_747); +lean_ctor_set(x_749, 1, x_748); +lean_ctor_set(x_749, 2, x_1); +x_750 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_750, 0, x_749); +lean_ctor_set(x_750, 1, x_3); lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 1, x_738); +lean_ctor_set(x_21, 1, x_750); lean_ctor_set(x_21, 0, x_9); -x_739 = lean_array_mk(x_21); -x_740 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_741 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_741, 0, x_735); -lean_ctor_set(x_741, 1, x_740); -lean_ctor_set(x_741, 2, x_739); -x_630 = x_741; -x_631 = x_627; -goto block_734; +x_751 = lean_array_mk(x_21); +x_752 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_753 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_753, 0, x_747); +lean_ctor_set(x_753, 1, x_752); +lean_ctor_set(x_753, 2, x_751); +x_640 = x_753; +x_641 = x_637; +goto block_746; } else { -lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; -x_742 = lean_ctor_get(x_8, 0); -lean_inc(x_742); +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; +x_754 = lean_ctor_get(x_8, 0); +lean_inc(x_754); lean_dec(x_8); -x_743 = l_Lean_mkIdentFrom(x_9, x_742, x_628); +x_755 = l_Lean_mkIdentFrom(x_9, x_754, x_638); lean_dec(x_9); -x_744 = lean_box(2); -x_745 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_746 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_746, 0, x_744); -lean_ctor_set(x_746, 1, x_745); -lean_ctor_set(x_746, 2, x_1); -x_747 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_747, 0, x_746); -lean_ctor_set(x_747, 1, x_3); +x_756 = lean_box(2); +x_757 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_758 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_758, 0, x_756); +lean_ctor_set(x_758, 1, x_757); +lean_ctor_set(x_758, 2, x_1); +x_759 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_759, 0, x_758); +lean_ctor_set(x_759, 1, x_3); lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 1, x_747); -lean_ctor_set(x_21, 0, x_743); -x_748 = lean_array_mk(x_21); -x_749 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_750 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_750, 0, x_744); -lean_ctor_set(x_750, 1, x_749); -lean_ctor_set(x_750, 2, x_748); -x_630 = x_750; -x_631 = x_627; -goto block_734; -} -block_734: -{ -lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; uint8_t x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; size_t x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; -x_632 = l_Lean_Elab_Command_getRef(x_12, x_13, x_631); -x_633 = lean_ctor_get(x_632, 0); -lean_inc(x_633); -x_634 = lean_ctor_get(x_632, 1); -lean_inc(x_634); -if (lean_is_exclusive(x_632)) { - lean_ctor_release(x_632, 0); - lean_ctor_release(x_632, 1); - x_635 = x_632; -} else { - lean_dec_ref(x_632); - x_635 = lean_box(0); -} -x_636 = l_Lean_replaceRef(x_4, x_633); -lean_dec(x_633); -x_637 = lean_ctor_get(x_12, 0); -lean_inc(x_637); -x_638 = lean_ctor_get(x_12, 1); -lean_inc(x_638); -x_639 = lean_ctor_get(x_12, 2); -lean_inc(x_639); -x_640 = lean_ctor_get(x_12, 3); -lean_inc(x_640); -x_641 = lean_ctor_get(x_12, 4); -lean_inc(x_641); -x_642 = lean_ctor_get(x_12, 5); -lean_inc(x_642); -x_643 = lean_ctor_get(x_12, 7); +lean_ctor_set(x_21, 1, x_759); +lean_ctor_set(x_21, 0, x_755); +x_760 = lean_array_mk(x_21); +x_761 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_762 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_762, 0, x_756); +lean_ctor_set(x_762, 1, x_761); +lean_ctor_set(x_762, 2, x_760); +x_640 = x_762; +x_641 = x_637; +goto block_746; +} +block_746: +{ +lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; uint8_t x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; size_t x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; +x_642 = l_Lean_Elab_Command_getRef(x_12, x_13, x_641); +x_643 = lean_ctor_get(x_642, 0); lean_inc(x_643); -x_644 = lean_ctor_get(x_12, 8); +x_644 = lean_ctor_get(x_642, 1); lean_inc(x_644); -x_645 = lean_ctor_get(x_12, 9); -lean_inc(x_645); -x_646 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); -x_647 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_647, 0, x_637); -lean_ctor_set(x_647, 1, x_638); -lean_ctor_set(x_647, 2, x_639); -lean_ctor_set(x_647, 3, x_640); -lean_ctor_set(x_647, 4, x_641); -lean_ctor_set(x_647, 5, x_642); -lean_ctor_set(x_647, 6, x_636); -lean_ctor_set(x_647, 7, x_643); -lean_ctor_set(x_647, 8, x_644); -lean_ctor_set(x_647, 9, x_645); -lean_ctor_set_uint8(x_647, sizeof(void*)*10, x_646); -x_648 = l_Lean_Elab_Command_getRef(x_647, x_13, x_634); -x_649 = lean_ctor_get(x_648, 0); +if (lean_is_exclusive(x_642)) { + lean_ctor_release(x_642, 0); + lean_ctor_release(x_642, 1); + x_645 = x_642; +} else { + lean_dec_ref(x_642); + x_645 = lean_box(0); +} +x_646 = l_Lean_replaceRef(x_4, x_643); +lean_dec(x_643); +x_647 = lean_ctor_get(x_12, 0); +lean_inc(x_647); +x_648 = lean_ctor_get(x_12, 1); +lean_inc(x_648); +x_649 = lean_ctor_get(x_12, 2); lean_inc(x_649); -x_650 = lean_ctor_get(x_648, 1); +x_650 = lean_ctor_get(x_12, 3); lean_inc(x_650); -if (lean_is_exclusive(x_648)) { - lean_ctor_release(x_648, 0); - lean_ctor_release(x_648, 1); - x_651 = x_648; -} else { - lean_dec_ref(x_648); - x_651 = lean_box(0); -} -x_652 = l_Lean_SourceInfo_fromRef(x_649, x_628); -lean_dec(x_649); -x_653 = l_Lean_Elab_Command_getCurrMacroScope(x_647, x_13, x_650); -lean_dec(x_647); -x_654 = lean_ctor_get(x_653, 1); -lean_inc(x_654); -if (lean_is_exclusive(x_653)) { - lean_ctor_release(x_653, 0); - lean_ctor_release(x_653, 1); - x_655 = x_653; -} else { - lean_dec_ref(x_653); - x_655 = lean_box(0); -} -x_656 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_654); -x_657 = lean_ctor_get(x_656, 1); -lean_inc(x_657); -if (lean_is_exclusive(x_656)) { - lean_ctor_release(x_656, 0); - lean_ctor_release(x_656, 1); - x_658 = x_656; -} else { - lean_dec_ref(x_656); - x_658 = lean_box(0); -} -x_659 = l_Lake_DSL_structVal___closed__3; +x_651 = lean_ctor_get(x_12, 4); +lean_inc(x_651); +x_652 = lean_ctor_get(x_12, 5); lean_inc(x_652); -if (lean_is_scalar(x_658)) { - x_660 = lean_alloc_ctor(2, 2, 0); +x_653 = lean_ctor_get(x_12, 7); +lean_inc(x_653); +x_654 = lean_ctor_get(x_12, 8); +lean_inc(x_654); +x_655 = lean_ctor_get(x_12, 9); +lean_inc(x_655); +x_656 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); +x_657 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_657, 0, x_647); +lean_ctor_set(x_657, 1, x_648); +lean_ctor_set(x_657, 2, x_649); +lean_ctor_set(x_657, 3, x_650); +lean_ctor_set(x_657, 4, x_651); +lean_ctor_set(x_657, 5, x_652); +lean_ctor_set(x_657, 6, x_646); +lean_ctor_set(x_657, 7, x_653); +lean_ctor_set(x_657, 8, x_654); +lean_ctor_set(x_657, 9, x_655); +lean_ctor_set_uint8(x_657, sizeof(void*)*10, x_656); +x_658 = l_Lean_Elab_Command_getRef(x_657, x_13, x_644); +x_659 = lean_ctor_get(x_658, 0); +lean_inc(x_659); +x_660 = lean_ctor_get(x_658, 1); +lean_inc(x_660); +if (lean_is_exclusive(x_658)) { + lean_ctor_release(x_658, 0); + lean_ctor_release(x_658, 1); + x_661 = x_658; +} else { + lean_dec_ref(x_658); + x_661 = lean_box(0); +} +x_662 = l_Lean_SourceInfo_fromRef(x_659, x_638); +lean_dec(x_659); +x_663 = l_Lean_Elab_Command_getCurrMacroScope(x_657, x_13, x_660); +lean_dec(x_657); +x_664 = lean_ctor_get(x_663, 1); +lean_inc(x_664); +if (lean_is_exclusive(x_663)) { + lean_ctor_release(x_663, 0); + lean_ctor_release(x_663, 1); + x_665 = x_663; +} else { + lean_dec_ref(x_663); + x_665 = lean_box(0); +} +x_666 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_664); +x_667 = lean_ctor_get(x_666, 1); +lean_inc(x_667); +if (lean_is_exclusive(x_666)) { + lean_ctor_release(x_666, 0); + lean_ctor_release(x_666, 1); + x_668 = x_666; +} else { + lean_dec_ref(x_666); + x_668 = lean_box(0); +} +x_669 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_662); +if (lean_is_scalar(x_668)) { + x_670 = lean_alloc_ctor(2, 2, 0); +} else { + x_670 = x_668; + lean_ctor_set_tag(x_670, 2); +} +lean_ctor_set(x_670, 0, x_662); +lean_ctor_set(x_670, 1, x_669); +x_671 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_672 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_662); +x_673 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_673, 0, x_662); +lean_ctor_set(x_673, 1, x_671); +lean_ctor_set(x_673, 2, x_672); +x_674 = lean_array_size(x_23); +x_675 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_674, x_17, x_23); +x_676 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_677 = l_Lean_mkSepArray(x_675, x_676); +lean_dec(x_675); +x_678 = l_Array_append___rarg(x_672, x_677); +lean_dec(x_677); +lean_inc(x_662); +x_679 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_679, 0, x_662); +lean_ctor_set(x_679, 1, x_671); +lean_ctor_set(x_679, 2, x_678); +x_680 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_662); +x_681 = l_Lean_Syntax_node1(x_662, x_680, x_679); +x_682 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_673); +lean_inc(x_662); +x_683 = l_Lean_Syntax_node1(x_662, x_682, x_673); +x_684 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_662); +if (lean_is_scalar(x_665)) { + x_685 = lean_alloc_ctor(2, 2, 0); +} else { + x_685 = x_665; + lean_ctor_set_tag(x_685, 2); +} +lean_ctor_set(x_685, 0, x_662); +lean_ctor_set(x_685, 1, x_684); +x_686 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_673); +x_687 = l_Lean_Syntax_node6(x_662, x_686, x_670, x_673, x_681, x_683, x_673, x_685); +x_688 = l_Lean_Elab_Command_getRef(x_12, x_13, x_667); +x_689 = lean_ctor_get(x_688, 0); +lean_inc(x_689); +x_690 = lean_ctor_get(x_688, 1); +lean_inc(x_690); +if (lean_is_exclusive(x_688)) { + lean_ctor_release(x_688, 0); + lean_ctor_release(x_688, 1); + x_691 = x_688; +} else { + lean_dec_ref(x_688); + x_691 = lean_box(0); +} +x_692 = l_Lean_SourceInfo_fromRef(x_689, x_638); +lean_dec(x_689); +x_693 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_690); +x_694 = lean_ctor_get(x_693, 1); +lean_inc(x_694); +if (lean_is_exclusive(x_693)) { + lean_ctor_release(x_693, 0); + lean_ctor_release(x_693, 1); + x_695 = x_693; } else { - x_660 = x_658; - lean_ctor_set_tag(x_660, 2); + lean_dec_ref(x_693); + x_695 = lean_box(0); +} +x_696 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_694); +x_697 = lean_ctor_get(x_696, 1); +lean_inc(x_697); +if (lean_is_exclusive(x_696)) { + lean_ctor_release(x_696, 0); + lean_ctor_release(x_696, 1); + x_698 = x_696; +} else { + lean_dec_ref(x_696); + x_698 = lean_box(0); } -lean_ctor_set(x_660, 0, x_652); -lean_ctor_set(x_660, 1, x_659); -x_661 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_662 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_652); -x_663 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_663, 0, x_652); -lean_ctor_set(x_663, 1, x_661); -lean_ctor_set(x_663, 2, x_662); -x_664 = lean_array_size(x_23); -x_665 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_664, x_17, x_23); -x_666 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_667 = l_Lean_mkSepArray(x_665, x_666); -lean_dec(x_665); -x_668 = l_Array_append___rarg(x_662, x_667); -lean_dec(x_667); -lean_inc(x_652); -x_669 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_669, 0, x_652); -lean_ctor_set(x_669, 1, x_661); -lean_ctor_set(x_669, 2, x_668); -x_670 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_663); -lean_inc(x_652); -x_671 = l_Lean_Syntax_node1(x_652, x_670, x_663); -x_672 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_652); -if (lean_is_scalar(x_655)) { - x_673 = lean_alloc_ctor(2, 2, 0); -} else { - x_673 = x_655; - lean_ctor_set_tag(x_673, 2); -} -lean_ctor_set(x_673, 0, x_652); -lean_ctor_set(x_673, 1, x_672); -x_674 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_663); -x_675 = l_Lean_Syntax_node6(x_652, x_674, x_660, x_663, x_669, x_671, x_663, x_673); -x_676 = l_Lean_Elab_Command_getRef(x_12, x_13, x_657); -x_677 = lean_ctor_get(x_676, 0); -lean_inc(x_677); -x_678 = lean_ctor_get(x_676, 1); -lean_inc(x_678); -if (lean_is_exclusive(x_676)) { - lean_ctor_release(x_676, 0); - lean_ctor_release(x_676, 1); - x_679 = x_676; -} else { - lean_dec_ref(x_676); - x_679 = lean_box(0); -} -x_680 = l_Lean_SourceInfo_fromRef(x_677, x_628); -lean_dec(x_677); -x_681 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_678); -x_682 = lean_ctor_get(x_681, 1); -lean_inc(x_682); -if (lean_is_exclusive(x_681)) { - lean_ctor_release(x_681, 0); - lean_ctor_release(x_681, 1); - x_683 = x_681; -} else { - lean_dec_ref(x_681); - x_683 = lean_box(0); -} -x_684 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_682); -x_685 = lean_ctor_get(x_684, 1); -lean_inc(x_685); -if (lean_is_exclusive(x_684)) { - lean_ctor_release(x_684, 0); - lean_ctor_release(x_684, 1); - x_686 = x_684; -} else { - lean_dec_ref(x_684); - x_686 = lean_box(0); -} -x_687 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_680); -if (lean_is_scalar(x_686)) { - x_688 = lean_alloc_ctor(2, 2, 0); -} else { - x_688 = x_686; - lean_ctor_set_tag(x_688, 2); -} -lean_ctor_set(x_688, 0, x_680); -lean_ctor_set(x_688, 1, x_687); -x_689 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_690 = l_Lean_Syntax_SepArray_ofElems(x_689, x_5); -x_691 = l_Array_append___rarg(x_662, x_690); -lean_dec(x_690); -lean_inc(x_680); -x_692 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_692, 0, x_680); -lean_ctor_set(x_692, 1, x_661); -lean_ctor_set(x_692, 2, x_691); -x_693 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_680); -if (lean_is_scalar(x_683)) { - x_694 = lean_alloc_ctor(2, 2, 0); -} else { - x_694 = x_683; - lean_ctor_set_tag(x_694, 2); -} -lean_ctor_set(x_694, 0, x_680); -lean_ctor_set(x_694, 1, x_693); -x_695 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_680); -x_696 = l_Lean_Syntax_node3(x_680, x_695, x_688, x_692, x_694); -lean_inc(x_680); -x_697 = l_Lean_Syntax_node1(x_680, x_661, x_696); -lean_inc(x_680); -x_698 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_698, 0, x_680); -lean_ctor_set(x_698, 1, x_661); -lean_ctor_set(x_698, 2, x_662); x_699 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_680); -if (lean_is_scalar(x_679)) { +lean_inc(x_692); +if (lean_is_scalar(x_698)) { x_700 = lean_alloc_ctor(2, 2, 0); } else { - x_700 = x_679; + x_700 = x_698; lean_ctor_set_tag(x_700, 2); } -lean_ctor_set(x_700, 0, x_680); +lean_ctor_set(x_700, 0, x_692); lean_ctor_set(x_700, 1, x_699); -x_701 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_680); -if (lean_is_scalar(x_651)) { - x_702 = lean_alloc_ctor(2, 2, 0); +x_701 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_702 = l_Lean_Syntax_SepArray_ofElems(x_701, x_5); +x_703 = l_Array_append___rarg(x_672, x_702); +lean_dec(x_702); +lean_inc(x_692); +x_704 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_704, 0, x_692); +lean_ctor_set(x_704, 1, x_671); +lean_ctor_set(x_704, 2, x_703); +x_705 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_692); +if (lean_is_scalar(x_695)) { + x_706 = lean_alloc_ctor(2, 2, 0); } else { - x_702 = x_651; - lean_ctor_set_tag(x_702, 2); + x_706 = x_695; + lean_ctor_set_tag(x_706, 2); } -lean_ctor_set(x_702, 0, x_680); -lean_ctor_set(x_702, 1, x_701); -x_703 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_680); -x_704 = l_Lean_Syntax_node2(x_680, x_703, x_702, x_629); -lean_inc(x_680); -x_705 = l_Lean_Syntax_node1(x_680, x_661, x_704); -x_706 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_698); -lean_inc(x_680); -x_707 = l_Lean_Syntax_node2(x_680, x_706, x_698, x_705); -x_708 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_680); -if (lean_is_scalar(x_635)) { - x_709 = lean_alloc_ctor(2, 2, 0); -} else { - x_709 = x_635; - lean_ctor_set_tag(x_709, 2); -} -lean_ctor_set(x_709, 0, x_680); -lean_ctor_set(x_709, 1, x_708); -x_710 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_698, 2); -lean_inc(x_680); -x_711 = l_Lean_Syntax_node2(x_680, x_710, x_698, x_698); -x_712 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_680); -x_713 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_713, 0, x_680); -lean_ctor_set(x_713, 1, x_661); -lean_ctor_set(x_713, 2, x_712); -x_714 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_713); -lean_inc(x_680); -x_715 = l_Lean_Syntax_node4(x_680, x_714, x_709, x_675, x_711, x_713); -x_716 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_680); -x_717 = l_Lean_Syntax_node4(x_680, x_716, x_700, x_630, x_707, x_715); +lean_ctor_set(x_706, 0, x_692); +lean_ctor_set(x_706, 1, x_705); +x_707 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_692); +x_708 = l_Lean_Syntax_node3(x_692, x_707, x_700, x_704, x_706); +lean_inc(x_692); +x_709 = l_Lean_Syntax_node1(x_692, x_671, x_708); +lean_inc(x_692); +x_710 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_710, 0, x_692); +lean_ctor_set(x_710, 1, x_671); +lean_ctor_set(x_710, 2, x_672); +x_711 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_692); +if (lean_is_scalar(x_691)) { + x_712 = lean_alloc_ctor(2, 2, 0); +} else { + x_712 = x_691; + lean_ctor_set_tag(x_712, 2); +} +lean_ctor_set(x_712, 0, x_692); +lean_ctor_set(x_712, 1, x_711); +x_713 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_692); +if (lean_is_scalar(x_661)) { + x_714 = lean_alloc_ctor(2, 2, 0); +} else { + x_714 = x_661; + lean_ctor_set_tag(x_714, 2); +} +lean_ctor_set(x_714, 0, x_692); +lean_ctor_set(x_714, 1, x_713); +x_715 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_692); +x_716 = l_Lean_Syntax_node2(x_692, x_715, x_714, x_639); +lean_inc(x_692); +x_717 = l_Lean_Syntax_node1(x_692, x_671, x_716); +x_718 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_710); +lean_inc(x_692); +x_719 = l_Lean_Syntax_node2(x_692, x_718, x_710, x_717); +x_720 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_692); +if (lean_is_scalar(x_645)) { + x_721 = lean_alloc_ctor(2, 2, 0); +} else { + x_721 = x_645; + lean_ctor_set_tag(x_721, 2); +} +lean_ctor_set(x_721, 0, x_692); +lean_ctor_set(x_721, 1, x_720); +x_722 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_710, 2); +lean_inc(x_692); +x_723 = l_Lean_Syntax_node2(x_692, x_722, x_710, x_710); +x_724 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_692); +x_725 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_725, 0, x_692); +lean_ctor_set(x_725, 1, x_671); +lean_ctor_set(x_725, 2, x_724); +x_726 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_725); +lean_inc(x_692); +x_727 = l_Lean_Syntax_node4(x_692, x_726, x_721, x_687, x_723, x_725); +x_728 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_692); +x_729 = l_Lean_Syntax_node4(x_692, x_728, x_712, x_640, x_719, x_727); if (lean_obj_tag(x_6) == 0) { -lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; -x_718 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_698, 3); -lean_inc(x_680); -x_719 = l_Lean_Syntax_node6(x_680, x_718, x_713, x_697, x_698, x_698, x_698, x_698); -x_720 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_721 = l_Lean_Syntax_node2(x_680, x_720, x_719, x_717); -lean_inc(x_721); -x_722 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_722, 0, x_721); -x_723 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_721, x_722, x_12, x_13, x_685); -return x_723; -} -else -{ -lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; -lean_dec(x_713); -x_724 = lean_ctor_get(x_6, 0); -lean_inc(x_724); -lean_dec(x_6); -x_725 = l_Array_mkArray1___rarg(x_724); -x_726 = l_Array_append___rarg(x_662, x_725); +lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +x_730 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_710, 3); +lean_inc(x_692); +x_731 = l_Lean_Syntax_node6(x_692, x_730, x_725, x_709, x_710, x_710, x_710, x_710); +x_732 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_733 = l_Lean_Syntax_node2(x_692, x_732, x_731, x_729); +lean_inc(x_733); +x_734 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_734, 0, x_733); +x_735 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_733, x_734, x_12, x_13, x_697); +return x_735; +} +else +{ +lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_dec(x_725); -lean_inc(x_680); -x_727 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_727, 0, x_680); -lean_ctor_set(x_727, 1, x_661); -lean_ctor_set(x_727, 2, x_726); -x_728 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_698, 3); -lean_inc(x_680); -x_729 = l_Lean_Syntax_node6(x_680, x_728, x_727, x_697, x_698, x_698, x_698, x_698); -x_730 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_731 = l_Lean_Syntax_node2(x_680, x_730, x_729, x_717); -lean_inc(x_731); -x_732 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_732, 0, x_731); -x_733 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_731, x_732, x_12, x_13, x_685); -return x_733; +x_736 = lean_ctor_get(x_6, 0); +lean_inc(x_736); +lean_dec(x_6); +x_737 = l_Array_mkArray1___rarg(x_736); +x_738 = l_Array_append___rarg(x_672, x_737); +lean_dec(x_737); +lean_inc(x_692); +x_739 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_739, 0, x_692); +lean_ctor_set(x_739, 1, x_671); +lean_ctor_set(x_739, 2, x_738); +x_740 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_710, 3); +lean_inc(x_692); +x_741 = l_Lean_Syntax_node6(x_692, x_740, x_739, x_709, x_710, x_710, x_710, x_710); +x_742 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_743 = l_Lean_Syntax_node2(x_692, x_742, x_741, x_729); +lean_inc(x_743); +x_744 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_744, 0, x_743); +x_745 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_743, x_744, x_12, x_13, x_697); +return x_745; } } } } else { -lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; uint8_t x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; -x_751 = lean_ctor_get(x_21, 0); -x_752 = lean_ctor_get(x_21, 1); -lean_inc(x_752); -lean_inc(x_751); +lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; uint8_t x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; +x_763 = lean_ctor_get(x_21, 0); +x_764 = lean_ctor_get(x_21, 1); +lean_inc(x_764); +lean_inc(x_763); lean_dec(x_21); -x_753 = l_Lean_Elab_Command_getRef(x_12, x_13, x_752); -x_754 = lean_ctor_get(x_753, 0); -lean_inc(x_754); -x_755 = lean_ctor_get(x_753, 1); -lean_inc(x_755); -if (lean_is_exclusive(x_753)) { - lean_ctor_release(x_753, 0); - lean_ctor_release(x_753, 1); - x_756 = x_753; +x_765 = l_Lean_Elab_Command_getRef(x_12, x_13, x_764); +x_766 = lean_ctor_get(x_765, 0); +lean_inc(x_766); +x_767 = lean_ctor_get(x_765, 1); +lean_inc(x_767); +if (lean_is_exclusive(x_765)) { + lean_ctor_release(x_765, 0); + lean_ctor_release(x_765, 1); + x_768 = x_765; } else { - lean_dec_ref(x_753); - x_756 = lean_box(0); + lean_dec_ref(x_765); + x_768 = lean_box(0); } -x_757 = 0; -x_758 = l_Lean_mkCIdentFrom(x_754, x_2, x_757); -lean_dec(x_754); +x_769 = 0; +x_770 = l_Lean_mkCIdentFrom(x_766, x_2, x_769); +lean_dec(x_766); if (lean_obj_tag(x_8) == 0) { -lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; -x_864 = lean_box(2); -x_865 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_866 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_866, 0, x_864); -lean_ctor_set(x_866, 1, x_865); -lean_ctor_set(x_866, 2, x_1); -if (lean_is_scalar(x_756)) { - x_867 = lean_alloc_ctor(1, 2, 0); -} else { - x_867 = x_756; - lean_ctor_set_tag(x_867, 1); -} -lean_ctor_set(x_867, 0, x_866); -lean_ctor_set(x_867, 1, x_3); -x_868 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_868, 0, x_9); -lean_ctor_set(x_868, 1, x_867); -x_869 = lean_array_mk(x_868); -x_870 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_871 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_871, 0, x_864); -lean_ctor_set(x_871, 1, x_870); -lean_ctor_set(x_871, 2, x_869); -x_759 = x_871; -x_760 = x_755; -goto block_863; +lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; +x_878 = lean_box(2); +x_879 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_880 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_880, 0, x_878); +lean_ctor_set(x_880, 1, x_879); +lean_ctor_set(x_880, 2, x_1); +if (lean_is_scalar(x_768)) { + x_881 = lean_alloc_ctor(1, 2, 0); +} else { + x_881 = x_768; + lean_ctor_set_tag(x_881, 1); +} +lean_ctor_set(x_881, 0, x_880); +lean_ctor_set(x_881, 1, x_3); +x_882 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_882, 0, x_9); +lean_ctor_set(x_882, 1, x_881); +x_883 = lean_array_mk(x_882); +x_884 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_885 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_885, 0, x_878); +lean_ctor_set(x_885, 1, x_884); +lean_ctor_set(x_885, 2, x_883); +x_771 = x_885; +x_772 = x_767; +goto block_877; } else { -lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; -x_872 = lean_ctor_get(x_8, 0); -lean_inc(x_872); +lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; +x_886 = lean_ctor_get(x_8, 0); +lean_inc(x_886); lean_dec(x_8); -x_873 = l_Lean_mkIdentFrom(x_9, x_872, x_757); +x_887 = l_Lean_mkIdentFrom(x_9, x_886, x_769); lean_dec(x_9); -x_874 = lean_box(2); -x_875 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_876 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_876, 0, x_874); -lean_ctor_set(x_876, 1, x_875); -lean_ctor_set(x_876, 2, x_1); -if (lean_is_scalar(x_756)) { - x_877 = lean_alloc_ctor(1, 2, 0); -} else { - x_877 = x_756; - lean_ctor_set_tag(x_877, 1); -} -lean_ctor_set(x_877, 0, x_876); -lean_ctor_set(x_877, 1, x_3); -x_878 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_878, 0, x_873); -lean_ctor_set(x_878, 1, x_877); -x_879 = lean_array_mk(x_878); -x_880 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; -x_881 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_881, 0, x_874); -lean_ctor_set(x_881, 1, x_880); -lean_ctor_set(x_881, 2, x_879); -x_759 = x_881; -x_760 = x_755; -goto block_863; -} -block_863: -{ -lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; uint8_t x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; size_t x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; -x_761 = l_Lean_Elab_Command_getRef(x_12, x_13, x_760); -x_762 = lean_ctor_get(x_761, 0); -lean_inc(x_762); -x_763 = lean_ctor_get(x_761, 1); -lean_inc(x_763); -if (lean_is_exclusive(x_761)) { - lean_ctor_release(x_761, 0); - lean_ctor_release(x_761, 1); - x_764 = x_761; +x_888 = lean_box(2); +x_889 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_890 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_890, 0, x_888); +lean_ctor_set(x_890, 1, x_889); +lean_ctor_set(x_890, 2, x_1); +if (lean_is_scalar(x_768)) { + x_891 = lean_alloc_ctor(1, 2, 0); +} else { + x_891 = x_768; + lean_ctor_set_tag(x_891, 1); +} +lean_ctor_set(x_891, 0, x_890); +lean_ctor_set(x_891, 1, x_3); +x_892 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_892, 0, x_887); +lean_ctor_set(x_892, 1, x_891); +x_893 = lean_array_mk(x_892); +x_894 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__24; +x_895 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_895, 0, x_888); +lean_ctor_set(x_895, 1, x_894); +lean_ctor_set(x_895, 2, x_893); +x_771 = x_895; +x_772 = x_767; +goto block_877; +} +block_877: +{ +lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; uint8_t x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; size_t x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; +x_773 = l_Lean_Elab_Command_getRef(x_12, x_13, x_772); +x_774 = lean_ctor_get(x_773, 0); +lean_inc(x_774); +x_775 = lean_ctor_get(x_773, 1); +lean_inc(x_775); +if (lean_is_exclusive(x_773)) { + lean_ctor_release(x_773, 0); + lean_ctor_release(x_773, 1); + x_776 = x_773; } else { - lean_dec_ref(x_761); - x_764 = lean_box(0); + lean_dec_ref(x_773); + x_776 = lean_box(0); } -x_765 = l_Lean_replaceRef(x_4, x_762); -lean_dec(x_762); -x_766 = lean_ctor_get(x_12, 0); -lean_inc(x_766); -x_767 = lean_ctor_get(x_12, 1); -lean_inc(x_767); -x_768 = lean_ctor_get(x_12, 2); -lean_inc(x_768); -x_769 = lean_ctor_get(x_12, 3); -lean_inc(x_769); -x_770 = lean_ctor_get(x_12, 4); -lean_inc(x_770); -x_771 = lean_ctor_get(x_12, 5); -lean_inc(x_771); -x_772 = lean_ctor_get(x_12, 7); -lean_inc(x_772); -x_773 = lean_ctor_get(x_12, 8); -lean_inc(x_773); -x_774 = lean_ctor_get(x_12, 9); -lean_inc(x_774); -x_775 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); -x_776 = lean_alloc_ctor(0, 10, 1); -lean_ctor_set(x_776, 0, x_766); -lean_ctor_set(x_776, 1, x_767); -lean_ctor_set(x_776, 2, x_768); -lean_ctor_set(x_776, 3, x_769); -lean_ctor_set(x_776, 4, x_770); -lean_ctor_set(x_776, 5, x_771); -lean_ctor_set(x_776, 6, x_765); -lean_ctor_set(x_776, 7, x_772); -lean_ctor_set(x_776, 8, x_773); -lean_ctor_set(x_776, 9, x_774); -lean_ctor_set_uint8(x_776, sizeof(void*)*10, x_775); -x_777 = l_Lean_Elab_Command_getRef(x_776, x_13, x_763); -x_778 = lean_ctor_get(x_777, 0); +x_777 = l_Lean_replaceRef(x_4, x_774); +lean_dec(x_774); +x_778 = lean_ctor_get(x_12, 0); lean_inc(x_778); -x_779 = lean_ctor_get(x_777, 1); +x_779 = lean_ctor_get(x_12, 1); lean_inc(x_779); -if (lean_is_exclusive(x_777)) { - lean_ctor_release(x_777, 0); - lean_ctor_release(x_777, 1); - x_780 = x_777; -} else { - lean_dec_ref(x_777); - x_780 = lean_box(0); -} -x_781 = l_Lean_SourceInfo_fromRef(x_778, x_757); -lean_dec(x_778); -x_782 = l_Lean_Elab_Command_getCurrMacroScope(x_776, x_13, x_779); -lean_dec(x_776); -x_783 = lean_ctor_get(x_782, 1); +x_780 = lean_ctor_get(x_12, 2); +lean_inc(x_780); +x_781 = lean_ctor_get(x_12, 3); +lean_inc(x_781); +x_782 = lean_ctor_get(x_12, 4); +lean_inc(x_782); +x_783 = lean_ctor_get(x_12, 5); lean_inc(x_783); -if (lean_is_exclusive(x_782)) { - lean_ctor_release(x_782, 0); - lean_ctor_release(x_782, 1); - x_784 = x_782; -} else { - lean_dec_ref(x_782); - x_784 = lean_box(0); -} -x_785 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_783); -x_786 = lean_ctor_get(x_785, 1); +x_784 = lean_ctor_get(x_12, 7); +lean_inc(x_784); +x_785 = lean_ctor_get(x_12, 8); +lean_inc(x_785); +x_786 = lean_ctor_get(x_12, 9); lean_inc(x_786); -if (lean_is_exclusive(x_785)) { - lean_ctor_release(x_785, 0); - lean_ctor_release(x_785, 1); - x_787 = x_785; -} else { - lean_dec_ref(x_785); - x_787 = lean_box(0); -} -x_788 = l_Lake_DSL_structVal___closed__3; -lean_inc(x_781); -if (lean_is_scalar(x_787)) { - x_789 = lean_alloc_ctor(2, 2, 0); -} else { - x_789 = x_787; - lean_ctor_set_tag(x_789, 2); -} -lean_ctor_set(x_789, 0, x_781); -lean_ctor_set(x_789, 1, x_788); -x_790 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; -x_791 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; -lean_inc(x_781); -x_792 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_792, 0, x_781); -lean_ctor_set(x_792, 1, x_790); -lean_ctor_set(x_792, 2, x_791); -x_793 = lean_array_size(x_751); -x_794 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_793, x_17, x_751); -x_795 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; -x_796 = l_Lean_mkSepArray(x_794, x_795); -lean_dec(x_794); -x_797 = l_Array_append___rarg(x_791, x_796); -lean_dec(x_796); -lean_inc(x_781); -x_798 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_798, 0, x_781); -lean_ctor_set(x_798, 1, x_790); -lean_ctor_set(x_798, 2, x_797); -x_799 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; -lean_inc(x_792); -lean_inc(x_781); -x_800 = l_Lean_Syntax_node1(x_781, x_799, x_792); -x_801 = l_Lake_DSL_structVal___closed__18; -lean_inc(x_781); -if (lean_is_scalar(x_784)) { - x_802 = lean_alloc_ctor(2, 2, 0); -} else { - x_802 = x_784; - lean_ctor_set_tag(x_802, 2); -} -lean_ctor_set(x_802, 0, x_781); -lean_ctor_set(x_802, 1, x_801); -x_803 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; -lean_inc(x_792); -x_804 = l_Lean_Syntax_node6(x_781, x_803, x_789, x_792, x_798, x_800, x_792, x_802); -x_805 = l_Lean_Elab_Command_getRef(x_12, x_13, x_786); -x_806 = lean_ctor_get(x_805, 0); -lean_inc(x_806); -x_807 = lean_ctor_get(x_805, 1); -lean_inc(x_807); -if (lean_is_exclusive(x_805)) { - lean_ctor_release(x_805, 0); - lean_ctor_release(x_805, 1); - x_808 = x_805; +x_787 = lean_ctor_get_uint8(x_12, sizeof(void*)*10); +x_788 = lean_alloc_ctor(0, 10, 1); +lean_ctor_set(x_788, 0, x_778); +lean_ctor_set(x_788, 1, x_779); +lean_ctor_set(x_788, 2, x_780); +lean_ctor_set(x_788, 3, x_781); +lean_ctor_set(x_788, 4, x_782); +lean_ctor_set(x_788, 5, x_783); +lean_ctor_set(x_788, 6, x_777); +lean_ctor_set(x_788, 7, x_784); +lean_ctor_set(x_788, 8, x_785); +lean_ctor_set(x_788, 9, x_786); +lean_ctor_set_uint8(x_788, sizeof(void*)*10, x_787); +x_789 = l_Lean_Elab_Command_getRef(x_788, x_13, x_775); +x_790 = lean_ctor_get(x_789, 0); +lean_inc(x_790); +x_791 = lean_ctor_get(x_789, 1); +lean_inc(x_791); +if (lean_is_exclusive(x_789)) { + lean_ctor_release(x_789, 0); + lean_ctor_release(x_789, 1); + x_792 = x_789; } else { - lean_dec_ref(x_805); - x_808 = lean_box(0); + lean_dec_ref(x_789); + x_792 = lean_box(0); } -x_809 = l_Lean_SourceInfo_fromRef(x_806, x_757); +x_793 = l_Lean_SourceInfo_fromRef(x_790, x_769); +lean_dec(x_790); +x_794 = l_Lean_Elab_Command_getCurrMacroScope(x_788, x_13, x_791); +lean_dec(x_788); +x_795 = lean_ctor_get(x_794, 1); +lean_inc(x_795); +if (lean_is_exclusive(x_794)) { + lean_ctor_release(x_794, 0); + lean_ctor_release(x_794, 1); + x_796 = x_794; +} else { + lean_dec_ref(x_794); + x_796 = lean_box(0); +} +x_797 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_795); +x_798 = lean_ctor_get(x_797, 1); +lean_inc(x_798); +if (lean_is_exclusive(x_797)) { + lean_ctor_release(x_797, 0); + lean_ctor_release(x_797, 1); + x_799 = x_797; +} else { + lean_dec_ref(x_797); + x_799 = lean_box(0); +} +x_800 = l_Lake_DSL_structVal___closed__3; +lean_inc(x_793); +if (lean_is_scalar(x_799)) { + x_801 = lean_alloc_ctor(2, 2, 0); +} else { + x_801 = x_799; + lean_ctor_set_tag(x_801, 2); +} +lean_ctor_set(x_801, 0, x_793); +lean_ctor_set(x_801, 1, x_800); +x_802 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__8; +x_803 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__5; +lean_inc(x_793); +x_804 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_804, 0, x_793); +lean_ctor_set(x_804, 1, x_802); +lean_ctor_set(x_804, 2, x_803); +x_805 = lean_array_size(x_763); +x_806 = l_Array_mapMUnsafe_map___at_Lake_DSL_elabConfigDecl___spec__5(x_805, x_17, x_763); +x_807 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__6; +x_808 = l_Lean_mkSepArray(x_806, x_807); lean_dec(x_806); -x_810 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_807); -x_811 = lean_ctor_get(x_810, 1); -lean_inc(x_811); -if (lean_is_exclusive(x_810)) { - lean_ctor_release(x_810, 0); - lean_ctor_release(x_810, 1); - x_812 = x_810; -} else { - lean_dec_ref(x_810); - x_812 = lean_box(0); -} -x_813 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_811); -x_814 = lean_ctor_get(x_813, 1); -lean_inc(x_814); -if (lean_is_exclusive(x_813)) { - lean_ctor_release(x_813, 0); - lean_ctor_release(x_813, 1); - x_815 = x_813; -} else { - lean_dec_ref(x_813); - x_815 = lean_box(0); -} -x_816 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__11; -lean_inc(x_809); -if (lean_is_scalar(x_815)) { - x_817 = lean_alloc_ctor(2, 2, 0); -} else { - x_817 = x_815; - lean_ctor_set_tag(x_817, 2); -} -lean_ctor_set(x_817, 0, x_809); -lean_ctor_set(x_817, 1, x_816); -x_818 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__3; -x_819 = l_Lean_Syntax_SepArray_ofElems(x_818, x_5); -x_820 = l_Array_append___rarg(x_791, x_819); -lean_dec(x_819); -lean_inc(x_809); -x_821 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_821, 0, x_809); -lean_ctor_set(x_821, 1, x_790); -lean_ctor_set(x_821, 2, x_820); -x_822 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; -lean_inc(x_809); -if (lean_is_scalar(x_812)) { - x_823 = lean_alloc_ctor(2, 2, 0); -} else { - x_823 = x_812; - lean_ctor_set_tag(x_823, 2); -} -lean_ctor_set(x_823, 0, x_809); -lean_ctor_set(x_823, 1, x_822); -x_824 = l_Lake_DSL_expandAttrs___closed__6; -lean_inc(x_809); -x_825 = l_Lean_Syntax_node3(x_809, x_824, x_817, x_821, x_823); -lean_inc(x_809); -x_826 = l_Lean_Syntax_node1(x_809, x_790, x_825); -lean_inc(x_809); -x_827 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_827, 0, x_809); -lean_ctor_set(x_827, 1, x_790); -lean_ctor_set(x_827, 2, x_791); -x_828 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; -lean_inc(x_809); -if (lean_is_scalar(x_808)) { - x_829 = lean_alloc_ctor(2, 2, 0); -} else { - x_829 = x_808; - lean_ctor_set_tag(x_829, 2); -} -lean_ctor_set(x_829, 0, x_809); -lean_ctor_set(x_829, 1, x_828); -x_830 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; -lean_inc(x_809); -if (lean_is_scalar(x_780)) { +x_809 = l_Array_append___rarg(x_803, x_808); +lean_dec(x_808); +lean_inc(x_793); +x_810 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_810, 0, x_793); +lean_ctor_set(x_810, 1, x_802); +lean_ctor_set(x_810, 2, x_809); +x_811 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__4; +lean_inc(x_793); +x_812 = l_Lean_Syntax_node1(x_793, x_811, x_810); +x_813 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; +lean_inc(x_804); +lean_inc(x_793); +x_814 = l_Lean_Syntax_node1(x_793, x_813, x_804); +x_815 = l_Lake_DSL_structVal___closed__18; +lean_inc(x_793); +if (lean_is_scalar(x_796)) { + x_816 = lean_alloc_ctor(2, 2, 0); +} else { + x_816 = x_796; + lean_ctor_set_tag(x_816, 2); +} +lean_ctor_set(x_816, 0, x_793); +lean_ctor_set(x_816, 1, x_815); +x_817 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__2; +lean_inc(x_804); +x_818 = l_Lean_Syntax_node6(x_793, x_817, x_801, x_804, x_812, x_814, x_804, x_816); +x_819 = l_Lean_Elab_Command_getRef(x_12, x_13, x_798); +x_820 = lean_ctor_get(x_819, 0); +lean_inc(x_820); +x_821 = lean_ctor_get(x_819, 1); +lean_inc(x_821); +if (lean_is_exclusive(x_819)) { + lean_ctor_release(x_819, 0); + lean_ctor_release(x_819, 1); + x_822 = x_819; +} else { + lean_dec_ref(x_819); + x_822 = lean_box(0); +} +x_823 = l_Lean_SourceInfo_fromRef(x_820, x_769); +lean_dec(x_820); +x_824 = l_Lean_Elab_Command_getCurrMacroScope(x_12, x_13, x_821); +x_825 = lean_ctor_get(x_824, 1); +lean_inc(x_825); +if (lean_is_exclusive(x_824)) { + lean_ctor_release(x_824, 0); + lean_ctor_release(x_824, 1); + x_826 = x_824; +} else { + lean_dec_ref(x_824); + x_826 = lean_box(0); +} +x_827 = l_Lean_Elab_Command_getMainModule___rarg(x_13, x_825); +x_828 = lean_ctor_get(x_827, 1); +lean_inc(x_828); +if (lean_is_exclusive(x_827)) { + lean_ctor_release(x_827, 0); + lean_ctor_release(x_827, 1); + x_829 = x_827; +} else { + lean_dec_ref(x_827); + x_829 = lean_box(0); +} +x_830 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__13; +lean_inc(x_823); +if (lean_is_scalar(x_829)) { x_831 = lean_alloc_ctor(2, 2, 0); } else { - x_831 = x_780; + x_831 = x_829; lean_ctor_set_tag(x_831, 2); } -lean_ctor_set(x_831, 0, x_809); +lean_ctor_set(x_831, 0, x_823); lean_ctor_set(x_831, 1, x_830); -x_832 = l_Lake_DSL_simpleDeclSig___closed__4; -lean_inc(x_809); -x_833 = l_Lean_Syntax_node2(x_809, x_832, x_831, x_758); -lean_inc(x_809); -x_834 = l_Lean_Syntax_node1(x_809, x_790, x_833); -x_835 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; -lean_inc(x_827); -lean_inc(x_809); -x_836 = l_Lean_Syntax_node2(x_809, x_835, x_827, x_834); -x_837 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; -lean_inc(x_809); -if (lean_is_scalar(x_764)) { - x_838 = lean_alloc_ctor(2, 2, 0); -} else { - x_838 = x_764; - lean_ctor_set_tag(x_838, 2); -} -lean_ctor_set(x_838, 0, x_809); -lean_ctor_set(x_838, 1, x_837); -x_839 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__19; -lean_inc_n(x_827, 2); -lean_inc(x_809); -x_840 = l_Lean_Syntax_node2(x_809, x_839, x_827, x_827); -x_841 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__20; -lean_inc(x_809); -x_842 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_842, 0, x_809); -lean_ctor_set(x_842, 1, x_790); -lean_ctor_set(x_842, 2, x_841); -x_843 = l_Lake_DSL_simpleDeclSig___closed__9; -lean_inc(x_842); -lean_inc(x_809); -x_844 = l_Lean_Syntax_node4(x_809, x_843, x_838, x_804, x_840, x_842); -x_845 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; -lean_inc(x_809); -x_846 = l_Lean_Syntax_node4(x_809, x_845, x_829, x_759, x_836, x_844); +x_832 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__5; +x_833 = l_Lean_Syntax_SepArray_ofElems(x_832, x_5); +x_834 = l_Array_append___rarg(x_803, x_833); +lean_dec(x_833); +lean_inc(x_823); +x_835 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_835, 0, x_823); +lean_ctor_set(x_835, 1, x_802); +lean_ctor_set(x_835, 2, x_834); +x_836 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__14; +lean_inc(x_823); +if (lean_is_scalar(x_826)) { + x_837 = lean_alloc_ctor(2, 2, 0); +} else { + x_837 = x_826; + lean_ctor_set_tag(x_837, 2); +} +lean_ctor_set(x_837, 0, x_823); +lean_ctor_set(x_837, 1, x_836); +x_838 = l_Lake_DSL_expandAttrs___closed__6; +lean_inc(x_823); +x_839 = l_Lean_Syntax_node3(x_823, x_838, x_831, x_835, x_837); +lean_inc(x_823); +x_840 = l_Lean_Syntax_node1(x_823, x_802, x_839); +lean_inc(x_823); +x_841 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_841, 0, x_823); +lean_ctor_set(x_841, 1, x_802); +lean_ctor_set(x_841, 2, x_803); +x_842 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__15; +lean_inc(x_823); +if (lean_is_scalar(x_822)) { + x_843 = lean_alloc_ctor(2, 2, 0); +} else { + x_843 = x_822; + lean_ctor_set_tag(x_843, 2); +} +lean_ctor_set(x_843, 0, x_823); +lean_ctor_set(x_843, 1, x_842); +x_844 = l_Lake_DSL_expandOptSimpleBinder___lambda__1___closed__6; +lean_inc(x_823); +if (lean_is_scalar(x_792)) { + x_845 = lean_alloc_ctor(2, 2, 0); +} else { + x_845 = x_792; + lean_ctor_set_tag(x_845, 2); +} +lean_ctor_set(x_845, 0, x_823); +lean_ctor_set(x_845, 1, x_844); +x_846 = l_Lake_DSL_simpleDeclSig___closed__4; +lean_inc(x_823); +x_847 = l_Lean_Syntax_node2(x_823, x_846, x_845, x_770); +lean_inc(x_823); +x_848 = l_Lean_Syntax_node1(x_823, x_802, x_847); +x_849 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__18; +lean_inc(x_841); +lean_inc(x_823); +x_850 = l_Lean_Syntax_node2(x_823, x_849, x_841, x_848); +x_851 = l_Lean_RBNode_foldM___at_Lake_DSL_elabConfigDecl___spec__4___closed__6; +lean_inc(x_823); +if (lean_is_scalar(x_776)) { + x_852 = lean_alloc_ctor(2, 2, 0); +} else { + x_852 = x_776; + lean_ctor_set_tag(x_852, 2); +} +lean_ctor_set(x_852, 0, x_823); +lean_ctor_set(x_852, 1, x_851); +x_853 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__21; +lean_inc_n(x_841, 2); +lean_inc(x_823); +x_854 = l_Lean_Syntax_node2(x_823, x_853, x_841, x_841); +x_855 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__22; +lean_inc(x_823); +x_856 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_856, 0, x_823); +lean_ctor_set(x_856, 1, x_802); +lean_ctor_set(x_856, 2, x_855); +x_857 = l_Lake_DSL_simpleDeclSig___closed__9; +lean_inc(x_856); +lean_inc(x_823); +x_858 = l_Lean_Syntax_node4(x_823, x_857, x_852, x_818, x_854, x_856); +x_859 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__16; +lean_inc(x_823); +x_860 = l_Lean_Syntax_node4(x_823, x_859, x_843, x_771, x_850, x_858); if (lean_obj_tag(x_6) == 0) { -lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; -x_847 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_827, 3); -lean_inc(x_809); -x_848 = l_Lean_Syntax_node6(x_809, x_847, x_842, x_826, x_827, x_827, x_827, x_827); -x_849 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_850 = l_Lean_Syntax_node2(x_809, x_849, x_848, x_846); -lean_inc(x_850); -x_851 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_851, 0, x_850); -x_852 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_850, x_851, x_12, x_13, x_814); -return x_852; +lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; +x_861 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_841, 3); +lean_inc(x_823); +x_862 = l_Lean_Syntax_node6(x_823, x_861, x_856, x_840, x_841, x_841, x_841, x_841); +x_863 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_864 = l_Lean_Syntax_node2(x_823, x_863, x_862, x_860); +lean_inc(x_864); +x_865 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_865, 0, x_864); +x_866 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_864, x_865, x_12, x_13, x_828); +return x_866; } else { -lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; -lean_dec(x_842); -x_853 = lean_ctor_get(x_6, 0); -lean_inc(x_853); +lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; +lean_dec(x_856); +x_867 = lean_ctor_get(x_6, 0); +lean_inc(x_867); lean_dec(x_6); -x_854 = l_Array_mkArray1___rarg(x_853); -x_855 = l_Array_append___rarg(x_791, x_854); -lean_dec(x_854); -lean_inc(x_809); -x_856 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_856, 0, x_809); -lean_ctor_set(x_856, 1, x_790); -lean_ctor_set(x_856, 2, x_855); -x_857 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; -lean_inc_n(x_827, 3); -lean_inc(x_809); -x_858 = l_Lean_Syntax_node6(x_809, x_857, x_856, x_826, x_827, x_827, x_827, x_827); -x_859 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__8; -x_860 = l_Lean_Syntax_node2(x_809, x_859, x_858, x_846); -lean_inc(x_860); -x_861 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); -lean_closure_set(x_861, 0, x_860); -x_862 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_860, x_861, x_12, x_13, x_814); -return x_862; +x_868 = l_Array_mkArray1___rarg(x_867); +x_869 = l_Array_append___rarg(x_803, x_868); +lean_dec(x_868); +lean_inc(x_823); +x_870 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_870, 0, x_823); +lean_ctor_set(x_870, 1, x_802); +lean_ctor_set(x_870, 2, x_869); +x_871 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__12; +lean_inc_n(x_841, 3); +lean_inc(x_823); +x_872 = l_Lean_Syntax_node6(x_823, x_871, x_870, x_840, x_841, x_841, x_841, x_841); +x_873 = l_Lake_DSL_elabConfigDecl___lambda__2___closed__10; +x_874 = l_Lean_Syntax_node2(x_823, x_873, x_872, x_860); +lean_inc(x_874); +x_875 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand), 4, 1); +lean_closure_set(x_875, 0, x_874); +x_876 = l_Lean_Elab_Command_withMacroExpansion___rarg(x_7, x_874, x_875, x_12, x_13, x_828); +return x_876; } } } } else { -uint8_t x_882; +uint8_t x_896; lean_dec(x_13); lean_dec(x_12); lean_dec(x_9); @@ -14354,23 +14442,23 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_882 = !lean_is_exclusive(x_18); -if (x_882 == 0) +x_896 = !lean_is_exclusive(x_18); +if (x_896 == 0) { return x_18; } else { -lean_object* x_883; lean_object* x_884; lean_object* x_885; -x_883 = lean_ctor_get(x_18, 0); -x_884 = lean_ctor_get(x_18, 1); -lean_inc(x_884); -lean_inc(x_883); +lean_object* x_897; lean_object* x_898; lean_object* x_899; +x_897 = lean_ctor_get(x_18, 0); +x_898 = lean_ctor_get(x_18, 1); +lean_inc(x_898); +lean_inc(x_897); lean_dec(x_18); -x_885 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_885, 0, x_883); -lean_ctor_set(x_885, 1, x_884); -return x_885; +x_899 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_899, 0, x_897); +lean_ctor_set(x_899, 1, x_898); +return x_899; } } } @@ -15360,6 +15448,10 @@ l_Lake_DSL_elabConfigDecl___lambda__2___closed__23 = _init_l_Lake_DSL_elabConfig lean_mark_persistent(l_Lake_DSL_elabConfigDecl___lambda__2___closed__23); l_Lake_DSL_elabConfigDecl___lambda__2___closed__24 = _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__24(); lean_mark_persistent(l_Lake_DSL_elabConfigDecl___lambda__2___closed__24); +l_Lake_DSL_elabConfigDecl___lambda__2___closed__25 = _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__25(); +lean_mark_persistent(l_Lake_DSL_elabConfigDecl___lambda__2___closed__25); +l_Lake_DSL_elabConfigDecl___lambda__2___closed__26 = _init_l_Lake_DSL_elabConfigDecl___lambda__2___closed__26(); +lean_mark_persistent(l_Lake_DSL_elabConfigDecl___lambda__2___closed__26); l_Lake_DSL_elabConfigDecl___lambda__3___closed__1 = _init_l_Lake_DSL_elabConfigDecl___lambda__3___closed__1(); lean_mark_persistent(l_Lake_DSL_elabConfigDecl___lambda__3___closed__1); l_Lake_DSL_elabConfigDecl___lambda__3___closed__2 = _init_l_Lake_DSL_elabConfigDecl___lambda__3___closed__2(); diff --git a/stage0/stdlib/Lake/DSL/Require.c b/stage0/stdlib/Lake/DSL/Require.c index 012a8d3999f4..371aaa65bfb7 100644 --- a/stage0/stdlib/Lake/DSL/Require.c +++ b/stage0/stdlib/Lake/DSL/Require.c @@ -91,6 +91,7 @@ LEAN_EXPORT lean_object* l_Lake_DSL_expandDepSpec___lambda__7___boxed(lean_objec static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__36; static lean_object* l_Lake_DSL_fromGit___closed__12; static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__28; +static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__78; static lean_object* l_Lake_DSL_fromGit___closed__20; static lean_object* l_Lake_DSL_requireDecl___closed__7; static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__16; @@ -158,6 +159,7 @@ LEAN_EXPORT lean_object* l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm__ lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__50; static lean_object* l_Lake_DSL_expandDepSpec___lambda__5___closed__2; +static lean_object* l_Lake_DSL_expandDepSpec___lambda__1___closed__77; static lean_object* l_Lake_DSL_verSpec___closed__4; static lean_object* l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lake_DSL_fromGit; @@ -1777,7 +1779,7 @@ static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -1797,31 +1799,51 @@ static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("Command", 7, 7); +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; +x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; +x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Command", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__15() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_unchecked("declaration", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__14() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__13; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__15() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -1829,19 +1851,19 @@ x_1 = lean_mk_string_unchecked("declModifiers", 13, 13); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__16() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__15; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__17() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -1849,19 +1871,19 @@ x_1 = lean_mk_string_unchecked("attributes", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__18() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__17; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__19() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__21() { _start: { lean_object* x_1; @@ -1869,7 +1891,7 @@ x_1 = lean_mk_string_unchecked("@[", 2, 2); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__20() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__22() { _start: { lean_object* x_1; @@ -1877,19 +1899,19 @@ x_1 = lean_mk_string_unchecked("attrInstance", 12, 12); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__21() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__20; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__22() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__24() { _start: { lean_object* x_1; @@ -1897,19 +1919,19 @@ x_1 = lean_mk_string_unchecked("attrKind", 8, 8); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__23() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__22; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__24; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__24() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__26() { _start: { lean_object* x_1; @@ -1917,7 +1939,7 @@ x_1 = lean_mk_string_unchecked("Attr", 4, 4); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__25() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__27() { _start: { lean_object* x_1; @@ -1925,19 +1947,19 @@ x_1 = lean_mk_string_unchecked("simple", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__26() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__24; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__25; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__26; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__27() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__29() { _start: { lean_object* x_1; @@ -1945,26 +1967,26 @@ x_1 = lean_mk_string_unchecked("package_dep", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__28() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__27; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__29; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__29() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__27; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__30() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__32() { _start: { lean_object* x_1; @@ -1972,7 +1994,7 @@ x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__31() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__33() { _start: { lean_object* x_1; @@ -1980,19 +2002,19 @@ x_1 = lean_mk_string_unchecked("definition", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__32() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__31; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__33; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__33() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__35() { _start: { lean_object* x_1; @@ -2000,7 +2022,7 @@ x_1 = lean_mk_string_unchecked("def", 3, 3); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__34() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__36() { _start: { lean_object* x_1; @@ -2008,19 +2030,19 @@ x_1 = lean_mk_string_unchecked("declId", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__35() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__34; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__36; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__36() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -2029,13 +2051,13 @@ x_2 = lean_array_mk(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__37() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__13; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__36; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__38; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2043,19 +2065,19 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__38() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__37; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__39; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__39() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__41() { _start: { lean_object* x_1; @@ -2063,19 +2085,19 @@ x_1 = lean_mk_string_unchecked("optDeclSig", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__40() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__39; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__41; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__41() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__43() { _start: { lean_object* x_1; @@ -2083,19 +2105,19 @@ x_1 = lean_mk_string_unchecked("typeSpec", 8, 8); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__42() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__41; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__43; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__43() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__45() { _start: { lean_object* x_1; @@ -2103,7 +2125,7 @@ x_1 = lean_mk_string_unchecked(":", 1, 1); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__44() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__46() { _start: { lean_object* x_1; @@ -2111,28 +2133,28 @@ x_1 = lean_mk_string_unchecked("Dependency", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__45() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lake_DSL_fromPath___closed__1; -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__44; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__46; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__46() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__45; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__47; x_3 = 0; x_4 = l_Lean_mkCIdentFrom(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__47() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__49() { _start: { lean_object* x_1; @@ -2140,19 +2162,19 @@ x_1 = lean_mk_string_unchecked("declValSimple", 13, 13); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__48() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__12; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__47; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__49; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__49() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__51() { _start: { lean_object* x_1; @@ -2160,7 +2182,7 @@ x_1 = lean_mk_string_unchecked(":=", 2, 2); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__50() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__52() { _start: { lean_object* x_1; @@ -2168,19 +2190,19 @@ x_1 = lean_mk_string_unchecked("structInstField", 15, 15); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__51() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__50; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__52; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__52() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__54() { _start: { lean_object* x_1; @@ -2188,19 +2210,19 @@ x_1 = lean_mk_string_unchecked("structInstLVal", 14, 14); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__53() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__52; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__54; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__54() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__56() { _start: { lean_object* x_1; @@ -2208,26 +2230,26 @@ x_1 = lean_mk_string_unchecked("name", 4, 4); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__55() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__54; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__56; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__56() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__58() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__54; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__56; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__57() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__59() { _start: { lean_object* x_1; @@ -2235,7 +2257,7 @@ x_1 = lean_mk_string_unchecked(",", 1, 1); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__58() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__60() { _start: { lean_object* x_1; @@ -2243,26 +2265,26 @@ x_1 = lean_mk_string_unchecked("scope", 5, 5); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__59() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__61() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__58; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__60; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__60() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__62() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__58; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__60; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__61() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__63() { _start: { lean_object* x_1; @@ -2270,26 +2292,26 @@ x_1 = lean_mk_string_unchecked("version\?", 8, 8); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__62() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__61; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__63; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__63() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__61; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__63; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__64() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__66() { _start: { lean_object* x_1; @@ -2297,26 +2319,26 @@ x_1 = lean_mk_string_unchecked("src\?", 4, 4); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__65() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__64; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__66; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__66() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__68() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__64; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__66; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__67() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__69() { _start: { lean_object* x_1; @@ -2324,26 +2346,26 @@ x_1 = lean_mk_string_unchecked("opts", 4, 4); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__68() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__70() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__67; +x_1 = l_Lake_DSL_expandDepSpec___lambda__1___closed__69; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__69() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__67; +x_2 = l_Lake_DSL_expandDepSpec___lambda__1___closed__69; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__70() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__72() { _start: { lean_object* x_1; @@ -2351,7 +2373,7 @@ x_1 = lean_mk_string_unchecked("Termination", 11, 11); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__71() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__73() { _start: { lean_object* x_1; @@ -2359,19 +2381,19 @@ x_1 = lean_mk_string_unchecked("suffix", 6, 6); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__72() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__74() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; -x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__70; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__71; +x_3 = l_Lake_DSL_expandDepSpec___lambda__1___closed__72; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__73; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__73() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__75() { _start: { lean_object* x_1; @@ -2379,19 +2401,19 @@ x_1 = lean_mk_string_unchecked("quotedName", 10, 10); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__74() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__1; x_2 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__2; x_3 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__3; -x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__73; +x_4 = l_Lake_DSL_expandDepSpec___lambda__1___closed__75; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__75() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__77() { _start: { lean_object* x_1; @@ -2399,7 +2421,7 @@ x_1 = lean_mk_string_unchecked(".", 1, 1); return x_1; } } -static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__76() { +static lean_object* _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__78() { _start: { lean_object* x_1; @@ -2414,68 +2436,68 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = l_Lake_DSL_expandIdentOrStrAsIdent(x_1); if (lean_obj_tag(x_5) == 0) { -lean_object* x_170; uint8_t x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_170 = lean_ctor_get(x_7, 5); -lean_inc(x_170); -x_171 = 0; -x_172 = l_Lean_SourceInfo_fromRef(x_170, x_171); -lean_dec(x_170); -x_173 = lean_ctor_get(x_7, 2); -lean_inc(x_173); -x_174 = lean_ctor_get(x_7, 1); +lean_object* x_174; uint8_t x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_174 = lean_ctor_get(x_7, 5); lean_inc(x_174); -x_175 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__3; -x_176 = l_Lean_addMacroScope(x_174, x_175, x_173); -x_177 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__2; -x_178 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__7; -x_179 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_179, 0, x_172); -lean_ctor_set(x_179, 1, x_177); -lean_ctor_set(x_179, 2, x_176); -lean_ctor_set(x_179, 3, x_178); -x_10 = x_179; +x_175 = 0; +x_176 = l_Lean_SourceInfo_fromRef(x_174, x_175); +lean_dec(x_174); +x_177 = lean_ctor_get(x_7, 2); +lean_inc(x_177); +x_178 = lean_ctor_get(x_7, 1); +lean_inc(x_178); +x_179 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__3; +x_180 = l_Lean_addMacroScope(x_178, x_179, x_177); +x_181 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__2; +x_182 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__2___closed__7; +x_183 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_183, 0, x_176); +lean_ctor_set(x_183, 1, x_181); +lean_ctor_set(x_183, 2, x_180); +lean_ctor_set(x_183, 3, x_182); +x_10 = x_183; x_11 = x_8; -goto block_169; +goto block_173; } else { -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_180 = lean_ctor_get(x_5, 0); -lean_inc(x_180); -lean_dec(x_5); -x_181 = lean_ctor_get(x_7, 5); -lean_inc(x_181); -x_182 = l_Lean_replaceRef(x_180, x_181); -lean_dec(x_181); -x_183 = lean_ctor_get(x_7, 1); -lean_inc(x_183); -x_184 = lean_ctor_get(x_7, 2); +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_184 = lean_ctor_get(x_5, 0); lean_inc(x_184); -x_185 = 0; -x_186 = l_Lean_SourceInfo_fromRef(x_182, x_185); -lean_dec(x_182); -x_187 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__8; -x_188 = l_Lean_addMacroScope(x_183, x_187, x_184); -x_189 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__7; -x_190 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__11; -lean_inc(x_186); -x_191 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_191, 0, x_186); -lean_ctor_set(x_191, 1, x_189); -lean_ctor_set(x_191, 2, x_188); -lean_ctor_set(x_191, 3, x_190); -x_192 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__13; -lean_inc(x_186); -x_193 = l_Lean_Syntax_node1(x_186, x_192, x_180); -x_194 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__5; -x_195 = l_Lean_Syntax_node2(x_186, x_194, x_191, x_193); -x_10 = x_195; +lean_dec(x_5); +x_185 = lean_ctor_get(x_7, 5); +lean_inc(x_185); +x_186 = l_Lean_replaceRef(x_184, x_185); +lean_dec(x_185); +x_187 = lean_ctor_get(x_7, 1); +lean_inc(x_187); +x_188 = lean_ctor_get(x_7, 2); +lean_inc(x_188); +x_189 = 0; +x_190 = l_Lean_SourceInfo_fromRef(x_186, x_189); +lean_dec(x_186); +x_191 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__8; +x_192 = l_Lean_addMacroScope(x_187, x_191, x_188); +x_193 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__7; +x_194 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__11; +lean_inc(x_190); +x_195 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_195, 0, x_190); +lean_ctor_set(x_195, 1, x_193); +lean_ctor_set(x_195, 2, x_192); +lean_ctor_set(x_195, 3, x_194); +x_196 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__13; +lean_inc(x_190); +x_197 = l_Lean_Syntax_node1(x_190, x_196, x_184); +x_198 = l___private_Lake_DSL_Require_0__Lake_DSL_quoteOptTerm___rarg___lambda__5___closed__5; +x_199 = l_Lean_Syntax_node2(x_190, x_198, x_195, x_197); +x_10 = x_199; x_11 = x_8; -goto block_169; +goto block_173; } -block_169: +block_173: { -lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_12 = lean_ctor_get(x_7, 5); lean_inc(x_12); x_13 = 0; @@ -2507,382 +2529,390 @@ x_24 = l_Lake_DSL_expandDepSpec___lambda__1___closed__11; lean_inc(x_23); lean_inc(x_14); x_25 = l_Lean_Syntax_node1(x_14, x_24, x_23); -x_26 = l_Lake_DSL_expandDepSpec___lambda__1___closed__8; +x_26 = l_Lake_DSL_expandDepSpec___lambda__1___closed__13; +lean_inc(x_23); +lean_inc(x_14); +x_27 = l_Lean_Syntax_node1(x_14, x_26, x_23); +x_28 = l_Lake_DSL_expandDepSpec___lambda__1___closed__8; lean_inc(x_18); -lean_inc(x_25); -lean_inc_n(x_23, 3); +lean_inc(x_27); +lean_inc_n(x_23, 2); lean_inc(x_16); lean_inc(x_14); -x_27 = l_Lean_Syntax_node6(x_14, x_26, x_16, x_23, x_23, x_25, x_23, x_18); -x_28 = l_Lake_DSL_expandDepSpec___lambda__1___closed__2; +x_29 = l_Lean_Syntax_node6(x_14, x_28, x_16, x_23, x_25, x_27, x_23, x_18); +x_30 = l_Lake_DSL_expandDepSpec___lambda__1___closed__2; lean_inc(x_14); -x_29 = l_Lean_Syntax_node2(x_14, x_28, x_20, x_27); -x_30 = lean_ctor_get(x_7, 2); -lean_inc(x_30); -x_31 = lean_ctor_get(x_7, 1); -lean_inc(x_31); +x_31 = l_Lean_Syntax_node2(x_14, x_30, x_20, x_29); +x_32 = lean_ctor_get(x_7, 2); +lean_inc(x_32); +x_33 = lean_ctor_get(x_7, 1); +lean_inc(x_33); lean_dec(x_7); -x_32 = l_Lake_DSL_expandDepSpec___lambda__1___closed__19; +x_34 = l_Lake_DSL_expandDepSpec___lambda__1___closed__21; lean_inc(x_14); -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_14); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lake_DSL_expandDepSpec___lambda__1___closed__23; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_14); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lake_DSL_expandDepSpec___lambda__1___closed__25; lean_inc(x_23); lean_inc(x_14); -x_35 = l_Lean_Syntax_node1(x_14, x_34, x_23); -x_36 = l_Lake_DSL_expandDepSpec___lambda__1___closed__29; -lean_inc(x_30); -lean_inc(x_31); -x_37 = l_Lean_addMacroScope(x_31, x_36, x_30); -x_38 = lean_box(0); -x_39 = l_Lake_DSL_expandDepSpec___lambda__1___closed__28; +x_37 = l_Lean_Syntax_node1(x_14, x_36, x_23); +x_38 = l_Lake_DSL_expandDepSpec___lambda__1___closed__31; +lean_inc(x_32); +lean_inc(x_33); +x_39 = l_Lean_addMacroScope(x_33, x_38, x_32); +x_40 = lean_box(0); +x_41 = l_Lake_DSL_expandDepSpec___lambda__1___closed__30; lean_inc(x_14); -x_40 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_40, 0, x_14); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_37); -lean_ctor_set(x_40, 3, x_38); -x_41 = l_Lake_DSL_expandDepSpec___lambda__1___closed__26; +x_42 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_42, 0, x_14); +lean_ctor_set(x_42, 1, x_41); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_42, 3, x_40); +x_43 = l_Lake_DSL_expandDepSpec___lambda__1___closed__28; lean_inc(x_23); lean_inc(x_14); -x_42 = l_Lean_Syntax_node2(x_14, x_41, x_40, x_23); -x_43 = l_Lake_DSL_expandDepSpec___lambda__1___closed__21; +x_44 = l_Lean_Syntax_node2(x_14, x_43, x_42, x_23); +x_45 = l_Lake_DSL_expandDepSpec___lambda__1___closed__23; lean_inc(x_14); -x_44 = l_Lean_Syntax_node2(x_14, x_43, x_35, x_42); +x_46 = l_Lean_Syntax_node2(x_14, x_45, x_37, x_44); lean_inc(x_14); -x_45 = l_Lean_Syntax_node1(x_14, x_21, x_44); -x_46 = l_Lake_DSL_expandDepSpec___lambda__1___closed__30; +x_47 = l_Lean_Syntax_node1(x_14, x_21, x_46); +x_48 = l_Lake_DSL_expandDepSpec___lambda__1___closed__32; lean_inc(x_14); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_14); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lake_DSL_expandDepSpec___lambda__1___closed__18; +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_14); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_Lake_DSL_expandDepSpec___lambda__1___closed__20; lean_inc(x_14); -x_49 = l_Lean_Syntax_node3(x_14, x_48, x_33, x_45, x_47); +x_51 = l_Lean_Syntax_node3(x_14, x_50, x_35, x_47, x_49); lean_inc(x_14); -x_50 = l_Lean_Syntax_node1(x_14, x_21, x_49); -x_51 = l_Lake_DSL_expandDepSpec___lambda__1___closed__33; +x_52 = l_Lean_Syntax_node1(x_14, x_21, x_51); +x_53 = l_Lake_DSL_expandDepSpec___lambda__1___closed__35; lean_inc(x_14); -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_14); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lake_DSL_expandDepSpec___lambda__1___closed__38; -lean_inc(x_9); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_9); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_14); lean_ctor_set(x_54, 1, x_53); -x_55 = lean_array_mk(x_54); -x_56 = lean_box(2); -x_57 = l_Lake_DSL_expandDepSpec___lambda__1___closed__35; -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -lean_ctor_set(x_58, 2, x_55); -x_59 = l_Lake_DSL_expandDepSpec___lambda__1___closed__43; -lean_inc(x_14); -x_60 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_60, 0, x_14); +x_55 = l_Lake_DSL_expandDepSpec___lambda__1___closed__40; +lean_inc(x_9); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_9); +lean_ctor_set(x_56, 1, x_55); +x_57 = lean_array_mk(x_56); +x_58 = lean_box(2); +x_59 = l_Lake_DSL_expandDepSpec___lambda__1___closed__37; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_58); lean_ctor_set(x_60, 1, x_59); -x_61 = l_Lake_DSL_expandDepSpec___lambda__1___closed__42; -x_62 = l_Lake_DSL_expandDepSpec___lambda__1___closed__46; +lean_ctor_set(x_60, 2, x_57); +x_61 = l_Lake_DSL_expandDepSpec___lambda__1___closed__45; lean_inc(x_14); -x_63 = l_Lean_Syntax_node2(x_14, x_61, x_60, x_62); +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_14); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lake_DSL_expandDepSpec___lambda__1___closed__44; +x_64 = l_Lake_DSL_expandDepSpec___lambda__1___closed__48; lean_inc(x_14); -x_64 = l_Lean_Syntax_node1(x_14, x_21, x_63); -x_65 = l_Lake_DSL_expandDepSpec___lambda__1___closed__40; +x_65 = l_Lean_Syntax_node2(x_14, x_63, x_62, x_64); +lean_inc(x_14); +x_66 = l_Lean_Syntax_node1(x_14, x_21, x_65); +x_67 = l_Lake_DSL_expandDepSpec___lambda__1___closed__42; lean_inc(x_23); lean_inc(x_14); -x_66 = l_Lean_Syntax_node2(x_14, x_65, x_23, x_64); -x_67 = l_Lake_DSL_expandDepSpec___lambda__1___closed__49; +x_68 = l_Lean_Syntax_node2(x_14, x_67, x_23, x_66); +x_69 = l_Lake_DSL_expandDepSpec___lambda__1___closed__51; lean_inc(x_14); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_14); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Lake_DSL_expandDepSpec___lambda__1___closed__56; -lean_inc(x_30); -lean_inc(x_31); -x_70 = l_Lean_addMacroScope(x_31, x_69, x_30); -x_71 = l_Lake_DSL_expandDepSpec___lambda__1___closed__55; +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_14); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Lake_DSL_expandDepSpec___lambda__1___closed__58; +lean_inc(x_32); +lean_inc(x_33); +x_72 = l_Lean_addMacroScope(x_33, x_71, x_32); +x_73 = l_Lake_DSL_expandDepSpec___lambda__1___closed__57; lean_inc(x_14); -x_72 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_72, 0, x_14); -lean_ctor_set(x_72, 1, x_71); -lean_ctor_set(x_72, 2, x_70); -lean_ctor_set(x_72, 3, x_38); -x_73 = l_Lake_DSL_expandDepSpec___lambda__1___closed__53; +x_74 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_74, 0, x_14); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set(x_74, 2, x_72); +lean_ctor_set(x_74, 3, x_40); +x_75 = l_Lake_DSL_expandDepSpec___lambda__1___closed__55; lean_inc(x_23); lean_inc(x_14); -x_74 = l_Lean_Syntax_node2(x_14, x_73, x_72, x_23); -x_75 = l_Lean_Syntax_getId(x_9); +x_76 = l_Lean_Syntax_node2(x_14, x_75, x_74, x_23); +x_77 = l_Lean_Syntax_getId(x_9); lean_dec(x_9); -lean_inc(x_75); -x_76 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_38, x_75); -x_77 = l_Lake_DSL_expandDepSpec___lambda__1___closed__57; +lean_inc(x_77); +x_78 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_40, x_77); +x_79 = l_Lake_DSL_expandDepSpec___lambda__1___closed__59; lean_inc(x_14); -x_78 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_78, 0, x_14); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lake_DSL_expandDepSpec___lambda__1___closed__60; -lean_inc(x_30); -lean_inc(x_31); -x_80 = l_Lean_addMacroScope(x_31, x_79, x_30); -x_81 = l_Lake_DSL_expandDepSpec___lambda__1___closed__59; +x_80 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_80, 0, x_14); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lake_DSL_expandDepSpec___lambda__1___closed__62; +lean_inc(x_32); +lean_inc(x_33); +x_82 = l_Lean_addMacroScope(x_33, x_81, x_32); +x_83 = l_Lake_DSL_expandDepSpec___lambda__1___closed__61; lean_inc(x_14); -x_82 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_82, 0, x_14); -lean_ctor_set(x_82, 1, x_81); -lean_ctor_set(x_82, 2, x_80); -lean_ctor_set(x_82, 3, x_38); +x_84 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_84, 0, x_14); +lean_ctor_set(x_84, 1, x_83); +lean_ctor_set(x_84, 2, x_82); +lean_ctor_set(x_84, 3, x_40); lean_inc(x_23); lean_inc(x_14); -x_83 = l_Lean_Syntax_node2(x_14, x_73, x_82, x_23); -x_84 = l_Lake_DSL_expandDepSpec___lambda__1___closed__51; -lean_inc(x_68); +x_85 = l_Lean_Syntax_node2(x_14, x_75, x_84, x_23); +x_86 = l_Lake_DSL_expandDepSpec___lambda__1___closed__53; +lean_inc(x_70); lean_inc(x_14); -x_85 = l_Lean_Syntax_node3(x_14, x_84, x_83, x_68, x_2); -x_86 = l_Lake_DSL_expandDepSpec___lambda__1___closed__63; -lean_inc(x_30); -lean_inc(x_31); -x_87 = l_Lean_addMacroScope(x_31, x_86, x_30); -x_88 = l_Lake_DSL_expandDepSpec___lambda__1___closed__62; +x_87 = l_Lean_Syntax_node3(x_14, x_86, x_85, x_70, x_2); +x_88 = l_Lake_DSL_expandDepSpec___lambda__1___closed__65; +lean_inc(x_32); +lean_inc(x_33); +x_89 = l_Lean_addMacroScope(x_33, x_88, x_32); +x_90 = l_Lake_DSL_expandDepSpec___lambda__1___closed__64; lean_inc(x_14); -x_89 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_89, 0, x_14); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set(x_89, 2, x_87); -lean_ctor_set(x_89, 3, x_38); +x_91 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_91, 0, x_14); +lean_ctor_set(x_91, 1, x_90); +lean_ctor_set(x_91, 2, x_89); +lean_ctor_set(x_91, 3, x_40); lean_inc(x_23); lean_inc(x_14); -x_90 = l_Lean_Syntax_node2(x_14, x_73, x_89, x_23); -lean_inc(x_68); +x_92 = l_Lean_Syntax_node2(x_14, x_75, x_91, x_23); +lean_inc(x_70); lean_inc(x_14); -x_91 = l_Lean_Syntax_node3(x_14, x_84, x_90, x_68, x_6); -x_92 = l_Lake_DSL_expandDepSpec___lambda__1___closed__66; -lean_inc(x_30); -lean_inc(x_31); -x_93 = l_Lean_addMacroScope(x_31, x_92, x_30); -x_94 = l_Lake_DSL_expandDepSpec___lambda__1___closed__65; +x_93 = l_Lean_Syntax_node3(x_14, x_86, x_92, x_70, x_6); +x_94 = l_Lake_DSL_expandDepSpec___lambda__1___closed__68; +lean_inc(x_32); +lean_inc(x_33); +x_95 = l_Lean_addMacroScope(x_33, x_94, x_32); +x_96 = l_Lake_DSL_expandDepSpec___lambda__1___closed__67; lean_inc(x_14); -x_95 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_95, 0, x_14); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_95, 2, x_93); -lean_ctor_set(x_95, 3, x_38); +x_97 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_97, 0, x_14); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); +lean_ctor_set(x_97, 3, x_40); lean_inc(x_23); lean_inc(x_14); -x_96 = l_Lean_Syntax_node2(x_14, x_73, x_95, x_23); -lean_inc(x_68); +x_98 = l_Lean_Syntax_node2(x_14, x_75, x_97, x_23); +lean_inc(x_70); lean_inc(x_14); -x_97 = l_Lean_Syntax_node3(x_14, x_84, x_96, x_68, x_10); -x_98 = l_Lake_DSL_expandDepSpec___lambda__1___closed__69; -x_99 = l_Lean_addMacroScope(x_31, x_98, x_30); -x_100 = l_Lake_DSL_expandDepSpec___lambda__1___closed__68; +x_99 = l_Lean_Syntax_node3(x_14, x_86, x_98, x_70, x_10); +x_100 = l_Lake_DSL_expandDepSpec___lambda__1___closed__71; +x_101 = l_Lean_addMacroScope(x_33, x_100, x_32); +x_102 = l_Lake_DSL_expandDepSpec___lambda__1___closed__70; lean_inc(x_14); -x_101 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_101, 0, x_14); -lean_ctor_set(x_101, 1, x_100); -lean_ctor_set(x_101, 2, x_99); -lean_ctor_set(x_101, 3, x_38); +x_103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_103, 0, x_14); +lean_ctor_set(x_103, 1, x_102); +lean_ctor_set(x_103, 2, x_101); +lean_ctor_set(x_103, 3, x_40); lean_inc(x_23); lean_inc(x_14); -x_102 = l_Lean_Syntax_node2(x_14, x_73, x_101, x_23); -lean_inc(x_78); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_78); -lean_ctor_set(x_103, 1, x_38); -x_104 = l_Lake_DSL_expandDepSpec___lambda__1___closed__72; +x_104 = l_Lean_Syntax_node2(x_14, x_75, x_103, x_23); +lean_inc(x_80); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_80); +lean_ctor_set(x_105, 1, x_40); +x_106 = l_Lake_DSL_expandDepSpec___lambda__1___closed__74; lean_inc_n(x_23, 2); lean_inc(x_14); -x_105 = l_Lean_Syntax_node2(x_14, x_104, x_23, x_23); +x_107 = l_Lean_Syntax_node2(x_14, x_106, x_23, x_23); if (lean_obj_tag(x_4) == 0) { -x_106 = x_22; -goto block_166; +x_108 = x_22; +goto block_170; } else { -lean_object* x_167; lean_object* x_168; -x_167 = lean_ctor_get(x_4, 0); -lean_inc(x_167); +lean_object* x_171; lean_object* x_172; +x_171 = lean_ctor_get(x_4, 0); +lean_inc(x_171); lean_dec(x_4); -x_168 = l_Array_mkArray1___rarg(x_167); -x_106 = x_168; -goto block_166; +x_172 = l_Array_mkArray1___rarg(x_171); +x_108 = x_172; +goto block_170; } -block_166: +block_170: { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_107 = l_Array_append___rarg(x_22, x_106); -lean_dec(x_106); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_109 = l_Array_append___rarg(x_22, x_108); +lean_dec(x_108); lean_inc(x_14); -x_108 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_108, 0, x_14); -lean_ctor_set(x_108, 1, x_21); -lean_ctor_set(x_108, 2, x_107); -x_109 = l_Lake_DSL_expandDepSpec___lambda__1___closed__16; +x_110 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_110, 0, x_14); +lean_ctor_set(x_110, 1, x_21); +lean_ctor_set(x_110, 2, x_109); +x_111 = l_Lake_DSL_expandDepSpec___lambda__1___closed__18; lean_inc_n(x_23, 4); lean_inc(x_14); -x_110 = l_Lean_Syntax_node6(x_14, x_109, x_108, x_50, x_23, x_23, x_23, x_23); -if (lean_obj_tag(x_76) == 0) +x_112 = l_Lean_Syntax_node6(x_14, x_111, x_110, x_52, x_23, x_23, x_23, x_23); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_155; -x_155 = l_Lean_quoteNameMk(x_75); -x_111 = x_155; -goto block_154; +lean_object* x_159; +x_159 = l_Lean_quoteNameMk(x_77); +x_113 = x_159; +goto block_158; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_75); -x_156 = lean_ctor_get(x_76, 0); -lean_inc(x_156); -lean_dec(x_76); -x_157 = l_Lake_DSL_expandDepSpec___lambda__1___closed__75; -x_158 = l_String_intercalate(x_157, x_156); -x_159 = l_Lake_DSL_expandDepSpec___lambda__1___closed__76; -x_160 = lean_string_append(x_159, x_158); -lean_dec(x_158); -x_161 = l_Lean_Syntax_mkNameLit(x_160, x_56); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_38); -x_163 = lean_array_mk(x_162); -x_164 = l_Lake_DSL_expandDepSpec___lambda__1___closed__74; -x_165 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_165, 0, x_56); -lean_ctor_set(x_165, 1, x_164); -lean_ctor_set(x_165, 2, x_163); -x_111 = x_165; -goto block_154; -} -block_154: -{ -lean_object* x_112; -lean_inc(x_68); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_77); +x_160 = lean_ctor_get(x_78, 0); +lean_inc(x_160); +lean_dec(x_78); +x_161 = l_Lake_DSL_expandDepSpec___lambda__1___closed__77; +x_162 = l_String_intercalate(x_161, x_160); +x_163 = l_Lake_DSL_expandDepSpec___lambda__1___closed__78; +x_164 = lean_string_append(x_163, x_162); +lean_dec(x_162); +x_165 = l_Lean_Syntax_mkNameLit(x_164, x_58); +x_166 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_40); +x_167 = lean_array_mk(x_166); +x_168 = l_Lake_DSL_expandDepSpec___lambda__1___closed__76; +x_169 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_169, 0, x_58); +lean_ctor_set(x_169, 1, x_168); +lean_ctor_set(x_169, 2, x_167); +x_113 = x_169; +goto block_158; +} +block_158: +{ +lean_object* x_114; +lean_inc(x_70); lean_inc(x_14); -x_112 = l_Lean_Syntax_node3(x_14, x_84, x_74, x_68, x_111); +x_114 = l_Lean_Syntax_node3(x_14, x_86, x_76, x_70, x_113); if (lean_obj_tag(x_3) == 0) { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_inc(x_68); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +lean_inc(x_70); lean_inc(x_14); -x_113 = l_Lean_Syntax_node3(x_14, x_84, x_102, x_68, x_29); -x_114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_103); -lean_inc(x_78); -x_115 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_115, 0, x_78); -lean_ctor_set(x_115, 1, x_114); +x_115 = l_Lean_Syntax_node3(x_14, x_86, x_104, x_70, x_31); x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_97); -lean_ctor_set(x_116, 1, x_115); -lean_inc(x_78); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_105); +lean_inc(x_80); x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_78); +lean_ctor_set(x_117, 0, x_80); lean_ctor_set(x_117, 1, x_116); x_118 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_118, 0, x_91); +lean_ctor_set(x_118, 0, x_99); lean_ctor_set(x_118, 1, x_117); -lean_inc(x_78); +lean_inc(x_80); x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_78); +lean_ctor_set(x_119, 0, x_80); lean_ctor_set(x_119, 1, x_118); x_120 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_120, 0, x_85); +lean_ctor_set(x_120, 0, x_93); lean_ctor_set(x_120, 1, x_119); +lean_inc(x_80); x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_78); +lean_ctor_set(x_121, 0, x_80); lean_ctor_set(x_121, 1, x_120); x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_112); +lean_ctor_set(x_122, 0, x_87); lean_ctor_set(x_122, 1, x_121); -x_123 = lean_array_mk(x_122); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_80); +lean_ctor_set(x_123, 1, x_122); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_114); +lean_ctor_set(x_124, 1, x_123); +x_125 = lean_array_mk(x_124); lean_inc(x_14); -x_124 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_124, 0, x_14); -lean_ctor_set(x_124, 1, x_21); -lean_ctor_set(x_124, 2, x_123); +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_14); +lean_ctor_set(x_126, 1, x_21); +lean_ctor_set(x_126, 2, x_125); +lean_inc(x_14); +x_127 = l_Lean_Syntax_node1(x_14, x_24, x_126); lean_inc_n(x_23, 2); lean_inc(x_14); -x_125 = l_Lean_Syntax_node6(x_14, x_26, x_16, x_23, x_124, x_25, x_23, x_18); -x_126 = l_Lake_DSL_expandDepSpec___lambda__1___closed__48; +x_128 = l_Lean_Syntax_node6(x_14, x_28, x_16, x_23, x_127, x_27, x_23, x_18); +x_129 = l_Lake_DSL_expandDepSpec___lambda__1___closed__50; lean_inc(x_23); lean_inc(x_14); -x_127 = l_Lean_Syntax_node4(x_14, x_126, x_68, x_125, x_105, x_23); -x_128 = l_Lake_DSL_expandDepSpec___lambda__1___closed__32; +x_130 = l_Lean_Syntax_node4(x_14, x_129, x_70, x_128, x_107, x_23); +x_131 = l_Lake_DSL_expandDepSpec___lambda__1___closed__34; lean_inc(x_14); -x_129 = l_Lean_Syntax_node5(x_14, x_128, x_52, x_58, x_66, x_127, x_23); -x_130 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; -x_131 = l_Lean_Syntax_node2(x_14, x_130, x_110, x_129); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_11); -return x_132; +x_132 = l_Lean_Syntax_node5(x_14, x_131, x_54, x_60, x_68, x_130, x_23); +x_133 = l_Lake_DSL_expandDepSpec___lambda__1___closed__16; +x_134 = l_Lean_Syntax_node2(x_14, x_133, x_112, x_132); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_11); +return x_135; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; -lean_dec(x_29); -x_133 = lean_ctor_get(x_3, 0); -lean_inc(x_133); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_31); +x_136 = lean_ctor_get(x_3, 0); +lean_inc(x_136); lean_dec(x_3); -lean_inc(x_68); +lean_inc(x_70); lean_inc(x_14); -x_134 = l_Lean_Syntax_node3(x_14, x_84, x_102, x_68, x_133); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_103); -lean_inc(x_78); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_78); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_137, 0, x_97); -lean_ctor_set(x_137, 1, x_136); -lean_inc(x_78); +x_137 = l_Lean_Syntax_node3(x_14, x_86, x_104, x_70, x_136); x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_78); -lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_105); +lean_inc(x_80); x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_91); +lean_ctor_set(x_139, 0, x_80); lean_ctor_set(x_139, 1, x_138); -lean_inc(x_78); x_140 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_140, 0, x_78); +lean_ctor_set(x_140, 0, x_99); lean_ctor_set(x_140, 1, x_139); +lean_inc(x_80); x_141 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_141, 0, x_85); +lean_ctor_set(x_141, 0, x_80); lean_ctor_set(x_141, 1, x_140); x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_78); +lean_ctor_set(x_142, 0, x_93); lean_ctor_set(x_142, 1, x_141); +lean_inc(x_80); x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_112); +lean_ctor_set(x_143, 0, x_80); lean_ctor_set(x_143, 1, x_142); -x_144 = lean_array_mk(x_143); +x_144 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_144, 0, x_87); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_145, 0, x_80); +lean_ctor_set(x_145, 1, x_144); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_114); +lean_ctor_set(x_146, 1, x_145); +x_147 = lean_array_mk(x_146); +lean_inc(x_14); +x_148 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_148, 0, x_14); +lean_ctor_set(x_148, 1, x_21); +lean_ctor_set(x_148, 2, x_147); lean_inc(x_14); -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_14); -lean_ctor_set(x_145, 1, x_21); -lean_ctor_set(x_145, 2, x_144); +x_149 = l_Lean_Syntax_node1(x_14, x_24, x_148); lean_inc_n(x_23, 2); lean_inc(x_14); -x_146 = l_Lean_Syntax_node6(x_14, x_26, x_16, x_23, x_145, x_25, x_23, x_18); -x_147 = l_Lake_DSL_expandDepSpec___lambda__1___closed__48; +x_150 = l_Lean_Syntax_node6(x_14, x_28, x_16, x_23, x_149, x_27, x_23, x_18); +x_151 = l_Lake_DSL_expandDepSpec___lambda__1___closed__50; lean_inc(x_23); lean_inc(x_14); -x_148 = l_Lean_Syntax_node4(x_14, x_147, x_68, x_146, x_105, x_23); -x_149 = l_Lake_DSL_expandDepSpec___lambda__1___closed__32; +x_152 = l_Lean_Syntax_node4(x_14, x_151, x_70, x_150, x_107, x_23); +x_153 = l_Lake_DSL_expandDepSpec___lambda__1___closed__34; lean_inc(x_14); -x_150 = l_Lean_Syntax_node5(x_14, x_149, x_52, x_58, x_66, x_148, x_23); -x_151 = l_Lake_DSL_expandDepSpec___lambda__1___closed__14; -x_152 = l_Lean_Syntax_node2(x_14, x_151, x_110, x_150); -x_153 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_11); -return x_153; +x_154 = l_Lean_Syntax_node5(x_14, x_153, x_54, x_60, x_68, x_152, x_23); +x_155 = l_Lake_DSL_expandDepSpec___lambda__1___closed__16; +x_156 = l_Lean_Syntax_node2(x_14, x_155, x_112, x_154); +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_11); +return x_157; } } } @@ -5549,6 +5579,10 @@ l_Lake_DSL_expandDepSpec___lambda__1___closed__75 = _init_l_Lake_DSL_expandDepSp lean_mark_persistent(l_Lake_DSL_expandDepSpec___lambda__1___closed__75); l_Lake_DSL_expandDepSpec___lambda__1___closed__76 = _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__76(); lean_mark_persistent(l_Lake_DSL_expandDepSpec___lambda__1___closed__76); +l_Lake_DSL_expandDepSpec___lambda__1___closed__77 = _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__77(); +lean_mark_persistent(l_Lake_DSL_expandDepSpec___lambda__1___closed__77); +l_Lake_DSL_expandDepSpec___lambda__1___closed__78 = _init_l_Lake_DSL_expandDepSpec___lambda__1___closed__78(); +lean_mark_persistent(l_Lake_DSL_expandDepSpec___lambda__1___closed__78); l_Lake_DSL_expandDepSpec___lambda__2___closed__1 = _init_l_Lake_DSL_expandDepSpec___lambda__2___closed__1(); lean_mark_persistent(l_Lake_DSL_expandDepSpec___lambda__2___closed__1); l_Lake_DSL_expandDepSpec___lambda__2___closed__2 = _init_l_Lake_DSL_expandDepSpec___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c index df576e54a4e4..10358bd32034 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c @@ -19,7 +19,6 @@ size_t lean_usize_shift_right(size_t, size_t); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); @@ -49,6 +48,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Compiler_LC lean_object* l_Lean_Compiler_LCNF_normalizeFVarIds(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_instInhabited(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___closed__2; +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(lean_object*, lean_object*); uint64_t l___private_Lean_Compiler_LCNF_DeclHash_0__Lean_Compiler_LCNF_hashDecl____x40_Lean_Compiler_LCNF_DeclHash___hyg_269_(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); @@ -128,7 +128,7 @@ else lean_object* x_9; uint8_t x_10; x_9 = lean_array_fget(x_1, x_4); lean_inc(x_5); -x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_5, x_9); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_5, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -200,7 +200,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_3, x_12); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_3, x_12); if (x_14 == 0) { lean_object* x_15; @@ -261,7 +261,7 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_3, x_27); +x_29 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_3, x_27); if (x_29 == 0) { lean_object* x_30; @@ -407,7 +407,7 @@ else lean_object* x_17; uint8_t x_18; x_17 = lean_array_fget(x_5, x_2); lean_inc(x_3); -x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_3, x_17); +x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_3, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -505,7 +505,7 @@ x_19 = lean_ctor_get(x_15, 0); x_20 = lean_ctor_get(x_15, 1); lean_inc(x_19); lean_inc(x_4); -x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_4, x_19); +x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_4, x_19); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -541,7 +541,7 @@ lean_inc(x_26); lean_dec(x_15); lean_inc(x_26); lean_inc(x_4); -x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_4, x_26); +x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_4, x_26); if (x_28 == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; @@ -664,7 +664,7 @@ if (lean_is_exclusive(x_57)) { } lean_inc(x_60); lean_inc(x_4); -x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_4, x_60); +x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_4, x_60); if (x_63 == 0) { lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c index 1c8d858b4937..780039169da8 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c @@ -23,6 +23,7 @@ static lean_object* l_Lean_Compiler_LCNF_instInhabitedParam___closed__2; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__2; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_noinlineAttr(lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Decl_alwaysInlineAttr(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp(lean_object*, lean_object*); extern lean_object* l_Lean_instFVarIdSetInhabited; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetDeclCoreImp(lean_object*, lean_object*, lean_object*); @@ -59,7 +60,6 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean lean_object* l_Lean_Expr_lit___override(lean_object*); static lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore___closed__1; uint8_t l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_40____spec__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(lean_object*, lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instHashableLetValue___closed__1; @@ -141,7 +141,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_markRecDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_isDecl(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_sizeLe___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_go(lean_object*, lean_object*); @@ -154,9 +153,9 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Param_toArg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_sizeLe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1(lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateConstImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqAlt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqDecl___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_CasesCore_getCtorNames___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); @@ -170,6 +169,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_alwaysInlineAttr___boxed(lean LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instantiateRangeArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instCode(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__1; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike(lean_object*, lean_object*, lean_object*, lean_object*); @@ -223,6 +223,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateParamCoreImp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedCasesCore(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_getArity___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_attachCodeDecls(lean_object*, lean_object*); lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, lean_object*); @@ -246,6 +247,7 @@ static lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Comp LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_844____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_1084_(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateReturnImp___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases___boxed(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instantiateRangeArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -291,7 +293,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_ static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltImp___closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_hashLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_341____boxed(lean_object*); -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_AltCore_getParams(lean_object*); @@ -326,10 +327,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParamsNoCache(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_inlineable___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Code_isReturnOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instLetDecl(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_instBEqCode___closed__1; @@ -370,7 +371,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_L static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_size_go___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_hasLocalInst___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLitValue____x40_Lean_Compiler_LCNF_Basic___hyg_262____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1(lean_object*); @@ -6618,7 +6618,7 @@ x_1 = l_Lean_Compiler_LCNF_instInhabitedDecl___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6665,7 +6665,7 @@ return x_10; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; uint8_t x_9; @@ -6706,7 +6706,7 @@ return x_17; } } } -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; uint8_t x_19; uint8_t x_26; @@ -6820,7 +6820,7 @@ return x_35; else { uint8_t x_36; -x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_32, lean_box(0)); +x_36 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2(x_6, x_14, lean_box(0), x_6, x_14, x_32, lean_box(0)); lean_dec(x_14); lean_dec(x_6); if (x_36 == 0) @@ -6906,7 +6906,7 @@ if (x_9 == 0) if (x_17 == 0) { uint8_t x_21; -x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1(x_10, x_18); +x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1(x_10, x_18); return x_21; } else @@ -6931,7 +6931,7 @@ return x_23; else { uint8_t x_24; -x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1(x_10, x_18); +x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1(x_10, x_18); return x_24; } } @@ -6939,20 +6939,20 @@ return x_24; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__1(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__1(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -6961,11 +6961,11 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_1, x_2); x_4 = lean_box(x_3); return x_4; } @@ -6974,7 +6974,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_instBEqDecl___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951____boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c b/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c index efe88bad0af2..26ffe32267d4 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Specialize.c @@ -18,6 +18,7 @@ lean_object* l_Lean_Compiler_LCNF_LetDecl_updateValue(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_findSpecCache_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__4; static double l_Lean_addTrace___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__5___closed__2; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5; static lean_object* l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_Specialize_isGround___spec__1___rarg___closed__1; lean_object* l_Lean_SMap_insert___at_Lean_Compiler_SpecState_addEntry___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_LetValue_forFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__3(lean_object*, lean_object*); @@ -35,7 +36,6 @@ lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Simp_withDiscrCtorIm LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__11; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11; static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__14; lean_object* l_Lean_Compiler_LCNF_Code_inferType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); @@ -47,6 +47,7 @@ lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15; static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Specialize_expandCodeDecls___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_shouldSpecialize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -74,54 +75,46 @@ static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__ static lean_object* l_Lean_Compiler_LCNF_specialize___closed__2; static lean_object* l_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM___closed__1; lean_object* l_Lean_SMap_find_x3f___at_Lean_Compiler_getCachedSpecialization___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277_(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_cacheSpec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Specialize_expandCodeDecls___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_mkSpecDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_FunDeclCore_toExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_Specialize_shouldSpecialize___spec__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_shouldSpecialize___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____lambda__1___closed__1; lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp_go___at_Lean_Compiler_LCNF_Specialize_visitCode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__5___closed__2; static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____lambda__1___closed__2; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____lambda__1___closed__6; lean_object* l_Lean_Compiler_LCNF_mkForallParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; static lean_object* l_Lean_Compiler_LCNF_Specialize_cacheSpec___closed__2; static lean_object* l_Lean_Compiler_LCNF_Specialize_instInhabitedCacheEntry___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_addFVarSubst___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Specialize_shouldSpecialize___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12; lean_object* l_Lean_Compiler_LCNF_eraseDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_abstract(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_setLevelParams(lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_mkSpecDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_shouldSpecialize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specCacheExt; lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5; lean_object* l_Lean_Compiler_LCNF_Decl_internalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_ToExpr_abstractM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_specialize___closed__1; @@ -149,6 +142,7 @@ lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Met uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Compiler_LCNF_LCtx_addFunDecl___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_specialize___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Compiler_LCNF_addFVarSubst___spec__2(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11; static lean_object* l_Lean_Compiler_LCNF_specialize___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Specialize_getRemainingArgs___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__3(lean_object*, lean_object*, lean_object*); @@ -160,6 +154,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_withLetDecl___rarg(lean LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM___spec__2(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_findSpecCache_x3f___closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_addEntry(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); @@ -196,7 +191,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_visitFunDecl(lean_objec LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__4(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__9; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3; static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__7; lean_object* l_Lean_Compiler_LCNF_LetValue_toExpr(lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__4___closed__1; @@ -208,7 +202,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_mkKey(lean_object*, lea LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_Specialize_Collector_collect___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Specialize_mkSpecDecl_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6; lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Compiler_LCNF_addFVarSubst___spec__5(lean_object*, lean_object*, lean_object*); lean_object* l_OptionT_instMonad___rarg(lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); @@ -235,7 +228,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_S LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Compiler_LCNF_Specialize_mkSpecDecl_go___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getSpecParamInfo_x3f___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__5; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Specialize_Collector_collect___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_cacheSpec___closed__1; static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__6; @@ -247,6 +239,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_expandCodeDecls_go___bo LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10; lean_object* l_List_mapTR_loop___at_Lean_Compiler_LCNF_saveSpecParamInfo___spec__10(lean_object*, lean_object*); lean_object* l_Lean_RBTree_contains___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -257,6 +250,7 @@ static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_ lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_getRemainingArgs(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4; LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__3___closed__10; lean_object* l_Lean_Compiler_LCNF_Arg_toExpr(lean_object*); @@ -267,9 +261,9 @@ lean_object* l_Lean_Compiler_LCNF_CodeDecl_fvarId(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Specialize_isGround___rarg___lambda__1(lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashable; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_Specialize_shouldSpecialize___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2; lean_object* l_Lean_Compiler_LCNF_allFVarM_go___at_Lean_Compiler_LCNF_allFVar___spec__2(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Decl_specialize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -301,6 +295,9 @@ static lean_object* l_Lean_Compiler_LCNF_Specialize_instMonadScopeSpecializeM___ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_mkKey___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__10; uint8_t l_Lean_Expr_hasFVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275_(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_withParams___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -311,11 +308,12 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_visitCode___lambda__1(l LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_getRemainingArgs___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SMap_instInhabited___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12; LEAN_EXPORT uint8_t l_Lean_Compiler_LCNF_Specialize_Collector_collect___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Expr_forFVarM___at_Lean_Compiler_LCNF_Specialize_withLetDecl___spec__5___closed__4; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8; lean_object* l_Lean_Expr_instantiateLevelParamsNoCache(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_ToExpr_withParams_go___at_Lean_Compiler_LCNF_Specialize_mkKey___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*); @@ -329,6 +327,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_isGround___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__7; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_visitCode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__2___closed__1; @@ -342,6 +341,7 @@ static lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_mkSpecDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_getSpecParamInfo_x3f___at_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_Collector_collect___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_specializeApp_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Specialize_isGround___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -11764,7 +11764,7 @@ lean_dec(x_1); return x_12; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11774,7 +11774,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11784,27 +11784,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5() { _start: { lean_object* x_1; @@ -11812,17 +11812,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7() { _start: { lean_object* x_1; @@ -11830,57 +11830,57 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11; x_2 = l_Lean_Compiler_LCNF_Specialize_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_49____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13() { _start: { lean_object* x_1; @@ -11888,33 +11888,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14; -x_2 = lean_unsigned_to_nat(4277u); +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14; +x_2 = lean_unsigned_to_nat(4275u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -12191,37 +12191,37 @@ l_Lean_Compiler_LCNF_specialize___closed__3 = _init_l_Lean_Compiler_LCNF_special lean_mark_persistent(l_Lean_Compiler_LCNF_specialize___closed__3); l_Lean_Compiler_LCNF_specialize = _init_l_Lean_Compiler_LCNF_specialize(); lean_mark_persistent(l_Lean_Compiler_LCNF_specialize); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__4); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__5); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__6); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__9); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__10); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__11); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__12); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__13); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__14); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277____closed__15); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4277_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__4); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__5); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__6); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__7); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__8); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__9); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__10); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__11); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__12); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__13); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__14); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275____closed__15); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Specialize___hyg_4275_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Testing.c b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c index 6d5e36feb6d3..1b4aac151681 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Testing.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Testing.c @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertReducesOrPreservesSi lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertSize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Testing_assertSize___lambda__1___closed__2; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Compiler_LCNF_Testing_assertIsAtFixPoint___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,6 +144,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Testing_assertAroundEachOccurrence(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Testing_0__Lean_Compiler_LCNF_Testing_assertAfterTest___elambda__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Compiler_LCNF_Testing_assertNoFun___spec__4___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Code_forM_go___at_Lean_Compiler_LCNF_Testing_assertNoFun___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -3275,7 +3275,7 @@ x_11 = lean_nat_sub(x_6, x_10); lean_dec(x_6); x_12 = lean_array_fget(x_4, x_11); x_13 = lean_array_fget(x_5, x_11); -x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5957_(x_12, x_13); +x_14 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5951_(x_12, x_13); if (x_14 == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data.c b/stage0/stdlib/Lean/Data.c index 042ea00c5b26..bba4f445d54f 100644 --- a/stage0/stdlib/Lean/Data.c +++ b/stage0/stdlib/Lean/Data.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Data -// Imports: Lean.Data.AssocList Lean.Data.Format Lean.Data.HashMap Lean.Data.HashSet Lean.Data.Json Lean.Data.JsonRpc Lean.Data.KVMap Lean.Data.LBool Lean.Data.LOption Lean.Data.Lsp Lean.Data.Name Lean.Data.NameMap Lean.Data.OpenDecl Lean.Data.Options Lean.Data.PersistentArray Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet Lean.Data.Position Lean.Data.PrefixTree Lean.Data.SMap Lean.Data.Trie Lean.Data.Xml Lean.Data.NameTrie Lean.Data.RBTree Lean.Data.RBMap Lean.Data.Rat Lean.Data.RArray +// Imports: Lean.Data.AssocList Lean.Data.Format Lean.Data.HashMap Lean.Data.HashSet Lean.Data.Json Lean.Data.JsonRpc Lean.Data.KVMap Lean.Data.LBool Lean.Data.LOption Lean.Data.Lsp Lean.Data.Name Lean.Data.NameMap Lean.Data.OpenDecl Lean.Data.Options Lean.Data.PersistentArray Lean.Data.PersistentHashMap Lean.Data.PersistentHashSet Lean.Data.Position Lean.Data.PrefixTree Lean.Data.SMap Lean.Data.Trie Lean.Data.Xml Lean.Data.NameTrie Lean.Data.RBTree Lean.Data.RBMap Lean.Data.RArray #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -38,7 +38,6 @@ lean_object* initialize_Lean_Data_Xml(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_NameTrie(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_RBTree(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_RBMap(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Rat(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_RArray(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Data(uint8_t builtin, lean_object* w) { @@ -120,9 +119,6 @@ lean_dec_ref(res); res = initialize_Lean_Data_RBMap(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Data_Rat(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Data_RArray(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Data/Lsp/Basic.c b/stage0/stdlib/Lean/Data/Lsp/Basic.c index 32a9dea6abde..fec3c40fc81d 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Basic.c +++ b/stage0/stdlib/Lean/Data/Lsp/Basic.c @@ -15,6 +15,7 @@ extern "C" { #endif lean_object* l_Lean_JsonNumber_fromNat(lean_object*); static lean_object* l_Lean_Lsp_instToJsonMarkupContent___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__9; static lean_object* l_Lean_Lsp_instBEqRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___lambda__1___boxed(lean_object*); @@ -28,18 +29,16 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit__ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__13; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2; lean_object* l_Lean_Json_getObj_x3f(lean_object*); static lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedLocation; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2795____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__16; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__2; @@ -58,15 +57,16 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532_(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonCancelParams___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3628____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__11; static lean_object* l_Lean_Lsp_instHashablePosition___closed__1; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentFilter; LEAN_EXPORT lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams(lean_object*); @@ -88,6 +88,8 @@ static lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__17; +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupContent; +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableMarkupKind; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__6; @@ -105,6 +107,7 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFi LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedCancelParams___closed__1; static lean_object* l_Lean_Lsp_instOrdRange___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__25___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__26(lean_object*, lean_object*); @@ -118,10 +121,12 @@ uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); static lean_object* l_Lean_Lsp_instToJsonTextDocumentItem___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7; static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__1; static lean_object* l_Lean_Lsp_instFromJsonApplyWorkspaceEditParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__16(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupContent; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__1; @@ -132,23 +137,27 @@ lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCreateFile___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2; uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Lsp_MarkupKind_ofNat(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__3; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__11; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3; static lean_object* l_Lean_Lsp_instToJsonPartialResultParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__3; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4; lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__4; @@ -165,16 +174,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnipp LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3904____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextEdit(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__15; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__21___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__9; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2349____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); @@ -185,19 +190,25 @@ LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__L static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__21; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__4; LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3397____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__3; static lean_object* l_Lean_Lsp_instBEqLocation___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__17; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3; +uint64_t lean_string_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__21___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__22(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__12___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRange; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__27___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1; static lean_object* l_Lean_Lsp_instOrdPosition___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__11; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__5; @@ -220,15 +231,16 @@ static lean_object* l_Lean_Lsp_instToStringTextDocumentPositionParams___closed__ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5496____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2349____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2438____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqLocation; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocation; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__14; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjVal_x3f(lean_object*, lean_object*); @@ -237,11 +249,9 @@ static lean_object* l_Lean_Lsp_instToJsonSnippetString___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instToJsonOptions; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__15; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRange___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__13; @@ -261,15 +271,16 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCommand; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403_(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordLocation____x40_Lean_Data_Lsp_Basic___hyg_1238_(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentChange(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__15___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__20; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEdit; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__11; @@ -290,6 +301,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__17; LEAN_EXPORT lean_object* l_Lean_Lsp_instLTRange; LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____spec__2(lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__10; static lean_object* l_Lean_Lsp_instToJsonTextDocumentPositionParams___closed__1; static lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___closed__2; @@ -298,8 +310,9 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__7; +static lean_object* l_Lean_Lsp_instHashableMarkupKind___closed__1; +LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupKind(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__17(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(lean_object*, lean_object*); @@ -310,6 +323,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_ofTextDocumentEdit(lean_object static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressOptions; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__2; @@ -326,7 +340,6 @@ lean_object* l_Lean_Json_setObjVal_x21(lean_object*, lean_object*, lean_object*) static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__2; static lean_object* l_Lean_Lsp_instToJsonTextDocumentIdentifier___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__9; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); @@ -334,7 +347,9 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__11; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026____boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__11; @@ -348,12 +363,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCha static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__6; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCancelParams; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7; LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__16; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1___boxed(lean_object*, lean_object*); @@ -368,11 +381,13 @@ static lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instCoeTextEditTextEditBatch(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonRange___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__3; @@ -381,8 +396,6 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFi LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__3; static lean_object* l_Lean_Lsp_instToJsonMarkupKind___closed__2; @@ -399,10 +412,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLocationLink; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceEdit; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_WorkspaceEdit_instEmptyCollection___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instAppendTextEditBatch; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__11; @@ -414,33 +424,34 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeleteFile; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3904____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2; static lean_object* l_Lean_Lsp_instToJsonTextDocumentEdit___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6712_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instHashableMarkupContent___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__17; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1840____closed__1; uint8_t l___private_Lean_Data_JsonRpc_0__Lean_JsonRpc_beqRequestID____x40_Lean_Data_JsonRpc___hyg_35_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPartialResultParams; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCancelParams; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9; +LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__9; @@ -469,12 +480,11 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocu LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__18; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonVersionedTextDocumentIdentifier; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3206____closed__1; @@ -485,6 +495,7 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange___ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonMarkupKind___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__7; @@ -497,15 +508,18 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instLEPosition; lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqPosition; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____spec__5(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2349____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonApplyWorkspaceEditParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDeleteFile___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2606_(lean_object*); @@ -525,17 +539,20 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocument static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5827_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080_(uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonMarkupKind(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeleteFile; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__10; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__23(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5718_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqRange; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier___closed__1; @@ -544,11 +561,12 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation lean_object* l_Array_append___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CreateFile_instToJsonOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonTextEditBatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__16; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__8___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__9(lean_object*, lean_object*); @@ -559,22 +577,22 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object* static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonStaticRegistrationOptions; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6494_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406____closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instLTPosition; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__7; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_ordPosition____x40_Lean_Data_Lsp_Basic___hyg_300____boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__15(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonStaticRegistrationOptions; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__4; LEAN_EXPORT lean_object* l_Lean_RBNode_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____spec__6(lean_object*); @@ -590,14 +608,13 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFi static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___closed__2; static lean_object* l_Lean_Lsp_instFromJsonLocationLink___closed__1; static lean_object* l_Lean_Lsp_instInhabitedPosition___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1; static lean_object* l_Lean_Lsp_instFromJsonPosition___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4799_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__9; static lean_object* l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__23___boxed(lean_object*, lean_object*, lean_object*); @@ -606,20 +623,23 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__1; static lean_object* l_Lean_Lsp_instInhabitedLocation___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupContent___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonTextDocumentRegistrationOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4403____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9; static lean_object* l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_DeleteFile_instFromJsonOptions; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__7; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3; +LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupKind___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__2; @@ -627,11 +647,10 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand_ LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedRange; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__8; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__14(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631_(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqLocation____x40_Lean_Data_Lsp_Basic___hyg_1006_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__6___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__7(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__14; @@ -639,7 +658,8 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_5063____closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4983____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__6; @@ -652,7 +672,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__19; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__19___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__7; @@ -669,8 +688,10 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__9; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(lean_object*, lean_object*); lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableRange; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__5; @@ -678,7 +699,6 @@ static lean_object* l_Lean_Lsp_instAppendTextEditBatch___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__4; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__12; @@ -693,33 +713,33 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocu static lean_object* l_Lean_Lsp_instFromJsonVersionedTextDocumentIdentifier___closed__1; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonWorkspaceEdit___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3206_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_DeleteFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3258____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__4; size_t lean_usize_add(size_t, size_t); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2349____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__17; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____closed__3; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_WorkspaceEdit_instAppend___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____closed__13; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2438____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressEnd; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__5; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2; +LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_ofNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentItem; static lean_object* l_Lean_Lsp_instFromJsonChangeAnnotation___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__19; @@ -735,12 +755,11 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRe size_t lean_array_size(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2795____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentFilter___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__8; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__14___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_toJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3022_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); @@ -749,11 +768,12 @@ static lean_object* l_Lean_Lsp_instToStringPosition___closed__3; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__10; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8; +LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupContent___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__16___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentItem____x40_Lean_Data_Lsp_Basic___hyg_4983_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026_(uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5496_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__15; static lean_object* l_Lean_Lsp_instToJsonChangeAnnotation___closed__1; @@ -766,32 +786,28 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCommand; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__25___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextDocumentItem; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDeleteFile____x40_Lean_Data_Lsp_Basic___hyg_3950____closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__23___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__24(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLocationLink; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRenameFile____x40_Lean_Data_Lsp_Basic___hyg_3688____closed__13; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4; static lean_object* l_Lean_Lsp_instHashableRange___closed__1; lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameFile; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__14; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_458____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__9; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDeleteFile___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____closed__4; lean_object* l_Lean_RBNode_fold___at_Lean_RBMap_mergeBy___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__4; @@ -809,28 +825,31 @@ static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand_ LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__4; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentEdit____x40_Lean_Data_Lsp_Basic___hyg_2658____closed__11; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonStaticRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5746____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2480____closed__9; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____closed__3; uint8_t lean_string_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__2; -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6808_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTextDocumentIdentifier; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonMarkupKind___closed__3; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____closed__16; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365____boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2438_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCreateFile____x40_Lean_Data_Lsp_Basic___hyg_3443____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__12; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentRegistrationOptions____x40_Lean_Data_Lsp_Basic___hyg_5855____closed__2; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3; static lean_object* l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__6; static lean_object* l_Lean_Lsp_instToJsonDocumentChange___closed__5; @@ -849,20 +868,19 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRange; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____closed__13; static lean_object* l_Lean_Lsp_instToStringPosition___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__19___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__20(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6398_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPosition___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonApplyWorkspaceEditParams____x40_Lean_Data_Lsp_Basic___hyg_4841____closed__13; static lean_object* l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1; static lean_object* l_Lean_Lsp_instFromJsonRenameFile___closed__1; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_CreateFile_fromJsonOptions____x40_Lean_Data_Lsp_Basic___hyg_3074____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonDocumentSelector___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLocationLink___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkspaceEdit____x40_Lean_Data_Lsp_Basic___hyg_4439____spec__25(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTextEditBatch(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6; +static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1; static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonDocumentFilter____x40_Lean_Data_Lsp_Basic___hyg_5532____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqCancelParams____x40_Lean_Data_Lsp_Basic___hyg_28____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____closed__3; @@ -22917,6 +22935,105 @@ x_6 = l_Lean_Lsp_MarkupKind_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } +LEAN_EXPORT uint8_t l_Lean_Lsp_MarkupKind_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +else +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_MarkupKind_ofNat___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Lsp_MarkupKind_ofNat(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupKind(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Lsp_MarkupKind_toCtorIdx(x_1); +x_4 = l_Lean_Lsp_MarkupKind_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupKind___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l_Lean_Lsp_instDecidableEqMarkupKind(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026_(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +uint64_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint64_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026____boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint64_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_2); +x_4 = lean_box_uint64(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableMarkupKind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableMarkupKind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instHashableMarkupKind___closed__1; +return x_1; +} +} static lean_object* _init_l_Lean_Lsp_instFromJsonMarkupKind___closed__1() { _start: { @@ -23071,7 +23188,7 @@ x_3 = l_Lean_Lsp_instToJsonMarkupKind(x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23083,19 +23200,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23107,19 +23224,19 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; @@ -23142,7 +23259,7 @@ lean_ctor_set(x_9, 1, x_3); if (x_2 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2; +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2; x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -23154,7 +23271,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4; x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); @@ -23165,11 +23282,11 @@ return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(x_1); lean_dec(x_1); return x_2; } @@ -23178,7 +23295,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____boxed), 1, 0); return x_1; } } @@ -23190,7 +23307,7 @@ x_1 = l_Lean_Lsp_instToJsonMarkupContent___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -23239,7 +23356,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1() { _start: { lean_object* x_1; @@ -23247,39 +23364,39 @@ x_1 = lean_mk_string_unchecked("MarkupContent", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23289,64 +23406,64 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonSnippetString____x40_Lean_Data_Lsp_Basic___hyg_1878____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; x_2 = l_Lean_Lsp_instToJsonDocumentChange___closed__3; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -23356,7 +23473,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -23368,7 +23485,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -23393,7 +23510,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -23405,7 +23522,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10; +x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -23448,11 +23565,11 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -23461,7 +23578,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_), 1, 0); return x_1; } } @@ -23473,7 +23590,100 @@ x_1 = l_Lean_Lsp_instFromJsonMarkupContent___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1() { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get_uint8(x_2, sizeof(void*)*1); +x_6 = lean_ctor_get(x_2, 0); +x_7 = l_Lean_Lsp_instDecidableEqMarkupKind(x_3, x_5); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +else +{ +uint8_t x_9; +x_9 = lean_string_dec_eq(x_4, x_6); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqMarkupContent(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqMarkupContent___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Lsp_instDecidableEqMarkupContent(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; uint64_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; +x_2 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +x_3 = lean_ctor_get(x_1, 0); +x_4 = 0; +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupKind____x40_Lean_Data_Lsp_Basic___hyg_6026_(x_2); +x_6 = lean_uint64_mix_hash(x_4, x_5); +x_7 = lean_string_hash(x_3); +x_8 = lean_uint64_mix_hash(x_6, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452____boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableMarkupContent___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableMarkupContent() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instHashableMarkupContent___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1() { _start: { lean_object* x_1; @@ -23481,7 +23691,7 @@ x_1 = lean_mk_string_unchecked("token", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -23493,7 +23703,7 @@ x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_6, 0, x_4); -x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1; +x_7 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1; lean_ctor_set(x_2, 1, x_6); lean_ctor_set(x_2, 0, x_7); x_8 = lean_box(0); @@ -23529,7 +23739,7 @@ lean_inc(x_19); lean_dec(x_2); x_21 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_21, 0, x_19); -x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1; +x_22 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -23558,11 +23768,11 @@ return x_34; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg), 2, 0); return x_2; } } @@ -23570,7 +23780,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonProgressParams___rarg(lean_object* _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -23583,7 +23793,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instToJsonProgressParams___rarg), 1, return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1() { _start: { lean_object* x_1; @@ -23591,7 +23801,7 @@ x_1 = lean_mk_string_unchecked("message", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2() { _start: { lean_object* x_1; @@ -23599,7 +23809,7 @@ x_1 = lean_mk_string_unchecked("cancellable", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3() { _start: { lean_object* x_1; @@ -23607,7 +23817,7 @@ x_1 = lean_mk_string_unchecked("percentage", 10, 10); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -23629,18 +23839,18 @@ x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1; x_12 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_11, x_3); x_13 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_13, 0, x_4); -x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2; +x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); -x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3; +x_17 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3; x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2438____spec__1(x_17, x_5); x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); @@ -23664,7 +23874,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631_), 1, 0); return x_1; } } @@ -23676,7 +23886,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6398_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6712_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -23700,18 +23910,18 @@ x_10 = lean_box(0); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1; +x_12 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1; x_13 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_12, x_4); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_5); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2; +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_10); -x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3; +x_18 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3; x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonVersionedTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2438____spec__1(x_18, x_6); x_20 = lean_ctor_get(x_1, 1); lean_inc(x_20); @@ -23750,7 +23960,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6398_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressBegin____x40_Lean_Data_Lsp_Basic___hyg_6712_), 1, 0); return x_1; } } @@ -23762,7 +23972,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressBegin___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6494_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6808_(lean_object* x_1) { _start: { uint8_t x_2; @@ -23781,7 +23991,7 @@ x_7 = lean_box(0); x_8 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1; x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_9, x_4); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -23812,7 +24022,7 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1; +x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1; x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_17); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -23831,7 +24041,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6494_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressEnd____x40_Lean_Data_Lsp_Basic___hyg_6808_), 1, 0); return x_1; } } @@ -23843,7 +24053,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -23890,7 +24100,7 @@ return x_12; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1() { _start: { lean_object* x_1; @@ -23898,12 +24108,12 @@ x_1 = lean_mk_string_unchecked("workDoneToken", 13, 13); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); @@ -23918,7 +24128,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864_), 1, 0); return x_1; } } @@ -23930,7 +24140,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -24100,7 +24310,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1() { _start: { lean_object* x_1; @@ -24108,39 +24318,39 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressParams", 22, 22); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5() { _start: { lean_object* x_1; @@ -24148,53 +24358,53 @@ x_1 = lean_mk_string_unchecked("workDoneToken\?", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24203,7 +24413,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24215,7 +24425,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24244,11 +24454,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -24257,7 +24467,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892_), 1, 0); return x_1; } } @@ -24269,7 +24479,7 @@ x_1 = l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1() { _start: { lean_object* x_1; @@ -24277,12 +24487,12 @@ x_1 = lean_mk_string_unchecked("partialResultToken", 18, 18); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_2, x_1); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); @@ -24297,7 +24507,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973_), 1, 0); return x_1; } } @@ -24309,7 +24519,7 @@ x_1 = l_Lean_Lsp_instToJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1() { _start: { lean_object* x_1; @@ -24317,39 +24527,39 @@ x_1 = lean_mk_string_unchecked("PartialResultParams", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5() { _start: { lean_object* x_1; @@ -24357,53 +24567,53 @@ x_1 = lean_mk_string_unchecked("partialResultToken\?", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24412,7 +24622,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24424,7 +24634,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24457,7 +24667,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001_), 1, 0); return x_1; } } @@ -24469,7 +24679,7 @@ x_1 = l_Lean_Lsp_instFromJsonPartialResultParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1() { _start: { lean_object* x_1; @@ -24477,13 +24687,13 @@ x_1 = lean_mk_string_unchecked("workDoneProgress", 16, 16); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -24500,13 +24710,13 @@ x_10 = l_Lean_Json_mkObj(x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080_(x_2); return x_3; } } @@ -24514,7 +24724,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____boxed), 1, 0); return x_1; } } @@ -24526,7 +24736,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1() { _start: { lean_object* x_1; @@ -24534,84 +24744,84 @@ x_1 = lean_mk_string_unchecked("WorkDoneProgressOptions", 23, 23); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__1; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__2; -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1; +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__5; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7; +x_1 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1; +x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -24621,7 +24831,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8; +x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24633,7 +24843,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8; +x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24666,7 +24876,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118_), 1, 0); return x_1; } } @@ -25723,6 +25933,10 @@ l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions = _init_l_Lean_Lsp_instFr lean_mark_persistent(l_Lean_Lsp_instFromJsonTextDocumentRegistrationOptions); l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1 = _init_l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1(); lean_mark_persistent(l_Lean_Lsp_MarkupKind_noConfusion___rarg___closed__1); +l_Lean_Lsp_instHashableMarkupKind___closed__1 = _init_l_Lean_Lsp_instHashableMarkupKind___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableMarkupKind___closed__1); +l_Lean_Lsp_instHashableMarkupKind = _init_l_Lean_Lsp_instHashableMarkupKind(); +lean_mark_persistent(l_Lean_Lsp_instHashableMarkupKind); l_Lean_Lsp_instFromJsonMarkupKind___closed__1 = _init_l_Lean_Lsp_instFromJsonMarkupKind___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonMarkupKind___closed__1); l_Lean_Lsp_instFromJsonMarkupKind___closed__2 = _init_l_Lean_Lsp_instFromJsonMarkupKind___closed__2(); @@ -25739,50 +25953,54 @@ l_Lean_Lsp_instToJsonMarkupKind___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupK lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__1); l_Lean_Lsp_instToJsonMarkupKind___closed__2 = _init_l_Lean_Lsp_instToJsonMarkupKind___closed__2(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupKind___closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149____closed__4); l_Lean_Lsp_instToJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instToJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent___closed__1); l_Lean_Lsp_instToJsonMarkupContent = _init_l_Lean_Lsp_instToJsonMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instToJsonMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__9); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073____closed__10); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201____closed__10); l_Lean_Lsp_instFromJsonMarkupContent___closed__1 = _init_l_Lean_Lsp_instFromJsonMarkupContent___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonMarkupContent___closed__1); l_Lean_Lsp_instFromJsonMarkupContent = _init_l_Lean_Lsp_instFromJsonMarkupContent(); lean_mark_persistent(l_Lean_Lsp_instFromJsonMarkupContent); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6207____rarg___closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6317____closed__3); +l_Lean_Lsp_instHashableMarkupContent___closed__1 = _init_l_Lean_Lsp_instHashableMarkupContent___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent___closed__1); +l_Lean_Lsp_instHashableMarkupContent = _init_l_Lean_Lsp_instHashableMarkupContent(); +lean_mark_persistent(l_Lean_Lsp_instHashableMarkupContent); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6521____rarg___closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressReport____x40_Lean_Data_Lsp_Basic___hyg_6631____closed__3); l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressReport___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressReport = _init_l_Lean_Lsp_instToJsonWorkDoneProgressReport(); @@ -25795,84 +26013,84 @@ l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1 = _init_l_Lean_Lsp_instToJs lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressEnd = _init_l_Lean_Lsp_instToJsonWorkDoneProgressEnd(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressEnd); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instToJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____closed__9); l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressParams = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6659____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6973____closed__1); l_Lean_Lsp_instToJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instToJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams___closed__1); l_Lean_Lsp_instToJsonPartialResultParams = _init_l_Lean_Lsp_instToJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__8); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_6687____closed__9); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPartialResultParams____x40_Lean_Data_Lsp_Basic___hyg_7001____closed__9); l_Lean_Lsp_instFromJsonPartialResultParams___closed__1 = _init_l_Lean_Lsp_instFromJsonPartialResultParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams___closed__1); l_Lean_Lsp_instFromJsonPartialResultParams = _init_l_Lean_Lsp_instFromJsonPartialResultParams(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPartialResultParams); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6766____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7080____closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instToJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instToJsonWorkDoneProgressOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkDoneProgressOptions); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__1); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__2); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__3); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__4); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__5); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__6); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__7); -l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_6804____closed__8); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__1); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__2); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__3); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__4); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__5); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__6); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__7); +l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8 = _init_l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressOptions____x40_Lean_Data_Lsp_Basic___hyg_7118____closed__8); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkDoneProgressOptions___closed__1); l_Lean_Lsp_instFromJsonWorkDoneProgressOptions = _init_l_Lean_Lsp_instFromJsonWorkDoneProgressOptions(); diff --git a/stage0/stdlib/Lean/Data/Lsp/Capabilities.c b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c index 3080c8128d59..90494593641e 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Capabilities.c +++ b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c @@ -32,7 +32,6 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonServerCapabilities; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_545____closed__4; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171____closed__3; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_794____spec__1___boxed(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_545____closed__5; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__13; @@ -48,6 +47,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_from static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_995____closed__6; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171____closed__10; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10053_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1460_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_758____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_545_(lean_object*); @@ -154,7 +154,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonW static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__63; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__34; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171____closed__8; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__8; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonShowDocumentClientCapabilities; @@ -163,7 +162,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonSer LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1460____spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionClientCapabilities; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonTextDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_280_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_794____closed__9; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__5; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____spec__2___boxed(lean_object*, lean_object*); @@ -194,6 +192,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonS LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_758____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_517_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonTextDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_280____closed__5; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1460____closed__15; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceEditClientCapabilities; @@ -216,6 +215,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJs LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_517____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__53; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__33; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemCapabilities; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonTextDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_248____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWorkspaceClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_963_(lean_object*); @@ -254,6 +254,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonTex static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonCompletionClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_157____closed__6; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonTextDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_280____closed__14; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1135____closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__25; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171____closed__14; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1644____closed__68; @@ -295,7 +296,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonW lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonShowDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_436____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceEditClientCapabilities; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9480_(lean_object*); lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1171____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1460____closed__11; @@ -4920,7 +4920,7 @@ lean_inc(x_4); lean_dec(x_2); x_5 = lean_unbox(x_4); lean_dec(x_4); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332_(x_5); +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905_(x_5); x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_1); lean_ctor_set(x_7, 1, x_6); @@ -4948,7 +4948,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9480_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10053_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -5656,7 +5656,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -5707,7 +5707,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_(x_3); +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -5826,7 +5826,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -5877,7 +5877,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_(x_3); +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { diff --git a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c index f7de0e2dc453..0e2e1cc1a880 100644 --- a/stage0/stdlib/Lean/Data/Lsp/CodeActions.c +++ b/stage0/stdlib/Lean/Data/Lsp/CodeActions.c @@ -201,7 +201,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Ls LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCodeActionContext; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_884____closed__11; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1204____spec__3___boxed(lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionDisabled____x40_Lean_Data_Lsp_CodeActions___hyg_698____closed__4; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__6; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_147____closed__4; @@ -276,6 +275,7 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCo static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_147____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionLiteralSupportValueSet____x40_Lean_Data_Lsp_CodeActions___hyg_1820_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_147____closed__20; +lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCodeActionLiteralSupport___closed__1; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1204____closed__19; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049____closed__8; @@ -299,7 +299,6 @@ LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeAction static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__5; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049____closed__38; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_884____closed__12; -lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_838____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1204____closed__27; static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049____closed__18; @@ -318,6 +317,7 @@ static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonRe static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionContext____x40_Lean_Data_Lsp_CodeActions___hyg_147____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeAction____x40_Lean_Data_Lsp_CodeActions___hyg_1130____closed__2; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonResolveSupport____x40_Lean_Data_Lsp_CodeActions___hyg_1702_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2049_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionOptions____x40_Lean_Data_Lsp_CodeActions___hyg_884____closed__18; @@ -2100,7 +2100,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -2138,7 +2138,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -2362,11 +2362,11 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_8); @@ -3525,11 +3525,11 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__1; -x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_3, x_2); +x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); x_6 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__11; -x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6550____spec__1(x_6, x_5); +x_7 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6864____spec__1(x_6, x_5); x_8 = lean_ctor_get(x_1, 2); lean_inc(x_8); x_9 = lean_alloc_ctor(3, 1, 0); @@ -4950,7 +4950,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJ lean_object* x_2; lean_object* x_3; x_2 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4988,7 +4988,7 @@ lean_inc(x_12); lean_dec(x_3); x_13 = l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_fromJsonCodeActionParams____x40_Lean_Data_Lsp_CodeActions___hyg_389____closed__11; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6578____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonWorkDoneProgressParams____x40_Lean_Data_Lsp_Basic___hyg_6892____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; diff --git a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c index 40b423c9bb62..81a15c1c541b 100644 --- a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c +++ b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c @@ -13,1434 +13,1472 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__79; lean_object* l_Lean_JsonNumber_fromNat(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__111; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokens; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5; -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5; static lean_object* l_Lean_Lsp_instToJsonCompletionParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__17; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__90; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__51; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__107; static lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__102; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3389_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceContext; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__17; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__36; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolInformation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817_(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceContext; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(uint8_t, uint8_t, lean_object*); uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat(uint8_t); static lean_object* l_Lean_Lsp_instToJsonCompletionItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHover; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__119; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332_(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348_(uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____boxed(lean_object*, lean_object*); -static uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__2; static lean_object* l_Lean_Lsp_instFromJsonCompletionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokenModifier; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4(lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4; static lean_object* l_Lean_Lsp_instToJsonFoldingRange___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__32; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16; +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1(uint64_t, lean_object*, size_t, size_t, uint64_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__54; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__28; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend; lean_object* l_Lean_Json_mkObj(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__21; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__50; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__135; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__47; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10719_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10053_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8; static lean_object* l_Lean_Lsp_instHashableCallHierarchyItem___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__37; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__35; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__9; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHoverParams; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); static lean_object* l_Lean_Lsp_instBEqSymbolKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__83; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__45; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6; +static lean_object* l_Lean_Lsp_instHashableCompletionItem___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37; static lean_object* l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedSymbolTag; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3899_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonReferenceParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__28; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__142; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__11; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2___boxed(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__103; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10091_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30; lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__106; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894____boxed(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6417_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7660_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_ofNat___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__15; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__127; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7; static lean_object* l_Lean_Lsp_instFromJsonReferenceParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__108; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__1; LEAN_EXPORT uint8_t l_Lean_Lsp_CompletionItemKind_ofNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_names; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__104; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__84; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__80; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__101; LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqCompletionItemKind___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556_(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__136; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2; static lean_object* l_Lean_Lsp_instToJsonDocumentSymbol___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7825_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__43; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__1; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__22; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqCompletionItemKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokens; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11; static lean_object* l_Lean_Lsp_instFromJsonHover___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_ofNat___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__39; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItem; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__3(lean_object*, lean_object*); uint64_t lean_string_hash(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__18; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5; static lean_object* l_Lean_Lsp_instToJsonCompletionItemTag___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__93; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9789_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHover; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__57; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2986_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7; +static lean_object* l_Lean_Lsp_instHashableCompletionItemTag___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__17; static lean_object* l_Lean_Lsp_instToJsonDeclarationParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945_(lean_object*); static lean_object* l_Lean_Lsp_instBEqSymbolTag___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__151; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__19; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_noConfusion___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__126; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__25; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__32; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__138; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4; static lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__27; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__67; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__41; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__125; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokenType; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11; uint8_t l___private_Lean_Data_Json_Basic_0__Lean_Json_beq_x27(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814_(uint8_t, uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20; static lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7; static lean_object* l_Lean_Lsp_instToJsonHover___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPrepareRenameParams; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__77; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSemanticTokenType; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__53; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__21; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__49; static lean_object* l_Lean_Lsp_instToJsonCompletionOptions___closed__1; lean_object* l_Lean_Json_getStr_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__30; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__76; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__22; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHoverParams; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____lambda__1___boxed(lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__72; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__62; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__150; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__18; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241_(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__25; uint64_t l___private_Lean_Data_Json_Basic_0__Lean_Json_hash_x27(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10198_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameOptions; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__43; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCompletionItemTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeclarationParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__96; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5844_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__74; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3961_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3; static lean_object* l_Lean_Lsp_instFromJsonCompletionItem___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__87; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8; lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__118; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__36; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8; LEAN_EXPORT lean_object* l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__45; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1; +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableInsertReplaceEdit; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19; lean_object* lean_nat_to_int(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__56; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__124; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__121; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__147; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__24; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__45; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__47; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__3; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__20; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__23; -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1(uint64_t, lean_object*, size_t, size_t, uint64_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolInformation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenModifier; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14; static lean_object* l_Lean_Lsp_instFromJsonSymbolTag___closed__2; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__9; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__131; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__115; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__41; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__82; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_toCtorIdx(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__146; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__39; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18; uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__25; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259_(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_ofNat(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokenType___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__58; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9255_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__32; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__24; static lean_object* l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toCtorIdx(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__24; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__95; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10535_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11108_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCompletionItemTag___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3; static lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__123; static lean_object* l_Lean_Lsp_instHashableSymbolTag___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__8; lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2816_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemTag(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10; static lean_object* l_Lean_Lsp_instFromJsonHoverParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__91; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPrepareRenameParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1; static lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__149; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__25; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__50; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3156_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__29; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1; static lean_object* l_Lean_Lsp_instHashableSymbolKind___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1399_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8; +static lean_object* l_Lean_Lsp_instHashableCompletionItemKind___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionList; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__37; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSemanticTokenType; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__105; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321_(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCallHierarchyItem; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCompletionItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__33; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__46; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqCompletionItemTag(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457_(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol_go(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__110; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__2; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_noConfusion___rarg___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqCompletionItemTag___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4771_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__4; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__59; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__98; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__35; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__30; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__33; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__86; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__55; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCompletionItem; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__65; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4378_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__29; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__99; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__117; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894_(uint8_t, uint8_t); static lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__46; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1; static lean_object* l_Lean_Lsp_instToJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__31; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRange; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__144; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__9; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__40; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemKind(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__71; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__97; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__132; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1543_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__122; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7; static lean_object* l_Lean_Lsp_instToJsonDefinitionParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableCompletionItemKind; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__85; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSymbolTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instReprCompletionItemKind; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__130; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__4(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8; static lean_object* l_Lean_Lsp_instFromJsonRenameOptions___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__145; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionList; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__148; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8; lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7; lean_object* lean_nat_abs(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instReprCompletionItemTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toNat(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__45; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__63; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__140; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3660_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__19; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__61; static lean_object* l_Lean_Lsp_instFromJsonSymbolTag___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__21; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__42; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__26; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9625_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__112; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__60; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256_(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedSymbolKind; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__66; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDefinitionParams; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__39; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7087_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__23; static lean_object* l_Lean_Lsp_instToJsonSemanticTokens___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9828_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__128; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__14; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_names; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__35; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__120; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6780_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41; +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolTag(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolInformation; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3805_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCompletionList___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__134; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__116; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__31; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__89; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__44; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4233_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__73; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__88; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__75; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__37; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolKind; static lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642_(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8; uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9; static lean_object* l_Lean_Lsp_instToJsonHoverParams___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__133; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11292_(lean_object*); lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893_(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__52; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__64; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__141; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__31; lean_object* l_Lean_Json_getNat_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolTag; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeclarationParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instBEqCompletionItem___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9; +uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609_(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__113; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__22; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__114; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9480_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__52; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__41; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__94; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4; lean_object* lean_array_mk(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7252_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonCompletionItemTag___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__48; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolResult(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11; static lean_object* l_Lean_Lsp_instHashableSemanticTokenType___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__38; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__16; -lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedCompletionItemKind; size_t lean_usize_add(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3729_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__68; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__78; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItem; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSymbolKind; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18; static lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1; lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toNat___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag(lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemTag___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__4; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__13; static lean_object* l_Lean_Lsp_instReprCompletionItemTag___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__15; static lean_object* l_Lean_Lsp_instFromJsonCompletionOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolInformation; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12; lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17; static lean_object* l_Lean_Lsp_instFromJsonRenameParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__81; static lean_object* l_Lean_Lsp_instToJsonRenameParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonCompletionList___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24; +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableCompletionItem; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__48; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453_(lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__92; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10664_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10362_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__20; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__152; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905____boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__129; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4534_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlight; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__100; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__139; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__109; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3326_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolKind(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__28; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__69; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__70; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDefinitionParams; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemTag(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1; +static lean_object* l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__36; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqInsertReplaceEdit; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__7; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__143; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__46; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4198_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4; +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1(uint64_t, lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2; static lean_object* l_Lean_Lsp_instReprCompletionItemKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__42; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39; +LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableCompletionItemTag; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Lsp_instBEqSemanticTokenType___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRenameOptions___closed__1; static lean_object* l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__137; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__51; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__7; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___boxed(lean_object*); +static uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7353_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__45; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__49; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__7; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21; static lean_object* l_Lean_Lsp_instToJsonPrepareRenameParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__17; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { @@ -5056,6 +5094,190 @@ x_1 = l_Lean_Lsp_instReprCompletionItemKind___closed__1; return x_1; } } +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +uint64_t x_2; +x_2 = 0; +return x_2; +} +case 1: +{ +uint64_t x_3; +x_3 = 1; +return x_3; +} +case 2: +{ +uint64_t x_4; +x_4 = 2; +return x_4; +} +case 3: +{ +uint64_t x_5; +x_5 = 3; +return x_5; +} +case 4: +{ +uint64_t x_6; +x_6 = 4; +return x_6; +} +case 5: +{ +uint64_t x_7; +x_7 = 5; +return x_7; +} +case 6: +{ +uint64_t x_8; +x_8 = 6; +return x_8; +} +case 7: +{ +uint64_t x_9; +x_9 = 7; +return x_9; +} +case 8: +{ +uint64_t x_10; +x_10 = 8; +return x_10; +} +case 9: +{ +uint64_t x_11; +x_11 = 9; +return x_11; +} +case 10: +{ +uint64_t x_12; +x_12 = 10; +return x_12; +} +case 11: +{ +uint64_t x_13; +x_13 = 11; +return x_13; +} +case 12: +{ +uint64_t x_14; +x_14 = 12; +return x_14; +} +case 13: +{ +uint64_t x_15; +x_15 = 13; +return x_15; +} +case 14: +{ +uint64_t x_16; +x_16 = 14; +return x_16; +} +case 15: +{ +uint64_t x_17; +x_17 = 15; +return x_17; +} +case 16: +{ +uint64_t x_18; +x_18 = 16; +return x_18; +} +case 17: +{ +uint64_t x_19; +x_19 = 17; +return x_19; +} +case 18: +{ +uint64_t x_20; +x_20 = 18; +return x_20; +} +case 19: +{ +uint64_t x_21; +x_21 = 19; +return x_21; +} +case 20: +{ +uint64_t x_22; +x_22 = 20; +return x_22; +} +case 21: +{ +uint64_t x_23; +x_23 = 21; +return x_23; +} +case 22: +{ +uint64_t x_24; +x_24 = 22; +return x_24; +} +case 23: +{ +uint64_t x_25; +x_25 = 23; +return x_25; +} +default: +{ +uint64_t x_26; +x_26 = 24; +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168____boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint64_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(x_2); +x_4 = lean_box_uint64(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItemKind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItemKind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instHashableCompletionItemKind___closed__1; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind(uint8_t x_1) { _start: { @@ -5140,7 +5362,7 @@ return x_17; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1() { _start: { lean_object* x_1; @@ -5148,7 +5370,7 @@ x_1 = lean_mk_string_unchecked("newText", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2() { _start: { lean_object* x_1; @@ -5156,80 +5378,80 @@ x_1 = lean_mk_string_unchecked("InsertReplaceEdit", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10() { _start: { lean_object* x_1; @@ -5237,48 +5459,48 @@ x_1 = lean_mk_string_unchecked("insert", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15() { _start: { lean_object* x_1; @@ -5286,52 +5508,52 @@ x_1 = lean_mk_string_unchecked("replace", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -5343,7 +5565,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5355,7 +5577,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5369,7 +5591,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -5382,7 +5604,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -5394,7 +5616,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -5408,7 +5630,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15; x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -5420,7 +5642,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -5432,7 +5654,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -5478,7 +5700,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398_), 1, 0); return x_1; } } @@ -5490,7 +5712,7 @@ x_1 = l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1399_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1543_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -5503,7 +5725,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -5512,7 +5734,7 @@ x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_3); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -5520,7 +5742,7 @@ x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_4); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -5546,7 +5768,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1399_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1543_), 1, 0); return x_1; } } @@ -5558,6 +5780,112 @@ x_1 = l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; return x_1; } } +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_2, 2); +x_9 = lean_string_dec_eq(x_3, x_6); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +uint8_t x_11; +x_11 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_4, x_7); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_627_(x_5, x_8); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqInsertReplaceEdit() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1; +return x_1; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_1, 2); +x_5 = 0; +x_6 = lean_string_hash(x_2); +x_7 = lean_uint64_mix_hash(x_5, x_6); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_3); +x_9 = lean_uint64_mix_hash(x_7, x_8); +x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_4); +x_11 = lean_uint64_mix_hash(x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699____boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableInsertReplaceEdit() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemTag_toCtorIdx(lean_object* x_1) { _start: { @@ -5653,7 +5981,7 @@ x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1() { _start: { lean_object* x_1; @@ -5661,33 +5989,33 @@ x_1 = lean_mk_string_unchecked("Lean.Lsp.CompletionItemTag.deprecated", 37, 37); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -5695,23 +6023,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_348____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -5719,7 +6047,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; @@ -5728,41 +6056,41 @@ x_3 = lean_nat_dec_le(x_2, x_1); if (x_3 == 0) { lean_object* x_4; lean_object* x_5; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4; x_5 = l_Repr_addAppParen(x_4, x_1); return x_5; } else { lean_object* x_6; lean_object* x_7; -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6; x_7 = l_Repr_addAppParen(x_6, x_1); return x_7; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___boxed), 1, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___boxed), 1, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845_(x_1); lean_dec(x_1); return x_2; } @@ -5771,7 +6099,7 @@ static lean_object* _init_l_Lean_Lsp_instReprCompletionItemTag___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____boxed), 1, 0); return x_1; } } @@ -5783,10 +6111,44 @@ x_1 = l_Lean_Lsp_instReprCompletionItemTag___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemTag___closed__1() { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897_(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; +uint64_t x_2; +x_2 = 0; +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897____boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897_(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItemTag___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1897____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItemTag() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instHashableCompletionItemTag___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemTag___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = lean_unsigned_to_nat(1u); x_2 = l_Lean_JsonNumber_fromNat(x_1); return x_2; @@ -5862,7 +6224,7 @@ return x_6; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5877,7 +6239,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(x_3); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -5928,7 +6290,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(x_3); +x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -6032,7 +6394,7 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6232,7 +6594,7 @@ return x_51; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6247,7 +6609,7 @@ return x_4; case 1: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398_(x_3); if (lean_obj_tag(x_5) == 0) { uint8_t x_6; @@ -6298,7 +6660,7 @@ return x_14; { lean_object* x_15; uint8_t x_16; lean_inc(x_3); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254_(x_3); +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398_(x_3); x_16 = !lean_is_exclusive(x_3); if (x_16 == 0) { @@ -6402,7 +6764,7 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6424,7 +6786,7 @@ return x_6; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -6478,7 +6840,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6514,7 +6876,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; @@ -6570,7 +6932,7 @@ lean_inc(x_24); lean_dec(x_3); x_25 = lean_array_size(x_24); x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6(x_25, x_26, x_24); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6(x_25, x_26, x_24); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -6653,7 +7015,7 @@ return x_47; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1() { _start: { lean_object* x_1; @@ -6661,7 +7023,7 @@ x_1 = lean_mk_string_unchecked("label", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2() { _start: { lean_object* x_1; @@ -6669,80 +7031,80 @@ x_1 = lean_mk_string_unchecked("CompletionItem", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10() { _start: { lean_object* x_1; @@ -6750,7 +7112,7 @@ x_1 = lean_mk_string_unchecked("detail", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11() { _start: { lean_object* x_1; @@ -6758,48 +7120,48 @@ x_1 = lean_mk_string_unchecked("detail\?", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16() { _start: { lean_object* x_1; @@ -6807,7 +7169,7 @@ x_1 = lean_mk_string_unchecked("documentation", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17() { _start: { lean_object* x_1; @@ -6815,48 +7177,48 @@ x_1 = lean_mk_string_unchecked("documentation\?", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22() { _start: { lean_object* x_1; @@ -6864,7 +7226,7 @@ x_1 = lean_mk_string_unchecked("kind", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23() { _start: { lean_object* x_1; @@ -6872,48 +7234,48 @@ x_1 = lean_mk_string_unchecked("kind\?", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28() { _start: { lean_object* x_1; @@ -6921,7 +7283,7 @@ x_1 = lean_mk_string_unchecked("textEdit", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29() { _start: { lean_object* x_1; @@ -6929,48 +7291,48 @@ x_1 = lean_mk_string_unchecked("textEdit\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34() { _start: { lean_object* x_1; @@ -6978,7 +7340,7 @@ x_1 = lean_mk_string_unchecked("sortText", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35() { _start: { lean_object* x_1; @@ -6986,48 +7348,48 @@ x_1 = lean_mk_string_unchecked("sortText\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40() { _start: { lean_object* x_1; @@ -7035,7 +7397,7 @@ x_1 = lean_mk_string_unchecked("data", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41() { _start: { lean_object* x_1; @@ -7043,48 +7405,48 @@ x_1 = lean_mk_string_unchecked("data\?", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46() { _start: { lean_object* x_1; @@ -7092,7 +7454,7 @@ x_1 = lean_mk_string_unchecked("tags", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47() { _start: { lean_object* x_1; @@ -7100,52 +7462,52 @@ x_1 = lean_mk_string_unchecked("tags\?", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -7157,7 +7519,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7169,7 +7531,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7183,7 +7545,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -7196,7 +7558,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7208,7 +7570,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7222,9 +7584,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -7236,7 +7598,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -7248,7 +7610,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -7262,9 +7624,9 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -7277,7 +7639,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -7289,7 +7651,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -7303,9 +7665,9 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28; +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28; lean_inc(x_1); -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3(x_1, x_46); +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -7319,7 +7681,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -7331,7 +7693,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -7345,7 +7707,7 @@ lean_object* x_56; lean_object* x_57; lean_object* x_58; x_56 = lean_ctor_get(x_47, 0); lean_inc(x_56); lean_dec(x_47); -x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34; +x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34; lean_inc(x_1); x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) @@ -7362,7 +7724,7 @@ if (x_59 == 0) { lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_58, 0); -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39; x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); lean_ctor_set(x_58, 0, x_62); @@ -7374,7 +7736,7 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; x_63 = lean_ctor_get(x_58, 0); lean_inc(x_63); lean_dec(x_58); -x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39; +x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39; x_65 = lean_string_append(x_64, x_63); lean_dec(x_63); x_66 = lean_alloc_ctor(0, 1, 0); @@ -7388,14 +7750,14 @@ lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean x_67 = lean_ctor_get(x_58, 0); lean_inc(x_67); lean_dec(x_58); -x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; +x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; lean_inc(x_1); -x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4(x_1, x_68); +x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4(x_1, x_68); x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); lean_dec(x_69); -x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; -x_72 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5(x_1, x_71); +x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; +x_72 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5(x_1, x_71); if (lean_obj_tag(x_72) == 0) { uint8_t x_73; @@ -7411,7 +7773,7 @@ if (x_73 == 0) { lean_object* x_74; lean_object* x_75; lean_object* x_76; x_74 = lean_ctor_get(x_72, 0); -x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51; +x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51; x_76 = lean_string_append(x_75, x_74); lean_dec(x_74); lean_ctor_set(x_72, 0, x_76); @@ -7423,7 +7785,7 @@ lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; x_77 = lean_ctor_get(x_72, 0); lean_inc(x_77); lean_dec(x_72); -x_78 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51; +x_78 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51; x_79 = lean_string_append(x_78, x_77); lean_dec(x_77); x_80 = lean_alloc_ctor(0, 1, 0); @@ -7479,43 +7841,43 @@ return x_86; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__3(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -7523,15 +7885,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__6(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__6(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__5(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__5(x_1, x_2); lean_dec(x_2); return x_3; } @@ -7540,7 +7902,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_), 1, 0); return x_1; } } @@ -7552,7 +7914,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7566,7 +7928,7 @@ else { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = lean_ctor_get(x_2, 0); -x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(x_4); +x_5 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -7578,7 +7940,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7641,7 +8003,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7657,7 +8019,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1399_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1543_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -7669,7 +8031,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -7693,7 +8055,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__4(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7709,194 +8071,1064 @@ uint8_t x_4; x_4 = !lean_is_exclusive(x_2); if (x_4 == 0) { -lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_array_size(x_5); -x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5(x_6, x_7, x_5); -lean_ctor_set_tag(x_2, 4); -lean_ctor_set(x_2, 0, x_8); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_2); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -return x_11; +lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_array_size(x_5); +x_7 = 0; +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5(x_6, x_7, x_5); +lean_ctor_set_tag(x_2, 4); +lean_ctor_set(x_2, 0, x_8); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_2); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_ctor_get(x_2, 0); +lean_inc(x_12); +lean_dec(x_2); +x_13 = lean_array_size(x_12); +x_14 = 0; +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5(x_13, x_14, x_12); +x_16 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_16, 0, x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 4); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 5); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 6); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 7); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_10, 0, x_2); +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_15, x_3); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; +x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1(x_17, x_4); +lean_dec(x_4); +x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; +x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__2(x_19, x_5); +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28; +x_22 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__3(x_21, x_6); +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34; +x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_7); +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; +x_26 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_25, x_8); +lean_dec(x_8); +x_27 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; +x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__4(x_27, x_9); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_13); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_22); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_18); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_16); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_14); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; +x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_36, x_37); +x_39 = l_Lean_Json_mkObj(x_38); +return x_39; +} +} +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404____spec__5(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonCompletionItem___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Lsp_instInhabitedCompletionItem___closed__1; +x_3 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_1); +lean_ctor_set(x_3, 4, x_1); +lean_ctor_set(x_3, 5, x_1); +lean_ctor_set(x_3, 6, x_1); +lean_ctor_set(x_3, 7, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instInhabitedCompletionItem___closed__2; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_string_dec_eq(x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_decEqMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6307_(x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +lean_dec(x_2); +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +lean_dec(x_1); +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +lean_dec(x_2); +x_8 = lean_unbox(x_6); +lean_dec(x_6); +x_9 = lean_unbox(x_7); +lean_dec(x_7); +x_10 = l_Lean_Lsp_instDecidableEqCompletionItemKind(x_8, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1609_(x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = l___private_Lean_Data_Json_Basic_0__Lean_Json_beq_x27(x_6, x_7); +return x_8; +} +} +} +} +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_eq(x_6, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_sub(x_6, x_10); +lean_dec(x_6); +x_3 = lean_box(0); +x_6 = x_11; +x_7 = lean_box(0); +goto _start; +} +else +{ +uint8_t x_13; +lean_dec(x_6); +x_13 = 1; +return x_13; +} +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_array_get_size(x_6); +x_9 = lean_array_get_size(x_7); +x_10 = lean_nat_dec_eq(x_8, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +uint8_t x_11; +lean_dec(x_8); +x_11 = 0; +return x_11; +} +else +{ +uint8_t x_12; +x_12 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7(x_6, x_7, lean_box(0), x_6, x_7, x_8, lean_box(0)); +return x_12; +} +} +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 3); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 4); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 5); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 6); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 7); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_2, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_2, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 4); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = lean_ctor_get(x_2, 6); +lean_inc(x_17); +x_18 = lean_ctor_get(x_2, 7); +lean_inc(x_18); +lean_dec(x_2); +x_19 = lean_string_dec_eq(x_3, x_11); +lean_dec(x_11); +lean_dec(x_3); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = 0; +return x_20; +} +else +{ +uint8_t x_21; +x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_4, x_12); +lean_dec(x_12); +lean_dec(x_4); +if (x_21 == 0) +{ +uint8_t x_22; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_22 = 0; +return x_22; +} +else +{ +uint8_t x_23; +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_5, x_13); +lean_dec(x_13); +lean_dec(x_5); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_24 = 0; +return x_24; +} +else +{ +uint8_t x_25; +x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_6, x_14); +if (x_25 == 0) +{ +uint8_t x_26; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_26 = 0; +return x_26; +} +else +{ +uint8_t x_27; +x_27 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_7, x_15); +lean_dec(x_15); +lean_dec(x_7); +if (x_27 == 0) +{ +uint8_t x_28; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_28 = 0; +return x_28; +} +else +{ +uint8_t x_29; +x_29 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_8, x_16); +lean_dec(x_16); +lean_dec(x_8); +if (x_29 == 0) +{ +uint8_t x_30; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_10); +lean_dec(x_9); +x_30 = 0; +return x_30; +} +else +{ +uint8_t x_31; +x_31 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5(x_9, x_17); +lean_dec(x_17); +lean_dec(x_9); +if (x_31 == 0) +{ +uint8_t x_32; +lean_dec(x_18); +lean_dec(x_10); +x_32 = 0; +return x_32; +} +else +{ +uint8_t x_33; +x_33 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_10, x_18); +lean_dec(x_18); +lean_dec(x_10); +return x_33; +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487_(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instBEqCompletionItem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instBEqCompletionItem___closed__1; +return x_1; +} +} +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1(uint64_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint64_t x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +uint64_t x_7; size_t x_8; size_t x_9; +x_7 = lean_uint64_mix_hash(x_5, x_1); +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +x_5 = x_7; +goto _start; +} +else +{ +return x_5; +} +} +} +static uint64_t _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 7; +x_2 = 13; +x_3 = lean_uint64_mix_hash(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 4); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 5); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 6); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 7); +lean_inc(x_9); +lean_dec(x_1); +x_10 = 0; +x_11 = lean_string_hash(x_2); +lean_dec(x_2); +x_12 = lean_uint64_mix_hash(x_10, x_11); +if (lean_obj_tag(x_3) == 0) +{ +uint64_t x_75; +x_75 = 11; +x_13 = x_75; +goto block_74; +} +else +{ +lean_object* x_76; uint64_t x_77; uint64_t x_78; uint64_t x_79; +x_76 = lean_ctor_get(x_3, 0); +lean_inc(x_76); +lean_dec(x_3); +x_77 = lean_string_hash(x_76); +lean_dec(x_76); +x_78 = 13; +x_79 = lean_uint64_mix_hash(x_77, x_78); +x_13 = x_79; +goto block_74; +} +block_74: +{ +uint64_t x_14; uint64_t x_15; +x_14 = lean_uint64_mix_hash(x_12, x_13); +if (lean_obj_tag(x_4) == 0) +{ +uint64_t x_69; +x_69 = 11; +x_15 = x_69; +goto block_68; +} +else +{ +lean_object* x_70; uint64_t x_71; uint64_t x_72; uint64_t x_73; +x_70 = lean_ctor_get(x_4, 0); +lean_inc(x_70); +lean_dec(x_4); +x_71 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_70); +lean_dec(x_70); +x_72 = 13; +x_73 = lean_uint64_mix_hash(x_71, x_72); +x_15 = x_73; +goto block_68; +} +block_68: +{ +uint64_t x_16; uint64_t x_17; +x_16 = lean_uint64_mix_hash(x_14, x_15); +if (lean_obj_tag(x_5) == 0) +{ +uint64_t x_62; +x_62 = 11; +x_17 = x_62; +goto block_61; +} +else +{ +lean_object* x_63; uint8_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; +x_63 = lean_ctor_get(x_5, 0); +lean_inc(x_63); +lean_dec(x_5); +x_64 = lean_unbox(x_63); +lean_dec(x_63); +x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(x_64); +x_66 = 13; +x_67 = lean_uint64_mix_hash(x_65, x_66); +x_17 = x_67; +goto block_61; +} +block_61: +{ +uint64_t x_18; uint64_t x_19; +x_18 = lean_uint64_mix_hash(x_16, x_17); +if (lean_obj_tag(x_6) == 0) +{ +uint64_t x_56; +x_56 = 11; +x_19 = x_56; +goto block_55; +} +else +{ +lean_object* x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; +x_57 = lean_ctor_get(x_6, 0); +lean_inc(x_57); +lean_dec(x_6); +x_58 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(x_57); +lean_dec(x_57); +x_59 = 13; +x_60 = lean_uint64_mix_hash(x_58, x_59); +x_19 = x_60; +goto block_55; +} +block_55: +{ +uint64_t x_20; uint64_t x_21; +x_20 = lean_uint64_mix_hash(x_18, x_19); +if (lean_obj_tag(x_7) == 0) +{ +uint64_t x_50; +x_50 = 11; +x_21 = x_50; +goto block_49; +} +else +{ +lean_object* x_51; uint64_t x_52; uint64_t x_53; uint64_t x_54; +x_51 = lean_ctor_get(x_7, 0); +lean_inc(x_51); +lean_dec(x_7); +x_52 = lean_string_hash(x_51); +lean_dec(x_51); +x_53 = 13; +x_54 = lean_uint64_mix_hash(x_52, x_53); +x_21 = x_54; +goto block_49; +} +block_49: +{ +uint64_t x_22; uint64_t x_23; +x_22 = lean_uint64_mix_hash(x_20, x_21); +if (lean_obj_tag(x_8) == 0) +{ +uint64_t x_44; +x_44 = 11; +x_23 = x_44; +goto block_43; +} +else +{ +lean_object* x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; +x_45 = lean_ctor_get(x_8, 0); +lean_inc(x_45); +lean_dec(x_8); +x_46 = l___private_Lean_Data_Json_Basic_0__Lean_Json_hash_x27(x_45); +lean_dec(x_45); +x_47 = 13; +x_48 = lean_uint64_mix_hash(x_46, x_47); +x_23 = x_48; +goto block_43; +} +block_43: +{ +uint64_t x_24; +x_24 = lean_uint64_mix_hash(x_22, x_23); +if (lean_obj_tag(x_9) == 0) +{ +uint64_t x_25; uint64_t x_26; +x_25 = 11; +x_26 = lean_uint64_mix_hash(x_24, x_25); +return x_26; +} +else +{ +lean_object* x_27; uint64_t x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_27 = lean_ctor_get(x_9, 0); +lean_inc(x_27); +lean_dec(x_9); +x_28 = 7; +x_29 = lean_array_get_size(x_27); +x_30 = lean_unsigned_to_nat(0u); +x_31 = lean_nat_dec_lt(x_30, x_29); +if (x_31 == 0) +{ +uint64_t x_32; uint64_t x_33; +lean_dec(x_29); +lean_dec(x_27); +x_32 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1; +x_33 = lean_uint64_mix_hash(x_24, x_32); +return x_33; +} +else +{ +uint64_t x_34; uint8_t x_35; +x_34 = 13; +x_35 = lean_nat_dec_le(x_29, x_29); +if (x_35 == 0) +{ +uint64_t x_36; uint64_t x_37; +lean_dec(x_29); +lean_dec(x_27); +x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1; +x_37 = lean_uint64_mix_hash(x_24, x_36); +return x_37; } else { -lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -lean_dec(x_2); -x_13 = lean_array_size(x_12); -x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5(x_13, x_14, x_12); -x_16 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_16, 0, x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_1); -lean_ctor_set(x_17, 1, x_16); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +size_t x_38; size_t x_39; uint64_t x_40; uint64_t x_41; uint64_t x_42; +x_38 = 0; +x_39 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_40 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1(x_10, x_27, x_38, x_39, x_28); +lean_dec(x_27); +x_41 = lean_uint64_mix_hash(x_40, x_34); +x_42 = lean_uint64_mix_hash(x_24, x_41); +return x_42; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 2); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 3); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 4); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 5); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 6); -lean_inc(x_8); -x_9 = lean_ctor_get(x_1, 7); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_10, 0, x_2); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_15, x_3); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; -x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1(x_17, x_4); -lean_dec(x_4); -x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; -x_20 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__2(x_19, x_5); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28; -x_22 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__3(x_21, x_6); -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34; -x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_7); -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; -x_26 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_25, x_8); -lean_dec(x_8); -x_27 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; -x_28 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__4(x_27, x_9); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_13); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_26); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_24); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_22); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_20); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_18); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_16); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_14); -lean_ctor_set(x_36, 1, x_35); -x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____closed__1; -x_38 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_36, x_37); -x_39 = l_Lean_Json_mkObj(x_38); -return x_39; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096____spec__5(x_4, x_5, x_3); -return x_6; } } -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem___closed__1() { +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_), 1, 0); -return x_1; +uint64_t x_6; size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; +x_6 = lean_unbox_uint64(x_1); +lean_dec(x_1); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = lean_unbox_uint64(x_5); +lean_dec(x_5); +x_10 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____spec__1(x_6, x_2, x_7, x_8, x_9); +lean_dec(x_2); +x_11 = lean_box_uint64(x_10); +return x_11; } } -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem() { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____boxed(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCompletionItem___closed__1; -return x_1; +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657_(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; } } -static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem___closed__1() { +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Lsp_instInhabitedCompletionItem___closed__1; -x_3 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -lean_ctor_set(x_3, 2, x_1); -lean_ctor_set(x_3, 3, x_1); -lean_ctor_set(x_3, 4, x_1); -lean_ctor_set(x_3, 5, x_1); -lean_ctor_set(x_3, 6, x_1); -lean_ctor_set(x_3, 7, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instInhabitedCompletionItem() { +static lean_object* _init_l_Lean_Lsp_instHashableCompletionItem() { _start: { lean_object* x_1; -x_1 = l_Lean_Lsp_instInhabitedCompletionItem___closed__2; +x_1 = l_Lean_Lsp_instHashableCompletionItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -7914,7 +9146,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_(x_6); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_(x_6); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -7951,7 +9183,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -7993,7 +9225,7 @@ lean_inc(x_18); lean_dec(x_3); x_19 = lean_array_size(x_18); x_20 = 0; -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2(x_19, x_20, x_18); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2(x_19, x_20, x_18); return x_21; } default: @@ -8034,7 +9266,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1() { _start: { lean_object* x_1; @@ -8042,7 +9274,7 @@ x_1 = lean_mk_string_unchecked("isIncomplete", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2() { _start: { lean_object* x_1; @@ -8050,80 +9282,80 @@ x_1 = lean_mk_string_unchecked("CompletionList", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10() { _start: { lean_object* x_1; @@ -8131,52 +9363,52 @@ x_1 = lean_mk_string_unchecked("items", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -8188,7 +9420,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8200,7 +9432,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8214,8 +9446,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -8225,7 +9457,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8237,7 +9469,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8280,7 +9512,7 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -8288,15 +9520,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__2(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -8305,7 +9537,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773_), 1, 0); return x_1; } } @@ -8317,7 +9549,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionList___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -8332,7 +9564,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -8342,7 +9574,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -8352,7 +9584,7 @@ lean_inc(x_3); lean_dec(x_1); x_4 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_4, 0, x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -8362,10 +9594,10 @@ lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); x_9 = lean_array_size(x_3); x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1(x_9, x_10, x_3); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1(x_9, x_10, x_3); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -8384,7 +9616,7 @@ x_20 = l_Lean_Json_mkObj(x_19); return x_20; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -8392,7 +9624,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879____spec__1(x_4, x_5, x_3); return x_6; } } @@ -8400,7 +9632,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_), 1, 0); return x_1; } } @@ -8412,7 +9644,7 @@ x_1 = l_Lean_Lsp_instToJsonCompletionList___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1() { _start: { lean_object* x_1; @@ -8420,7 +9652,7 @@ x_1 = lean_mk_string_unchecked("textDocument", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2() { _start: { lean_object* x_1; @@ -8428,80 +9660,80 @@ x_1 = lean_mk_string_unchecked("CompletionParams", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10() { _start: { lean_object* x_1; @@ -8509,52 +9741,52 @@ x_1 = lean_mk_string_unchecked("position", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -8566,7 +9798,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8578,7 +9810,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8592,7 +9824,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -8603,7 +9835,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8615,7 +9847,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8658,7 +9890,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_), 1, 0); return x_1; } } @@ -8670,14 +9902,14 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8689,7 +9921,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -8712,7 +9944,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_), 1, 0); return x_1; } } @@ -8724,7 +9956,7 @@ x_1 = l_Lean_Lsp_instToJsonCompletionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1() { _start: { lean_object* x_1; @@ -8732,7 +9964,7 @@ x_1 = lean_mk_string_unchecked("contents", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2() { _start: { lean_object* x_1; @@ -8740,15 +9972,15 @@ x_1 = lean_mk_string_unchecked("range", 5, 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6021_(x_2); +x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6149_(x_2); lean_dec(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8759,7 +9991,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340____spec__1(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -8777,7 +10009,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123_), 1, 0); return x_1; } } @@ -8789,16 +10021,16 @@ x_1 = l_Lean_Lsp_instToJsonHover___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6073_(x_3); +x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6201_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1() { _start: { lean_object* x_1; @@ -8806,80 +10038,80 @@ x_1 = lean_mk_string_unchecked("Hover", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9() { _start: { lean_object* x_1; @@ -8887,54 +10119,54 @@ x_1 = lean_mk_string_unchecked("range\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -8944,7 +10176,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8956,7 +10188,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8970,7 +10202,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1410____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -8981,7 +10213,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8993,7 +10225,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9032,11 +10264,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9045,7 +10277,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165_), 1, 0); return x_1; } } @@ -9057,7 +10289,7 @@ x_1 = l_Lean_Lsp_instFromJsonHover___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1() { _start: { lean_object* x_1; @@ -9065,83 +10297,83 @@ x_1 = lean_mk_string_unchecked("HoverParams", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -9153,7 +10385,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9165,7 +10397,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9179,7 +10411,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9190,7 +10422,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9202,7 +10434,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9245,7 +10477,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283_), 1, 0); return x_1; } } @@ -9257,14 +10489,14 @@ x_1 = l_Lean_Lsp_instFromJsonHoverParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2816_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3389_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -9276,7 +10508,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -9299,7 +10531,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2816_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3389_), 1, 0); return x_1; } } @@ -9311,7 +10543,7 @@ x_1 = l_Lean_Lsp_instToJsonHoverParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1() { _start: { lean_object* x_1; @@ -9319,83 +10551,83 @@ x_1 = lean_mk_string_unchecked("DeclarationParams", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -9407,7 +10639,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9419,7 +10651,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9433,7 +10665,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9444,7 +10676,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9456,7 +10688,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9499,7 +10731,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDeclarationParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453_), 1, 0); return x_1; } } @@ -9511,14 +10743,14 @@ x_1 = l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2986_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -9530,7 +10762,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -9553,7 +10785,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDeclarationParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2986_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3559_), 1, 0); return x_1; } } @@ -9565,7 +10797,7 @@ x_1 = l_Lean_Lsp_instToJsonDeclarationParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1() { _start: { lean_object* x_1; @@ -9573,83 +10805,83 @@ x_1 = lean_mk_string_unchecked("DefinitionParams", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -9661,7 +10893,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9673,7 +10905,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9687,7 +10919,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9698,7 +10930,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9710,7 +10942,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9753,7 +10985,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDefinitionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623_), 1, 0); return x_1; } } @@ -9765,14 +10997,14 @@ x_1 = l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3156_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3729_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -9784,7 +11016,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -9807,7 +11039,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDefinitionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3156_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3729_), 1, 0); return x_1; } } @@ -9819,7 +11051,7 @@ x_1 = l_Lean_Lsp_instToJsonDefinitionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1() { _start: { lean_object* x_1; @@ -9827,83 +11059,83 @@ x_1 = lean_mk_string_unchecked("TypeDefinitionParams", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -9915,7 +11147,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9927,7 +11159,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9941,7 +11173,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9952,7 +11184,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9964,7 +11196,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -10007,7 +11239,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793_), 1, 0); return x_1; } } @@ -10019,14 +11251,14 @@ x_1 = l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3326_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3899_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -10038,7 +11270,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -10061,7 +11293,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3326_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3899_), 1, 0); return x_1; } } @@ -10073,7 +11305,7 @@ x_1 = l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1() { _start: { lean_object* x_1; @@ -10081,7 +11313,7 @@ x_1 = lean_mk_string_unchecked("includeDeclaration", 18, 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2() { _start: { lean_object* x_1; @@ -10089,84 +11321,84 @@ x_1 = lean_mk_string_unchecked("ReferenceContext", 16, 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10176,7 +11408,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -10188,7 +11420,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -10221,7 +11453,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963_), 1, 0); return x_1; } } @@ -10233,13 +11465,13 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -10256,13 +11488,13 @@ x_10 = l_Lean_Json_mkObj(x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030_(x_2); return x_3; } } @@ -10270,7 +11502,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceContext___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030____boxed), 1, 0); return x_1; } } @@ -10282,16 +11514,16 @@ x_1 = l_Lean_Lsp_instToJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1() { _start: { lean_object* x_1; @@ -10299,79 +11531,79 @@ x_1 = lean_mk_string_unchecked("ReferenceParams", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9() { _start: { lean_object* x_1; @@ -10379,52 +11611,52 @@ x_1 = lean_mk_string_unchecked("context", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -10436,7 +11668,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -10448,7 +11680,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -10462,7 +11694,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -10475,7 +11707,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -10487,7 +11719,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -10501,8 +11733,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -10513,7 +11745,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -10525,7 +11757,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -10571,11 +11803,11 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -10584,7 +11816,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088_), 1, 0); return x_1; } } @@ -10596,7 +11828,7 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3660_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4233_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; @@ -10605,7 +11837,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -10617,7 +11849,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -10629,8 +11861,8 @@ lean_inc(x_14); lean_dec(x_1); x_15 = lean_unbox(x_14); lean_dec(x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3457_(x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4030_(x_15); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -10656,7 +11888,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3660_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4233_), 1, 0); return x_1; } } @@ -10668,7 +11900,7 @@ x_1 = l_Lean_Lsp_instToJsonReferenceParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1() { _start: { lean_object* x_1; @@ -10676,7 +11908,7 @@ x_1 = lean_mk_string_unchecked("query", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2() { _start: { lean_object* x_1; @@ -10684,84 +11916,84 @@ x_1 = lean_mk_string_unchecked("WorkspaceSymbolParams", 21, 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10771,7 +12003,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -10783,7 +12015,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -10816,7 +12048,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311_), 1, 0); return x_1; } } @@ -10828,13 +12060,13 @@ x_1 = l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3805_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4378_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -10855,7 +12087,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3805_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4378_), 1, 0); return x_1; } } @@ -10867,7 +12099,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1() { _start: { lean_object* x_1; @@ -10875,83 +12107,83 @@ x_1 = lean_mk_string_unchecked("DocumentHighlightParams", 23, 23); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -10963,7 +12195,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -10975,7 +12207,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -10989,7 +12221,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -11000,7 +12232,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -11012,7 +12244,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -11055,7 +12287,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428_), 1, 0); return x_1; } } @@ -11067,14 +12299,14 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3961_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4534_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -11086,7 +12318,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -11109,7 +12341,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlightParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3961_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4534_), 1, 0); return x_1; } } @@ -11257,7 +12489,7 @@ x_3 = l_Lean_Lsp_instToJsonDocumentHighlightKind(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -11314,14 +12546,14 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -11332,8 +12564,8 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -11347,11 +12579,11 @@ x_15 = l_Lean_Json_mkObj(x_14); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -11360,7 +12592,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlight___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645_), 1, 0); return x_1; } } @@ -11372,7 +12604,7 @@ x_1 = l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1() { _start: { lean_object* x_1; @@ -11380,63 +12612,63 @@ x_1 = lean_mk_string_unchecked("DocumentSymbolParams", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -11446,7 +12678,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -11458,7 +12690,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -11491,7 +12723,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704_), 1, 0); return x_1; } } @@ -11503,12 +12735,12 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4198_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4771_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -11529,7 +12761,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4198_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4771_), 1, 0); return x_1; } } @@ -11742,7 +12974,7 @@ x_6 = l_Lean_Lsp_SymbolKind_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -11754,7 +12986,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -11762,7 +12994,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241_(x_3, x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -11771,7 +13003,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSymbolKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814____boxed), 2, 0); return x_1; } } @@ -11783,7 +13015,7 @@ x_1 = l_Lean_Lsp_instBEqSymbolKind___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832_(uint8_t x_1) { _start: { switch (x_1) { @@ -11946,13 +13178,13 @@ return x_27; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832____boxed(lean_object* x_1) { _start: { uint8_t x_2; uint64_t x_3; lean_object* x_4; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -11961,7 +13193,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSymbolKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832____boxed), 1, 0); return x_1; } } @@ -16603,7 +17835,7 @@ x_3 = l_Lean_Lsp_instToJsonSymbolKind(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -20098,7 +21330,7 @@ return x_1168; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -20156,15 +21388,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; @@ -20202,7 +21434,7 @@ lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; x_14 = lean_ctor_get(x_4, 0); x_15 = lean_array_size(x_14); x_16 = 0; -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg(x_1, x_15, x_16, x_14); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg(x_1, x_15, x_16, x_14); if (lean_obj_tag(x_17) == 0) { uint8_t x_18; @@ -20258,7 +21490,7 @@ lean_inc(x_25); lean_dec(x_4); x_26 = lean_array_size(x_25); x_27 = 0; -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg(x_1, x_26, x_27, x_25); +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg(x_1, x_26, x_27, x_25); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; @@ -20342,15 +21574,15 @@ return x_48; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg___boxed), 3, 0); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1() { _start: { lean_object* x_1; @@ -20358,7 +21590,7 @@ x_1 = lean_mk_string_unchecked("name", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2() { _start: { lean_object* x_1; @@ -20366,182 +21598,182 @@ x_1 = lean_mk_string_unchecked("DocumentSymbolAux", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20() { _start: { lean_object* x_1; @@ -20549,48 +21781,48 @@ x_1 = lean_mk_string_unchecked("selectionRange", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25() { _start: { lean_object* x_1; @@ -20598,7 +21830,7 @@ x_1 = lean_mk_string_unchecked("children", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26() { _start: { lean_object* x_1; @@ -20606,52 +21838,52 @@ x_1 = lean_mk_string_unchecked("children\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; lean_inc(x_2); x_4 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_2, x_3); if (lean_obj_tag(x_4) == 0) @@ -20664,7 +21896,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_4, 0); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); lean_ctor_set(x_4, 0, x_8); @@ -20676,7 +21908,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_9 = lean_ctor_get(x_4, 0); lean_inc(x_9); lean_dec(x_4); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9; x_11 = lean_string_append(x_10, x_9); lean_dec(x_9); x_12 = lean_alloc_ctor(0, 1, 0); @@ -20690,7 +21922,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_4, 0); lean_inc(x_13); lean_dec(x_4); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; lean_inc(x_2); x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_2, x_14); if (lean_obj_tag(x_15) == 0) @@ -20704,7 +21936,7 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; x_17 = lean_ctor_get(x_15, 0); -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11; +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11; x_19 = lean_string_append(x_18, x_17); lean_dec(x_17); lean_ctor_set(x_15, 0, x_19); @@ -20716,7 +21948,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_20 = lean_ctor_get(x_15, 0); lean_inc(x_20); lean_dec(x_15); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11; x_22 = lean_string_append(x_21, x_20); lean_dec(x_20); x_23 = lean_alloc_ctor(0, 1, 0); @@ -20730,9 +21962,9 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = lean_ctor_get(x_15, 0); lean_inc(x_24); lean_dec(x_15); -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; lean_inc(x_2); -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(x_2, x_25); +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(x_2, x_25); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; @@ -20745,7 +21977,7 @@ if (x_27 == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; x_28 = lean_ctor_get(x_26, 0); -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15; x_30 = lean_string_append(x_29, x_28); lean_dec(x_28); lean_ctor_set(x_26, 0, x_30); @@ -20757,7 +21989,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_31 = lean_ctor_get(x_26, 0); lean_inc(x_31); lean_dec(x_26); -x_32 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15; +x_32 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15; x_33 = lean_string_append(x_32, x_31); lean_dec(x_31); x_34 = lean_alloc_ctor(0, 1, 0); @@ -20771,7 +22003,7 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = lean_ctor_get(x_26, 0); lean_inc(x_35); lean_dec(x_26); -x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; lean_inc(x_2); x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_36); if (lean_obj_tag(x_37) == 0) @@ -20787,7 +22019,7 @@ if (x_38 == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; x_39 = lean_ctor_get(x_37, 0); -x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19; +x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19; x_41 = lean_string_append(x_40, x_39); lean_dec(x_39); lean_ctor_set(x_37, 0, x_41); @@ -20799,7 +22031,7 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_42 = lean_ctor_get(x_37, 0); lean_inc(x_42); lean_dec(x_37); -x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19; +x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19; x_44 = lean_string_append(x_43, x_42); lean_dec(x_42); x_45 = lean_alloc_ctor(0, 1, 0); @@ -20813,7 +22045,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_37, 0); lean_inc(x_46); lean_dec(x_37); -x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; lean_inc(x_2); x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) @@ -20830,7 +22062,7 @@ if (x_49 == 0) { lean_object* x_50; lean_object* x_51; lean_object* x_52; x_50 = lean_ctor_get(x_48, 0); -x_51 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24; +x_51 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24; x_52 = lean_string_append(x_51, x_50); lean_dec(x_50); lean_ctor_set(x_48, 0, x_52); @@ -20842,7 +22074,7 @@ lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; x_53 = lean_ctor_get(x_48, 0); lean_inc(x_53); lean_dec(x_48); -x_54 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24; +x_54 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24; x_55 = lean_string_append(x_54, x_53); lean_dec(x_53); x_56 = lean_alloc_ctor(0, 1, 0); @@ -20856,8 +22088,8 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_48, 0); lean_inc(x_57); lean_dec(x_48); -x_58 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25; -x_59 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg(x_1, x_2, x_58); +x_58 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25; +x_59 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg(x_1, x_2, x_58); if (lean_obj_tag(x_59) == 0) { uint8_t x_60; @@ -20871,7 +22103,7 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; x_61 = lean_ctor_get(x_59, 0); -x_62 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30; +x_62 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30; x_63 = lean_string_append(x_62, x_61); lean_dec(x_61); lean_ctor_set(x_59, 0, x_63); @@ -20883,7 +22115,7 @@ lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; x_64 = lean_ctor_get(x_59, 0); lean_inc(x_64); lean_dec(x_59); -x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30; +x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30; x_66 = lean_string_append(x_65, x_64); lean_dec(x_64); x_67 = lean_alloc_ctor(0, 1, 0); @@ -20938,24 +22170,24 @@ return x_75; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -20963,15 +22195,15 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__3___rarg(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__3___rarg(x_1, x_5, x_6, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__2___rarg(x_1, x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__2___rarg(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -20980,7 +22212,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg(lean_ob _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -20993,7 +22225,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -21020,15 +22252,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -21049,7 +22281,7 @@ lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; l x_6 = lean_ctor_get(x_3, 0); x_7 = lean_array_size(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg(x_1, x_7, x_8, x_6); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg(x_1, x_7, x_8, x_6); lean_ctor_set_tag(x_3, 4); lean_ctor_set(x_3, 0, x_9); x_10 = lean_alloc_ctor(0, 2, 0); @@ -21069,7 +22301,7 @@ lean_inc(x_13); lean_dec(x_3); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg(x_1, x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg(x_1, x_14, x_15, x_13); x_17 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_17, 0, x_16); x_18 = lean_alloc_ctor(0, 2, 0); @@ -21084,15 +22316,15 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -21110,7 +22342,7 @@ lean_inc(x_8); lean_dec(x_2); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -21118,10 +22350,10 @@ x_12 = lean_box(0); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_14, x_4); x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -21129,15 +22361,15 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_12); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25; -x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__1___rarg(x_1, x_24, x_8); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25; +x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__1___rarg(x_1, x_24, x_8); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_12); @@ -21334,7 +22566,7 @@ goto block_39; block_39: { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -21357,15 +22589,15 @@ return x_38; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____rarg), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -21373,7 +22605,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____spec__2___rarg(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____spec__2___rarg(x_1, x_5, x_6, x_4); return x_7; } } @@ -21381,7 +22613,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_obje _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -21474,7 +22706,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -21492,7 +22724,7 @@ lean_inc(x_7); lean_dec(x_1); x_8 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_8, 0, x_2); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -21500,10 +22732,10 @@ x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_13, x_3); x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_5); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -21511,14 +22743,14 @@ x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_11); x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_6); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_11); -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25; x_24 = l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(x_23, x_7); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -21716,7 +22948,7 @@ goto block_38; block_38: { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); @@ -21746,7 +22978,7 @@ lean_object* x_2; lean_object* x_3; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5055____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5628____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); return x_3; } } @@ -21841,7 +23073,7 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -21849,11 +23081,11 @@ x_3 = 1; return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -21864,7 +23096,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSymbolTag___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5244____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5817____boxed), 2, 0); return x_1; } } @@ -21876,7 +23108,7 @@ x_1 = l_Lean_Lsp_instBEqSymbolTag___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835_(lean_object* x_1) { _start: { uint64_t x_2; @@ -21884,11 +23116,11 @@ x_2 = 0; return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -21898,7 +23130,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSymbolTag___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5262____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5835____boxed), 1, 0); return x_1; } } @@ -22196,7 +23428,7 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -22475,7 +23707,7 @@ return x_107; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -22517,7 +23749,7 @@ lean_inc(x_18); lean_dec(x_3); x_19 = lean_array_size(x_18); x_20 = 0; -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(x_19, x_20, x_18); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(x_19, x_20, x_18); return x_21; } default: @@ -22558,7 +23790,7 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -22567,7 +23799,7 @@ x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_D return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1() { _start: { lean_object* x_1; @@ -22575,120 +23807,120 @@ x_1 = lean_mk_string_unchecked("SymbolInformation", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13() { _start: { lean_object* x_1; @@ -22696,48 +23928,48 @@ x_1 = lean_mk_string_unchecked("location", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18() { _start: { lean_object* x_1; @@ -22745,7 +23977,7 @@ x_1 = lean_mk_string_unchecked("containerName", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19() { _start: { lean_object* x_1; @@ -22753,52 +23985,52 @@ x_1 = lean_mk_string_unchecked("containerName\?", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -22810,7 +24042,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22822,7 +24054,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22836,9 +24068,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -22849,7 +24081,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -22861,7 +24093,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -22875,9 +24107,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -22889,7 +24121,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -22901,7 +24133,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -22915,9 +24147,9 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13; lean_inc(x_1); -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3(x_1, x_35); +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -22930,7 +24162,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -22942,7 +24174,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -22956,7 +24188,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18; +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18; x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_46); if (lean_obj_tag(x_47) == 0) { @@ -22970,7 +24202,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -22982,7 +24214,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -23034,7 +24266,7 @@ return x_63; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -23042,24 +24274,24 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__3(x_1, x_2); lean_dec(x_2); return x_3; } @@ -23068,7 +24300,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSymbolInformation___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992_), 1, 0); return x_1; } } @@ -23080,7 +24312,7 @@ x_1 = l_Lean_Lsp_instFromJsonSymbolInformation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -23104,7 +24336,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -23120,7 +24352,7 @@ lean_inc(x_6); lean_dec(x_1); x_7 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_7, 0, x_2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -23130,10 +24362,10 @@ lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); x_12 = lean_array_size(x_4); x_13 = 0; -x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(x_12, x_13, x_4); +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(x_12, x_13, x_4); x_15 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_15, 0, x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -23141,14 +24373,14 @@ x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_10); x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(x_5); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_10); -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18; x_24 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_23, x_6); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -23346,7 +24578,7 @@ goto block_37; block_37: { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); @@ -23366,7 +24598,7 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -23374,7 +24606,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(x_4, x_5, x_3); return x_6; } } @@ -23382,7 +24614,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSymbolInformation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215_), 1, 0); return x_1; } } @@ -23394,7 +24626,7 @@ x_1 = l_Lean_Lsp_instToJsonSymbolInformation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1() { _start: { lean_object* x_1; @@ -23402,83 +24634,83 @@ x_1 = lean_mk_string_unchecked("CallHierarchyPrepareParams", 26, 26); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -23490,7 +24722,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -23502,7 +24734,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -23516,7 +24748,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -23527,7 +24759,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -23539,7 +24771,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -23582,7 +24814,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311_), 1, 0); return x_1; } } @@ -23594,14 +24826,14 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5844_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6417_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -23613,7 +24845,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -23636,7 +24868,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5844_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6417_), 1, 0); return x_1; } } @@ -23648,7 +24880,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -23684,7 +24916,7 @@ lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_3, 0); x_14 = lean_array_size(x_13); x_15 = 0; -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(x_14, x_15, x_13); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { uint8_t x_17; @@ -23740,7 +24972,7 @@ lean_inc(x_24); lean_dec(x_3); x_25 = lean_array_size(x_24); x_26 = 0; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____spec__2(x_25, x_26, x_24); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____spec__2(x_25, x_26, x_24); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -23823,7 +25055,7 @@ return x_47; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1() { _start: { lean_object* x_1; @@ -23831,119 +25063,119 @@ x_1 = lean_mk_string_unchecked("CallHierarchyItem", 17, 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13() { _start: { lean_object* x_1; @@ -23951,112 +25183,112 @@ x_1 = lean_mk_string_unchecked("uri", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -24068,7 +25300,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24080,7 +25312,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24094,9 +25326,9 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; lean_inc(x_1); -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____spec__1(x_1, x_13); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24107,7 +25339,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -24119,7 +25351,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -24133,9 +25365,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -24147,7 +25379,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -24159,7 +25391,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -24173,7 +25405,7 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; lean_inc(x_1); x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) @@ -24188,7 +25420,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -24200,7 +25432,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -24214,7 +25446,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13; +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13; lean_inc(x_1); x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__1(x_1, x_46); if (lean_obj_tag(x_47) == 0) @@ -24230,7 +25462,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -24242,7 +25474,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -24256,7 +25488,7 @@ lean_object* x_56; lean_object* x_57; lean_object* x_58; x_56 = lean_ctor_get(x_47, 0); lean_inc(x_56); lean_dec(x_47); -x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; lean_inc(x_1); x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) @@ -24273,7 +25505,7 @@ if (x_59 == 0) { lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_58, 0); -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19; x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); lean_ctor_set(x_58, 0, x_62); @@ -24285,7 +25517,7 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; x_63 = lean_ctor_get(x_58, 0); lean_inc(x_63); lean_dec(x_58); -x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19; +x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19; x_65 = lean_string_append(x_64, x_63); lean_dec(x_63); x_66 = lean_alloc_ctor(0, 1, 0); @@ -24299,7 +25531,7 @@ lean_object* x_67; lean_object* x_68; lean_object* x_69; x_67 = lean_ctor_get(x_58, 0); lean_inc(x_67); lean_dec(x_58); -x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; lean_inc(x_1); x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_68); if (lean_obj_tag(x_69) == 0) @@ -24317,7 +25549,7 @@ if (x_70 == 0) { lean_object* x_71; lean_object* x_72; lean_object* x_73; x_71 = lean_ctor_get(x_69, 0); -x_72 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21; +x_72 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21; x_73 = lean_string_append(x_72, x_71); lean_dec(x_71); lean_ctor_set(x_69, 0, x_73); @@ -24329,7 +25561,7 @@ lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; x_74 = lean_ctor_get(x_69, 0); lean_inc(x_74); lean_dec(x_69); -x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21; +x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21; x_76 = lean_string_append(x_75, x_74); lean_dec(x_74); x_77 = lean_alloc_ctor(0, 1, 0); @@ -24343,8 +25575,8 @@ lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; x_78 = lean_ctor_get(x_69, 0); lean_inc(x_78); lean_dec(x_69); -x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; -x_80 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____spec__4(x_1, x_79); +x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; +x_80 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____spec__4(x_1, x_79); x_81 = !lean_is_exclusive(x_80); if (x_81 == 0) { @@ -24394,11 +25626,11 @@ return x_88; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -24407,7 +25639,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546_), 1, 0); return x_1; } } @@ -24419,7 +25651,7 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -24439,7 +25671,7 @@ lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; le x_5 = lean_ctor_get(x_2, 0); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(x_6, x_7, x_5); lean_ctor_set_tag(x_2, 4); lean_ctor_set(x_2, 0, x_8); x_9 = lean_alloc_ctor(0, 2, 0); @@ -24459,7 +25691,7 @@ lean_inc(x_12); lean_dec(x_2); x_13 = lean_array_size(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642____spec__1(x_13, x_14, x_12); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215____spec__1(x_13, x_14, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(0, 2, 0); @@ -24474,7 +25706,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; @@ -24496,7 +25728,7 @@ lean_inc(x_9); lean_dec(x_1); x_10 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_10, 0, x_2); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -24504,13 +25736,13 @@ x_13 = lean_box(0); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313____spec__1(x_15, x_4); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46; +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886____spec__1(x_15, x_4); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10; x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_17, x_5); x_19 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_19, 0, x_6); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); @@ -24518,7 +25750,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_13); x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_7); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); @@ -24526,14 +25758,14 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_13); x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_8); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20; x_29 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_27); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_13); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; x_32 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_31, x_9); lean_dec(x_9); x_33 = lean_alloc_ctor(1, 2, 0); @@ -24741,7 +25973,7 @@ goto block_48; block_48: { lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; +x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; x_41 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_39); @@ -24765,7 +25997,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_), 1, 0); return x_1; } } @@ -24777,7 +26009,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; uint8_t x_9; @@ -24803,7 +26035,7 @@ return x_13; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -24848,88 +26080,14 @@ return x_11; else { uint8_t x_12; -x_12 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2(x_6, x_7, lean_box(0), x_6, x_7, x_8, lean_box(0)); +x_12 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2(x_6, x_7, lean_box(0), x_6, x_7, x_8, lean_box(0)); return x_12; } } } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_string_dec_eq(x_6, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Data_Json_Basic_0__Lean_Json_beq_x27(x_6, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -24959,7 +26117,7 @@ return x_20; else { uint8_t x_21; -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4241_(x_4, x_12); +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4814_(x_4, x_12); if (x_21 == 0) { uint8_t x_22; @@ -24969,7 +26127,7 @@ return x_22; else { uint8_t x_23; -x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(x_5, x_13); +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1(x_5, x_13); if (x_23 == 0) { uint8_t x_24; @@ -24979,7 +26137,7 @@ return x_24; else { uint8_t x_25; -x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(x_6, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_6, x_14); if (x_25 == 0) { uint8_t x_26; @@ -25019,7 +26177,7 @@ return x_32; else { uint8_t x_33; -x_33 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4(x_10, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__5(x_10, x_18); return x_33; } } @@ -25030,11 +26188,11 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -25043,44 +26201,22 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__4(x_1, x_2); +x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -25091,7 +26227,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqCallHierarchyItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992____boxed), 2, 0); return x_1; } } @@ -25103,7 +26239,7 @@ x_1 = l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1(uint64_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint64_t x_5) { +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1(uint64_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint64_t x_5) { _start: { uint8_t x_6; @@ -25124,17 +26260,7 @@ return x_5; } } } -static uint64_t _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1() { -_start: -{ -uint64_t x_1; uint64_t x_2; uint64_t x_3; -x_1 = 7; -x_2 = 13; -x_3 = lean_uint64_mix_hash(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; @@ -25149,7 +26275,7 @@ x_9 = lean_ctor_get(x_1, 6); x_10 = 0; x_11 = lean_string_hash(x_2); x_12 = lean_uint64_mix_hash(x_10, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4259_(x_3); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4832_(x_3); x_14 = lean_uint64_mix_hash(x_12, x_13); x_15 = lean_string_hash(x_6); x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_701_(x_7); @@ -25173,7 +26299,7 @@ if (x_51 == 0) { uint64_t x_52; lean_dec(x_49); -x_52 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1; +x_52 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1; x_18 = x_52; goto block_45; } @@ -25186,7 +26312,7 @@ if (x_54 == 0) { uint64_t x_55; lean_dec(x_49); -x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1; +x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1; x_18 = x_55; goto block_45; } @@ -25196,7 +26322,7 @@ size_t x_56; size_t x_57; uint64_t x_58; uint64_t x_59; x_56 = 0; x_57 = lean_usize_of_nat(x_49); lean_dec(x_49); -x_58 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1(x_10, x_47, x_56, x_57, x_48); +x_58 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1(x_10, x_47, x_56, x_57, x_48); x_59 = lean_uint64_mix_hash(x_58, x_53); x_18 = x_59; goto block_45; @@ -25263,7 +26389,7 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint64_t x_6; size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; @@ -25275,17 +26401,17 @@ x_8 = lean_unbox_usize(x_4); lean_dec(x_4); x_9 = lean_unbox_uint64(x_5); lean_dec(x_5); -x_10 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____spec__1(x_6, x_2, x_7, x_8, x_9); +x_10 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____spec__1(x_6, x_2, x_7, x_8, x_9); lean_dec(x_2); x_11 = lean_box_uint64(x_10); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -25295,7 +26421,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableCallHierarchyItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162____boxed), 1, 0); return x_1; } } @@ -25357,16 +26483,16 @@ x_1 = l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1() { _start: { lean_object* x_1; @@ -25374,7 +26500,7 @@ x_1 = lean_mk_string_unchecked("item", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2() { _start: { lean_object* x_1; @@ -25382,85 +26508,85 @@ x_1 = lean_mk_string_unchecked("CallHierarchyIncomingCallsParams", 32, 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -25469,7 +26595,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -25481,7 +26607,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -25510,11 +26636,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -25523,7 +26649,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParam _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286_), 1, 0); return x_1; } } @@ -25535,12 +26661,12 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6780_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7353_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(x_1); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -25561,7 +26687,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6780_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7353_), 1, 0); return x_1; } } @@ -25573,7 +26699,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -25628,7 +26754,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -25670,7 +26796,7 @@ lean_inc(x_18); lean_dec(x_3); x_19 = lean_array_size(x_18); x_20 = 0; -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2(x_19, x_20, x_18); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2(x_19, x_20, x_18); return x_21; } default: @@ -25711,7 +26837,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1() { _start: { lean_object* x_1; @@ -25719,7 +26845,7 @@ x_1 = lean_mk_string_unchecked("from", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2() { _start: { lean_object* x_1; @@ -25727,80 +26853,80 @@ x_1 = lean_mk_string_unchecked("CallHierarchyIncomingCall", 25, 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10() { _start: { lean_object* x_1; @@ -25808,54 +26934,54 @@ x_1 = lean_mk_string_unchecked("fromRanges", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -25865,7 +26991,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -25877,7 +27003,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -25891,8 +27017,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -25902,7 +27028,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -25914,7 +27040,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -25953,7 +27079,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -25961,15 +27087,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__2(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -25978,7 +27104,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412_), 1, 0); return x_1; } } @@ -25990,7 +27116,7 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -26015,14 +27141,14 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -26035,10 +27161,10 @@ lean_inc(x_8); lean_dec(x_1); x_9 = lean_array_size(x_8); x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1(x_9, x_10, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1(x_9, x_10, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -26057,7 +27183,7 @@ x_20 = l_Lean_Json_mkObj(x_19); return x_20; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -26065,7 +27191,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1(x_4, x_5, x_3); return x_6; } } @@ -26073,7 +27199,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518_), 1, 0); return x_1; } } @@ -26114,7 +27240,7 @@ x_1 = l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__2; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1() { _start: { lean_object* x_1; @@ -26122,64 +27248,64 @@ x_1 = lean_mk_string_unchecked("CallHierarchyOutgoingCallsParams", 32, 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -26188,7 +27314,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -26200,7 +27326,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -26233,7 +27359,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParam _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593_), 1, 0); return x_1; } } @@ -26245,12 +27371,12 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7087_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7660_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(x_1); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -26271,7 +27397,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7087_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7660_), 1, 0); return x_1; } } @@ -26283,7 +27409,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1() { _start: { lean_object* x_1; @@ -26291,7 +27417,7 @@ x_1 = lean_mk_string_unchecked("to", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2() { _start: { lean_object* x_1; @@ -26299,106 +27425,106 @@ x_1 = lean_mk_string_unchecked("CallHierarchyOutgoingCall", 25, 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -26408,7 +27534,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -26420,7 +27546,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -26434,8 +27560,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -26445,7 +27571,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -26457,7 +27583,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -26500,7 +27626,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719_), 1, 0); return x_1; } } @@ -26512,14 +27638,14 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7252_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7825_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -26532,10 +27658,10 @@ lean_inc(x_8); lean_dec(x_1); x_9 = lean_array_size(x_8); x_10 = 0; -x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945____spec__1(x_9, x_10, x_8); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518____spec__1(x_9, x_10, x_8); x_12 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_12, 0, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -26558,7 +27684,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7252_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7825_), 1, 0); return x_1; } } @@ -26779,7 +27905,7 @@ x_6 = l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1() { _start: { lean_object* x_1; @@ -26787,17 +27913,17 @@ x_1 = lean_mk_string_unchecked("keyword", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3() { _start: { lean_object* x_1; @@ -26805,17 +27931,17 @@ x_1 = lean_mk_string_unchecked("variable", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5() { _start: { lean_object* x_1; @@ -26823,17 +27949,17 @@ x_1 = lean_mk_string_unchecked("property", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7() { _start: { lean_object* x_1; @@ -26841,17 +27967,17 @@ x_1 = lean_mk_string_unchecked("function", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9() { _start: { lean_object* x_1; @@ -26859,17 +27985,17 @@ x_1 = lean_mk_string_unchecked("namespace", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11() { _start: { lean_object* x_1; @@ -26877,17 +28003,17 @@ x_1 = lean_mk_string_unchecked("type", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13() { _start: { lean_object* x_1; @@ -26895,17 +28021,17 @@ x_1 = lean_mk_string_unchecked("class", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15() { _start: { lean_object* x_1; @@ -26913,17 +28039,17 @@ x_1 = lean_mk_string_unchecked("enum", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17() { _start: { lean_object* x_1; @@ -26931,17 +28057,17 @@ x_1 = lean_mk_string_unchecked("interface", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19() { _start: { lean_object* x_1; @@ -26949,17 +28075,17 @@ x_1 = lean_mk_string_unchecked("struct", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21() { _start: { lean_object* x_1; @@ -26967,17 +28093,17 @@ x_1 = lean_mk_string_unchecked("typeParameter", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23() { _start: { lean_object* x_1; @@ -26985,17 +28111,17 @@ x_1 = lean_mk_string_unchecked("parameter", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25() { _start: { lean_object* x_1; @@ -27003,17 +28129,17 @@ x_1 = lean_mk_string_unchecked("enumMember", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27() { _start: { lean_object* x_1; @@ -27021,17 +28147,17 @@ x_1 = lean_mk_string_unchecked("event", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29() { _start: { lean_object* x_1; @@ -27039,17 +28165,17 @@ x_1 = lean_mk_string_unchecked("method", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31() { _start: { lean_object* x_1; @@ -27057,17 +28183,17 @@ x_1 = lean_mk_string_unchecked("macro", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33() { _start: { lean_object* x_1; @@ -27075,17 +28201,17 @@ x_1 = lean_mk_string_unchecked("modifier", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35() { _start: { lean_object* x_1; @@ -27093,17 +28219,17 @@ x_1 = lean_mk_string_unchecked("comment", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37() { _start: { lean_object* x_1; @@ -27111,17 +28237,17 @@ x_1 = lean_mk_string_unchecked("string", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39() { _start: { lean_object* x_1; @@ -27129,17 +28255,17 @@ x_1 = lean_mk_string_unchecked("number", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41() { _start: { lean_object* x_1; @@ -27147,17 +28273,17 @@ x_1 = lean_mk_string_unchecked("regexp", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43() { _start: { lean_object* x_1; @@ -27165,17 +28291,17 @@ x_1 = lean_mk_string_unchecked("operator", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45() { _start: { lean_object* x_1; @@ -27183,17 +28309,17 @@ x_1 = lean_mk_string_unchecked("decorator", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47() { _start: { lean_object* x_1; @@ -27201,174 +28327,174 @@ x_1 = lean_mk_string_unchecked("leanSorryLike", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893_(uint8_t x_1) { _start: { switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2; return x_2; } case 1: { lean_object* x_3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4; return x_3; } case 2: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6; return x_4; } case 3: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8; return x_5; } case 4: { lean_object* x_6; -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10; return x_6; } case 5: { lean_object* x_7; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12; return x_7; } case 6: { lean_object* x_8; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14; return x_8; } case 7: { lean_object* x_9; -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16; return x_9; } case 8: { lean_object* x_10; -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18; return x_10; } case 9: { lean_object* x_11; -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20; return x_11; } case 10: { lean_object* x_12; -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22; return x_12; } case 11: { lean_object* x_13; -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24; return x_13; } case 12: { lean_object* x_14; -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26; return x_14; } case 13: { lean_object* x_15; -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28; return x_15; } case 14: { lean_object* x_16; -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30; return x_16; } case 15: { lean_object* x_17; -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32; return x_17; } case 16: { lean_object* x_18; -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34; +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34; return x_18; } case 17: { lean_object* x_19; -x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36; +x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36; return x_19; } case 18: { lean_object* x_20; -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38; return x_20; } case 19: { lean_object* x_21; -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40; return x_21; } case 20: { lean_object* x_22; -x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42; +x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42; return x_22; } case 21: { lean_object* x_23; -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44; return x_23; } case 22: { lean_object* x_24; -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46; return x_24; } default: { lean_object* x_25; -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48; +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48; return x_25; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893_(x_2); return x_3; } } @@ -27376,7 +28502,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokenType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____boxed), 1, 0); return x_1; } } @@ -27388,7 +28514,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokenType___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -27396,33 +28522,33 @@ x_1 = lean_mk_string_unchecked("no inductive constructor matched", 32, 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2; return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___boxed), 1, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27433,21 +28559,21 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_3 = l_Except_orElseLazy___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); if (lean_obj_tag(x_6) == 0) @@ -27457,7 +28583,7 @@ x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_9 = l_Except_orElseLazy___rarg(x_6, x_8); return x_9; } @@ -27469,7 +28595,7 @@ lean_inc(x_10); lean_dec(x_6); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_13 = l_Except_orElseLazy___rarg(x_11, x_12); return x_13; } @@ -27478,12 +28604,12 @@ else { lean_object* x_14; lean_dec(x_6); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3; return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27494,15 +28620,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27531,13 +28657,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27548,15 +28674,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27585,13 +28711,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27602,15 +28728,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27639,13 +28765,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27656,15 +28782,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27693,13 +28819,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27710,15 +28836,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27747,13 +28873,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27764,15 +28890,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27801,13 +28927,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27818,15 +28944,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27855,13 +28981,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27872,15 +28998,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27909,13 +29035,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27926,15 +29052,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27963,13 +29089,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27980,15 +29106,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28017,13 +29143,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28034,15 +29160,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28071,13 +29197,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28088,15 +29214,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28125,13 +29251,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28142,15 +29268,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28179,13 +29305,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28196,15 +29322,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28233,13 +29359,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28250,15 +29376,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28287,13 +29413,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28304,15 +29430,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28341,13 +29467,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28358,15 +29484,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28395,13 +29521,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28412,15 +29538,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28449,13 +29575,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28466,15 +29592,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28503,13 +29629,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28520,15 +29646,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28557,13 +29683,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28574,15 +29700,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28611,13 +29737,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28628,15 +29754,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28665,13 +29791,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -28681,7 +29807,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28692,16 +29818,16 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47; x_3 = lean_unsigned_to_nat(0u); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1; lean_inc(x_1); x_5 = l_Lean_Json_parseTagged(x_1, x_2, x_3, x_4); -x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___boxed), 3, 2); +x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___boxed), 3, 2); lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_4); if (lean_obj_tag(x_5) == 0) @@ -28730,225 +29856,225 @@ else { lean_object* x_12; lean_object* x_13; lean_dec(x_5); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2; x_13 = l_Except_orElseLazy___rarg(x_12, x_6); return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -28957,7 +30083,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058_), 1, 0); return x_1; } } @@ -28969,7 +30095,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -28981,7 +30107,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -28989,7 +30115,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321_(x_3, x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -28998,7 +30124,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSemanticTokenType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894____boxed), 2, 0); return x_1; } } @@ -29010,7 +30136,7 @@ x_1 = l_Lean_Lsp_instBEqSemanticTokenType___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912_(uint8_t x_1) { _start: { switch (x_1) { @@ -29161,13 +30287,13 @@ return x_25; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912____boxed(lean_object* x_1) { _start: { uint8_t x_2; uint64_t x_3; lean_object* x_4; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -29176,7 +30302,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSemanticTokenType___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912____boxed), 1, 0); return x_1; } } @@ -29193,7 +30319,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -29204,7 +30330,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29216,7 +30342,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29228,7 +30354,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__3; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29240,7 +30366,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29252,7 +30378,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29264,7 +30390,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29276,7 +30402,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29288,7 +30414,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29300,7 +30426,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29312,7 +30438,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__10; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29324,7 +30450,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__11; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29336,7 +30462,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__12; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29348,7 +30474,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__13; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29360,7 +30486,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__14; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29372,7 +30498,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__15; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29384,7 +30510,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__16; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29396,7 +30522,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__17; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29408,7 +30534,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__18; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29420,7 +30546,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__19; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29432,7 +30558,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__20; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29444,7 +30570,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__21; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29456,7 +30582,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__22; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29468,7 +30594,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1; x_2 = l_Lean_Lsp_SemanticTokenType_names___closed__23; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29616,7 +30742,7 @@ x_6 = l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1() { _start: { lean_object* x_1; @@ -29624,17 +30750,17 @@ x_1 = lean_mk_string_unchecked("declaration", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3() { _start: { lean_object* x_1; @@ -29642,17 +30768,17 @@ x_1 = lean_mk_string_unchecked("definition", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5() { _start: { lean_object* x_1; @@ -29660,17 +30786,17 @@ x_1 = lean_mk_string_unchecked("readonly", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7() { _start: { lean_object* x_1; @@ -29678,17 +30804,17 @@ x_1 = lean_mk_string_unchecked("static", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9() { _start: { lean_object* x_1; @@ -29696,17 +30822,17 @@ x_1 = lean_mk_string_unchecked("deprecated", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11() { _start: { lean_object* x_1; @@ -29714,17 +30840,17 @@ x_1 = lean_mk_string_unchecked("abstract", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13() { _start: { lean_object* x_1; @@ -29732,17 +30858,17 @@ x_1 = lean_mk_string_unchecked("async", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15() { _start: { lean_object* x_1; @@ -29750,27 +30876,27 @@ x_1 = lean_mk_string_unchecked("modification", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18() { _start: { lean_object* x_1; @@ -29778,90 +30904,90 @@ x_1 = lean_mk_string_unchecked("defaultLibrary", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129_(uint8_t x_1) { _start: { switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2; return x_2; } case 1: { lean_object* x_3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4; return x_3; } case 2: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6; return x_4; } case 3: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8; return x_5; } case 4: { lean_object* x_6; -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10; return x_6; } case 5: { lean_object* x_7; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12; return x_7; } case 6: { lean_object* x_8; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14; return x_8; } case 7: { lean_object* x_9; -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16; return x_9; } case 8: { lean_object* x_10; -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17; return x_10; } default: { lean_object* x_11; -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19; return x_11; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129_(x_2); return x_3; } } @@ -29869,7 +30995,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____boxed), 1, 0); return x_1; } } @@ -29881,7 +31007,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -29892,21 +31018,21 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_3 = l_Except_orElseLazy___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); if (lean_obj_tag(x_6) == 0) @@ -29916,7 +31042,7 @@ x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_9 = l_Except_orElseLazy___rarg(x_6, x_8); return x_9; } @@ -29928,7 +31054,7 @@ lean_inc(x_10); lean_dec(x_6); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1; x_13 = l_Except_orElseLazy___rarg(x_11, x_12); return x_13; } @@ -29937,12 +31063,12 @@ else { lean_object* x_14; lean_dec(x_6); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2; return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -29953,15 +31079,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -29990,13 +31116,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30007,15 +31133,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30044,13 +31170,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30061,15 +31187,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30098,13 +31224,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30115,15 +31241,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30152,13 +31278,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30169,15 +31295,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30206,13 +31332,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30223,15 +31349,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30260,13 +31386,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30277,15 +31403,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30314,13 +31440,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30331,15 +31457,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1; x_5 = lean_unsigned_to_nat(0u); lean_inc(x_1); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -30368,13 +31494,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -30385,16 +31511,16 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18; x_3 = lean_unsigned_to_nat(0u); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1; lean_inc(x_1); x_5 = l_Lean_Json_parseTagged(x_1, x_2, x_3, x_4); -x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___boxed), 3, 2); +x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___boxed), 3, 2); lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_4); if (lean_obj_tag(x_5) == 0) @@ -30423,90 +31549,90 @@ else { lean_object* x_12; lean_object* x_13; lean_dec(x_5); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1; x_13 = l_Except_orElseLazy___rarg(x_12, x_6); return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -30515,7 +31641,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210_), 1, 0); return x_1; } } @@ -30532,7 +31658,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -30543,7 +31669,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30555,7 +31681,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30567,7 +31693,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__3; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30579,7 +31705,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30591,7 +31717,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30603,7 +31729,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30615,7 +31741,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30627,7 +31753,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30639,7 +31765,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__10() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1; x_2 = l_Lean_Lsp_SemanticTokenModifier_names___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30682,7 +31808,7 @@ x_3 = l_Lean_Lsp_SemanticTokenModifier_toNat(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -30765,7 +31891,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1() { _start: { lean_object* x_1; @@ -30773,7 +31899,7 @@ x_1 = lean_mk_string_unchecked("tokenTypes", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2() { _start: { lean_object* x_1; @@ -30781,80 +31907,80 @@ x_1 = lean_mk_string_unchecked("SemanticTokensLegend", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10() { _start: { lean_object* x_1; @@ -30862,54 +31988,54 @@ x_1 = lean_mk_string_unchecked("tokenModifiers", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -30919,7 +32045,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -30931,7 +32057,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -30945,8 +32071,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -30956,7 +32082,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -30968,7 +32094,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -31007,11 +32133,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -31020,7 +32146,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722_), 1, 0); return x_1; } } @@ -31032,7 +32158,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9255_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9828_(lean_object* x_1) { _start: { lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -31043,7 +32169,7 @@ x_4 = 0; x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2(x_3, x_4, x_2); x_6 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_6, 0, x_5); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -31058,7 +32184,7 @@ x_12 = lean_array_size(x_11); x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_189____spec__2(x_12, x_4, x_11); x_14 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_14, 0, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -31081,7 +32207,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9255_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9828_), 1, 0); return x_1; } } @@ -31093,16 +32219,16 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1() { _start: { lean_object* x_1; @@ -31110,7 +32236,7 @@ x_1 = lean_mk_string_unchecked("legend", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2() { _start: { lean_object* x_1; @@ -31118,100 +32244,100 @@ x_1 = lean_mk_string_unchecked("SemanticTokensOptions", 21, 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12() { _start: { lean_object* x_1; @@ -31219,54 +32345,54 @@ x_1 = lean_mk_string_unchecked("full", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -31276,7 +32402,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -31288,7 +32414,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31302,7 +32428,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -31315,7 +32441,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -31327,7 +32453,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -31341,7 +32467,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12; x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -31353,7 +32479,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -31365,7 +32491,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -31415,11 +32541,11 @@ return x_43; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -31428,7 +32554,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_), 1, 0); return x_1; } } @@ -31440,14 +32566,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9480_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10053_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9255_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9828_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -31458,7 +32584,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -31469,7 +32595,7 @@ x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*1 + 1); lean_dec(x_1); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -31495,7 +32621,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9480_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10053_), 1, 0); return x_1; } } @@ -31507,7 +32633,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1() { _start: { lean_object* x_1; @@ -31515,63 +32641,63 @@ x_1 = lean_mk_string_unchecked("SemanticTokensParams", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -31581,7 +32707,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -31593,7 +32719,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31626,7 +32752,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131_), 1, 0); return x_1; } } @@ -31638,12 +32764,12 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9625_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10198_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -31664,7 +32790,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9625_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10198_), 1, 0); return x_1; } } @@ -31676,7 +32802,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1() { _start: { lean_object* x_1; @@ -31684,83 +32810,83 @@ x_1 = lean_mk_string_unchecked("SemanticTokensRangeParams", 25, 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -31772,7 +32898,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -31784,7 +32910,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31798,7 +32924,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1132____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -31809,7 +32935,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -31821,7 +32947,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -31864,7 +32990,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256_), 1, 0); return x_1; } } @@ -31876,14 +33002,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9789_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10362_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -31895,7 +33021,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_742_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -31918,7 +33044,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensRangeParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9789_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10362_), 1, 0); return x_1; } } @@ -31930,7 +33056,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -31985,7 +33111,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -32027,7 +33153,7 @@ lean_inc(x_18); lean_dec(x_3); x_19 = lean_array_size(x_18); x_20 = 0; -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2(x_19, x_20, x_18); +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2(x_19, x_20, x_18); return x_21; } default: @@ -32068,7 +33194,7 @@ return x_34; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1() { _start: { lean_object* x_1; @@ -32076,7 +33202,7 @@ x_1 = lean_mk_string_unchecked("resultId", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2() { _start: { lean_object* x_1; @@ -32084,39 +33210,39 @@ x_1 = lean_mk_string_unchecked("SemanticTokens", 14, 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6() { _start: { lean_object* x_1; @@ -32124,93 +33250,93 @@ x_1 = lean_mk_string_unchecked("resultId\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_2045____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -32222,7 +33348,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -32234,7 +33360,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -32248,8 +33374,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -32259,7 +33385,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -32271,7 +33397,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -32310,7 +33436,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -32318,15 +33444,15 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__2(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -32335,7 +33461,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437_), 1, 0); return x_1; } } @@ -32347,7 +33473,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -32374,23 +33500,23 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1; x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1985____spec__2(x_3, x_2); x_5 = lean_ctor_get(x_1, 1); lean_inc(x_5); lean_dec(x_1); x_6 = lean_array_size(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1(x_6, x_7, x_5); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1(x_6, x_7, x_5); x_9 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -32410,7 +33536,7 @@ x_18 = l_Lean_Json_mkObj(x_17); return x_18; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -32418,7 +33544,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543____spec__1(x_4, x_5, x_3); return x_6; } } @@ -32426,7 +33552,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543_), 1, 0); return x_1; } } @@ -32438,7 +33564,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokens___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1() { _start: { lean_object* x_1; @@ -32446,63 +33572,63 @@ x_1 = lean_mk_string_unchecked("FoldingRangeParams", 18, 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -32512,7 +33638,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -32524,7 +33650,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -32557,7 +33683,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597_), 1, 0); return x_1; } } @@ -32569,12 +33695,12 @@ x_1 = l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10091_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10664_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -32595,7 +33721,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10091_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10664_), 1, 0); return x_1; } } @@ -32713,7 +33839,7 @@ switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36; return x_2; } case 1: @@ -32741,7 +33867,7 @@ x_3 = l_Lean_Lsp_instToJsonFoldingRangeKind(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -32761,7 +33887,7 @@ switch (x_6) { case 0: { lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); @@ -32798,7 +33924,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1() { _start: { lean_object* x_1; @@ -32806,7 +33932,7 @@ x_1 = lean_mk_string_unchecked("startLine", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2() { _start: { lean_object* x_1; @@ -32814,7 +33940,7 @@ x_1 = lean_mk_string_unchecked("endLine", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -32823,7 +33949,7 @@ lean_inc(x_2); x_3 = l_Lean_JsonNumber_fromNat(x_2); x_4 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -32836,7 +33962,7 @@ lean_inc(x_9); x_10 = l_Lean_JsonNumber_fromNat(x_9); x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -32846,8 +33972,8 @@ lean_ctor_set(x_14, 1, x_7); x_15 = lean_ctor_get(x_1, 2); lean_inc(x_15); lean_dec(x_1); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22; -x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1(x_16, x_15); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22; +x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1(x_16, x_15); lean_dec(x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -32864,11 +33990,11 @@ x_23 = l_Lean_Json_mkObj(x_22); return x_23; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -32877,7 +34003,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769_), 1, 0); return x_1; } } @@ -32889,7 +34015,7 @@ x_1 = l_Lean_Lsp_instToJsonFoldingRange___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1() { _start: { lean_object* x_1; @@ -32897,7 +34023,7 @@ x_1 = lean_mk_string_unchecked("prepareProvider", 15, 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2() { _start: { lean_object* x_1; @@ -32905,84 +34031,84 @@ x_1 = lean_mk_string_unchecked("RenameOptions", 13, 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2851____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -32992,7 +34118,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -33004,7 +34130,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -33037,7 +34163,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameOptions___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838_), 1, 0); return x_1; } } @@ -33049,13 +34175,13 @@ x_1 = l_Lean_Lsp_instFromJsonRenameOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -33072,13 +34198,13 @@ x_10 = l_Lean_Json_mkObj(x_9); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905_(x_2); return x_3; } } @@ -33086,7 +34212,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameOptions___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10332____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10905____boxed), 1, 0); return x_1; } } @@ -33098,7 +34224,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1() { _start: { lean_object* x_1; @@ -33106,79 +34232,79 @@ x_1 = lean_mk_string_unchecked("RenameParams", 12, 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9() { _start: { lean_object* x_1; @@ -33186,52 +34312,52 @@ x_1 = lean_mk_string_unchecked("newName", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -33243,7 +34369,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -33255,7 +34381,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -33269,7 +34395,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -33282,7 +34408,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -33294,7 +34420,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -33308,7 +34434,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9; x_25 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -33320,7 +34446,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -33332,7 +34458,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -33382,7 +34508,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963_), 1, 0); return x_1; } } @@ -33394,7 +34520,7 @@ x_1 = l_Lean_Lsp_instFromJsonRenameParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10535_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11108_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -33403,7 +34529,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -33415,7 +34541,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -33427,7 +34553,7 @@ lean_inc(x_14); lean_dec(x_1); x_15 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_15, 0, x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -33453,7 +34579,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10535_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11108_), 1, 0); return x_1; } } @@ -33465,7 +34591,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1() { _start: { lean_object* x_1; @@ -33473,83 +34599,83 @@ x_1 = lean_mk_string_unchecked("PrepareRenameParams", 19, 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2; x_2 = 1; x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__6; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_44____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -33561,7 +34687,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -33573,7 +34699,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -33587,7 +34713,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -33598,7 +34724,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -33610,7 +34736,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -33653,7 +34779,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186_), 1, 0); return x_1; } } @@ -33665,14 +34791,14 @@ x_1 = l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10719_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11292_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2311_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -33684,7 +34810,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -33707,7 +34833,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPrepareRenameParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10719_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11292_), 1, 0); return x_1; } } @@ -34111,44 +35237,48 @@ l_Lean_Lsp_instReprCompletionItemKind___closed__1 = _init_l_Lean_Lsp_instReprCom lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind___closed__1); l_Lean_Lsp_instReprCompletionItemKind = _init_l_Lean_Lsp_instReprCompletionItemKind(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1254____closed__19); +l_Lean_Lsp_instHashableCompletionItemKind___closed__1 = _init_l_Lean_Lsp_instHashableCompletionItemKind___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItemKind___closed__1); +l_Lean_Lsp_instHashableCompletionItemKind = _init_l_Lean_Lsp_instHashableCompletionItemKind(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItemKind); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1398____closed__19); l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1); l_Lean_Lsp_instFromJsonInsertReplaceEdit = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit(); @@ -34157,132 +35287,144 @@ l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1); l_Lean_Lsp_instToJsonInsertReplaceEdit = _init_l_Lean_Lsp_instToJsonInsertReplaceEdit(); lean_mark_persistent(l_Lean_Lsp_instToJsonInsertReplaceEdit); +l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instBEqInsertReplaceEdit___closed__1); +l_Lean_Lsp_instBEqInsertReplaceEdit = _init_l_Lean_Lsp_instBEqInsertReplaceEdit(); +lean_mark_persistent(l_Lean_Lsp_instBEqInsertReplaceEdit); +l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableInsertReplaceEdit___closed__1); +l_Lean_Lsp_instHashableInsertReplaceEdit = _init_l_Lean_Lsp_instHashableInsertReplaceEdit(); +lean_mark_persistent(l_Lean_Lsp_instHashableInsertReplaceEdit); l_Lean_Lsp_instInhabitedCompletionItemTag = _init_l_Lean_Lsp_instInhabitedCompletionItemTag(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItemTag); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1561____rarg___closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1845____rarg___closed__6); l_Lean_Lsp_instReprCompletionItemTag___closed__1 = _init_l_Lean_Lsp_instReprCompletionItemTag___closed__1(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemTag___closed__1); l_Lean_Lsp_instReprCompletionItemTag = _init_l_Lean_Lsp_instReprCompletionItemTag(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemTag); +l_Lean_Lsp_instHashableCompletionItemTag___closed__1 = _init_l_Lean_Lsp_instHashableCompletionItemTag___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItemTag___closed__1); +l_Lean_Lsp_instHashableCompletionItemTag = _init_l_Lean_Lsp_instHashableCompletionItemTag(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItemTag); l_Lean_Lsp_instToJsonCompletionItemTag___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionItemTag___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemTag___closed__1); l_Lean_Lsp_instToJsonCompletionItemTag___closed__2 = _init_l_Lean_Lsp_instToJsonCompletionItemTag___closed__2(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemTag___closed__2); l_Lean_Lsp_instFromJsonCompletionItemTag___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItemTag___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemTag___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__45); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__47); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__48); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__49); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__50); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756____closed__51); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__46); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__47); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__48); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__49); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__50); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064____closed__51); l_Lean_Lsp_instFromJsonCompletionItem___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItem___closed__1); l_Lean_Lsp_instFromJsonCompletionItem = _init_l_Lean_Lsp_instFromJsonCompletionItem(); @@ -34297,34 +35439,43 @@ l_Lean_Lsp_instInhabitedCompletionItem___closed__2 = _init_l_Lean_Lsp_instInhabi lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem___closed__2); l_Lean_Lsp_instInhabitedCompletionItem = _init_l_Lean_Lsp_instInhabitedCompletionItem(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2200____closed__14); +l_Lean_Lsp_instBEqCompletionItem___closed__1 = _init_l_Lean_Lsp_instBEqCompletionItem___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instBEqCompletionItem___closed__1); +l_Lean_Lsp_instBEqCompletionItem = _init_l_Lean_Lsp_instBEqCompletionItem(); +lean_mark_persistent(l_Lean_Lsp_instBEqCompletionItem); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2657____closed__1(); +l_Lean_Lsp_instHashableCompletionItem___closed__1 = _init_l_Lean_Lsp_instHashableCompletionItem___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItem___closed__1); +l_Lean_Lsp_instHashableCompletionItem = _init_l_Lean_Lsp_instHashableCompletionItem(); +lean_mark_persistent(l_Lean_Lsp_instHashableCompletionItem); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2773____closed__14); l_Lean_Lsp_instFromJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionList___closed__1); l_Lean_Lsp_instFromJsonCompletionList = _init_l_Lean_Lsp_instFromJsonCompletionList(); @@ -34333,34 +35484,34 @@ l_Lean_Lsp_instToJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instToJsonCom lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList___closed__1); l_Lean_Lsp_instToJsonCompletionList = _init_l_Lean_Lsp_instToJsonCompletionList(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943____closed__14); l_Lean_Lsp_instFromJsonCompletionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionParams___closed__1); l_Lean_Lsp_instFromJsonCompletionParams = _init_l_Lean_Lsp_instFromJsonCompletionParams(); @@ -34369,60 +35520,60 @@ l_Lean_Lsp_instToJsonCompletionParams___closed__1 = _init_l_Lean_Lsp_instToJsonC lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionParams___closed__1); l_Lean_Lsp_instToJsonCompletionParams = _init_l_Lean_Lsp_instToJsonCompletionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123____closed__2); l_Lean_Lsp_instToJsonHover___closed__1 = _init_l_Lean_Lsp_instToJsonHover___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonHover___closed__1); l_Lean_Lsp_instToJsonHover = _init_l_Lean_Lsp_instToJsonHover(); lean_mark_persistent(l_Lean_Lsp_instToJsonHover); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2592____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3165____closed__13); l_Lean_Lsp_instFromJsonHover___closed__1 = _init_l_Lean_Lsp_instFromJsonHover___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHover___closed__1); l_Lean_Lsp_instFromJsonHover = _init_l_Lean_Lsp_instFromJsonHover(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHover); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283____closed__8); l_Lean_Lsp_instFromJsonHoverParams___closed__1 = _init_l_Lean_Lsp_instFromJsonHoverParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHoverParams___closed__1); l_Lean_Lsp_instFromJsonHoverParams = _init_l_Lean_Lsp_instFromJsonHoverParams(); @@ -34431,22 +35582,22 @@ l_Lean_Lsp_instToJsonHoverParams___closed__1 = _init_l_Lean_Lsp_instToJsonHoverP lean_mark_persistent(l_Lean_Lsp_instToJsonHoverParams___closed__1); l_Lean_Lsp_instToJsonHoverParams = _init_l_Lean_Lsp_instToJsonHoverParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonHoverParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2880____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3453____closed__8); l_Lean_Lsp_instFromJsonDeclarationParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDeclarationParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDeclarationParams___closed__1); l_Lean_Lsp_instFromJsonDeclarationParams = _init_l_Lean_Lsp_instFromJsonDeclarationParams(); @@ -34455,22 +35606,22 @@ l_Lean_Lsp_instToJsonDeclarationParams___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonDeclarationParams___closed__1); l_Lean_Lsp_instToJsonDeclarationParams = _init_l_Lean_Lsp_instToJsonDeclarationParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonDeclarationParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3050____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3623____closed__8); l_Lean_Lsp_instFromJsonDefinitionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDefinitionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDefinitionParams___closed__1); l_Lean_Lsp_instFromJsonDefinitionParams = _init_l_Lean_Lsp_instFromJsonDefinitionParams(); @@ -34479,22 +35630,22 @@ l_Lean_Lsp_instToJsonDefinitionParams___closed__1 = _init_l_Lean_Lsp_instToJsonD lean_mark_persistent(l_Lean_Lsp_instToJsonDefinitionParams___closed__1); l_Lean_Lsp_instToJsonDefinitionParams = _init_l_Lean_Lsp_instToJsonDefinitionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonDefinitionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3220____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3793____closed__8); l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1); l_Lean_Lsp_instFromJsonTypeDefinitionParams = _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams(); @@ -34503,24 +35654,24 @@ l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1); l_Lean_Lsp_instToJsonTypeDefinitionParams = _init_l_Lean_Lsp_instToJsonTypeDefinitionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3390____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3963____closed__9); l_Lean_Lsp_instFromJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceContext___closed__1); l_Lean_Lsp_instFromJsonReferenceContext = _init_l_Lean_Lsp_instFromJsonReferenceContext(); @@ -34529,32 +35680,32 @@ l_Lean_Lsp_instToJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instToJsonR lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext___closed__1); l_Lean_Lsp_instToJsonReferenceContext = _init_l_Lean_Lsp_instToJsonReferenceContext(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088____closed__13); l_Lean_Lsp_instFromJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceParams___closed__1); l_Lean_Lsp_instFromJsonReferenceParams = _init_l_Lean_Lsp_instFromJsonReferenceParams(); @@ -34563,24 +35714,24 @@ l_Lean_Lsp_instToJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instToJsonRe lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams___closed__1); l_Lean_Lsp_instToJsonReferenceParams = _init_l_Lean_Lsp_instToJsonReferenceParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311____closed__9); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams(); @@ -34589,22 +35740,22 @@ l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1 = _init_l_Lean_Lsp_instTo lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1); l_Lean_Lsp_instToJsonWorkspaceSymbolParams = _init_l_Lean_Lsp_instToJsonWorkspaceSymbolParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceSymbolParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428____closed__8); l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1); l_Lean_Lsp_instFromJsonDocumentHighlightParams = _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams(); @@ -34625,18 +35776,18 @@ l_Lean_Lsp_instToJsonDocumentHighlight___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentHighlight___closed__1); l_Lean_Lsp_instToJsonDocumentHighlight = _init_l_Lean_Lsp_instToJsonDocumentHighlight(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentHighlight); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704____closed__6); l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1); l_Lean_Lsp_instFromJsonDocumentSymbolParams = _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams(); @@ -34850,66 +36001,66 @@ l_Lean_Lsp_instToJsonSymbolKind___closed__45 = _init_l_Lean_Lsp_instToJsonSymbol lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolKind___closed__45); l_Lean_Lsp_instToJsonSymbolKind___closed__46 = _init_l_Lean_Lsp_instToJsonSymbolKind___closed__46(); lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolKind___closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4776____rarg___closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5349____rarg___closed__30); l_Lean_Lsp_instToJsonDocumentSymbol___closed__1 = _init_l_Lean_Lsp_instToJsonDocumentSymbol___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentSymbol___closed__1); l_Lean_Lsp_instToJsonDocumentSymbol = _init_l_Lean_Lsp_instToJsonDocumentSymbol(); @@ -34928,52 +36079,52 @@ l_Lean_Lsp_instFromJsonSymbolTag___closed__1 = _init_l_Lean_Lsp_instFromJsonSymb lean_mark_persistent(l_Lean_Lsp_instFromJsonSymbolTag___closed__1); l_Lean_Lsp_instFromJsonSymbolTag___closed__2 = _init_l_Lean_Lsp_instFromJsonSymbolTag___closed__2(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSymbolTag___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5419____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5992____closed__23); l_Lean_Lsp_instFromJsonSymbolInformation___closed__1 = _init_l_Lean_Lsp_instFromJsonSymbolInformation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSymbolInformation___closed__1); l_Lean_Lsp_instFromJsonSymbolInformation = _init_l_Lean_Lsp_instFromJsonSymbolInformation(); @@ -34982,22 +36133,22 @@ l_Lean_Lsp_instToJsonSymbolInformation___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolInformation___closed__1); l_Lean_Lsp_instToJsonSymbolInformation = _init_l_Lean_Lsp_instToJsonSymbolInformation(); lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolInformation); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311____closed__8); l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams(); @@ -35006,52 +36157,52 @@ l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1 = _init_l_Lean_Lsp_i lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1); l_Lean_Lsp_instToJsonCallHierarchyPrepareParams = _init_l_Lean_Lsp_instToJsonCallHierarchyPrepareParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyPrepareParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5973____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6546____closed__23); l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyItem = _init_l_Lean_Lsp_instFromJsonCallHierarchyItem(); @@ -35064,7 +36215,6 @@ l_Lean_Lsp_instBEqCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instBEqCallHi lean_mark_persistent(l_Lean_Lsp_instBEqCallHierarchyItem___closed__1); l_Lean_Lsp_instBEqCallHierarchyItem = _init_l_Lean_Lsp_instBEqCallHierarchyItem(); lean_mark_persistent(l_Lean_Lsp_instBEqCallHierarchyItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589____closed__1(); l_Lean_Lsp_instHashableCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instHashableCallHierarchyItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instHashableCallHierarchyItem___closed__1); l_Lean_Lsp_instHashableCallHierarchyItem = _init_l_Lean_Lsp_instHashableCallHierarchyItem(); @@ -35077,24 +36227,24 @@ l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3 = _init_l_Lean_Lsp_instInh lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3); l_Lean_Lsp_instInhabitedCallHierarchyItem = _init_l_Lean_Lsp_instInhabitedCallHierarchyItem(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286____closed__9); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams(); @@ -35103,34 +36253,34 @@ l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1 = _init_l_Lean lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1); l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams = _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6839____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7412____closed__14); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall(); @@ -35145,18 +36295,18 @@ l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__2 = _init_l_Lean_Lsp lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__2); l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall = _init_l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593____closed__6); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams(); @@ -35165,28 +36315,28 @@ l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1 = _init_l_Lean lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1); l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams = _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7146____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7719____closed__11); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall(); @@ -35199,164 +36349,164 @@ l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1 = _init_l_Lean_Lsp lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1); l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall = _init_l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__45); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__47); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320____closed__48); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__46); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__47); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893____closed__48); l_Lean_Lsp_instToJsonSemanticTokenType___closed__1 = _init_l_Lean_Lsp_instToJsonSemanticTokenType___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenType___closed__1); l_Lean_Lsp_instToJsonSemanticTokenType = _init_l_Lean_Lsp_instToJsonSemanticTokenType(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenType); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__1___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__2___closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__3___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__4___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__5___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__6___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__7___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__8___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__9___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__10___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__11___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__12___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__13___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__14___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__15___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__16___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__17___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__18___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__19___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__20___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__21___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__22___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__23___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____lambda__24___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__1___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__2___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__3___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__4___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__5___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__6___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__7___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__8___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__9___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__10___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__11___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__12___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__13___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__14___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__15___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__16___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__17___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__18___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__19___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__20___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__21___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__22___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__23___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____lambda__24___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058____closed__2); l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1); l_Lean_Lsp_instFromJsonSemanticTokenType = _init_l_Lean_Lsp_instFromJsonSemanticTokenType(); @@ -35421,70 +36571,70 @@ l_Lean_Lsp_SemanticTokenType_names___closed__25 = _init_l_Lean_Lsp_SemanticToken lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names___closed__25); l_Lean_Lsp_SemanticTokenType_names = _init_l_Lean_Lsp_SemanticTokenType_names(); lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8556____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9129____closed__19); l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1 = _init_l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1); l_Lean_Lsp_instToJsonSemanticTokenModifier = _init_l_Lean_Lsp_instToJsonSemanticTokenModifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenModifier); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__1___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__2___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__3___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__4___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__5___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__6___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__7___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__8___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____lambda__9___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8637____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__1___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__2___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__3___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__4___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__5___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__6___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__7___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__8___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____lambda__9___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9210____closed__1); l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1); l_Lean_Lsp_instFromJsonSemanticTokenModifier = _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier(); @@ -35513,34 +36663,34 @@ l_Lean_Lsp_SemanticTokenModifier_names___closed__11 = _init_l_Lean_Lsp_SemanticT lean_mark_persistent(l_Lean_Lsp_SemanticTokenModifier_names___closed__11); l_Lean_Lsp_SemanticTokenModifier_names = _init_l_Lean_Lsp_SemanticTokenModifier_names(); lean_mark_persistent(l_Lean_Lsp_SemanticTokenModifier_names); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9722____closed__14); l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensLegend = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend(); @@ -35549,38 +36699,38 @@ l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instToJsonSemanticTokensLegend = _init_l_Lean_Lsp_instToJsonSemanticTokensLegend(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9335____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__16); l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensOptions = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions(); @@ -35589,18 +36739,18 @@ l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1 = _init_l_Lean_Lsp_instTo lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1); l_Lean_Lsp_instToJsonSemanticTokensOptions = _init_l_Lean_Lsp_instToJsonSemanticTokensOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensOptions); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131____closed__6); l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensParams = _init_l_Lean_Lsp_instFromJsonSemanticTokensParams(); @@ -35609,22 +36759,22 @@ l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1); l_Lean_Lsp_instToJsonSemanticTokensParams = _init_l_Lean_Lsp_instToJsonSemanticTokensParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256____closed__8); l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensRangeParams = _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams(); @@ -35633,34 +36783,34 @@ l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1 = _init_l_Lean_Lsp_in lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1); l_Lean_Lsp_instToJsonSemanticTokensRangeParams = _init_l_Lean_Lsp_instToJsonSemanticTokensRangeParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensRangeParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9864____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10437____closed__14); l_Lean_Lsp_instFromJsonSemanticTokens___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokens___closed__1); l_Lean_Lsp_instFromJsonSemanticTokens = _init_l_Lean_Lsp_instFromJsonSemanticTokens(); @@ -35669,18 +36819,18 @@ l_Lean_Lsp_instToJsonSemanticTokens___closed__1 = _init_l_Lean_Lsp_instToJsonSem lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokens___closed__1); l_Lean_Lsp_instToJsonSemanticTokens = _init_l_Lean_Lsp_instToJsonSemanticTokens(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokens); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597____closed__6); l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1 = _init_l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1); l_Lean_Lsp_instFromJsonFoldingRangeParams = _init_l_Lean_Lsp_instFromJsonFoldingRangeParams(); @@ -35697,32 +36847,32 @@ l_Lean_Lsp_instToJsonFoldingRangeKind___closed__3 = _init_l_Lean_Lsp_instToJsonF lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRangeKind___closed__3); l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4 = _init_l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769____closed__2); l_Lean_Lsp_instToJsonFoldingRange___closed__1 = _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRange___closed__1); l_Lean_Lsp_instToJsonFoldingRange = _init_l_Lean_Lsp_instToJsonFoldingRange(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRange); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10265____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10838____closed__9); l_Lean_Lsp_instFromJsonRenameOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameOptions___closed__1); l_Lean_Lsp_instFromJsonRenameOptions = _init_l_Lean_Lsp_instFromJsonRenameOptions(); @@ -35731,32 +36881,32 @@ l_Lean_Lsp_instToJsonRenameOptions___closed__1 = _init_l_Lean_Lsp_instToJsonRena lean_mark_persistent(l_Lean_Lsp_instToJsonRenameOptions___closed__1); l_Lean_Lsp_instToJsonRenameOptions = _init_l_Lean_Lsp_instToJsonRenameOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameOptions); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963____closed__13); l_Lean_Lsp_instFromJsonRenameParams___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameParams___closed__1); l_Lean_Lsp_instFromJsonRenameParams = _init_l_Lean_Lsp_instFromJsonRenameParams(); @@ -35765,22 +36915,22 @@ l_Lean_Lsp_instToJsonRenameParams___closed__1 = _init_l_Lean_Lsp_instToJsonRenam lean_mark_persistent(l_Lean_Lsp_instToJsonRenameParams___closed__1); l_Lean_Lsp_instToJsonRenameParams = _init_l_Lean_Lsp_instToJsonRenameParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186____closed__8); l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1 = _init_l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1); l_Lean_Lsp_instFromJsonPrepareRenameParams = _init_l_Lean_Lsp_instFromJsonPrepareRenameParams(); diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index cf1fed8e2379..178031992fd1 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -33,13 +33,13 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shou lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1; lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange__1___closed__7; lean_object* l_Lean_getStructureParentInfo(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__1; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__30; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice__1(lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__10; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -50,7 +50,6 @@ static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getArgExpectedType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_ElabElim_finalize___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -78,19 +77,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_ElabElim_revertArgs___spec__16___rarg___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__4___at_Lean_Elab_Term_getElabElimExprInfo___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent__1(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__1___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_Term_hasElabWithoutExpectedType(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMCtxImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_setMotive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -111,6 +107,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Term_instToStringNamedArg___lambda__1(lean_objec lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__9; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__29; +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDotIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__3; @@ -121,7 +118,6 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1(lean_obje static lean_object* l_Lean_Elab_Term_instToStringNamedArg___closed__1; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__8___closed__2; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9; lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__6___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,6 +136,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_mkMotive___spec__2___closed__2; lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__16; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); @@ -151,7 +148,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__5___closed__4; lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__18; @@ -167,6 +163,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_setMotive___boxed(lean_object LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instToStringNamedArg___closed__4; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at_Lean_Elab_Term_ElabElim_revertArgs___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__10; @@ -182,6 +179,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_A static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__12___closed__7; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__14; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_ElabElim_finalize___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12; lean_object* l_Lean_Meta_forallMetaTelescope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabExplicitUniv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -232,7 +230,6 @@ static lean_object* l_Lean_Elab_Term_ElabElim_finalize___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getArgExpectedType(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -256,15 +253,14 @@ lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instToStringNamedArg___closed__2; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instReprElabElimInfo; -LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_exceptionToSorry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody___spec__1(lean_object*); lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -329,10 +325,12 @@ lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_mkMotive___spec__2___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice_declRange__1___closed__2; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(size_t, size_t, lean_object*); uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_ppGoal_ppVars___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp__1___closed__5; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__11; @@ -359,7 +357,7 @@ lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__19(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_elabAppArgs___spec__4___at_Lean_Elab_Term_elabAppArgs___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__5; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_elabAppArgs___spec__4___at_Lean_Elab_Term_elabAppArgs___spec__5___at_Lean_Elab_Term_elabAppArgs___spec__6(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); @@ -380,12 +378,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_getElabElimExprInfo___lambda__2___boxe uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____lambda__2___closed__11; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4; lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId_toName(lean_object*); +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__7; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__10___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__25; @@ -395,7 +395,6 @@ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_App_0__ static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj__1___closed__1; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___spec__2___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_getNextArg_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult_isOutParamOfLocalInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -408,6 +407,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_ElabEl LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplicitArg___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__4; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__11___closed__1; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___closed__6; @@ -417,6 +417,8 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_fina LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ElabElim_revertArgs___spec__12(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9(lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -433,14 +435,15 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addProjTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeAppInstMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTermEnsuringType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__22___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4; static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__4; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -462,6 +465,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj_declRange__1___closed lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedElabElimInfo___closed__1; lean_object* l_Lean_getPathToBaseStructure_x3f(lean_object*, lean_object*, lean_object*); @@ -469,6 +473,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__ LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at_Lean_Elab_Term_ElabElim_revertArgs___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Elab_instInhabitedInfoTree; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringArg(lean_object*); @@ -482,7 +487,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux( LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadControlTOfMonadControl___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_elabExplicitUnivs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__23; @@ -536,10 +540,10 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_ lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__23___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__1(lean_object*); +static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addInstMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__20(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__5; @@ -560,7 +564,7 @@ uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__2; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___lambda__1___closed__1; -static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); @@ -601,7 +605,6 @@ lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(le lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_209____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__1___closed__7; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___lambda__3___boxed(lean_object**); lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(size_t, size_t, lean_object*); @@ -669,6 +672,7 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_ LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit__1___closed__3; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__21; +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_hasElabWithoutExpectedType___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(lean_object*, lean_object*, lean_object*); @@ -678,6 +682,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ElabElim_rev LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_ElabElim_revertArgs___spec__16___rarg___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1; lean_object* l_Lean_Exception_getRef(lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -691,6 +696,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj__1(lean_obje LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getElabElimInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_annotateIfRec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -704,7 +710,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ElabElim_re static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__1; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(lean_object*); lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkRecAppWithSyntax(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___spec__5(lean_object*, lean_object*, lean_object*); @@ -719,6 +724,7 @@ static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult_hasLocalInstaceWithOutParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent_declRange__1___closed__3; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__16(lean_object*, lean_object*, size_t, size_t); @@ -736,7 +742,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at_Lean_Elab_ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ElabElim_revertArgs___spec__8___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__1; lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -774,7 +779,6 @@ uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processImplicitArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4; lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange__1___closed__3; @@ -786,6 +790,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_App_ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_getElabElimExprInfo___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_eraseNamedArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -806,7 +811,6 @@ uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj_declRange__1___closed__6; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2; lean_object* lean_expr_consume_type_annotations(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv__1___closed__1; @@ -846,6 +850,7 @@ static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__12___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751_(lean_object*); static lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedElabElimInfo; @@ -872,6 +877,8 @@ lean_object* l_Lean_LocalDecl_fvarId(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__1___closed__7; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_ElabElim_mkMotive___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -893,13 +900,14 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange__1(lean lean_object* l_Lean_registerTagAttribute(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitPost___at_Lean_Elab_Term_ElabElim_revertArgs___spec__3(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_eraseIdx___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_nextArgHole_x3f___boxed(lean_object*); extern lean_object* l_Std_Format_defWidth; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__5; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabPipeProj___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -962,6 +970,7 @@ lean_object* l_instMonadControlTOfPure___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__3___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -972,7 +981,6 @@ lean_object* l_List_tail_x21___rarg(lean_object*); extern lean_object* l_Lean_structureResolutionExt; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDotIdent__1(lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLVal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processInstImplicitArg_mkInstMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -983,6 +991,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at_Lean_Elab_Term_ElabElim_revertArgs___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_pushInfoTree___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__4___lambda__1___boxed(lean_object*); @@ -1007,6 +1016,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLVal LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Term_ElabElim_revertArgs___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_synthesizePendingAndNormalizeFunType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ElabElim_revertArgs___spec__7___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLValsAux_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1015,6 +1025,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_main___lambda__1(lean_object* LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj__1___closed__3; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addTermInfo_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_hasArgsToProcess___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1030,6 +1041,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice__1___closed__2; lean_object* l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__22; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__12___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15; lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_ElabElim_mkMotive___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__7___closed__1; @@ -1052,8 +1064,11 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0_ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethodAlias_x3f___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getArgExpectedType___boxed(lean_object*); +static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__1; static lean_object* l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___lambda__1___closed__2; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); @@ -1068,6 +1083,7 @@ LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_Term_eraseNamedArg_ static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj__1___closed__2; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1075,6 +1091,7 @@ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_App_0__ static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__1___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabProj__1(lean_object*); +lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__1___closed__3; lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3025_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_fTypeHasOptAutoParams___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1082,7 +1099,6 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj_declRange__1___closed__5; extern lean_object* l_Lean_brecOnSuffix; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3; @@ -1135,7 +1151,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRange__1_ static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processInstImplicitArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__1___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__4___closed__2; @@ -1160,16 +1175,17 @@ static lean_object* l_Lean_Elab_Term_getElabElimExprInfo___lambda__1___closed__1 static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560_(lean_object*); static lean_object* l_Lean_Elab_Term_getElabElimExprInfo___closed__3; uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__5___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__8; lean_object* l_Lean_LocalDecl_type(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addNewArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_getElabElimExprInfo___closed__2; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___closed__2; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at_Lean_Elab_Term_ElabElim_revertArgs___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____lambda__2___closed__15; @@ -1189,7 +1205,6 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_ LEAN_EXPORT lean_object* l_Lean_isRec___at___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_hasArgsToProcess___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_indexOfAux___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1253,6 +1268,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ensureArgTy uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadCoreM; +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_typeMatchesBaseName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange__1___closed__4; static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; @@ -1271,7 +1287,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabAppArgs_main(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_mkMotive___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_Term_eraseNamedArg___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplicitArg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1283,6 +1298,7 @@ lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkCoe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_get_x21___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg_go___spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1325,6 +1341,7 @@ lean_object* l_Lean_Expr_fvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__6; size_t lean_array_size(lean_object*); +lean_object* l_Lean_instInhabitedPersistentArray(lean_object*); lean_object* l_Lean_Elab_Term_addTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__3___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent__1___closed__1; @@ -1355,7 +1372,6 @@ lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___lambda__1___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10; LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Elab_Term_getElabElimExprInfo___spec__2___at_Lean_Elab_Term_getElabElimExprInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_ElabElim_finalize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabIdent_declRange__1___closed__1; @@ -1363,10 +1379,8 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwInvalidNamed LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__9___boxed(lean_object**); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabAppArgs_synthesizeAppInstMVars___boxed(lean_object*); -static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform___at_Lean_Elab_Term_ElabElim_revertArgs___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDotIdent_declRange__1(lean_object*); @@ -1375,11 +1389,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAsElim; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_normalizeFunType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_getElabElimExprInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1; uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); @@ -1396,6 +1410,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_synthesizePendingAndNormalizeFunType___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ElabElim_revertArgs___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_elabAppArgs___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_8219____lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_ElabElim_finalize___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_ElabElim_revertArgs___spec__16___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1443,6 +1459,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__2; static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange__1___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__12(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice_declRange__1(lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); @@ -1452,14 +1469,15 @@ lean_object* l_Lean_Expr_bindingInfo_x21(lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_findNamedArgDependsOnCurrent_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_registerMVarArgName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName___closed__1; lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__5___closed__2; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__8; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__7___closed__2; @@ -1479,6 +1497,7 @@ lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwInvalidNamedArg(lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__13___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__3(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplicitArg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); @@ -1487,7 +1506,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNe LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___spec__3(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1; static lean_object* l_Lean_Elab_Term_elabPipeProj___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1515,11 +1534,12 @@ size_t lean_usize_land(size_t, size_t); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_reprElabElimInfo____x40_Lean_Elab_App___hyg_7108____closed__31; lean_object* l_Lean_Elab_Term_isMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_resolveName_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_withForallBody___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addInstMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1532,6 +1552,7 @@ lean_object* l_instMonadControlReaderT___lambda__3___boxed(lean_object*, lean_ob LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_findMethod_x3f___spec__13(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabAppArgs_eraseNamedArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at_Lean_Elab_Term_ElabElim_revertArgs___spec__2___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: @@ -55097,7 +55118,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean x_21 = lean_ctor_get(x_7, 1); lean_inc(x_21); lean_dec(x_7); -x_22 = lean_ctor_get(x_17, 3); +x_22 = lean_ctor_get(x_17, 0); lean_inc(x_22); lean_inc(x_3); lean_inc(x_6); @@ -64720,7 +64741,651 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_instInhabitedPersistentArray(lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__8; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1; +x_3 = lean_unsigned_to_nat(1653u); +x_4 = lean_unsigned_to_nat(84u); +x_5 = l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___lambda__1___closed__2; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1; +x_11 = l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3(x_10); +x_12 = lean_array_uset(x_7, x_2, x_11); +x_2 = x_9; +x_3 = x_12; +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +lean_dec(x_5); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_ctor_get(x_17, 6); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_array_uset(x_7, x_2, x_19); +x_2 = x_9; +x_3 = x_20; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; size_t x_10; size_t x_11; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get(x_6, 2); +lean_inc(x_7); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +lean_dec(x_7); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +if (x_9 == 0) +{ +lean_dec(x_6); +x_2 = x_11; +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = l_Lean_Elab_instInhabitedInfoTree; +x_14 = l_Lean_PersistentArray_get_x21___rarg(x_13, x_6, x_8); +x_15 = lean_array_push(x_4, x_14); +x_2 = x_11; +x_4 = x_15; +goto _start; +} +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = l_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___closed__1; +return x_5; +} +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_le(x_3, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +x_8 = l_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___closed__1; +return x_8; +} +else +{ +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_usize_of_nat(x_2); +x_10 = lean_usize_of_nat(x_3); +x_11 = l_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___closed__1; +x_12 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(x_1, x_9, x_10, x_11); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_5, x_4); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_6); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; +lean_dec(x_6); +x_16 = lean_array_uget(x_3, x_5); +x_17 = l_Lean_Elab_pushInfoTree___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__3(x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = 1; +x_20 = lean_usize_add(x_5, x_19); +x_21 = lean_box(0); +x_5 = x_20; +x_6 = x_21; +x_13 = x_18; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_st_ref_get(x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 6); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_dec(x_10); +x_17 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(x_8, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_20 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_st_ref_get(x_8, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_24, 6); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +lean_inc(x_8); +x_28 = lean_apply_8(x_2, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_st_ref_take(x_8, x_30); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_32, 6); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = !lean_is_exclusive(x_32); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; +x_36 = lean_ctor_get(x_32, 6); +lean_dec(x_36); +x_37 = !lean_is_exclusive(x_33); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_33, 1); +lean_dec(x_38); +x_39 = l_Lean_PersistentArray_push___rarg(x_18, x_29); +lean_ctor_set(x_33, 1, x_39); +x_40 = lean_st_ref_set(x_8, x_32, x_34); +lean_dec(x_8); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +lean_ctor_set(x_40, 0, x_21); +return x_40; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_21); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +} +else +{ +uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_45 = lean_ctor_get_uint8(x_33, sizeof(void*)*2); +x_46 = lean_ctor_get(x_33, 0); +lean_inc(x_46); +lean_dec(x_33); +x_47 = l_Lean_PersistentArray_push___rarg(x_18, x_29); +x_48 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +lean_ctor_set_uint8(x_48, sizeof(void*)*2, x_45); +lean_ctor_set(x_32, 6, x_48); +x_49 = lean_st_ref_set(x_8, x_32, x_34); +lean_dec(x_8); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_51 = x_49; +} else { + lean_dec_ref(x_49); + x_51 = lean_box(0); +} +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_51; +} +lean_ctor_set(x_52, 0, x_21); +lean_ctor_set(x_52, 1, x_50); +return x_52; +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_53 = lean_ctor_get(x_32, 0); +x_54 = lean_ctor_get(x_32, 1); +x_55 = lean_ctor_get(x_32, 2); +x_56 = lean_ctor_get(x_32, 3); +x_57 = lean_ctor_get(x_32, 4); +x_58 = lean_ctor_get(x_32, 5); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_32); +x_59 = lean_ctor_get_uint8(x_33, sizeof(void*)*2); +x_60 = lean_ctor_get(x_33, 0); +lean_inc(x_60); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_61 = x_33; +} else { + lean_dec_ref(x_33); + x_61 = lean_box(0); +} +x_62 = l_Lean_PersistentArray_push___rarg(x_18, x_29); +if (lean_is_scalar(x_61)) { + x_63 = lean_alloc_ctor(0, 2, 1); +} else { + x_63 = x_61; +} +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_62); +lean_ctor_set_uint8(x_63, sizeof(void*)*2, x_59); +x_64 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_64, 0, x_53); +lean_ctor_set(x_64, 1, x_54); +lean_ctor_set(x_64, 2, x_55); +lean_ctor_set(x_64, 3, x_56); +lean_ctor_set(x_64, 4, x_57); +lean_ctor_set(x_64, 5, x_58); +lean_ctor_set(x_64, 6, x_63); +x_65 = lean_st_ref_set(x_8, x_64, x_34); +lean_dec(x_8); +x_66 = lean_ctor_get(x_65, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_67 = x_65; +} else { + lean_dec_ref(x_65); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_21); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +else +{ +uint8_t x_69; +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_8); +x_69 = !lean_is_exclusive(x_28); +if (x_69 == 0) +{ +return x_28; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_28, 0); +x_71 = lean_ctor_get(x_28, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_28); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} +} +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_73 = lean_ctor_get(x_20, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_20, 1); +lean_inc(x_74); +lean_dec(x_20); +x_75 = lean_st_ref_get(x_8, x_74); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_78 = lean_ctor_get(x_76, 6); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +lean_inc(x_8); +x_80 = lean_apply_8(x_2, x_79, x_3, x_4, x_5, x_6, x_7, x_8, x_77); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +lean_dec(x_80); +x_83 = lean_st_ref_take(x_8, x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_84, 6); +lean_inc(x_85); +x_86 = lean_ctor_get(x_83, 1); +lean_inc(x_86); +lean_dec(x_83); +x_87 = !lean_is_exclusive(x_84); +if (x_87 == 0) +{ +lean_object* x_88; uint8_t x_89; +x_88 = lean_ctor_get(x_84, 6); +lean_dec(x_88); +x_89 = !lean_is_exclusive(x_85); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_90 = lean_ctor_get(x_85, 1); +lean_dec(x_90); +x_91 = l_Lean_PersistentArray_push___rarg(x_18, x_81); +lean_ctor_set(x_85, 1, x_91); +x_92 = lean_st_ref_set(x_8, x_84, x_86); +lean_dec(x_8); +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); +lean_dec(x_94); +lean_ctor_set_tag(x_92, 1); +lean_ctor_set(x_92, 0, x_73); +return x_92; +} +else +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +lean_dec(x_92); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_73); +lean_ctor_set(x_96, 1, x_95); +return x_96; +} +} +else +{ +uint8_t x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_97 = lean_ctor_get_uint8(x_85, sizeof(void*)*2); +x_98 = lean_ctor_get(x_85, 0); +lean_inc(x_98); +lean_dec(x_85); +x_99 = l_Lean_PersistentArray_push___rarg(x_18, x_81); +x_100 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +lean_ctor_set_uint8(x_100, sizeof(void*)*2, x_97); +lean_ctor_set(x_84, 6, x_100); +x_101 = lean_st_ref_set(x_8, x_84, x_86); +lean_dec(x_8); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_103 = x_101; +} else { + lean_dec_ref(x_101); + x_103 = lean_box(0); +} +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); +} else { + x_104 = x_103; + lean_ctor_set_tag(x_104, 1); +} +lean_ctor_set(x_104, 0, x_73); +lean_ctor_set(x_104, 1, x_102); +return x_104; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_105 = lean_ctor_get(x_84, 0); +x_106 = lean_ctor_get(x_84, 1); +x_107 = lean_ctor_get(x_84, 2); +x_108 = lean_ctor_get(x_84, 3); +x_109 = lean_ctor_get(x_84, 4); +x_110 = lean_ctor_get(x_84, 5); +lean_inc(x_110); +lean_inc(x_109); +lean_inc(x_108); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_84); +x_111 = lean_ctor_get_uint8(x_85, sizeof(void*)*2); +x_112 = lean_ctor_get(x_85, 0); +lean_inc(x_112); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_113 = x_85; +} else { + lean_dec_ref(x_85); + x_113 = lean_box(0); +} +x_114 = l_Lean_PersistentArray_push___rarg(x_18, x_81); +if (lean_is_scalar(x_113)) { + x_115 = lean_alloc_ctor(0, 2, 1); +} else { + x_115 = x_113; +} +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set_uint8(x_115, sizeof(void*)*2, x_111); +x_116 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_116, 0, x_105); +lean_ctor_set(x_116, 1, x_106); +lean_ctor_set(x_116, 2, x_107); +lean_ctor_set(x_116, 3, x_108); +lean_ctor_set(x_116, 4, x_109); +lean_ctor_set(x_116, 5, x_110); +lean_ctor_set(x_116, 6, x_115); +x_117 = lean_st_ref_set(x_8, x_116, x_86); +lean_dec(x_8); +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_119 = x_117; +} else { + lean_dec_ref(x_117); + x_119 = lean_box(0); +} +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_119; + lean_ctor_set_tag(x_120, 1); +} +lean_ctor_set(x_120, 0, x_73); +lean_ctor_set(x_120, 1, x_118); +return x_120; +} +} +else +{ +uint8_t x_121; +lean_dec(x_73); +lean_dec(x_18); +lean_dec(x_8); +x_121 = !lean_is_exclusive(x_80); +if (x_121 == 0) +{ +return x_80; +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_80, 0); +x_123 = lean_ctor_get(x_80, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_80); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; @@ -64749,15 +65414,15 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = lean_alloc_closure((void*)(l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg___boxed), 3, 0); +x_5 = lean_alloc_closure((void*)(l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg___boxed), 3, 0); return x_5; } } -static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1() { +static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1() { _start: { lean_object* x_1; @@ -64765,16 +65430,16 @@ x_1 = lean_mk_string_unchecked(":", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2() { +static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1; +x_1 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3() { +static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3() { _start: { lean_object* x_1; @@ -64782,20 +65447,20 @@ x_1 = lean_mk_string_unchecked(" ", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4() { +static lean_object* _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3; +x_1 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg(x_6, x_7, x_8); +x_9 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg(x_6, x_7, x_8); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { @@ -64842,7 +65507,7 @@ x_24 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__8; x_25 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2; +x_26 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2; x_27 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); @@ -64856,7 +65521,7 @@ x_31 = l_Lean_MessageData_ofFormat(x_30); x_32 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_32, 0, x_27); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4; +x_33 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4; x_34 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_34, 0, x_32); lean_ctor_set(x_34, 1, x_33); @@ -64907,7 +65572,7 @@ x_47 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__8; x_48 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_48, 0, x_47); lean_ctor_set(x_48, 1, x_46); -x_49 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2; +x_49 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2; x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); @@ -64921,7 +65586,7 @@ x_54 = l_Lean_MessageData_ofFormat(x_53); x_55 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_55, 0, x_50); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4; +x_56 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4; x_57 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); @@ -65007,7 +65672,7 @@ x_78 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__8; x_79 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_79, 0, x_78); lean_ctor_set(x_79, 1, x_77); -x_80 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2; +x_80 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2; x_81 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_81, 0, x_79); lean_ctor_set(x_81, 1, x_80); @@ -65021,7 +65686,7 @@ x_85 = l_Lean_MessageData_ofFormat(x_84); x_86 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_86, 0, x_81); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4; +x_87 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4; x_88 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_88, 0, x_86); lean_ctor_set(x_88, 1, x_87); @@ -65053,7 +65718,7 @@ return x_94; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -65074,7 +65739,7 @@ x_13 = lean_array_uget(x_3, x_2); x_14 = lean_unsigned_to_nat(0u); x_15 = lean_array_uset(x_3, x_2, x_14); lean_inc(x_8); -x_16 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_16 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -65090,7 +65755,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -65172,15 +65837,15 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg___boxed), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg___boxed), 8, 0); return x_2; } } -static lean_object* _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1() { +static lean_object* _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1() { _start: { lean_object* x_1; @@ -65188,23 +65853,23 @@ x_1 = lean_mk_string_unchecked(", errors ", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2() { +static lean_object* _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1; +x_1 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { size_t x_10; size_t x_11; lean_object* x_12; uint8_t x_13; x_10 = lean_array_size(x_2); x_11 = 0; lean_inc(x_7); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(x_10, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12(x_10, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); x_13 = !lean_is_exclusive(x_12); if (x_13 == 0) { @@ -65215,7 +65880,7 @@ x_16 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__8; lean_ctor_set_tag(x_12, 7); lean_ctor_set(x_12, 1, x_1); lean_ctor_set(x_12, 0, x_16); -x_17 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2; +x_17 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2; x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_12); lean_ctor_set(x_18, 1, x_17); @@ -65226,7 +65891,7 @@ lean_ctor_set(x_20, 1, x_19); x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_16); -x_22 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +x_22 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_15); lean_dec(x_7); return x_22; } @@ -65242,7 +65907,7 @@ x_25 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mkProjAndCheck___closed__8; x_26 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_1); -x_27 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2; +x_27 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2; x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -65253,20 +65918,61 @@ lean_ctor_set(x_30, 1, x_29); x_31 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_25); -x_32 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg(x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +x_32 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg(x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_24); lean_dec(x_7); return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___boxed), 9, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___boxed), 9, 0); return x_2; } } +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_box(0); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(x_1, x_2, x_1, x_3, x_4, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_12); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_2); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1() { _start: { @@ -65294,23 +66000,105 @@ x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } +static lean_object* _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1() { +_start: +{ +size_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_box_usize(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; x_9 = lean_array_size(x_1); x_10 = 0; +lean_inc(x_1); x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2(x_9, x_10, x_1); -x_12 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3; -x_13 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg(x_12, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_13; +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(x_9, x_10, x_1); +x_13 = lean_array_get_size(x_12); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(x_12, x_14, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_16 = lean_ctor_get(x_6, 5); +lean_inc(x_16); +x_17 = lean_box(0); +x_18 = lean_array_size(x_15); +x_19 = lean_box_usize(x_18); +x_20 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1; +x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1___boxed), 11, 4); +lean_closure_set(x_21, 0, x_15); +lean_closure_set(x_21, 1, x_17); +lean_closure_set(x_21, 2, x_19); +lean_closure_set(x_21, 3, x_20); +x_22 = lean_box(0); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_16); +x_24 = lean_alloc_ctor(13, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2___boxed), 9, 1); +lean_closure_set(x_25, 0, x_24); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_26 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__8(x_21, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3; +x_29 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg(x_28, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_26); +if (x_30 == 0) +{ +return x_26; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_26, 0); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_26); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} } } LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed), 8, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg), 8, 0); return x_2; } } @@ -65326,21 +66114,78 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_merg return x_6; } } -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg(x_1, x_2, x_3); +x_4 = l_Array_filterMapM___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(x_1, x_2, x_3); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7(x_1, x_2, x_3, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__11(x_1, x_2, x_3, x_4); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -65348,11 +66193,11 @@ lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); @@ -65361,7 +66206,7 @@ lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -65369,7 +66214,7 @@ x_11 = lean_unbox_usize(x_1); lean_dec(x_1); x_12 = lean_unbox_usize(x_2); lean_dec(x_2); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__6(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__12(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -65378,11 +66223,11 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__7___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__13___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -65391,11 +66236,11 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -65403,16 +66248,38 @@ lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__1(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_9; +return x_10; } } LEAN_EXPORT lean_object* l_Lean_setEnv___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -65858,7 +66725,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0_ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Term_ElabElim_finalize___lambda__9___closed__8; x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__1; -x_3 = lean_unsigned_to_nat(1669u); +x_3 = lean_unsigned_to_nat(1677u); x_4 = lean_unsigned_to_nat(21u); x_5 = l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___lambda__1___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -66186,10 +67053,6 @@ lean_dec(x_29); lean_dec(x_1); lean_ctor_set(x_10, 5, x_30); x_31 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg(x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_24); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); return x_31; } else @@ -66241,10 +67104,6 @@ lean_ctor_set(x_47, 11, x_44); lean_ctor_set_uint8(x_47, sizeof(void*)*12, x_43); lean_ctor_set_uint8(x_47, sizeof(void*)*12 + 1, x_45); x_48 = l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg(x_17, x_6, x_7, x_8, x_9, x_47, x_11, x_24); -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); return x_48; } } @@ -68613,7 +69472,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68623,7 +69482,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68633,27 +69492,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5() { _start: { lean_object* x_1; @@ -68661,17 +69520,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7() { _start: { lean_object* x_1; @@ -68679,37 +69538,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_4____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11() { _start: { lean_object* x_1; @@ -68717,17 +69576,17 @@ x_1 = lean_mk_string_unchecked("App", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13() { _start: { lean_object* x_1; @@ -68735,33 +69594,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14; -x_2 = lean_unsigned_to_nat(24560u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14; +x_2 = lean_unsigned_to_nat(24751u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -69609,24 +70468,30 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailu lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__2); -l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1(); -lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1); -l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2(); -lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__2); -l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3(); -lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3); -l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4(); -lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4); -l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1 = _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__1); -l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2 = _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___rarg___closed__2); +l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1 = _init_l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__3___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__1); +l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1(); +lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__1); +l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2(); +lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__2); +l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3(); +lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__3); +l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4 = _init_l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4(); +lean_mark_persistent(l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__10___closed__4); +l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1 = _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__1); +l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2 = _init_l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__9___rarg___closed__2); l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1); l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__2 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__2(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__2); l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3(); lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__3); +l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1 = _init_l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1(); +lean_mark_persistent(l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___boxed__const__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__2(); @@ -69885,37 +70750,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabProj_declRange__1___close if (builtin) {res = l___regBuiltin_Lean_Elab_Term_elabProj_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24560_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_24751_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index c4520d720e36..bf27cb4f9833 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -309,7 +309,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange__1___clo static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_elabClear___spec__14___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabOpen___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabHole_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Lean_popScope___at_Lean_Elab_Term_elabOpen___spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSyntheticHole__1___closed__3; @@ -469,7 +468,6 @@ lean_object* l_Lean_Elab_Term_mkSaveInfoAnnotation(lean_object*); static lean_object* l_Lean_Elab_Term_elabLetMVar___closed__3; lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Elab_Term_elabOpen___closed__2; -lean_object* l_Lean_Elab_Term_mkTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Term_elabOpen___spec__34___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange__1___closed__4; static lean_object* l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1; @@ -648,6 +646,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabOp LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_elabClear___spec__11___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabClear_declRange__1___closed__1; static lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_evalFilePathUnsafe___closed__2; +lean_object* l_Lean_Elab_Term_withTermInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabOpen_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabEnsureExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -850,7 +849,6 @@ uint8_t l_Lean_DataValue_sameCtor(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabClear__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabIncludeStr__1(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange__1___closed__7; -lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabSetOption___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_getMVarFromUserName___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNumLit__1___closed__3; @@ -1078,7 +1076,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabIncludeStr__1___closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__52(size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCompletion_declRange__1___closed__4; static lean_object* l_Lean_Elab_Term_elabWaitIfTypeMVar___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__2; static lean_object* l_Lean_Elab_Term_elabScientificLit___closed__10; LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Term_elabOpen___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2245,86 +2242,85 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_14 = l_Lean_Elab_Term_saveState___rarg(x_4, x_5, x_6, x_7, x_8, x_9); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_box(0); -x_18 = 1; +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +x_15 = l_Lean_Syntax_getId(x_11); +x_16 = 1; +lean_inc(x_2); +lean_inc(x_1); +x_17 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_15); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_2); +lean_ctor_set_uint8(x_17, sizeof(void*)*4, x_16); +x_18 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l_Lean_Elab_Term_saveState___rarg(x_4, x_5, x_6, x_7, x_8, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_11); -x_19 = l_Lean_Elab_Term_elabTerm(x_11, x_17, x_18, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_16); -if (lean_obj_tag(x_19) == 0) +x_24 = l_Lean_Elab_Term_elabTerm(x_11, x_23, x_16, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_15); -lean_dec(x_11); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_21); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); lean_inc(x_1); -x_22 = l_Lean_Elab_Term_addDotCompletionInfo(x_1, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_21); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_unsigned_to_nat(1u); -x_25 = l_Lean_Syntax_getArg(x_1, x_24); +x_27 = l_Lean_Elab_Term_addDotCompletionInfo(x_1, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_26); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_unsigned_to_nat(1u); +x_30 = l_Lean_Syntax_getArg(x_1, x_29); lean_dec(x_1); -x_26 = l_Lean_Elab_Term_elabPipeCompletion___lambda__1___closed__2; -x_27 = l_Lean_throwErrorAt___at_Lean_Elab_Term_elabPipeCompletion___spec__1(x_25, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_23); +x_31 = l_Lean_Elab_Term_elabPipeCompletion___lambda__1___closed__2; +x_32 = l_Lean_throwErrorAt___at_Lean_Elab_Term_elabPipeCompletion___spec__1(x_30, x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_28); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_25); -return x_27; +lean_dec(x_30); +return x_32; } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_19); -if (x_28 == 0) +uint8_t x_33; +lean_dec(x_2); +x_33 = !lean_is_exclusive(x_24); +if (x_33 == 0) { -lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_29 = lean_ctor_get(x_19, 0); -x_30 = lean_ctor_get(x_19, 1); -x_31 = l_Lean_Exception_isInterrupt(x_29); -if (x_31 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +x_36 = l_Lean_Exception_isInterrupt(x_34); +if (x_36 == 0) { -uint8_t x_32; -x_32 = l_Lean_Exception_isRuntime(x_29); -if (x_32 == 0) +uint8_t x_37; +x_37 = l_Lean_Exception_isRuntime(x_34); +if (x_37 == 0) { -uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_free_object(x_19); -lean_dec(x_29); -x_33 = 0; -x_34 = l_Lean_Elab_Term_SavedState_restore(x_15, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_30); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); +uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_free_object(x_24); lean_dec(x_34); -x_36 = lean_ctor_get(x_5, 1); -lean_inc(x_36); -x_37 = l_Lean_Syntax_getId(x_11); -lean_dec(x_11); -lean_inc(x_1); -x_38 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_38, 0, x_1); -lean_ctor_set(x_38, 1, x_37); -lean_ctor_set(x_38, 2, x_36); -lean_ctor_set(x_38, 3, x_2); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_18); -x_39 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_35); +x_38 = 0; +x_39 = l_Lean_Elab_Term_SavedState_restore(x_21, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_35); x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); lean_dec(x_39); @@ -2342,42 +2338,38 @@ return x_44; } else { -lean_dec(x_15); -lean_dec(x_11); +lean_dec(x_21); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -return x_19; +return x_24; } } else { -lean_dec(x_15); -lean_dec(x_11); +lean_dec(x_21); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -return x_19; +return x_24; } } else { lean_object* x_45; lean_object* x_46; uint8_t x_47; -x_45 = lean_ctor_get(x_19, 0); -x_46 = lean_ctor_get(x_19, 1); +x_45 = lean_ctor_get(x_24, 0); +x_46 = lean_ctor_get(x_24, 1); lean_inc(x_46); lean_inc(x_45); -lean_dec(x_19); +lean_dec(x_24); x_47 = l_Lean_Exception_isInterrupt(x_45); if (x_47 == 0) { @@ -2385,76 +2377,57 @@ uint8_t x_48; x_48 = l_Lean_Exception_isRuntime(x_45); if (x_48 == 0) { -uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_dec(x_45); x_49 = 0; -x_50 = l_Lean_Elab_Term_SavedState_restore(x_15, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_46); +x_50 = l_Lean_Elab_Term_SavedState_restore(x_21, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_46); x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = lean_ctor_get(x_5, 1); -lean_inc(x_52); -x_53 = l_Lean_Syntax_getId(x_11); -lean_dec(x_11); -lean_inc(x_1); -x_54 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_54, 0, x_1); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_52); -lean_ctor_set(x_54, 3, x_2); -lean_ctor_set_uint8(x_54, sizeof(void*)*4, x_18); -x_55 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_54, x_3, x_4, x_5, x_6, x_7, x_8, x_51); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -lean_dec(x_55); -x_57 = lean_unsigned_to_nat(1u); -x_58 = l_Lean_Syntax_getArg(x_1, x_57); +x_52 = lean_unsigned_to_nat(1u); +x_53 = l_Lean_Syntax_getArg(x_1, x_52); lean_dec(x_1); -x_59 = l_Lean_Elab_Term_elabPipeCompletion___lambda__1___closed__2; -x_60 = l_Lean_throwErrorAt___at_Lean_Elab_Term_elabPipeCompletion___spec__1(x_58, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_56); +x_54 = l_Lean_Elab_Term_elabPipeCompletion___lambda__1___closed__2; +x_55 = l_Lean_throwErrorAt___at_Lean_Elab_Term_elabPipeCompletion___spec__1(x_53, x_54, x_3, x_4, x_5, x_6, x_7, x_8, x_51); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_58); -return x_60; +lean_dec(x_53); +return x_55; } else { -lean_object* x_61; -lean_dec(x_15); -lean_dec(x_11); +lean_object* x_56; +lean_dec(x_21); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_45); -lean_ctor_set(x_61, 1, x_46); -return x_61; +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_45); +lean_ctor_set(x_56, 1, x_46); +return x_56; } } else { -lean_object* x_62; -lean_dec(x_15); -lean_dec(x_11); +lean_object* x_57; +lean_dec(x_21); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_45); -lean_ctor_set(x_62, 1, x_46); -return x_62; +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_45); +lean_ctor_set(x_57, 1, x_46); +return x_57; } } } @@ -24124,7 +24097,7 @@ x_28 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_28, 0, x_1); lean_ctor_set(x_28, 1, x_23); lean_ctor_set(x_28, 2, x_27); -x_29 = lean_alloc_ctor(4, 1, 0); +x_29 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_29, 0, x_28); x_30 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_26); x_31 = !lean_is_exclusive(x_30); @@ -24520,7 +24493,7 @@ x_120 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_120, 0, x_1); lean_ctor_set(x_120, 1, x_115); lean_ctor_set(x_120, 2, x_119); -x_121 = lean_alloc_ctor(4, 1, 0); +x_121 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_121, 0, x_120); x_122 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(x_121, x_3, x_4, x_5, x_6, x_7, x_8, x_118); x_123 = lean_ctor_get(x_122, 1); @@ -25630,17 +25603,6 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_11 = lean_box(0); -x_12 = lean_box(0); -x_13 = 0; -x_14 = l_Lean_Elab_Term_mkTermInfo(x_12, x_1, x_3, x_2, x_11, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_14; -} -} static lean_object* _init_l_Lean_Elab_Term_elabWithAnnotateTerm___closed__1() { _start: { @@ -25682,7 +25644,7 @@ return x_12; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; x_13 = lean_unsigned_to_nat(1u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); x_15 = lean_unsigned_to_nat(2u); @@ -25697,27 +25659,12 @@ lean_closure_set(x_20, 0, x_16); lean_closure_set(x_20, 1, x_2); lean_closure_set(x_20, 2, x_18); lean_closure_set(x_20, 3, x_19); -lean_inc(x_14); -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1___boxed), 10, 2); -lean_closure_set(x_21, 0, x_14); -lean_closure_set(x_21, 1, x_2); -x_22 = l_Lean_Elab_Term_withInfoContext_x27(x_14, x_20, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_22; -} -} +x_21 = lean_box(0); +x_22 = lean_box(0); +x_23 = 0; +x_24 = l_Lean_Elab_Term_withTermInfoContext_x27(x_22, x_14, x_20, x_2, x_21, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_24; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm__1___closed__1() { diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index a3cce508f79d..180c28b3c759 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -15215,7 +15215,7 @@ lean_dec(x_11); x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_1); lean_ctor_set(x_15, 1, x_2); -x_16 = lean_alloc_ctor(2, 1, 0); +x_16 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_16, 0, x_15); x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_16); @@ -15267,7 +15267,7 @@ lean_dec(x_29); x_33 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_33, 0, x_1); lean_ctor_set(x_33, 1, x_2); -x_34 = lean_alloc_ctor(2, 1, 0); +x_34 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_34, 0, x_33); x_35 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_35, 0, x_34); @@ -18250,7 +18250,7 @@ x_9 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_1); lean_ctor_set(x_9, 2, x_2); -x_10 = lean_alloc_ctor(3, 1, 0); +x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__2___boxed), 5, 1); lean_closure_set(x_11, 0, x_10); @@ -36800,7 +36800,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn _start: { lean_object* x_5; lean_object* x_6; -x_5 = lean_alloc_ctor(6, 1, 0); +x_5 = lean_alloc_ctor(7, 1, 0); lean_ctor_set(x_5, 0, x_1); x_6 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_5, x_2, x_3, x_4); return x_6; @@ -37190,7 +37190,7 @@ x_26 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_26, 0, x_1); lean_ctor_set(x_26, 1, x_21); lean_ctor_set(x_26, 2, x_25); -x_27 = lean_alloc_ctor(4, 1, 0); +x_27 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_27, 0, x_26); x_28 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_27, x_3, x_4, x_24); x_29 = !lean_is_exclusive(x_28); @@ -37580,7 +37580,7 @@ x_118 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_118, 0, x_1); lean_ctor_set(x_118, 1, x_113); lean_ctor_set(x_118, 2, x_117); -x_119 = lean_alloc_ctor(4, 1, 0); +x_119 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_119, 0, x_118); x_120 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_119, x_3, x_4, x_116); x_121 = lean_ctor_get(x_120, 1); diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index 8c69c5b33845..609d3cc9afee 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -40,7 +40,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_From static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__22; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__14; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__14; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__5; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__16; static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__4; @@ -51,6 +50,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJs static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__20; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__6; @@ -58,10 +58,8 @@ lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__23; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__6; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__5; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__17; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -70,17 +68,16 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__5; +static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__5; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__13; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__7; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lambda__1___closed__13; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__15; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lambda__1___closed__9; lean_object* l_Lean_Elab_Deriving_mkInstanceCmds(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__11; @@ -136,7 +133,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader(lean_obj static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lambda__1___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBody(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__5___boxed(lean_object**); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__37; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -152,6 +148,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJ static lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__1; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__24; @@ -164,6 +161,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__10; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__20; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___lambda__1___closed__12; @@ -188,6 +186,7 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromT static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__11; static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__31; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__9; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__10; @@ -214,6 +213,7 @@ lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_ob lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___closed__1; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8; lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__4; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(size_t, size_t, lean_object*); @@ -242,7 +242,6 @@ lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__23; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -250,6 +249,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_FromToJson_mk static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__9; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader___closed__2; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; @@ -258,7 +258,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJs static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__3___closed__1; static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__13; lean_object* lean_array_to_list(lean_object*); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -268,12 +267,12 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkT static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__8; lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__6; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__16; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__4; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__9; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__7; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__14; lean_object* l_Lean_Expr_constName_x21(lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__3; @@ -281,14 +280,14 @@ extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField___closed__1; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6; static lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__7; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Deriving_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField___closed__2; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__10; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__8___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__5; @@ -301,9 +300,11 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__8; LEAN_EXPORT lean_object* l_Substring_takeRightWhileAux___at_Lean_Elab_Deriving_FromToJson_mkJsonField___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16; lean_object* lean_string_utf8_prev(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__10; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663_(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__23; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -323,17 +324,14 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJso static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___closed__5; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__9(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9; lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__1___closed__3; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonMutualBlock___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__32; @@ -342,10 +340,12 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__4; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__5; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__13; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonMutualBlock(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; +static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14; lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -379,6 +379,7 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3___closed__5; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__8; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__2; LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__8___lambda__1(lean_object*, lean_object*); @@ -386,10 +387,8 @@ lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__10; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660_(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__9; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___closed__18; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__3; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__23; @@ -433,6 +432,7 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___cl lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9; static lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___closed__2; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -466,7 +466,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField(lean_object lean_object* l_String_intercalate(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___closed__1; @@ -507,6 +510,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mk LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__4; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14; static lean_object* l_Array_qsort_sort___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__8___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__28; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct___closed__7; @@ -522,7 +526,6 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1___closed__7; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonAuxFunction___lambda__1___closed__24; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___lambda__3___closed__14; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__25; @@ -530,7 +533,6 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonAuxFunction___lamb lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForStruct___spec__1___closed__6; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__38; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16; lean_object* l_String_toSubstring_x27(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -540,12 +542,12 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__1 lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1___closed__35; -static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8; static lean_object* l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__6; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct___closed__3; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3___closed__4; +static lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__12; @@ -9132,7 +9134,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruc _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -9152,6 +9154,26 @@ static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruc _start: { lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader___closed__1; +x_2 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__3; +x_3 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__4; +x_4 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } @@ -9197,7 +9219,7 @@ x_24 = lean_st_ref_get(x_7, x_20); x_25 = !lean_is_exclusive(x_24); if (x_25 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_object* x_26; lean_object* x_27; lean_object* x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; x_26 = lean_ctor_get(x_24, 0); lean_dec(x_26); x_27 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; @@ -9246,311 +9268,320 @@ lean_ctor_set(x_45, 0, x_23); lean_ctor_set(x_45, 1, x_30); lean_ctor_set(x_45, 2, x_44); x_46 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12; +lean_inc(x_23); +x_47 = l_Lean_Syntax_node1(x_23, x_46, x_45); +x_48 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14; lean_inc(x_38); lean_inc(x_23); -x_47 = l_Lean_Syntax_node1(x_23, x_46, x_38); -x_48 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13; +x_49 = l_Lean_Syntax_node1(x_23, x_48, x_38); +x_50 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15; lean_inc(x_23); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_23); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_23); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; lean_inc_n(x_38, 2); lean_inc(x_23); -x_51 = l_Lean_Syntax_node6(x_23, x_50, x_37, x_38, x_45, x_47, x_38, x_49); +x_53 = l_Lean_Syntax_node6(x_23, x_52, x_37, x_38, x_47, x_49, x_38, x_51); lean_inc(x_23); -x_52 = l_Lean_Syntax_node1(x_23, x_30, x_51); -x_53 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; +x_54 = l_Lean_Syntax_node1(x_23, x_30, x_53); +x_55 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; lean_inc(x_23); -x_54 = l_Lean_Syntax_node2(x_23, x_53, x_35, x_52); -x_55 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; +x_56 = l_Lean_Syntax_node2(x_23, x_55, x_35, x_54); +x_57 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; lean_inc(x_23); -x_56 = l_Lean_Syntax_node2(x_23, x_55, x_54, x_38); -x_57 = lean_array_push(x_33, x_56); +x_58 = l_Lean_Syntax_node2(x_23, x_57, x_56, x_38); +x_59 = lean_array_push(x_33, x_58); lean_inc(x_23); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_23); -lean_ctor_set(x_58, 1, x_30); -lean_ctor_set(x_58, 2, x_57); -x_59 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_23); +lean_ctor_set(x_60, 1, x_30); +lean_ctor_set(x_60, 2, x_59); +x_61 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; lean_inc(x_23); -x_60 = l_Lean_Syntax_node1(x_23, x_59, x_58); -x_61 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; -x_62 = l_Lean_Syntax_node2(x_23, x_61, x_9, x_60); -lean_ctor_set(x_24, 0, x_62); +x_62 = l_Lean_Syntax_node1(x_23, x_61, x_60); +x_63 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; +x_64 = l_Lean_Syntax_node2(x_23, x_63, x_9, x_62); +lean_ctor_set(x_24, 0, x_64); return x_24; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; size_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_63 = lean_ctor_get(x_24, 1); -lean_inc(x_63); +lean_object* x_65; lean_object* x_66; lean_object* x_67; size_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; size_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_65 = lean_ctor_get(x_24, 1); +lean_inc(x_65); lean_dec(x_24); -x_64 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; +x_66 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; lean_inc(x_23); lean_ctor_set_tag(x_9, 2); -lean_ctor_set(x_9, 1, x_64); +lean_ctor_set(x_9, 1, x_66); lean_ctor_set(x_9, 0, x_23); -x_65 = l_Array_zip___rarg(x_21, x_19); +x_67 = l_Array_zip___rarg(x_21, x_19); lean_dec(x_19); -x_66 = lean_array_size(x_65); -x_67 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; -x_68 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; -lean_inc(x_23); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2(x_23, x_67, x_68, x_66, x_17, x_65); -x_70 = l_Array_append___rarg(x_68, x_69); -lean_dec(x_69); -x_71 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; +x_68 = lean_array_size(x_67); +x_69 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; +x_70 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; lean_inc(x_23); -x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_23); -lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__10; +x_71 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2(x_23, x_69, x_70, x_68, x_17, x_67); +x_72 = l_Array_append___rarg(x_70, x_71); +lean_dec(x_71); +x_73 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; lean_inc(x_23); x_74 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_74, 0, x_23); lean_ctor_set(x_74, 1, x_73); +x_75 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__10; lean_inc(x_23); -x_75 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_75, 0, x_23); -lean_ctor_set(x_75, 1, x_67); -lean_ctor_set(x_75, 2, x_68); -x_76 = l_Array_zip___rarg(x_21, x_21); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_23); +lean_ctor_set(x_76, 1, x_75); +lean_inc(x_23); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_23); +lean_ctor_set(x_77, 1, x_69); +lean_ctor_set(x_77, 2, x_70); +x_78 = l_Array_zip___rarg(x_21, x_21); lean_dec(x_21); -x_77 = lean_array_size(x_76); -lean_inc(x_75); +x_79 = lean_array_size(x_78); +lean_inc(x_77); lean_inc(x_23); -x_78 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3(x_23, x_75, x_77, x_17, x_76); -x_79 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1___closed__8; -x_80 = l_Lean_mkSepArray(x_78, x_79); -lean_dec(x_78); -x_81 = l_Array_append___rarg(x_68, x_80); +x_80 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3(x_23, x_77, x_79, x_17, x_78); +x_81 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1___closed__8; +x_82 = l_Lean_mkSepArray(x_80, x_81); lean_dec(x_80); +x_83 = l_Array_append___rarg(x_70, x_82); +lean_dec(x_82); lean_inc(x_23); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_23); -lean_ctor_set(x_82, 1, x_67); -lean_ctor_set(x_82, 2, x_81); -x_83 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12; -lean_inc(x_75); +x_84 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_84, 0, x_23); +lean_ctor_set(x_84, 1, x_69); +lean_ctor_set(x_84, 2, x_83); +x_85 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12; lean_inc(x_23); -x_84 = l_Lean_Syntax_node1(x_23, x_83, x_75); -x_85 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13; +x_86 = l_Lean_Syntax_node1(x_23, x_85, x_84); +x_87 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14; +lean_inc(x_77); lean_inc(x_23); -x_86 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_86, 0, x_23); -lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; -lean_inc_n(x_75, 2); +x_88 = l_Lean_Syntax_node1(x_23, x_87, x_77); +x_89 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15; lean_inc(x_23); -x_88 = l_Lean_Syntax_node6(x_23, x_87, x_74, x_75, x_82, x_84, x_75, x_86); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_23); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; +lean_inc_n(x_77, 2); lean_inc(x_23); -x_89 = l_Lean_Syntax_node1(x_23, x_67, x_88); -x_90 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; +x_92 = l_Lean_Syntax_node6(x_23, x_91, x_76, x_77, x_86, x_88, x_77, x_90); lean_inc(x_23); -x_91 = l_Lean_Syntax_node2(x_23, x_90, x_72, x_89); -x_92 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; +x_93 = l_Lean_Syntax_node1(x_23, x_69, x_92); +x_94 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; lean_inc(x_23); -x_93 = l_Lean_Syntax_node2(x_23, x_92, x_91, x_75); -x_94 = lean_array_push(x_70, x_93); +x_95 = l_Lean_Syntax_node2(x_23, x_94, x_74, x_93); +x_96 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; lean_inc(x_23); -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_23); -lean_ctor_set(x_95, 1, x_67); -lean_ctor_set(x_95, 2, x_94); -x_96 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; +x_97 = l_Lean_Syntax_node2(x_23, x_96, x_95, x_77); +x_98 = lean_array_push(x_72, x_97); +lean_inc(x_23); +x_99 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_99, 0, x_23); +lean_ctor_set(x_99, 1, x_69); +lean_ctor_set(x_99, 2, x_98); +x_100 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; lean_inc(x_23); -x_97 = l_Lean_Syntax_node1(x_23, x_96, x_95); -x_98 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; -x_99 = l_Lean_Syntax_node2(x_23, x_98, x_9, x_97); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_63); -return x_100; +x_101 = l_Lean_Syntax_node1(x_23, x_100, x_99); +x_102 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; +x_103 = l_Lean_Syntax_node2(x_23, x_102, x_9, x_101); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_65); +return x_104; } } else { -uint8_t x_101; +uint8_t x_105; lean_dec(x_15); lean_free_object(x_9); lean_dec(x_6); -x_101 = !lean_is_exclusive(x_18); -if (x_101 == 0) +x_105 = !lean_is_exclusive(x_18); +if (x_105 == 0) { return x_18; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_18, 0); -x_103 = lean_ctor_get(x_18, 1); -lean_inc(x_103); -lean_inc(x_102); +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_18, 0); +x_107 = lean_ctor_get(x_18, 1); +lean_inc(x_107); +lean_inc(x_106); lean_dec(x_18); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; size_t x_110; size_t x_111; lean_object* x_112; -x_105 = lean_ctor_get(x_9, 0); -x_106 = lean_ctor_get(x_9, 1); -lean_inc(x_106); -lean_inc(x_105); +lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; size_t x_114; size_t x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_9, 0); +x_110 = lean_ctor_get(x_9, 1); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_9); -x_107 = lean_ctor_get(x_105, 0); -lean_inc(x_107); -lean_dec(x_105); -x_108 = 0; +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +lean_dec(x_109); +x_112 = 0; lean_inc(x_1); -x_109 = l_Lean_getStructureFieldsFlattened(x_107, x_1, x_108); -lean_dec(x_107); -x_110 = lean_array_size(x_109); -x_111 = 0; +x_113 = l_Lean_getStructureFieldsFlattened(x_111, x_1, x_112); +lean_dec(x_111); +x_114 = lean_array_size(x_113); +x_115 = 0; lean_inc(x_6); -lean_inc(x_109); -x_112 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1(x_1, x_110, x_111, x_109, x_2, x_3, x_4, x_5, x_6, x_7, x_106); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; size_t x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; size_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -lean_dec(x_112); -x_115 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(x_110, x_111, x_109); -x_116 = lean_ctor_get(x_6, 5); -lean_inc(x_116); -lean_dec(x_6); -x_117 = l_Lean_SourceInfo_fromRef(x_116, x_108); +x_116 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__1(x_1, x_114, x_115, x_113, x_2, x_3, x_4, x_5, x_6, x_7, x_110); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; size_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; size_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); lean_dec(x_116); -x_118 = lean_st_ref_get(x_7, x_114); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_120 = x_118; +x_119 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkInductiveApp___spec__1(x_114, x_115, x_113); +x_120 = lean_ctor_get(x_6, 5); +lean_inc(x_120); +lean_dec(x_6); +x_121 = l_Lean_SourceInfo_fromRef(x_120, x_112); +lean_dec(x_120); +x_122 = lean_st_ref_get(x_7, x_118); +x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_122)) { + lean_ctor_release(x_122, 0); + lean_ctor_release(x_122, 1); + x_124 = x_122; } else { - lean_dec_ref(x_118); - x_120 = lean_box(0); + lean_dec_ref(x_122); + x_124 = lean_box(0); } -x_121 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; -lean_inc(x_117); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_117); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Array_zip___rarg(x_115, x_113); -lean_dec(x_113); -x_124 = lean_array_size(x_123); -x_125 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; -x_126 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; -lean_inc(x_117); -x_127 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2(x_117, x_125, x_126, x_124, x_111, x_123); -x_128 = l_Array_append___rarg(x_126, x_127); -lean_dec(x_127); -x_129 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; -lean_inc(x_117); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_117); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__10; -lean_inc(x_117); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_117); -lean_ctor_set(x_132, 1, x_131); -lean_inc(x_117); -x_133 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_133, 0, x_117); -lean_ctor_set(x_133, 1, x_125); -lean_ctor_set(x_133, 2, x_126); -x_134 = l_Array_zip___rarg(x_115, x_115); -lean_dec(x_115); -x_135 = lean_array_size(x_134); -lean_inc(x_133); -lean_inc(x_117); -x_136 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3(x_117, x_133, x_135, x_111, x_134); -x_137 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1___closed__8; -x_138 = l_Lean_mkSepArray(x_136, x_137); -lean_dec(x_136); -x_139 = l_Array_append___rarg(x_126, x_138); -lean_dec(x_138); -lean_inc(x_117); -x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_117); -lean_ctor_set(x_140, 1, x_125); -lean_ctor_set(x_140, 2, x_139); -x_141 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12; -lean_inc(x_133); -lean_inc(x_117); -x_142 = l_Lean_Syntax_node1(x_117, x_141, x_133); -x_143 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13; -lean_inc(x_117); -x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_117); -lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; -lean_inc_n(x_133, 2); -lean_inc(x_117); -x_146 = l_Lean_Syntax_node6(x_117, x_145, x_132, x_133, x_140, x_142, x_133, x_144); -lean_inc(x_117); -x_147 = l_Lean_Syntax_node1(x_117, x_125, x_146); -x_148 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; -lean_inc(x_117); -x_149 = l_Lean_Syntax_node2(x_117, x_148, x_130, x_147); -x_150 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; -lean_inc(x_117); -x_151 = l_Lean_Syntax_node2(x_117, x_150, x_149, x_133); -x_152 = lean_array_push(x_128, x_151); -lean_inc(x_117); -x_153 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_153, 0, x_117); -lean_ctor_set(x_153, 1, x_125); -lean_ctor_set(x_153, 2, x_152); -x_154 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; -lean_inc(x_117); -x_155 = l_Lean_Syntax_node1(x_117, x_154, x_153); -x_156 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; -x_157 = l_Lean_Syntax_node2(x_117, x_156, x_122, x_155); -if (lean_is_scalar(x_120)) { - x_158 = lean_alloc_ctor(0, 2, 0); +x_125 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__1; +lean_inc(x_121); +x_126 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_126, 0, x_121); +lean_ctor_set(x_126, 1, x_125); +x_127 = l_Array_zip___rarg(x_119, x_117); +lean_dec(x_117); +x_128 = lean_array_size(x_127); +x_129 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__9; +x_130 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__25; +lean_inc(x_121); +x_131 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2(x_121, x_129, x_130, x_128, x_115, x_127); +x_132 = l_Array_append___rarg(x_130, x_131); +lean_dec(x_131); +x_133 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__7; +lean_inc(x_121); +x_134 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_134, 0, x_121); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__10; +lean_inc(x_121); +x_136 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_136, 0, x_121); +lean_ctor_set(x_136, 1, x_135); +lean_inc(x_121); +x_137 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_137, 0, x_121); +lean_ctor_set(x_137, 1, x_129); +lean_ctor_set(x_137, 2, x_130); +x_138 = l_Array_zip___rarg(x_119, x_119); +lean_dec(x_119); +x_139 = lean_array_size(x_138); +lean_inc(x_137); +lean_inc(x_121); +x_140 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__3(x_121, x_137, x_139, x_115, x_138); +x_141 = l_List_forIn_x27_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonBodyForInduct_mkAlts___spec__6___lambda__1___closed__8; +x_142 = l_Lean_mkSepArray(x_140, x_141); +lean_dec(x_140); +x_143 = l_Array_append___rarg(x_130, x_142); +lean_dec(x_142); +lean_inc(x_121); +x_144 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_144, 0, x_121); +lean_ctor_set(x_144, 1, x_129); +lean_ctor_set(x_144, 2, x_143); +x_145 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12; +lean_inc(x_121); +x_146 = l_Lean_Syntax_node1(x_121, x_145, x_144); +x_147 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14; +lean_inc(x_137); +lean_inc(x_121); +x_148 = l_Lean_Syntax_node1(x_121, x_147, x_137); +x_149 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15; +lean_inc(x_121); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_121); +lean_ctor_set(x_150, 1, x_149); +x_151 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__9; +lean_inc_n(x_137, 2); +lean_inc(x_121); +x_152 = l_Lean_Syntax_node6(x_121, x_151, x_136, x_137, x_146, x_148, x_137, x_150); +lean_inc(x_121); +x_153 = l_Lean_Syntax_node1(x_121, x_129, x_152); +x_154 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__6; +lean_inc(x_121); +x_155 = l_Lean_Syntax_node2(x_121, x_154, x_134, x_153); +x_156 = l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___spec__2___closed__2; +lean_inc(x_121); +x_157 = l_Lean_Syntax_node2(x_121, x_156, x_155, x_137); +x_158 = lean_array_push(x_132, x_157); +lean_inc(x_121); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_121); +lean_ctor_set(x_159, 1, x_129); +lean_ctor_set(x_159, 2, x_158); +x_160 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__4; +lean_inc(x_121); +x_161 = l_Lean_Syntax_node1(x_121, x_160, x_159); +x_162 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__2; +x_163 = l_Lean_Syntax_node2(x_121, x_162, x_126, x_161); +if (lean_is_scalar(x_124)) { + x_164 = lean_alloc_ctor(0, 2, 0); } else { - x_158 = x_120; + x_164 = x_124; } -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_119); -return x_158; +lean_ctor_set(x_164, 0, x_163); +lean_ctor_set(x_164, 1, x_123); +return x_164; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -lean_dec(x_109); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +lean_dec(x_113); lean_dec(x_6); -x_159 = lean_ctor_get(x_112, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_112, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_161 = x_112; +x_165 = lean_ctor_get(x_116, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_116, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_167 = x_116; } else { - lean_dec_ref(x_112); - x_161 = lean_box(0); + lean_dec_ref(x_116); + x_167 = lean_box(0); } -if (lean_is_scalar(x_161)) { - x_162 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); } else { - x_162 = x_161; + x_168 = x_167; } -lean_ctor_set(x_162, 0, x_159); -lean_ctor_set(x_162, 1, x_160); -return x_162; +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; } } } @@ -18215,7 +18246,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1() { _start: { lean_object* x_1; @@ -18223,7 +18254,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanc return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2() { _start: { lean_object* x_1; @@ -18231,7 +18262,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInsta return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -18241,27 +18272,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3; x_2 = l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4; x_2 = l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6() { _start: { lean_object* x_1; @@ -18269,17 +18300,17 @@ x_1 = lean_mk_string_unchecked("FromToJson", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5; -x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5; +x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8() { _start: { lean_object* x_1; @@ -18287,17 +18318,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7; -x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7; +x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10() { _start: { lean_object* x_1; @@ -18305,57 +18336,57 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9; -x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9; +x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11; x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12; x_2 = l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13; x_2 = l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14; -x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14; +x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16() { _start: { lean_object* x_1; @@ -18363,32 +18394,32 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15; -x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16; +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15; +x_2 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18() { +static lean_object* _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17; -x_2 = lean_unsigned_to_nat(6660u); +x_1 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17; +x_2 = lean_unsigned_to_nat(6663u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_FromToJson_mkToJsonHeader___closed__3; -x_3 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1; +x_3 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -18397,7 +18428,7 @@ x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); x_6 = l_Lean_Elab_Deriving_FromToJson_mkFromJsonHeader___closed__2; -x_7 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2; +x_7 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2; x_8 = l_Lean_Elab_registerDerivingHandler(x_6, x_7, x_5); if (lean_obj_tag(x_8) == 0) { @@ -18407,7 +18438,7 @@ lean_inc(x_9); lean_dec(x_8); x_10 = l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkToJsonInstance___closed__3; x_11 = 0; -x_12 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18; +x_12 = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18; x_13 = l_Lean_registerTraceClass(x_10, x_11, x_12, x_9); if (lean_obj_tag(x_13) == 0) { @@ -18874,6 +18905,10 @@ l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12 = _init_l_L lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__12); l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13 = _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13(); lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__13); +l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14 = _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__14); +l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15 = _init_l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForStruct___closed__15); l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__1 = _init_l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__1(); lean_mark_persistent(l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__1); l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__2 = _init_l_Array_mapFinIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonBodyForInduct_mkAlts___spec__3___closed__2(); @@ -19092,43 +19127,43 @@ l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFro lean_mark_persistent(l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___closed__1); l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___closed__2 = _init_l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Deriving_FromToJson_0__Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___closed__2); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__1); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__2); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__3); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__4); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__5); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__6); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__7); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__8); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__9); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__10); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__11); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__12); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__13); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__14); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__15); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__16); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__17); -l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18(); -lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660____closed__18); -if (builtin) {res = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6660_(lean_io_mk_world()); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__1); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__2); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__3); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__4); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__5); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__6); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__7); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__8); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__9); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__10); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__11); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__12); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__13); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__14); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__15); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__16); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__17); +l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18 = _init_l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18(); +lean_mark_persistent(l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663____closed__18); +if (builtin) {res = l_Lean_Elab_Deriving_FromToJson_initFn____x40_Lean_Elab_Deriving_FromToJson___hyg_6663_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Extra.c b/stage0/stdlib/Lean/Elab/Extra.c index 73bf0cc57395..54d7b50e4de2 100644 --- a/stage0/stdlib/Lean/Elab/Extra.c +++ b/stage0/stdlib/Lean/Elab/Extra.c @@ -22,7 +22,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__1___boxed( static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkUnOp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083_(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange__1___closed__6; @@ -100,6 +99,7 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__7; uint8_t l_Lean_Elab_Term_hasCDot(lean_object*); @@ -124,7 +124,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpK uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinOpLazy(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHomogeneousInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__8; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6; @@ -143,6 +142,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact__1___closed__1; static lean_object* l_Lean_Elab_Term_elabForIn___closed__25; lean_object* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy__1___closed__3; @@ -151,10 +151,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__2(lean_obj static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange__1___closed__6; static lean_object* l_Lean_Elab_Term_elabForIn___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange__1___closed__3; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkBinOp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withPushMacroExpansionStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__13; @@ -162,6 +162,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinOp(lean_object*, lean_object static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1___closed__6; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange__1___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -179,7 +180,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp__1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_throwForInFailure___closed__1; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__7; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy__1___closed__2; @@ -188,6 +188,7 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange__1___closed__7; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__5; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange__1___closed__2; lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact__1___closed__3; @@ -203,7 +204,6 @@ static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange__1___closed__4; @@ -211,7 +211,6 @@ lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion(lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__1___closed__7; @@ -221,6 +220,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree static lean_object* l_Lean_Elab_Term_elabForIn___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty__1___closed__4; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,11 +231,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstances___at___private_Lean_Ela static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__8; static lean_object* l_Lean_Elab_Term_elabForIn___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange__1___closed__3; -lean_object* l_Lean_Elab_Term_mkTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1; lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp__1(lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -250,6 +248,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary(ui lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange__1___closed__7; extern lean_object* l_Lean_levelZero; @@ -263,8 +262,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange__1___c static lean_object* l_Lean_Elab_Term_elabForIn___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -278,6 +275,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange__1___ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn__1(lean_object*); lean_object* l_Lean_Meta_mkFunUnit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabUnOp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange__1___closed__4; @@ -292,7 +290,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__2___boxed( static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Op_elabBinRelCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__23; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -308,6 +305,8 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange__1(lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__8; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4; +lean_object* l_Lean_Elab_Term_withTermInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_isUnknown___boxed(lean_object*); static lean_object* l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__1___closed__1; @@ -339,13 +338,11 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declR LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1___closed__5; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___closed__3; static lean_object* l_Lean_Elab_Term_elabForIn___closed__10; static lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -360,7 +357,9 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processBinOp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__17; static lean_object* l_Lean_Elab_Term_elabForIn___closed__13; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange__1___closed__4; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11; static lean_object* l_Lean_Elab_Term_Op_instBEqBinOpKind___closed__1; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__3; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -385,20 +384,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___boxed(lean_object* lean_object* l_Lean_Meta_coerceSimple_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkBinOp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty__1___closed__3; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2; lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___closed__4; static lean_object* l_Lean_Elab_Term_elabForIn___closed__20; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processLeaf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -420,7 +418,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declR lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange__1___closed__3; static lean_object* l_Lean_Elab_Term_elabForIn___closed__5; -lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange__1___closed__1; lean_object* l_Lean_Elab_Term_mkConst(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__15; @@ -449,6 +446,7 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp__1(lean_object*); static lean_object* l_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067_(lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__18; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27__1___closed__2; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -472,7 +470,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__1___cl LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processBinOp___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27__1___closed__1; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange__1___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -504,14 +501,13 @@ lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabRightact(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__9; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__1___closed__3; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8; lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange__1___closed__7; lean_object* l_Lean_Elab_Term_tryPostpone(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27__1___closed__3; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact__1___closed__1; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__2___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp__1___closed__2; @@ -519,7 +515,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp__1___closed__2; static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__1() { @@ -11588,17 +11583,6 @@ return x_28; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; -x_10 = lean_box(0); -x_11 = lean_box(0); -x_12 = 0; -x_13 = l_Lean_Elab_Term_mkTermInfo(x_11, x_1, x_2, x_10, x_10, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ lean_object* x_10; lean_inc(x_8); lean_inc(x_7); @@ -11649,16 +11633,6 @@ return x_17; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; lean_object* x_13; -x_11 = lean_box(0); -x_12 = 0; -x_13 = l_Lean_Elab_Term_mkTermInfo(x_1, x_2, x_3, x_11, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { @@ -11847,38 +11821,40 @@ lean_closure_set(x_57, 0, x_54); lean_closure_set(x_57, 1, x_55); lean_closure_set(x_57, 2, x_56); lean_closure_set(x_57, 3, x_53); -lean_inc(x_51); -x_58 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2___boxed), 9, 1); -lean_closure_set(x_58, 0, x_51); +x_58 = lean_box(0); x_59 = !lean_is_exclusive(x_6); if (x_59 == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; x_60 = lean_ctor_get(x_6, 5); x_61 = l_Lean_replaceRef(x_51, x_60); lean_dec(x_60); lean_ctor_set(x_6, 5, x_61); -x_62 = l_Lean_Elab_Term_withInfoContext_x27(x_51, x_57, x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_62; +x_62 = lean_box(0); +x_63 = 0; +x_64 = l_Lean_Elab_Term_withTermInfoContext_x27(x_62, x_51, x_57, x_58, x_58, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_64; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_63 = lean_ctor_get(x_6, 0); -x_64 = lean_ctor_get(x_6, 1); -x_65 = lean_ctor_get(x_6, 2); -x_66 = lean_ctor_get(x_6, 3); -x_67 = lean_ctor_get(x_6, 4); -x_68 = lean_ctor_get(x_6, 5); -x_69 = lean_ctor_get(x_6, 6); -x_70 = lean_ctor_get(x_6, 7); -x_71 = lean_ctor_get(x_6, 8); -x_72 = lean_ctor_get(x_6, 9); -x_73 = lean_ctor_get(x_6, 10); -x_74 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); -x_75 = lean_ctor_get(x_6, 11); -x_76 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_object* x_83; +x_65 = lean_ctor_get(x_6, 0); +x_66 = lean_ctor_get(x_6, 1); +x_67 = lean_ctor_get(x_6, 2); +x_68 = lean_ctor_get(x_6, 3); +x_69 = lean_ctor_get(x_6, 4); +x_70 = lean_ctor_get(x_6, 5); +x_71 = lean_ctor_get(x_6, 6); +x_72 = lean_ctor_get(x_6, 7); +x_73 = lean_ctor_get(x_6, 8); +x_74 = lean_ctor_get(x_6, 9); +x_75 = lean_ctor_get(x_6, 10); +x_76 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); +x_77 = lean_ctor_get(x_6, 11); +x_78 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_inc(x_77); lean_inc(x_75); +lean_inc(x_74); lean_inc(x_73); lean_inc(x_72); lean_inc(x_71); @@ -11888,191 +11864,192 @@ lean_inc(x_68); lean_inc(x_67); lean_inc(x_66); lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); lean_dec(x_6); -x_77 = l_Lean_replaceRef(x_51, x_68); -lean_dec(x_68); -x_78 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_78, 0, x_63); -lean_ctor_set(x_78, 1, x_64); -lean_ctor_set(x_78, 2, x_65); -lean_ctor_set(x_78, 3, x_66); -lean_ctor_set(x_78, 4, x_67); -lean_ctor_set(x_78, 5, x_77); -lean_ctor_set(x_78, 6, x_69); -lean_ctor_set(x_78, 7, x_70); -lean_ctor_set(x_78, 8, x_71); -lean_ctor_set(x_78, 9, x_72); -lean_ctor_set(x_78, 10, x_73); -lean_ctor_set(x_78, 11, x_75); -lean_ctor_set_uint8(x_78, sizeof(void*)*12, x_74); -lean_ctor_set_uint8(x_78, sizeof(void*)*12 + 1, x_76); -x_79 = l_Lean_Elab_Term_withInfoContext_x27(x_51, x_57, x_58, x_2, x_3, x_4, x_5, x_78, x_7, x_8); -return x_79; +x_79 = l_Lean_replaceRef(x_51, x_70); +lean_dec(x_70); +x_80 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_80, 0, x_65); +lean_ctor_set(x_80, 1, x_66); +lean_ctor_set(x_80, 2, x_67); +lean_ctor_set(x_80, 3, x_68); +lean_ctor_set(x_80, 4, x_69); +lean_ctor_set(x_80, 5, x_79); +lean_ctor_set(x_80, 6, x_71); +lean_ctor_set(x_80, 7, x_72); +lean_ctor_set(x_80, 8, x_73); +lean_ctor_set(x_80, 9, x_74); +lean_ctor_set(x_80, 10, x_75); +lean_ctor_set(x_80, 11, x_77); +lean_ctor_set_uint8(x_80, sizeof(void*)*12, x_76); +lean_ctor_set_uint8(x_80, sizeof(void*)*12 + 1, x_78); +x_81 = lean_box(0); +x_82 = 0; +x_83 = l_Lean_Elab_Term_withTermInfoContext_x27(x_81, x_51, x_57, x_58, x_58, x_82, x_2, x_3, x_4, x_5, x_80, x_7, x_8); +return x_83; } } case 2: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_80 = lean_ctor_get(x_1, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_1, 1); -lean_inc(x_81); -x_82 = lean_ctor_get(x_1, 2); -lean_inc(x_82); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_84 = lean_ctor_get(x_1, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_1, 1); +lean_inc(x_85); +x_86 = lean_ctor_get(x_1, 2); +lean_inc(x_86); lean_dec(x_1); -x_83 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__3), 9, 2); -lean_closure_set(x_83, 0, x_82); -lean_closure_set(x_83, 1, x_81); -lean_inc(x_80); -x_84 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2___boxed), 9, 1); -lean_closure_set(x_84, 0, x_80); -x_85 = !lean_is_exclusive(x_6); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_6, 5); -x_87 = l_Lean_replaceRef(x_80, x_86); -lean_dec(x_86); -lean_ctor_set(x_6, 5, x_87); -x_88 = l_Lean_Elab_Term_withInfoContext_x27(x_80, x_83, x_84, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_88; +x_87 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2), 9, 2); +lean_closure_set(x_87, 0, x_86); +lean_closure_set(x_87, 1, x_85); +x_88 = lean_box(0); +x_89 = !lean_is_exclusive(x_6); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; +x_90 = lean_ctor_get(x_6, 5); +x_91 = l_Lean_replaceRef(x_84, x_90); +lean_dec(x_90); +lean_ctor_set(x_6, 5, x_91); +x_92 = lean_box(0); +x_93 = 0; +x_94 = l_Lean_Elab_Term_withTermInfoContext_x27(x_92, x_84, x_87, x_88, x_88, x_93, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_94; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_89 = lean_ctor_get(x_6, 0); -x_90 = lean_ctor_get(x_6, 1); -x_91 = lean_ctor_get(x_6, 2); -x_92 = lean_ctor_get(x_6, 3); -x_93 = lean_ctor_get(x_6, 4); -x_94 = lean_ctor_get(x_6, 5); -x_95 = lean_ctor_get(x_6, 6); -x_96 = lean_ctor_get(x_6, 7); -x_97 = lean_ctor_get(x_6, 8); -x_98 = lean_ctor_get(x_6, 9); -x_99 = lean_ctor_get(x_6, 10); -x_100 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); -x_101 = lean_ctor_get(x_6, 11); -x_102 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; +x_95 = lean_ctor_get(x_6, 0); +x_96 = lean_ctor_get(x_6, 1); +x_97 = lean_ctor_get(x_6, 2); +x_98 = lean_ctor_get(x_6, 3); +x_99 = lean_ctor_get(x_6, 4); +x_100 = lean_ctor_get(x_6, 5); +x_101 = lean_ctor_get(x_6, 6); +x_102 = lean_ctor_get(x_6, 7); +x_103 = lean_ctor_get(x_6, 8); +x_104 = lean_ctor_get(x_6, 9); +x_105 = lean_ctor_get(x_6, 10); +x_106 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); +x_107 = lean_ctor_get(x_6, 11); +x_108 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_inc(x_107); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); lean_inc(x_101); +lean_inc(x_100); lean_inc(x_99); lean_inc(x_98); lean_inc(x_97); lean_inc(x_96); lean_inc(x_95); -lean_inc(x_94); -lean_inc(x_93); -lean_inc(x_92); -lean_inc(x_91); -lean_inc(x_90); -lean_inc(x_89); lean_dec(x_6); -x_103 = l_Lean_replaceRef(x_80, x_94); -lean_dec(x_94); -x_104 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_104, 0, x_89); -lean_ctor_set(x_104, 1, x_90); -lean_ctor_set(x_104, 2, x_91); -lean_ctor_set(x_104, 3, x_92); -lean_ctor_set(x_104, 4, x_93); -lean_ctor_set(x_104, 5, x_103); -lean_ctor_set(x_104, 6, x_95); -lean_ctor_set(x_104, 7, x_96); -lean_ctor_set(x_104, 8, x_97); -lean_ctor_set(x_104, 9, x_98); -lean_ctor_set(x_104, 10, x_99); -lean_ctor_set(x_104, 11, x_101); -lean_ctor_set_uint8(x_104, sizeof(void*)*12, x_100); -lean_ctor_set_uint8(x_104, sizeof(void*)*12 + 1, x_102); -x_105 = l_Lean_Elab_Term_withInfoContext_x27(x_80, x_83, x_84, x_2, x_3, x_4, x_5, x_104, x_7, x_8); -return x_105; +x_109 = l_Lean_replaceRef(x_84, x_100); +lean_dec(x_100); +x_110 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_110, 0, x_95); +lean_ctor_set(x_110, 1, x_96); +lean_ctor_set(x_110, 2, x_97); +lean_ctor_set(x_110, 3, x_98); +lean_ctor_set(x_110, 4, x_99); +lean_ctor_set(x_110, 5, x_109); +lean_ctor_set(x_110, 6, x_101); +lean_ctor_set(x_110, 7, x_102); +lean_ctor_set(x_110, 8, x_103); +lean_ctor_set(x_110, 9, x_104); +lean_ctor_set(x_110, 10, x_105); +lean_ctor_set(x_110, 11, x_107); +lean_ctor_set_uint8(x_110, sizeof(void*)*12, x_106); +lean_ctor_set_uint8(x_110, sizeof(void*)*12 + 1, x_108); +x_111 = lean_box(0); +x_112 = 0; +x_113 = l_Lean_Elab_Term_withTermInfoContext_x27(x_111, x_84, x_87, x_88, x_88, x_112, x_2, x_3, x_4, x_5, x_110, x_7, x_8); +return x_113; } } default: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; -x_106 = lean_ctor_get(x_1, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_1, 1); -lean_inc(x_107); -x_108 = lean_ctor_get(x_1, 2); -lean_inc(x_108); -x_109 = lean_ctor_get(x_1, 3); -lean_inc(x_109); +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_114 = lean_ctor_get(x_1, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_1, 1); +lean_inc(x_115); +x_116 = lean_ctor_get(x_1, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_1, 3); +lean_inc(x_117); lean_dec(x_1); -x_110 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore), 8, 1); -lean_closure_set(x_110, 0, x_109); -lean_inc(x_107); -x_111 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withMacroExpansion___rarg), 10, 3); -lean_closure_set(x_111, 0, x_107); -lean_closure_set(x_111, 1, x_108); -lean_closure_set(x_111, 2, x_110); -lean_inc(x_107); -x_112 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4___boxed), 10, 2); -lean_closure_set(x_112, 0, x_106); -lean_closure_set(x_112, 1, x_107); -x_113 = !lean_is_exclusive(x_6); -if (x_113 == 0) -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_6, 5); -x_115 = l_Lean_replaceRef(x_107, x_114); -lean_dec(x_114); -lean_ctor_set(x_6, 5, x_115); -x_116 = l_Lean_Elab_Term_withInfoContext_x27(x_107, x_111, x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_116; -} -else +x_118 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore), 8, 1); +lean_closure_set(x_118, 0, x_117); +lean_inc(x_115); +x_119 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withMacroExpansion___rarg), 10, 3); +lean_closure_set(x_119, 0, x_115); +lean_closure_set(x_119, 1, x_116); +lean_closure_set(x_119, 2, x_118); +x_120 = lean_box(0); +x_121 = !lean_is_exclusive(x_6); +if (x_121 == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_117 = lean_ctor_get(x_6, 0); -x_118 = lean_ctor_get(x_6, 1); -x_119 = lean_ctor_get(x_6, 2); -x_120 = lean_ctor_get(x_6, 3); -x_121 = lean_ctor_get(x_6, 4); +lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; x_122 = lean_ctor_get(x_6, 5); -x_123 = lean_ctor_get(x_6, 6); -x_124 = lean_ctor_get(x_6, 7); -x_125 = lean_ctor_get(x_6, 8); -x_126 = lean_ctor_get(x_6, 9); -x_127 = lean_ctor_get(x_6, 10); -x_128 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); -x_129 = lean_ctor_get(x_6, 11); -x_130 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +x_123 = l_Lean_replaceRef(x_115, x_122); +lean_dec(x_122); +lean_ctor_set(x_6, 5, x_123); +x_124 = 0; +x_125 = l_Lean_Elab_Term_withTermInfoContext_x27(x_114, x_115, x_119, x_120, x_120, x_124, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; +x_126 = lean_ctor_get(x_6, 0); +x_127 = lean_ctor_get(x_6, 1); +x_128 = lean_ctor_get(x_6, 2); +x_129 = lean_ctor_get(x_6, 3); +x_130 = lean_ctor_get(x_6, 4); +x_131 = lean_ctor_get(x_6, 5); +x_132 = lean_ctor_get(x_6, 6); +x_133 = lean_ctor_get(x_6, 7); +x_134 = lean_ctor_get(x_6, 8); +x_135 = lean_ctor_get(x_6, 9); +x_136 = lean_ctor_get(x_6, 10); +x_137 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); +x_138 = lean_ctor_get(x_6, 11); +x_139 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_inc(x_138); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +lean_inc(x_133); +lean_inc(x_132); +lean_inc(x_131); +lean_inc(x_130); lean_inc(x_129); +lean_inc(x_128); lean_inc(x_127); lean_inc(x_126); -lean_inc(x_125); -lean_inc(x_124); -lean_inc(x_123); -lean_inc(x_122); -lean_inc(x_121); -lean_inc(x_120); -lean_inc(x_119); -lean_inc(x_118); -lean_inc(x_117); lean_dec(x_6); -x_131 = l_Lean_replaceRef(x_107, x_122); -lean_dec(x_122); -x_132 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_132, 0, x_117); -lean_ctor_set(x_132, 1, x_118); -lean_ctor_set(x_132, 2, x_119); -lean_ctor_set(x_132, 3, x_120); -lean_ctor_set(x_132, 4, x_121); -lean_ctor_set(x_132, 5, x_131); -lean_ctor_set(x_132, 6, x_123); -lean_ctor_set(x_132, 7, x_124); -lean_ctor_set(x_132, 8, x_125); -lean_ctor_set(x_132, 9, x_126); -lean_ctor_set(x_132, 10, x_127); -lean_ctor_set(x_132, 11, x_129); -lean_ctor_set_uint8(x_132, sizeof(void*)*12, x_128); -lean_ctor_set_uint8(x_132, sizeof(void*)*12 + 1, x_130); -x_133 = l_Lean_Elab_Term_withInfoContext_x27(x_107, x_111, x_112, x_2, x_3, x_4, x_5, x_132, x_7, x_8); -return x_133; +x_140 = l_Lean_replaceRef(x_115, x_131); +lean_dec(x_131); +x_141 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_141, 0, x_126); +lean_ctor_set(x_141, 1, x_127); +lean_ctor_set(x_141, 2, x_128); +lean_ctor_set(x_141, 3, x_129); +lean_ctor_set(x_141, 4, x_130); +lean_ctor_set(x_141, 5, x_140); +lean_ctor_set(x_141, 6, x_132); +lean_ctor_set(x_141, 7, x_133); +lean_ctor_set(x_141, 8, x_134); +lean_ctor_set(x_141, 9, x_135); +lean_ctor_set(x_141, 10, x_136); +lean_ctor_set(x_141, 11, x_138); +lean_ctor_set_uint8(x_141, sizeof(void*)*12, x_137); +lean_ctor_set_uint8(x_141, sizeof(void*)*12 + 1, x_139); +x_142 = 0; +x_143 = l_Lean_Elab_Term_withTermInfoContext_x27(x_114, x_115, x_119, x_120, x_120, x_142, x_2, x_3, x_4, x_5, x_141, x_7, x_8); +return x_143; } } } @@ -12088,34 +12065,6 @@ x_13 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__1(x return x_13; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; -} -} static lean_object* _init_l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__1___closed__1() { _start: { @@ -20124,7 +20073,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20134,37 +20083,37 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1; x_2 = l___regBuiltin_Lean_Elab_Term_elabForIn__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2; x_2 = l_Lean_Elab_Term_elabForIn___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3; x_2 = l___regBuiltin_Lean_Elab_Term_Op_elabBinOp__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5() { _start: { lean_object* x_1; @@ -20172,17 +20121,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7() { _start: { lean_object* x_1; @@ -20190,37 +20139,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8; x_2 = l_Lean_Elab_Term_elabForIn___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9; x_2 = l___regBuiltin_Lean_Elab_Term_elabForIn__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11() { _start: { lean_object* x_1; @@ -20228,17 +20177,17 @@ x_1 = lean_mk_string_unchecked("Extra", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13() { _start: { lean_object* x_1; @@ -20246,33 +20195,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14; -x_2 = lean_unsigned_to_nat(7083u); +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14; +x_2 = lean_unsigned_to_nat(7067u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15; +x_4 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -20854,37 +20803,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_decl if (builtin) {res = l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__1); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__2); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__3); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__4); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__5); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__6); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__7); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__8); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__9); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__10); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__11); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__12); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__13); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__14); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083____closed__15); -if (builtin) {res = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7083_(lean_io_mk_world()); +}l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__1); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__2); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__3); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__4); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__5); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__6); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__7); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__8); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__9); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__10); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__11); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__12); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__13); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__14); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067____closed__15); +if (builtin) {res = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_7067_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/GuardMsgs.c b/stage0/stdlib/Lean/Elab/GuardMsgs.c index 55b8fac5506c..8f564b7c5df6 100644 --- a/stage0/stdlib/Lean/Elab/GuardMsgs.c +++ b/stage0/stdlib/Lean/Elab/GuardMsgs.c @@ -81,7 +81,6 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_El LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__5; static lean_object* l___private_Lean_Elab_GuardMsgs_0__Lean_Elab_Tactic_GuardMsgs_messageToStringWithoutPos___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace(lean_object*); @@ -151,7 +150,6 @@ lean_object* lean_string_utf8_next(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__36___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_MessageOrdering_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210_(lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs__1___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__24___boxed(lean_object*, lean_object*, lean_object*); @@ -242,6 +240,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_SpecResult_noConfusion___r LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_removeTrailingWhitespaceMarker(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs__1___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___lambda__7(lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__36(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___lambda__8___closed__9; @@ -280,7 +279,6 @@ extern lean_object* l_Lean_NameSet_empty; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs_declRange__1___closed__2; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___lambda__8___closed__10; @@ -295,7 +293,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Guar LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_initFn____x40_Lean_Elab_GuardMsgs___hyg_6____closed__2; static lean_object* l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__12; @@ -308,6 +305,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___lambda__1( static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___lambda__8___closed__11; uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__23(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Diff_Histogram_addLeft___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___lambda__5(lean_object*, lean_object*); @@ -372,6 +370,7 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Ela LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_SpecResult_toCtorIdx(uint8_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_initFn____x40_Lean_Elab_GuardMsgs___hyg_6____closed__1; +static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_instImpl____x40_Lean_Elab_GuardMsgs___hyg_1821____closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__2___closed__2; @@ -418,6 +417,7 @@ lean_object* l_Subarray_split___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_GuardMsgs_revealTrailingWhitespace___closed__1; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_GuardMsgs_elabGuardMsgs___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_GuardMsgs_WhitespaceMode_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_GuardMsgs_parseGuardMsgsSpec___spec__1(lean_object*, lean_object*); @@ -9324,7 +9324,7 @@ x_103 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; lean_ctor_set(x_28, 1, x_67); lean_ctor_set(x_28, 0, x_103); lean_ctor_set(x_100, 1, x_28); -x_104 = lean_alloc_ctor(8, 1, 0); +x_104 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_104, 0, x_100); x_105 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_104, x_4, x_5, x_102); lean_dec(x_5); @@ -9345,7 +9345,7 @@ lean_ctor_set(x_28, 0, x_108); x_109 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_109, 0, x_106); lean_ctor_set(x_109, 1, x_28); -x_110 = lean_alloc_ctor(8, 1, 0); +x_110 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_110, 0, x_109); x_111 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_110, x_4, x_5, x_107); lean_dec(x_5); @@ -9392,7 +9392,7 @@ x_128 = l_Lean_Elab_Tactic_GuardMsgs_instTypeNameGuardMsgFailure; lean_ctor_set(x_28, 1, x_67); lean_ctor_set(x_28, 0, x_128); lean_ctor_set(x_125, 1, x_28); -x_129 = lean_alloc_ctor(8, 1, 0); +x_129 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_129, 0, x_125); x_130 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_129, x_4, x_5, x_127); lean_dec(x_5); @@ -9413,7 +9413,7 @@ lean_ctor_set(x_28, 0, x_133); x_134 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_134, 0, x_131); lean_ctor_set(x_134, 1, x_28); -x_135 = lean_alloc_ctor(8, 1, 0); +x_135 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_135, 0, x_134); x_136 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_135, x_4, x_5, x_132); lean_dec(x_5); @@ -9486,7 +9486,7 @@ if (lean_is_scalar(x_156)) { } lean_ctor_set(x_158, 0, x_154); lean_ctor_set(x_158, 1, x_28); -x_159 = lean_alloc_ctor(8, 1, 0); +x_159 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_159, 0, x_158); x_160 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_159, x_4, x_5, x_155); lean_dec(x_5); @@ -9545,7 +9545,7 @@ if (lean_is_scalar(x_178)) { } lean_ctor_set(x_180, 0, x_176); lean_ctor_set(x_180, 1, x_28); -x_181 = lean_alloc_ctor(8, 1, 0); +x_181 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_181, 0, x_180); x_182 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_181, x_4, x_5, x_177); lean_dec(x_5); @@ -9634,7 +9634,7 @@ if (lean_is_scalar(x_206)) { } lean_ctor_set(x_208, 0, x_204); lean_ctor_set(x_208, 1, x_28); -x_209 = lean_alloc_ctor(8, 1, 0); +x_209 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_209, 0, x_208); x_210 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_209, x_4, x_5, x_205); lean_dec(x_5); @@ -9698,7 +9698,7 @@ if (lean_is_scalar(x_229)) { } lean_ctor_set(x_231, 0, x_227); lean_ctor_set(x_231, 1, x_28); -x_232 = lean_alloc_ctor(8, 1, 0); +x_232 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_232, 0, x_231); x_233 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_232, x_4, x_5, x_228); lean_dec(x_5); @@ -9825,7 +9825,7 @@ if (lean_is_scalar(x_268)) { } lean_ctor_set(x_270, 0, x_266); lean_ctor_set(x_270, 1, x_28); -x_271 = lean_alloc_ctor(8, 1, 0); +x_271 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_271, 0, x_270); x_272 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_271, x_4, x_5, x_267); lean_dec(x_5); @@ -9894,7 +9894,7 @@ if (lean_is_scalar(x_291)) { } lean_ctor_set(x_293, 0, x_289); lean_ctor_set(x_293, 1, x_28); -x_294 = lean_alloc_ctor(8, 1, 0); +x_294 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_294, 0, x_293); x_295 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_294, x_4, x_5, x_290); lean_dec(x_5); @@ -10310,7 +10310,7 @@ if (lean_is_scalar(x_413)) { } lean_ctor_set(x_415, 0, x_411); lean_ctor_set(x_415, 1, x_28); -x_416 = lean_alloc_ctor(8, 1, 0); +x_416 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_416, 0, x_415); x_417 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_416, x_4, x_5, x_412); lean_dec(x_5); @@ -10379,7 +10379,7 @@ if (lean_is_scalar(x_436)) { } lean_ctor_set(x_438, 0, x_434); lean_ctor_set(x_438, 1, x_28); -x_439 = lean_alloc_ctor(8, 1, 0); +x_439 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_439, 0, x_438); x_440 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_439, x_4, x_5, x_435); lean_dec(x_5); @@ -10808,7 +10808,7 @@ if (lean_is_scalar(x_555)) { } lean_ctor_set(x_558, 0, x_553); lean_ctor_set(x_558, 1, x_557); -x_559 = lean_alloc_ctor(8, 1, 0); +x_559 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_559, 0, x_558); x_560 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_559, x_4, x_5, x_554); lean_dec(x_5); @@ -10878,7 +10878,7 @@ if (lean_is_scalar(x_579)) { } lean_ctor_set(x_582, 0, x_577); lean_ctor_set(x_582, 1, x_581); -x_583 = lean_alloc_ctor(8, 1, 0); +x_583 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_583, 0, x_582); x_584 = l_Lean_Elab_pushInfoLeaf___at_Lean_withSetOptionIn___spec__3(x_583, x_4, x_5, x_578); lean_dec(x_5); @@ -11923,7 +11923,7 @@ lean_object* x_16; lean_object* x_17; x_16 = lean_ctor_get(x_14, 0); x_17 = lean_ctor_get(x_14, 1); lean_dec(x_17); -if (lean_obj_tag(x_16) == 8) +if (lean_obj_tag(x_16) == 9) { uint8_t x_18; x_18 = !lean_is_exclusive(x_16); @@ -12111,7 +12111,7 @@ lean_object* x_52; x_52 = lean_ctor_get(x_14, 0); lean_inc(x_52); lean_dec(x_14); -if (lean_obj_tag(x_52) == 8) +if (lean_obj_tag(x_52) == 9) { lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; x_53 = lean_ctor_get(x_52, 0); @@ -12332,7 +12332,7 @@ lean_object* x_16; lean_object* x_17; x_16 = lean_ctor_get(x_14, 0); x_17 = lean_ctor_get(x_14, 1); lean_dec(x_17); -if (lean_obj_tag(x_16) == 8) +if (lean_obj_tag(x_16) == 9) { uint8_t x_18; x_18 = !lean_is_exclusive(x_16); @@ -12520,7 +12520,7 @@ lean_object* x_52; x_52 = lean_ctor_get(x_14, 0); lean_inc(x_52); lean_dec(x_14); -if (lean_obj_tag(x_52) == 8) +if (lean_obj_tag(x_52) == 9) { lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; x_53 = lean_ctor_get(x_52, 0); @@ -14920,7 +14920,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -14932,16 +14932,16 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1; +x_1 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1; x_2 = lean_array_mk(x_1); return x_2; } } -static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3() { +static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3() { _start: { lean_object* x_1; @@ -14949,12 +14949,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeActio return x_1; } } -LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2; -x_3 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3; +x_2 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2; +x_3 = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3; x_4 = l_Lean_CodeAction_insertBuiltin(x_2, x_3, x_1); return x_4; } @@ -15238,13 +15238,13 @@ l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17 = _init_l_L lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__17); l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18 = _init_l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18(); lean_mark_persistent(l_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction___rarg___closed__18); -l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__1); -l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__2); -l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3(); -lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210____closed__3); -if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3210_(lean_io_mk_world()); +l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__1); +l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__2); +l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3 = _init_l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3(); +lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212____closed__3); +if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_GuardMsgs_guardMsgsCodeAction_declare__1____x40_Lean_Elab_GuardMsgs___hyg_3212_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/InfoTree/Main.c b/stage0/stdlib/Lean/Elab/InfoTree/Main.c index a66503ad894e..3e564eac7f76 100644 --- a/stage0/stdlib/Lean/Elab/InfoTree/Main.c +++ b/stage0/stdlib/Lean/Elab/InfoTree/Main.c @@ -26,6 +26,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_CompletionInfo_format(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_InfoTree_format___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstWithInfos___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_PartialTermInfo_format___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); @@ -45,7 +46,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa LEAN_EXPORT lean_object* l_Lean_Elab_withInfoHole(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_TermInfo_format___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_TermInfo_runMetaM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CompletionInfo_format___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__5(lean_object*); static lean_object* l_Lean_PersistentArray_findSomeMAux___at_Lean_Elab_InfoTree_findInfo_x3f___spec__2___closed__1; @@ -74,6 +75,7 @@ static lean_object* l_Lean_Elab_OmissionInfo_format___closed__4; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_realizeGlobalConstWithInfos___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_InfoTree_substitute___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_PartialTermInfo_format(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRange_fmtPos___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___rarg___lambda__1(lean_object*, lean_object*); @@ -99,7 +101,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___rarg___lambda__2(lean lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoHole___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_ppGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRange_fmtPos___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext(lean_object*); @@ -136,7 +138,7 @@ static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRa static lean_object* l_Lean_Elab_InfoTree_format___closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_InfoTree_findInfo_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_InfoTree_format___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___rarg___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_TermInfo_format___lambda__1___closed__5; @@ -200,6 +202,7 @@ static lean_object* l_Lean_Elab_OptionInfo_format___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_ChoiceInfo_format(lean_object*, lean_object*); static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_OptionInfo_format(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_assignInfoHoleId___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -271,6 +274,7 @@ static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRa static lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___rarg___lambda__1___closed__3; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_PartialTermInfo_format___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_UserWidgetInfo_format(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__3___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -323,7 +327,7 @@ static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Elab_InfoTree_substitut LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_realizeGlobalConstWithInfos___spec__3___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_InfoTree_findInfo_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_assignInfoHoleId(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_InfoTree_substitute___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__3; @@ -336,7 +340,7 @@ lean_object* l_Lean_ppTerm(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__3; static uint8_t l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13; -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__1(lean_object*); @@ -348,6 +352,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalCo uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoHole___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_ChoiceInfo_format___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_assignInfoHoleId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -374,7 +379,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_In static lean_object* l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___closed__3; static lean_object* l_Lean_Elab_TermInfo_format___lambda__1___closed__3; static lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_ContextInfo_ppGoals___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_format(lean_object*, lean_object*, lean_object*); @@ -418,6 +423,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_MacroExpansionInfo_format(lean_object*, lea static lean_object* l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_ppGoals(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_ChoiceInfo_format___closed__2; static lean_object* l_Lean_Elab_instToFormatCustomInfo___closed__1; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); @@ -430,7 +436,7 @@ static lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRa lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_findInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_modifyInfoTrees___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_assignInfoHoleId___spec__5(lean_object*); @@ -4959,6 +4965,43 @@ lean_dec(x_4); return x_9; } } +static lean_object* _init_l_Lean_Elab_PartialTermInfo_format___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Partial term @ ", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_PartialTermInfo_format___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_PartialTermInfo_format___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_PartialTermInfo_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +lean_dec(x_2); +x_4 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatElabInfo(x_1, x_3); +x_5 = l_Lean_Elab_PartialTermInfo_format___closed__2; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRange_fmtPos___closed__7; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} static lean_object* _init_l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__1() { _start: { @@ -7287,6 +7330,40 @@ return x_32; } } } +static lean_object* _init_l_Lean_Elab_ChoiceInfo_format___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Choice @ ", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_ChoiceInfo_format___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_ChoiceInfo_format___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_ChoiceInfo_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatElabInfo(x_1, x_2); +x_4 = l_Lean_Elab_ChoiceInfo_format___closed__2; +x_5 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +x_6 = l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_formatStxRange_fmtPos___closed__7; +x_7 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Info_format(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -7311,61 +7388,60 @@ return x_7; } case 2: { -lean_object* x_8; lean_object* x_9; +lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_ctor_get(x_2, 0); lean_inc(x_8); lean_dec(x_2); -x_9 = l_Lean_Elab_CommandInfo_format(x_1, x_8, x_3); -return x_9; +x_9 = l_Lean_Elab_PartialTermInfo_format(x_1, x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_3); +return x_10; } case 3: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_2, 0); -lean_inc(x_10); +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); lean_dec(x_2); -x_11 = l_Lean_Elab_MacroExpansionInfo_format(x_1, x_10, x_3); -lean_dec(x_1); -return x_11; +x_12 = l_Lean_Elab_CommandInfo_format(x_1, x_11, x_3); +return x_12; } case 4: { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); lean_dec(x_2); -x_13 = l_Lean_Elab_OptionInfo_format(x_1, x_12, x_3); -return x_13; +x_14 = l_Lean_Elab_MacroExpansionInfo_format(x_1, x_13, x_3); +lean_dec(x_1); +return x_14; } case 5: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_2, 0); +lean_inc(x_15); lean_dec(x_2); -x_15 = l_Lean_Elab_FieldInfo_format(x_1, x_14, x_3); -return x_15; +x_16 = l_Lean_Elab_OptionInfo_format(x_1, x_15, x_3); +return x_16; } case 6: { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_2, 0); -lean_inc(x_16); +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_2, 0); +lean_inc(x_17); lean_dec(x_2); -x_17 = l_Lean_Elab_CompletionInfo_format(x_1, x_16, x_3); -return x_17; +x_18 = l_Lean_Elab_FieldInfo_format(x_1, x_17, x_3); +return x_18; } case 7: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_1); -x_18 = lean_ctor_get(x_2, 0); -lean_inc(x_18); +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_2, 0); +lean_inc(x_19); lean_dec(x_2); -x_19 = l_Lean_Elab_UserWidgetInfo_format(x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_3); +x_20 = l_Lean_Elab_CompletionInfo_format(x_1, x_19, x_3); return x_20; } case 8: @@ -7375,7 +7451,7 @@ lean_dec(x_1); x_21 = lean_ctor_get(x_2, 0); lean_inc(x_21); lean_dec(x_2); -x_22 = l_Lean_Elab_CustomInfo_format(x_21); +x_22 = l_Lean_Elab_UserWidgetInfo_format(x_21); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_3); @@ -7388,7 +7464,7 @@ lean_dec(x_1); x_24 = lean_ctor_get(x_2, 0); lean_inc(x_24); lean_dec(x_2); -x_25 = l_Lean_Elab_FVarAliasInfo_format(x_24); +x_25 = l_Lean_Elab_CustomInfo_format(x_24); x_26 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_3); @@ -7397,24 +7473,49 @@ return x_26; case 10: { lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); x_27 = lean_ctor_get(x_2, 0); lean_inc(x_27); lean_dec(x_2); -x_28 = l_Lean_Elab_FieldRedeclInfo_format(x_1, x_27); -lean_dec(x_27); +x_28 = l_Lean_Elab_FVarAliasInfo_format(x_27); x_29 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_3); return x_29; } -default: +case 11: { -lean_object* x_30; lean_object* x_31; +lean_object* x_30; lean_object* x_31; lean_object* x_32; x_30 = lean_ctor_get(x_2, 0); lean_inc(x_30); lean_dec(x_2); -x_31 = l_Lean_Elab_OmissionInfo_format(x_1, x_30, x_3); -return x_31; +x_31 = l_Lean_Elab_FieldRedeclInfo_format(x_1, x_30); +lean_dec(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_3); +return x_32; +} +case 12: +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_2, 0); +lean_inc(x_33); +lean_dec(x_2); +x_34 = l_Lean_Elab_OmissionInfo_format(x_1, x_33, x_3); +return x_34; +} +default: +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_2, 0); +lean_inc(x_35); +lean_dec(x_2); +x_36 = l_Lean_Elab_ChoiceInfo_format(x_1, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_3); +return x_37; } } } @@ -7486,61 +7587,110 @@ uint8_t x_14; x_14 = !lean_is_exclusive(x_1); if (x_14 == 0) { +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_1, 0); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_16); return x_1; } else { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_1, 0); -lean_inc(x_15); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); lean_dec(x_1); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -return x_16; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +return x_19; } } -case 11: +case 3: { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_1); -if (x_17 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_1); +if (x_20 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_1, 0); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); lean_ctor_set_tag(x_1, 1); -lean_ctor_set(x_1, 0, x_20); return x_1; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_1, 0); lean_inc(x_21); lean_dec(x_1); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec(x_21); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -return x_24; +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +case 12: +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_1); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +lean_ctor_set_tag(x_1, 1); +lean_ctor_set(x_1, 0, x_26); +return x_1; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +return x_30; +} +} +case 13: +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_1); +if (x_31 == 0) +{ +lean_ctor_set_tag(x_1, 1); +return x_1; +} +else +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_1, 0); +lean_inc(x_32); +lean_dec(x_1); +x_33 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_33, 0, x_32); +return x_33; } } default: { -lean_object* x_25; +lean_object* x_34; lean_dec(x_1); -x_25 = lean_box(0); -return x_25; +x_34 = lean_box(0); +return x_34; } } } @@ -8804,7 +8954,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___rarg(lean_object* x_1, _start: { lean_object* x_4; lean_object* x_5; -x_4 = lean_alloc_ctor(6, 1, 0); +x_4 = lean_alloc_ctor(7, 1, 0); lean_ctor_set(x_4, 0, x_3); x_5 = l_Lean_Elab_pushInfoLeaf___rarg(x_1, x_2, x_4); return x_5; @@ -10478,35 +10628,7 @@ lean_inc(x_2); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; -x_4 = lean_ctor_get(x_2, 1); -lean_dec(x_4); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -else -{ -uint8_t x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_1); -lean_ctor_set_uint8(x_7, sizeof(void*)*2, x_5); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10614,48 +10736,64 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_1, 1); lean_inc(x_4); lean_dec(x_1); -x_5 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__3), 3, 2); +x_5 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__2), 3, 2); lean_closure_set(x_5, 0, x_3); lean_closure_set(x_5, 1, x_2); x_6 = lean_apply_1(x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_dec(x_4); -lean_dec(x_3); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); lean_dec(x_1); -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__2), 2, 1); -lean_closure_set(x_7, 0, x_2); -x_8 = lean_apply_1(x_6, x_7); +x_6 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_6, 0, x_4); +x_7 = lean_apply_2(x_5, lean_box(0), x_6); +x_8 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_7, x_2); return x_8; } +} +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__3), 3, 2); +lean_closure_set(x_8, 0, x_1); +lean_closure_set(x_8, 1, x_2); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_6); +lean_inc(x_4); +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__4), 4, 3); +lean_closure_set(x_9, 0, x_3); +lean_closure_set(x_9, 1, x_8); +lean_closure_set(x_9, 2, x_4); +x_10 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_5, x_9); +return x_10; +} else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_dec(x_5); -x_10 = lean_apply_1(x_3, x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__4), 3, 2); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_2); -x_12 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_10, x_11); -return x_12; +lean_dec(x_3); +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); +lean_dec(x_7); +x_12 = lean_apply_1(x_6, x_11); +x_13 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_12, x_8); +return x_13; } } } @@ -10667,37 +10805,39 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -lean_dec(x_1); -x_9 = lean_ctor_get(x_8, 0); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); -lean_dec(x_8); +lean_dec(x_1); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__5), 5, 4); -lean_closure_set(x_11, 0, x_2); -lean_closure_set(x_11, 1, x_7); -lean_closure_set(x_11, 2, x_3); -lean_closure_set(x_11, 3, x_4); -x_12 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_6, x_11); -x_13 = l_Lean_Elab_withInfoContext_x27___rarg___lambda__6___closed__1; -x_14 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_13, x_12); -return x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__5), 7, 6); +lean_closure_set(x_12, 0, x_2); +lean_closure_set(x_12, 1, x_8); +lean_closure_set(x_12, 2, x_9); +lean_closure_set(x_12, 3, x_3); +lean_closure_set(x_12, 4, x_4); +lean_closure_set(x_12, 5, x_5); +x_13 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_12); +x_14 = l_Lean_Elab_withInfoContext_x27___rarg___lambda__6___closed__1; +x_15 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_14, x_13); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_8; -x_8 = lean_ctor_get_uint8(x_7, sizeof(void*)*2); -if (x_8 == 0) +uint8_t x_9; +x_9 = lean_ctor_get_uint8(x_8, sizeof(void*)*2); +if (x_9 == 0) { +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -10707,48 +10847,50 @@ return x_1; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_inc(x_3); lean_inc(x_2); -x_9 = l_Lean_Elab_getResetInfoTrees___rarg(x_2, x_3); -lean_inc(x_5); -x_10 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__6), 7, 6); -lean_closure_set(x_10, 0, x_2); -lean_closure_set(x_10, 1, x_3); -lean_closure_set(x_10, 2, x_4); -lean_closure_set(x_10, 3, x_5); -lean_closure_set(x_10, 4, x_6); -lean_closure_set(x_10, 5, x_1); -x_11 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_9, x_10); -return x_11; +x_10 = l_Lean_Elab_getResetInfoTrees___rarg(x_2, x_3); +lean_inc(x_4); +x_11 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__6), 8, 7); +lean_closure_set(x_11, 0, x_2); +lean_closure_set(x_11, 1, x_3); +lean_closure_set(x_11, 2, x_4); +lean_closure_set(x_11, 3, x_5); +lean_closure_set(x_11, 4, x_6); +lean_closure_set(x_11, 5, x_7); +lean_closure_set(x_11, 6, x_1); +x_12 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_10, x_11); +return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = lean_ctor_get(x_2, 0); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); -lean_inc(x_7); -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed), 7, 6); -lean_closure_set(x_9, 0, x_5); -lean_closure_set(x_9, 1, x_1); -lean_closure_set(x_9, 2, x_2); -lean_closure_set(x_9, 3, x_6); -lean_closure_set(x_9, 4, x_7); -lean_closure_set(x_9, 5, x_4); -x_10 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_9); -return x_10; +x_9 = lean_ctor_get(x_2, 0); +lean_inc(x_9); +lean_inc(x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed), 8, 7); +lean_closure_set(x_10, 0, x_5); +lean_closure_set(x_10, 1, x_1); +lean_closure_set(x_10, 2, x_2); +lean_closure_set(x_10, 3, x_8); +lean_closure_set(x_10, 4, x_7); +lean_closure_set(x_10, 5, x_6); +lean_closure_set(x_10, 6, x_4); +x_11 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_9, x_10); +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg), 6, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoContext_x27___rarg), 7, 0); return x_2; } } @@ -10761,13 +10903,13 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___rarg___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; -x_8 = l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_7); -return x_8; +lean_object* x_9; +x_9 = l_Lean_Elab_withInfoContext_x27___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +return x_9; } } LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -12521,7 +12663,7 @@ static lean_object* _init_l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___clos _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Elab.InfoTree.Main._hyg.4432.0 ).isNone\n ", 62, 62); +x_1 = lean_mk_string_unchecked("( __do_lift._@.Lean.Elab.InfoTree.Main._hyg.4574.0 ).isNone\n ", 62, 62); return x_1; } } @@ -12549,7 +12691,7 @@ static lean_object* _init_l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f___closed__1; x_2 = l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___closed__4; -x_3 = lean_unsigned_to_nat(402u); +x_3 = lean_unsigned_to_nat(425u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12688,7 +12830,7 @@ x_7 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_7, 0, x_4); lean_ctor_set(x_7, 1, x_2); lean_ctor_set(x_7, 2, x_3); -x_8 = lean_alloc_ctor(3, 1, 0); +x_8 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_8, 0, x_7); x_9 = lean_apply_2(x_6, lean_box(0), x_8); return x_9; @@ -13269,6 +13411,10 @@ l_Lean_Elab_TermInfo_format___lambda__2___closed__1 = _init_l_Lean_Elab_TermInfo lean_mark_persistent(l_Lean_Elab_TermInfo_format___lambda__2___closed__1); l_Lean_Elab_TermInfo_format___lambda__2___closed__2 = _init_l_Lean_Elab_TermInfo_format___lambda__2___closed__2(); lean_mark_persistent(l_Lean_Elab_TermInfo_format___lambda__2___closed__2); +l_Lean_Elab_PartialTermInfo_format___closed__1 = _init_l_Lean_Elab_PartialTermInfo_format___closed__1(); +lean_mark_persistent(l_Lean_Elab_PartialTermInfo_format___closed__1); +l_Lean_Elab_PartialTermInfo_format___closed__2 = _init_l_Lean_Elab_PartialTermInfo_format___closed__2(); +lean_mark_persistent(l_Lean_Elab_PartialTermInfo_format___closed__2); l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__1 = _init_l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__1(); lean_mark_persistent(l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__1); l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__2 = _init_l_Option_format___at_Lean_Elab_CompletionInfo_format___spec__1___closed__2(); @@ -13353,6 +13499,10 @@ l_Lean_Elab_OmissionInfo_format___closed__3 = _init_l_Lean_Elab_OmissionInfo_for lean_mark_persistent(l_Lean_Elab_OmissionInfo_format___closed__3); l_Lean_Elab_OmissionInfo_format___closed__4 = _init_l_Lean_Elab_OmissionInfo_format___closed__4(); lean_mark_persistent(l_Lean_Elab_OmissionInfo_format___closed__4); +l_Lean_Elab_ChoiceInfo_format___closed__1 = _init_l_Lean_Elab_ChoiceInfo_format___closed__1(); +lean_mark_persistent(l_Lean_Elab_ChoiceInfo_format___closed__1); +l_Lean_Elab_ChoiceInfo_format___closed__2 = _init_l_Lean_Elab_ChoiceInfo_format___closed__2(); +lean_mark_persistent(l_Lean_Elab_ChoiceInfo_format___closed__2); l_Lean_Elab_InfoTree_format___closed__1 = _init_l_Lean_Elab_InfoTree_format___closed__1(); lean_mark_persistent(l_Lean_Elab_InfoTree_format___closed__1); l_Lean_Elab_InfoTree_format___closed__2 = _init_l_Lean_Elab_InfoTree_format___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/InfoTree/Types.c b/stage0/stdlib/Lean/Elab/InfoTree/Types.c index 39c4ec7db816..ea36dc1916b2 100644 --- a/stage0/stdlib/Lean/Elab/InfoTree/Types.c +++ b/stage0/stdlib/Lean/Elab/InfoTree/Types.c @@ -20,6 +20,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_setInfoState(lean_object*); static lean_object* l_Lean_Elab_instInhabitedTermInfo___closed__6; lean_object* l_Lean_Expr_bvar___override(lean_object*); size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Elab_instInhabitedPartialTermInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedMacroExpansionInfo; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedTermInfo; static lean_object* l_Lean_Elab_instInhabitedFieldInfo___closed__1; @@ -27,6 +28,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedInfoTree; static lean_object* l_Lean_Elab_instInhabitedTermInfo___closed__2; static lean_object* l_Lean_Elab_instInhabitedInfoTree___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_setInfoState___rarg___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedPartialTermInfo; LEAN_EXPORT lean_object* l_Lean_Elab_instInhabitedInfoState; static lean_object* l_Lean_Elab_instInhabitedTacticInfo___closed__1; static lean_object* l_Lean_Elab_instInhabitedTermInfo___closed__8; @@ -180,6 +182,28 @@ x_1 = l_Lean_Elab_instInhabitedTermInfo___closed__9; return x_1; } } +static lean_object* _init_l_Lean_Elab_instInhabitedPartialTermInfo___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_instInhabitedElabInfo___closed__1; +x_3 = l_Lean_Elab_instInhabitedTermInfo___closed__7; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 2, x_1); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_instInhabitedPartialTermInfo() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_instInhabitedPartialTermInfo___closed__1; +return x_1; +} +} static lean_object* _init_l_Lean_Elab_instInhabitedCommandInfo() { _start: { @@ -468,6 +492,10 @@ l_Lean_Elab_instInhabitedTermInfo___closed__9 = _init_l_Lean_Elab_instInhabitedT lean_mark_persistent(l_Lean_Elab_instInhabitedTermInfo___closed__9); l_Lean_Elab_instInhabitedTermInfo = _init_l_Lean_Elab_instInhabitedTermInfo(); lean_mark_persistent(l_Lean_Elab_instInhabitedTermInfo); +l_Lean_Elab_instInhabitedPartialTermInfo___closed__1 = _init_l_Lean_Elab_instInhabitedPartialTermInfo___closed__1(); +lean_mark_persistent(l_Lean_Elab_instInhabitedPartialTermInfo___closed__1); +l_Lean_Elab_instInhabitedPartialTermInfo = _init_l_Lean_Elab_instInhabitedPartialTermInfo(); +lean_mark_persistent(l_Lean_Elab_instInhabitedPartialTermInfo); l_Lean_Elab_instInhabitedCommandInfo = _init_l_Lean_Elab_instInhabitedCommandInfo(); lean_mark_persistent(l_Lean_Elab_instInhabitedCommandInfo); l_Lean_Elab_instInhabitedFieldInfo___closed__1 = _init_l_Lean_Elab_instInhabitedFieldInfo___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/LetRec.c b/stage0/stdlib/Lean/Elab/LetRec.c index 5c65f844d8a1..f6f6ebedcf56 100644 --- a/stage0/stdlib/Lean/Elab/LetRec.c +++ b/stage0/stdlib/Lean/Elab/LetRec.c @@ -76,6 +76,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +lean_object* l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__22___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__27___lambda__2___boxed(lean_object**); @@ -268,7 +269,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_LetR static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18___closed__5; lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___lambda__3___closed__2; -lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_withAuxLocalDecls_loop(lean_object*); static lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__3; @@ -7387,14 +7388,16 @@ return x_22; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_Elab_Term_mkBodyInfo(x_1, x_2); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_2); +x_11 = l_Lean_Elab_Term_mkBodyInfo(x_1, x_10); +x_12 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -7421,7 +7424,7 @@ lean_dec(x_15); lean_dec(x_1); if (lean_obj_tag(x_17) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); lean_dec(x_17); @@ -7442,16 +7445,21 @@ lean_closure_set(x_23, 3, x_4); lean_inc(x_20); x_24 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__2___lambda__2___boxed), 9, 1); lean_closure_set(x_24, 0, x_20); -x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withInfoContext_x27), 10, 3); -lean_closure_set(x_25, 0, x_20); -lean_closure_set(x_25, 1, x_23); -lean_closure_set(x_25, 2, x_24); -x_26 = l_Lean_Elab_Term_withDeclName___rarg(x_19, x_25, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -return x_26; +lean_inc(x_20); +x_25 = l_Lean_Elab_Term_mkBodyInfo(x_20, x_22); +x_26 = lean_alloc_closure((void*)(l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed), 8, 1); +lean_closure_set(x_26, 0, x_25); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withInfoContext_x27), 11, 4); +lean_closure_set(x_27, 0, x_20); +lean_closure_set(x_27, 1, x_23); +lean_closure_set(x_27, 2, x_24); +lean_closure_set(x_27, 3, x_26); +x_28 = l_Lean_Elab_Term_withDeclName___rarg(x_19, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +return x_28; } else { -uint8_t x_27; +uint8_t x_29; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -7461,23 +7469,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_27 = !lean_is_exclusive(x_17); -if (x_27 == 0) +x_29 = !lean_is_exclusive(x_17); +if (x_29 == 0) { return x_17; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_17, 0); -x_29 = lean_ctor_get(x_17, 1); -lean_inc(x_29); -lean_inc(x_28); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_17, 0); +x_31 = lean_ctor_get(x_17, 1); +lean_inc(x_31); +lean_inc(x_30); lean_dec(x_17); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index 03eea69cc763..a996b7ec5f74 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -22,7 +22,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Matc LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_precheckMatch___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___lambda__1___closed__1; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main_pack___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch__1___closed__4; LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -60,6 +59,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPa lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_waitExpectedTypeAndDiscrs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__1___closed__1; @@ -83,7 +83,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at_Lean_Elab_Term_ToDepElimPattern_main___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5; lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -101,7 +100,7 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Match_0__Le extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_Meta_resetMVarUserNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabNoMatch_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___closed__1; @@ -125,15 +124,16 @@ static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_ static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1; static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__6___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__11___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize___closed__2; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__12___boxed(lean_object**); @@ -146,7 +146,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962_(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_containsFVar___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -158,6 +157,7 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ToDepElimPa static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___lambda__1___closed__1; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__9___lambda__1___closed__3; @@ -167,6 +167,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lea LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__2___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4; lean_object* l_Lean_Meta_getFVarsToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -187,7 +188,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoFun_declRange__1___close lean_object* l_Lean_Elab_Term_getPatternsVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__2(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getIndexToInclude_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_pack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -316,7 +316,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__12___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -347,6 +346,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar( LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoFun__1___closed__1; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__3(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch__1___closed__3; size_t lean_ptr_addr(lean_object*); @@ -355,11 +355,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchG LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Term_elabNoFun___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__8___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__6___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Elab_Term_Quotation_withNewLocals___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -413,10 +414,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Term static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange__1___closed__7; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitPost___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__3(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___closed__1; lean_object* l_Lean_mkInaccessible(lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -436,6 +437,7 @@ lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___lambda__1___closed__2; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* l_Lean_Elab_Term_runPendingTacticsAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -469,6 +471,7 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_packMatchTypePatterns___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__5___lambda__1___boxed(lean_object**); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_kabstract(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_isAtomicDiscr___closed__4; @@ -520,7 +523,6 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Match_0__Lean_ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__9(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Closure_mkBinding___spec__1(size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_outOfBounds___rarg(lean_object*); @@ -536,6 +538,7 @@ static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__2___ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__2___lambda__1(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withDepElimPatterns___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -548,6 +551,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch__1(lean_objec static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__5___closed__2; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14; lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_markIsDep(lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); @@ -577,6 +581,7 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__11___boxed(lean_object**); lean_object* l_Lean_throwError___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__50(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ToDepElimPattern_main___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); @@ -587,6 +592,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_elabMatch_el lean_object* l_Lean_isPatternWithRef___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_waitExpectedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main_pack___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__5; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2; @@ -598,7 +604,6 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_ToDepElimPattern_n lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___boxed(lean_object**); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Match_Pattern_hasExprMVar(lean_object*); static lean_object* l_Lean_Elab_Term_elabNoFun___closed__10; @@ -636,9 +641,7 @@ lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11; lean_object* l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -678,6 +681,7 @@ uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__16___rarg___closed__2; lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__10___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3; LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -705,8 +709,9 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible__1___closed__ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2; extern lean_object* l_Lean_Elab_abortTermExceptionId; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -719,7 +724,6 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Match_0__Lean_Elab_Ter LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__16___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -763,13 +767,13 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPatte lean_object* l_Lean_Meta_Match_isNamedPattern_x3f(lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__5___closed__1; static lean_object* l_Lean_Elab_Term_elabNoFun___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__11; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__16(lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__10___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__6(lean_object*, lean_object*, lean_object*); @@ -841,9 +845,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Ma lean_object* l_Lean_Expr_replaceFVars(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831_(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_mkPartialTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at_Lean_Elab_Term_ToDepElimPattern_main___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ToDepElimPattern_normalize_processInaccessible___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -856,6 +862,7 @@ uint8_t l_Lean_Name_isAtomic(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getIndexToInclude_x3f_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_withMVar(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___lambda__2(lean_object*, lean_object*, lean_object*); @@ -889,7 +896,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDep lean_object* l_Lean_Meta_setMVarUserNamesAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabMatch_elabMatchDefault___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -938,6 +944,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange__1___close lean_object* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch__1___closed__2; lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12; lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_erase(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_normalize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -952,7 +959,6 @@ lean_object* l_Lean_localDeclDependsOn___at_Lean_Elab_Term_addAutoBoundImplicits LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -983,6 +989,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Matc lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_ToDepElimPattern_normalize___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1004,7 +1011,6 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Matc LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1106,10 +1112,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatch lean_object* l_Lean_MVarId_setTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch__1___closed__2; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10; lean_object* l_Lean_Meta_mkFreshTypeMVar(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__6___boxed(lean_object**); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__7(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__5; @@ -1121,6 +1125,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElim LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ToDepElimPattern_main___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar___closed__1; @@ -1137,7 +1142,6 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at_Lean_Elab_Term_ToDepElimPattern_main___spec__8(lean_object*); @@ -1163,12 +1167,12 @@ static lean_object* l_Lean_Elab_Term_isAtomicDiscr___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoMatch___spec__3(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConst(lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); @@ -1192,6 +1196,7 @@ static lean_object* l_Lean_Elab_Term_precheckMatch___closed__2; lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__11(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_isAtomicDiscr___spec__1___rarg___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_unpack_go(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__16___rarg___closed__6; @@ -27712,396 +27717,389 @@ x_7 = lean_alloc_closure((void*)(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Te return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_st_ref_get(x_9, x_10); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_st_ref_take(x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_12, 6); lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); -lean_dec(x_13); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_11); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_2); +lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_11, 1); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_box(x_3); -x_17 = lean_apply_8(x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_dec(x_11); -x_19 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg(x_9, x_18); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_box(x_3); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_23 = lean_apply_8(x_1, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_box(x_3); -lean_inc(x_9); -lean_inc(x_24); -x_27 = lean_apply_9(x_2, x_24, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_st_ref_take(x_9, x_29); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_31, 6); -lean_inc(x_32); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_33; -x_33 = !lean_is_exclusive(x_30); -if (x_33 == 0) -{ -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = lean_ctor_get(x_30, 1); -x_35 = lean_ctor_get(x_30, 0); -lean_dec(x_35); -x_36 = !lean_is_exclusive(x_31); -if (x_36 == 0) +x_16 = lean_ctor_get(x_11, 0); +lean_dec(x_16); +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) { -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_31, 6); -lean_dec(x_37); -x_38 = !lean_is_exclusive(x_32); -if (x_38 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_12, 6); +lean_dec(x_18); +x_19 = !lean_is_exclusive(x_13); +if (x_19 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_39 = lean_ctor_get(x_32, 1); -x_40 = lean_ctor_get(x_28, 0); -lean_inc(x_40); -lean_dec(x_28); -lean_ctor_set_tag(x_30, 1); -lean_ctor_set(x_30, 1, x_39); -lean_ctor_set(x_30, 0, x_40); -x_41 = l_Lean_PersistentArray_push___rarg(x_20, x_30); -lean_ctor_set(x_32, 1, x_41); -x_42 = lean_st_ref_set(x_9, x_31, x_34); -lean_dec(x_9); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_ctor_get(x_13, 1); +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +lean_dec(x_2); +lean_ctor_set_tag(x_11, 1); +lean_ctor_set(x_11, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +x_22 = l_Lean_PersistentArray_push___rarg(x_1, x_11); +lean_ctor_set(x_13, 1, x_22); +x_23 = lean_st_ref_set(x_9, x_12, x_15); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_44; -x_44 = lean_ctor_get(x_42, 0); -lean_dec(x_44); -lean_ctor_set(x_42, 0, x_24); -return x_42; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_23, 0, x_26); +return x_23; } else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_24); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_47 = lean_ctor_get_uint8(x_32, sizeof(void*)*2); -x_48 = lean_ctor_get(x_32, 0); -x_49 = lean_ctor_get(x_32, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_32); -x_50 = lean_ctor_get(x_28, 0); -lean_inc(x_50); -lean_dec(x_28); -lean_ctor_set_tag(x_30, 1); -lean_ctor_set(x_30, 1, x_49); -lean_ctor_set(x_30, 0, x_50); -x_51 = l_Lean_PersistentArray_push___rarg(x_20, x_30); -x_52 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_52, 0, x_48); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set_uint8(x_52, sizeof(void*)*2, x_47); -lean_ctor_set(x_31, 6, x_52); -x_53 = lean_st_ref_set(x_9, x_31, x_34); -lean_dec(x_9); -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_53)) { - lean_ctor_release(x_53, 0); - lean_ctor_release(x_53, 1); - x_55 = x_53; +uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_30 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +x_31 = lean_ctor_get(x_13, 0); +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_13); +x_33 = lean_ctor_get(x_2, 0); +lean_inc(x_33); +lean_dec(x_2); +lean_ctor_set_tag(x_11, 1); +lean_ctor_set(x_11, 1, x_32); +lean_ctor_set(x_11, 0, x_33); +x_34 = l_Lean_PersistentArray_push___rarg(x_1, x_11); +x_35 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_35, 0, x_31); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set_uint8(x_35, sizeof(void*)*2, x_30); +lean_ctor_set(x_12, 6, x_35); +x_36 = lean_st_ref_set(x_9, x_12, x_15); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_38 = x_36; } else { - lean_dec_ref(x_53); - x_55 = lean_box(0); + lean_dec_ref(x_36); + x_38 = lean_box(0); } -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(0, 2, 0); +x_39 = lean_box(0); +if (lean_is_scalar(x_38)) { + x_40 = lean_alloc_ctor(0, 2, 0); } else { - x_56 = x_55; + x_40 = x_38; } -lean_ctor_set(x_56, 0, x_24); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_37); +return x_40; } } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_57 = lean_ctor_get(x_31, 0); -x_58 = lean_ctor_get(x_31, 1); -x_59 = lean_ctor_get(x_31, 2); -x_60 = lean_ctor_get(x_31, 3); -x_61 = lean_ctor_get(x_31, 4); -x_62 = lean_ctor_get(x_31, 5); -lean_inc(x_62); -lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_31); -x_63 = lean_ctor_get_uint8(x_32, sizeof(void*)*2); -x_64 = lean_ctor_get(x_32, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_32, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_66 = x_32; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_41 = lean_ctor_get(x_12, 0); +x_42 = lean_ctor_get(x_12, 1); +x_43 = lean_ctor_get(x_12, 2); +x_44 = lean_ctor_get(x_12, 3); +x_45 = lean_ctor_get(x_12, 4); +x_46 = lean_ctor_get(x_12, 5); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_12); +x_47 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +x_48 = lean_ctor_get(x_13, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_13, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_50 = x_13; } else { - lean_dec_ref(x_32); - x_66 = lean_box(0); + lean_dec_ref(x_13); + x_50 = lean_box(0); } -x_67 = lean_ctor_get(x_28, 0); -lean_inc(x_67); -lean_dec(x_28); -lean_ctor_set_tag(x_30, 1); -lean_ctor_set(x_30, 1, x_65); -lean_ctor_set(x_30, 0, x_67); -x_68 = l_Lean_PersistentArray_push___rarg(x_20, x_30); -if (lean_is_scalar(x_66)) { - x_69 = lean_alloc_ctor(0, 2, 1); +x_51 = lean_ctor_get(x_2, 0); +lean_inc(x_51); +lean_dec(x_2); +lean_ctor_set_tag(x_11, 1); +lean_ctor_set(x_11, 1, x_49); +lean_ctor_set(x_11, 0, x_51); +x_52 = l_Lean_PersistentArray_push___rarg(x_1, x_11); +if (lean_is_scalar(x_50)) { + x_53 = lean_alloc_ctor(0, 2, 1); } else { - x_69 = x_66; + x_53 = x_50; } -lean_ctor_set(x_69, 0, x_64); -lean_ctor_set(x_69, 1, x_68); -lean_ctor_set_uint8(x_69, sizeof(void*)*2, x_63); -x_70 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_70, 0, x_57); -lean_ctor_set(x_70, 1, x_58); -lean_ctor_set(x_70, 2, x_59); -lean_ctor_set(x_70, 3, x_60); -lean_ctor_set(x_70, 4, x_61); -lean_ctor_set(x_70, 5, x_62); -lean_ctor_set(x_70, 6, x_69); -x_71 = lean_st_ref_set(x_9, x_70, x_34); -lean_dec(x_9); -x_72 = lean_ctor_get(x_71, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_73 = x_71; +lean_ctor_set(x_53, 0, x_48); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set_uint8(x_53, sizeof(void*)*2, x_47); +x_54 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_54, 0, x_41); +lean_ctor_set(x_54, 1, x_42); +lean_ctor_set(x_54, 2, x_43); +lean_ctor_set(x_54, 3, x_44); +lean_ctor_set(x_54, 4, x_45); +lean_ctor_set(x_54, 5, x_46); +lean_ctor_set(x_54, 6, x_53); +x_55 = lean_st_ref_set(x_9, x_54, x_15); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_57 = x_55; } else { - lean_dec_ref(x_71); - x_73 = lean_box(0); + lean_dec_ref(x_55); + x_57 = lean_box(0); } -if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(0, 2, 0); +x_58 = lean_box(0); +if (lean_is_scalar(x_57)) { + x_59 = lean_alloc_ctor(0, 2, 0); } else { - x_74 = x_73; + x_59 = x_57; } -lean_ctor_set(x_74, 0, x_24); -lean_ctor_set(x_74, 1, x_72); -return x_74; +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_56); +return x_59; } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_75 = lean_ctor_get(x_30, 1); -lean_inc(x_75); -lean_dec(x_30); -x_76 = lean_ctor_get(x_31, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_31, 1); -lean_inc(x_77); -x_78 = lean_ctor_get(x_31, 2); -lean_inc(x_78); -x_79 = lean_ctor_get(x_31, 3); -lean_inc(x_79); -x_80 = lean_ctor_get(x_31, 4); -lean_inc(x_80); -x_81 = lean_ctor_get(x_31, 5); -lean_inc(x_81); -if (lean_is_exclusive(x_31)) { - lean_ctor_release(x_31, 0); - lean_ctor_release(x_31, 1); - lean_ctor_release(x_31, 2); - lean_ctor_release(x_31, 3); - lean_ctor_release(x_31, 4); - lean_ctor_release(x_31, 5); - lean_ctor_release(x_31, 6); - x_82 = x_31; -} else { - lean_dec_ref(x_31); - x_82 = lean_box(0); -} -x_83 = lean_ctor_get_uint8(x_32, sizeof(void*)*2); -x_84 = lean_ctor_get(x_32, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_32, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_86 = x_32; -} else { - lean_dec_ref(x_32); - x_86 = lean_box(0); -} -x_87 = lean_ctor_get(x_28, 0); -lean_inc(x_87); -lean_dec(x_28); -x_88 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_85); -x_89 = l_Lean_PersistentArray_push___rarg(x_20, x_88); -if (lean_is_scalar(x_86)) { - x_90 = lean_alloc_ctor(0, 2, 1); +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_60 = lean_ctor_get(x_11, 1); +lean_inc(x_60); +lean_dec(x_11); +x_61 = lean_ctor_get(x_12, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_12, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_12, 2); +lean_inc(x_63); +x_64 = lean_ctor_get(x_12, 3); +lean_inc(x_64); +x_65 = lean_ctor_get(x_12, 4); +lean_inc(x_65); +x_66 = lean_ctor_get(x_12, 5); +lean_inc(x_66); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + lean_ctor_release(x_12, 2); + lean_ctor_release(x_12, 3); + lean_ctor_release(x_12, 4); + lean_ctor_release(x_12, 5); + lean_ctor_release(x_12, 6); + x_67 = x_12; +} else { + lean_dec_ref(x_12); + x_67 = lean_box(0); +} +x_68 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +x_69 = lean_ctor_get(x_13, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_13, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_71 = x_13; } else { - x_90 = x_86; + lean_dec_ref(x_13); + x_71 = lean_box(0); } -lean_ctor_set(x_90, 0, x_84); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set_uint8(x_90, sizeof(void*)*2, x_83); -if (lean_is_scalar(x_82)) { - x_91 = lean_alloc_ctor(0, 7, 0); +x_72 = lean_ctor_get(x_2, 0); +lean_inc(x_72); +lean_dec(x_2); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_70); +x_74 = l_Lean_PersistentArray_push___rarg(x_1, x_73); +if (lean_is_scalar(x_71)) { + x_75 = lean_alloc_ctor(0, 2, 1); } else { - x_91 = x_82; + x_75 = x_71; } -lean_ctor_set(x_91, 0, x_76); -lean_ctor_set(x_91, 1, x_77); -lean_ctor_set(x_91, 2, x_78); -lean_ctor_set(x_91, 3, x_79); -lean_ctor_set(x_91, 4, x_80); -lean_ctor_set(x_91, 5, x_81); -lean_ctor_set(x_91, 6, x_90); -x_92 = lean_st_ref_set(x_9, x_91, x_75); -lean_dec(x_9); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; +lean_ctor_set(x_75, 0, x_69); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set_uint8(x_75, sizeof(void*)*2, x_68); +if (lean_is_scalar(x_67)) { + x_76 = lean_alloc_ctor(0, 7, 0); +} else { + x_76 = x_67; +} +lean_ctor_set(x_76, 0, x_61); +lean_ctor_set(x_76, 1, x_62); +lean_ctor_set(x_76, 2, x_63); +lean_ctor_set(x_76, 3, x_64); +lean_ctor_set(x_76, 4, x_65); +lean_ctor_set(x_76, 5, x_66); +lean_ctor_set(x_76, 6, x_75); +x_77 = lean_st_ref_set(x_9, x_76, x_60); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_79 = x_77; } else { - lean_dec_ref(x_92); - x_94 = lean_box(0); + lean_dec_ref(x_77); + x_79 = lean_box(0); } -if (lean_is_scalar(x_94)) { - x_95 = lean_alloc_ctor(0, 2, 0); +x_80 = lean_box(0); +if (lean_is_scalar(x_79)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_95 = x_94; + x_81 = x_79; } -lean_ctor_set(x_95, 0, x_24); -lean_ctor_set(x_95, 1, x_93); -return x_95; +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +return x_81; } } else { -lean_object* x_96; uint8_t x_97; -x_96 = lean_ctor_get(x_30, 1); -lean_inc(x_96); -lean_dec(x_30); -x_97 = !lean_is_exclusive(x_31); -if (x_97 == 0) +lean_object* x_82; uint8_t x_83; +x_82 = lean_ctor_get(x_11, 1); +lean_inc(x_82); +lean_dec(x_11); +x_83 = !lean_is_exclusive(x_12); +if (x_83 == 0) { -lean_object* x_98; uint8_t x_99; -x_98 = lean_ctor_get(x_31, 6); -lean_dec(x_98); -x_99 = !lean_is_exclusive(x_32); -if (x_99 == 0) +lean_object* x_84; uint8_t x_85; +x_84 = lean_ctor_get(x_12, 6); +lean_dec(x_84); +x_85 = !lean_is_exclusive(x_13); +if (x_85 == 0) { -lean_object* x_100; uint8_t x_101; -x_100 = lean_ctor_get(x_32, 1); -lean_dec(x_100); -x_101 = !lean_is_exclusive(x_28); -if (x_101 == 0) +lean_object* x_86; uint8_t x_87; +x_86 = lean_ctor_get(x_13, 1); +lean_dec(x_86); +x_87 = !lean_is_exclusive(x_2); +if (x_87 == 0) { -lean_object* x_102; lean_object* x_103; uint8_t x_104; -lean_ctor_set_tag(x_28, 2); -x_102 = l_Lean_PersistentArray_push___rarg(x_20, x_28); -lean_ctor_set(x_32, 1, x_102); -x_103 = lean_st_ref_set(x_9, x_31, x_96); -lean_dec(x_9); -x_104 = !lean_is_exclusive(x_103); -if (x_104 == 0) +lean_object* x_88; lean_object* x_89; uint8_t x_90; +lean_ctor_set_tag(x_2, 2); +x_88 = l_Lean_PersistentArray_push___rarg(x_1, x_2); +lean_ctor_set(x_13, 1, x_88); +x_89 = lean_st_ref_set(x_9, x_12, x_82); +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) { -lean_object* x_105; -x_105 = lean_ctor_get(x_103, 0); -lean_dec(x_105); -lean_ctor_set(x_103, 0, x_24); -return x_103; +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_89, 0); +lean_dec(x_91); +x_92 = lean_box(0); +lean_ctor_set(x_89, 0, x_92); +return x_89; } else { -lean_object* x_106; lean_object* x_107; -x_106 = lean_ctor_get(x_103, 1); -lean_inc(x_106); -lean_dec(x_103); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_24); -lean_ctor_set(x_107, 1, x_106); -return x_107; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_89, 1); +lean_inc(x_93); +lean_dec(x_89); +x_94 = lean_box(0); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_108 = lean_ctor_get(x_28, 0); -lean_inc(x_108); -lean_dec(x_28); -x_109 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_109, 0, x_108); -x_110 = l_Lean_PersistentArray_push___rarg(x_20, x_109); -lean_ctor_set(x_32, 1, x_110); -x_111 = lean_st_ref_set(x_9, x_31, x_96); -lean_dec(x_9); +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_96 = lean_ctor_get(x_2, 0); +lean_inc(x_96); +lean_dec(x_2); +x_97 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_98 = l_Lean_PersistentArray_push___rarg(x_1, x_97); +lean_ctor_set(x_13, 1, x_98); +x_99 = lean_st_ref_set(x_9, x_12, x_82); +x_100 = lean_ctor_get(x_99, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_101 = x_99; +} else { + lean_dec_ref(x_99); + x_101 = lean_box(0); +} +x_102 = lean_box(0); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); +} else { + x_103 = x_101; +} +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +return x_103; +} +} +else +{ +uint8_t x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_104 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +x_105 = lean_ctor_get(x_13, 0); +lean_inc(x_105); +lean_dec(x_13); +x_106 = lean_ctor_get(x_2, 0); +lean_inc(x_106); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_107 = x_2; +} else { + lean_dec_ref(x_2); + x_107 = lean_box(0); +} +if (lean_is_scalar(x_107)) { + x_108 = lean_alloc_ctor(2, 1, 0); +} else { + x_108 = x_107; + lean_ctor_set_tag(x_108, 2); +} +lean_ctor_set(x_108, 0, x_106); +x_109 = l_Lean_PersistentArray_push___rarg(x_1, x_108); +x_110 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_110, 0, x_105); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set_uint8(x_110, sizeof(void*)*2, x_104); +lean_ctor_set(x_12, 6, x_110); +x_111 = lean_st_ref_set(x_9, x_12, x_82); x_112 = lean_ctor_get(x_111, 1); lean_inc(x_112); if (lean_is_exclusive(x_111)) { @@ -28112,334 +28110,315 @@ if (lean_is_exclusive(x_111)) { lean_dec_ref(x_111); x_113 = lean_box(0); } +x_114 = lean_box(0); if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(0, 2, 0); + x_115 = lean_alloc_ctor(0, 2, 0); } else { - x_114 = x_113; + x_115 = x_113; } -lean_ctor_set(x_114, 0, x_24); -lean_ctor_set(x_114, 1, x_112); -return x_114; +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_112); +return x_115; } } else { -uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_115 = lean_ctor_get_uint8(x_32, sizeof(void*)*2); -x_116 = lean_ctor_get(x_32, 0); -lean_inc(x_116); -lean_dec(x_32); -x_117 = lean_ctor_get(x_28, 0); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_116 = lean_ctor_get(x_12, 0); +x_117 = lean_ctor_get(x_12, 1); +x_118 = lean_ctor_get(x_12, 2); +x_119 = lean_ctor_get(x_12, 3); +x_120 = lean_ctor_get(x_12, 4); +x_121 = lean_ctor_get(x_12, 5); +lean_inc(x_121); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); lean_inc(x_117); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - x_118 = x_28; -} else { - lean_dec_ref(x_28); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(2, 1, 0); -} else { - x_119 = x_118; - lean_ctor_set_tag(x_119, 2); -} -lean_ctor_set(x_119, 0, x_117); -x_120 = l_Lean_PersistentArray_push___rarg(x_20, x_119); -x_121 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_121, 0, x_116); -lean_ctor_set(x_121, 1, x_120); -lean_ctor_set_uint8(x_121, sizeof(void*)*2, x_115); -lean_ctor_set(x_31, 6, x_121); -x_122 = lean_st_ref_set(x_9, x_31, x_96); -lean_dec(x_9); -x_123 = lean_ctor_get(x_122, 1); +lean_inc(x_116); +lean_dec(x_12); +x_122 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +x_123 = lean_ctor_get(x_13, 0); lean_inc(x_123); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_124 = x_122; +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_124 = x_13; } else { - lean_dec_ref(x_122); + lean_dec_ref(x_13); x_124 = lean_box(0); } -if (lean_is_scalar(x_124)) { - x_125 = lean_alloc_ctor(0, 2, 0); -} else { - x_125 = x_124; -} -lean_ctor_set(x_125, 0, x_24); -lean_ctor_set(x_125, 1, x_123); -return x_125; -} -} -else -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_126 = lean_ctor_get(x_31, 0); -x_127 = lean_ctor_get(x_31, 1); -x_128 = lean_ctor_get(x_31, 2); -x_129 = lean_ctor_get(x_31, 3); -x_130 = lean_ctor_get(x_31, 4); -x_131 = lean_ctor_get(x_31, 5); -lean_inc(x_131); -lean_inc(x_130); -lean_inc(x_129); -lean_inc(x_128); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_31); -x_132 = lean_ctor_get_uint8(x_32, sizeof(void*)*2); -x_133 = lean_ctor_get(x_32, 0); -lean_inc(x_133); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_134 = x_32; +x_125 = lean_ctor_get(x_2, 0); +lean_inc(x_125); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_126 = x_2; } else { - lean_dec_ref(x_32); - x_134 = lean_box(0); + lean_dec_ref(x_2); + x_126 = lean_box(0); } -x_135 = lean_ctor_get(x_28, 0); -lean_inc(x_135); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - x_136 = x_28; +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(2, 1, 0); } else { - lean_dec_ref(x_28); - x_136 = lean_box(0); + x_127 = x_126; + lean_ctor_set_tag(x_127, 2); } -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_127, 0, x_125); +x_128 = l_Lean_PersistentArray_push___rarg(x_1, x_127); +if (lean_is_scalar(x_124)) { + x_129 = lean_alloc_ctor(0, 2, 1); +} else { + x_129 = x_124; +} +lean_ctor_set(x_129, 0, x_123); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set_uint8(x_129, sizeof(void*)*2, x_122); +x_130 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_130, 0, x_116); +lean_ctor_set(x_130, 1, x_117); +lean_ctor_set(x_130, 2, x_118); +lean_ctor_set(x_130, 3, x_119); +lean_ctor_set(x_130, 4, x_120); +lean_ctor_set(x_130, 5, x_121); +lean_ctor_set(x_130, 6, x_129); +x_131 = lean_st_ref_set(x_9, x_130, x_82); +x_132 = lean_ctor_get(x_131, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_133 = x_131; } else { - x_137 = x_136; - lean_ctor_set_tag(x_137, 2); + lean_dec_ref(x_131); + x_133 = lean_box(0); } -lean_ctor_set(x_137, 0, x_135); -x_138 = l_Lean_PersistentArray_push___rarg(x_20, x_137); -if (lean_is_scalar(x_134)) { - x_139 = lean_alloc_ctor(0, 2, 1); +x_134 = lean_box(0); +if (lean_is_scalar(x_133)) { + x_135 = lean_alloc_ctor(0, 2, 0); } else { - x_139 = x_134; + x_135 = x_133; } -lean_ctor_set(x_139, 0, x_133); -lean_ctor_set(x_139, 1, x_138); -lean_ctor_set_uint8(x_139, sizeof(void*)*2, x_132); -x_140 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_140, 0, x_126); -lean_ctor_set(x_140, 1, x_127); -lean_ctor_set(x_140, 2, x_128); -lean_ctor_set(x_140, 3, x_129); -lean_ctor_set(x_140, 4, x_130); -lean_ctor_set(x_140, 5, x_131); -lean_ctor_set(x_140, 6, x_139); -x_141 = lean_st_ref_set(x_9, x_140, x_96); -lean_dec(x_9); -x_142 = lean_ctor_get(x_141, 1); -lean_inc(x_142); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_143 = x_141; -} else { - lean_dec_ref(x_141); - x_143 = lean_box(0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_132); +return x_135; } -if (lean_is_scalar(x_143)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_143; } -lean_ctor_set(x_144, 0, x_24); -lean_ctor_set(x_144, 1, x_142); -return x_144; } } +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_st_ref_get(x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_13, 6); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get_uint8(x_14, sizeof(void*)*2); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_box(x_4); +x_18 = lean_apply_8(x_1, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_18; } else { -uint8_t x_145; -lean_dec(x_24); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg(x_10, x_19); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); lean_dec(x_20); +x_23 = lean_box(x_4); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_24 = lean_apply_8(x_1, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_3); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_box(x_4); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_25); +x_28 = lean_apply_9(x_2, x_25, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_26); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1(x_21, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_10); lean_dec(x_9); -x_145 = !lean_is_exclusive(x_27); -if (x_145 == 0) +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -return x_27; +lean_object* x_33; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +lean_ctor_set(x_31, 0, x_25); +return x_31; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_27, 0); -x_147 = lean_ctor_get(x_27, 1); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_27); -x_148 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_148, 0, x_146); -lean_ctor_set(x_148, 1, x_147); -return x_148; -} +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_25); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +uint8_t x_36; +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_149 = lean_ctor_get(x_23, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_23, 1); -lean_inc(x_150); -lean_dec(x_23); -x_151 = lean_st_ref_take(x_9, x_150); -x_152 = lean_ctor_get(x_151, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_152, 6); -lean_inc(x_153); -x_154 = lean_ctor_get(x_151, 1); -lean_inc(x_154); -lean_dec(x_151); -x_155 = !lean_is_exclusive(x_152); -if (x_155 == 0) -{ -lean_object* x_156; uint8_t x_157; -x_156 = lean_ctor_get(x_152, 6); -lean_dec(x_156); -x_157 = !lean_is_exclusive(x_153); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; uint8_t x_160; -x_158 = lean_ctor_get(x_153, 1); -lean_dec(x_158); -lean_ctor_set(x_153, 1, x_20); -x_159 = lean_st_ref_set(x_9, x_152, x_154); -lean_dec(x_9); -x_160 = !lean_is_exclusive(x_159); -if (x_160 == 0) +x_36 = !lean_is_exclusive(x_28); +if (x_36 == 0) { -lean_object* x_161; -x_161 = lean_ctor_get(x_159, 0); -lean_dec(x_161); -lean_ctor_set_tag(x_159, 1); -lean_ctor_set(x_159, 0, x_149); -return x_159; +return x_28; } else { -lean_object* x_162; lean_object* x_163; -x_162 = lean_ctor_get(x_159, 1); -lean_inc(x_162); -lean_dec(x_159); -x_163 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_163, 0, x_149); -lean_ctor_set(x_163, 1, x_162); -return x_163; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_28, 0); +x_38 = lean_ctor_get(x_28, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_28); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} } } else { -uint8_t x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_164 = lean_ctor_get_uint8(x_153, sizeof(void*)*2); -x_165 = lean_ctor_get(x_153, 0); -lean_inc(x_165); -lean_dec(x_153); -x_166 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_20); -lean_ctor_set_uint8(x_166, sizeof(void*)*2, x_164); -lean_ctor_set(x_152, 6, x_166); -x_167 = lean_st_ref_set(x_9, x_152, x_154); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_2); +x_40 = lean_ctor_get(x_24, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_24, 1); +lean_inc(x_41); +lean_dec(x_24); +x_42 = lean_box(x_4); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_43 = lean_apply_8(x_3, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_41); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_46, 0, x_44); +x_47 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1(x_21, x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_45); +lean_dec(x_10); lean_dec(x_9); -x_168 = lean_ctor_get(x_167, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_167)) { - lean_ctor_release(x_167, 0); - lean_ctor_release(x_167, 1); - x_169 = x_167; -} else { - lean_dec_ref(x_167); - x_169 = lean_box(0); -} -if (lean_is_scalar(x_169)) { - x_170 = lean_alloc_ctor(1, 2, 0); -} else { - x_170 = x_169; - lean_ctor_set_tag(x_170, 1); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) +{ +lean_object* x_49; +x_49 = lean_ctor_get(x_47, 0); +lean_dec(x_49); +lean_ctor_set_tag(x_47, 1); +lean_ctor_set(x_47, 0, x_40); +return x_47; } -lean_ctor_set(x_170, 0, x_149); -lean_ctor_set(x_170, 1, x_168); -return x_170; +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); +lean_dec(x_47); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_40); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; uint8_t x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_171 = lean_ctor_get(x_152, 0); -x_172 = lean_ctor_get(x_152, 1); -x_173 = lean_ctor_get(x_152, 2); -x_174 = lean_ctor_get(x_152, 3); -x_175 = lean_ctor_get(x_152, 4); -x_176 = lean_ctor_get(x_152, 5); -lean_inc(x_176); -lean_inc(x_175); -lean_inc(x_174); -lean_inc(x_173); -lean_inc(x_172); -lean_inc(x_171); -lean_dec(x_152); -x_177 = lean_ctor_get_uint8(x_153, sizeof(void*)*2); -x_178 = lean_ctor_get(x_153, 0); -lean_inc(x_178); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_179 = x_153; -} else { - lean_dec_ref(x_153); - x_179 = lean_box(0); -} -if (lean_is_scalar(x_179)) { - x_180 = lean_alloc_ctor(0, 2, 1); -} else { - x_180 = x_179; -} -lean_ctor_set(x_180, 0, x_178); -lean_ctor_set(x_180, 1, x_20); -lean_ctor_set_uint8(x_180, sizeof(void*)*2, x_177); -x_181 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_181, 0, x_171); -lean_ctor_set(x_181, 1, x_172); -lean_ctor_set(x_181, 2, x_173); -lean_ctor_set(x_181, 3, x_174); -lean_ctor_set(x_181, 4, x_175); -lean_ctor_set(x_181, 5, x_176); -lean_ctor_set(x_181, 6, x_180); -x_182 = lean_st_ref_set(x_9, x_181, x_154); -lean_dec(x_9); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_184 = x_182; -} else { - lean_dec_ref(x_182); - x_184 = lean_box(0); +uint8_t x_52; +lean_dec(x_40); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_52 = !lean_is_exclusive(x_43); +if (x_52 == 0) +{ +return x_43; } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); -} else { - x_185 = x_184; - lean_ctor_set_tag(x_185, 1); +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_43, 0); +x_54 = lean_ctor_get(x_43, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_43); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } -lean_ctor_set(x_185, 0, x_149); -lean_ctor_set(x_185, 1, x_183); -return x_185; } } } @@ -28604,6 +28583,16 @@ return x_21; } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +lean_inc(x_2); +x_12 = l_Lean_Elab_Term_mkPartialTermInfo(x_11, x_1, x_2, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -29021,7 +29010,7 @@ return x_95; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_dec(x_81); lean_dec(x_80); x_96 = lean_ctor_get(x_83, 0); @@ -29034,186 +29023,191 @@ lean_inc(x_98); lean_dec(x_96); x_99 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___boxed), 9, 1); lean_closure_set(x_99, 0, x_98); +lean_inc(x_97); x_100 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__3___boxed), 10, 1); lean_closure_set(x_100, 0, x_97); -x_101 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_99, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_101; +x_101 = lean_box(0); +x_102 = lean_alloc_closure((void*)(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4___boxed), 10, 2); +lean_closure_set(x_102, 0, x_97); +lean_closure_set(x_102, 1, x_101); +x_103 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_99, x_100, x_102, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_103; } } else { -uint8_t x_102; lean_object* x_103; +uint8_t x_104; lean_object* x_105; lean_dec(x_82); lean_dec(x_1); -x_102 = 1; -x_103 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_81, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_103) == 0) +x_104 = 1; +x_105 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_81, x_104, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_105) == 0) { -uint8_t x_104; -x_104 = !lean_is_exclusive(x_103); -if (x_104 == 0) +uint8_t x_106; +x_106 = !lean_is_exclusive(x_105); +if (x_106 == 0) { -lean_object* x_105; lean_object* x_106; -x_105 = lean_ctor_get(x_103, 0); -x_106 = l_Lean_Expr_mdata___override(x_80, x_105); -lean_ctor_set(x_103, 0, x_106); -return x_103; +lean_object* x_107; lean_object* x_108; +x_107 = lean_ctor_get(x_105, 0); +x_108 = l_Lean_Expr_mdata___override(x_80, x_107); +lean_ctor_set(x_105, 0, x_108); +return x_105; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_107 = lean_ctor_get(x_103, 0); -x_108 = lean_ctor_get(x_103, 1); -lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_103); -x_109 = l_Lean_Expr_mdata___override(x_80, x_107); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_108); -return x_110; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_109 = lean_ctor_get(x_105, 0); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_105); +x_111 = l_Lean_Expr_mdata___override(x_80, x_109); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +return x_112; } } else { -uint8_t x_111; +uint8_t x_113; lean_dec(x_80); -x_111 = !lean_is_exclusive(x_103); -if (x_111 == 0) +x_113 = !lean_is_exclusive(x_105); +if (x_113 == 0) { -return x_103; +return x_105; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_112 = lean_ctor_get(x_103, 0); -x_113 = lean_ctor_get(x_103, 1); -lean_inc(x_113); -lean_inc(x_112); -lean_dec(x_103); -x_114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_114, 0, x_112); -lean_ctor_set(x_114, 1, x_113); -return x_114; +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_105, 0); +x_115 = lean_ctor_get(x_105, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_105); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } } case 11: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = lean_ctor_get(x_1, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_1, 1); -lean_inc(x_116); -x_117 = lean_ctor_get(x_1, 2); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_117 = lean_ctor_get(x_1, 0); lean_inc(x_117); -lean_inc(x_117); -x_118 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_117, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_118) == 0) +x_118 = lean_ctor_get(x_1, 1); +lean_inc(x_118); +x_119 = lean_ctor_get(x_1, 2); +lean_inc(x_119); +lean_inc(x_119); +x_120 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(x_119, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_120) == 0) { -uint8_t x_119; -x_119 = !lean_is_exclusive(x_118); -if (x_119 == 0) +uint8_t x_121; +x_121 = !lean_is_exclusive(x_120); +if (x_121 == 0) { -lean_object* x_120; size_t x_121; size_t x_122; uint8_t x_123; -x_120 = lean_ctor_get(x_118, 0); -x_121 = lean_ptr_addr(x_117); -lean_dec(x_117); -x_122 = lean_ptr_addr(x_120); -x_123 = lean_usize_dec_eq(x_121, x_122); -if (x_123 == 0) +lean_object* x_122; size_t x_123; size_t x_124; uint8_t x_125; +x_122 = lean_ctor_get(x_120, 0); +x_123 = lean_ptr_addr(x_119); +lean_dec(x_119); +x_124 = lean_ptr_addr(x_122); +x_125 = lean_usize_dec_eq(x_123, x_124); +if (x_125 == 0) { -lean_object* x_124; +lean_object* x_126; lean_dec(x_1); -x_124 = l_Lean_Expr_proj___override(x_115, x_116, x_120); -lean_ctor_set(x_118, 0, x_124); -return x_118; +x_126 = l_Lean_Expr_proj___override(x_117, x_118, x_122); +lean_ctor_set(x_120, 0, x_126); +return x_120; } else { -lean_dec(x_120); -lean_dec(x_116); -lean_dec(x_115); -lean_ctor_set(x_118, 0, x_1); -return x_118; +lean_dec(x_122); +lean_dec(x_118); +lean_dec(x_117); +lean_ctor_set(x_120, 0, x_1); +return x_120; } } else { -lean_object* x_125; lean_object* x_126; size_t x_127; size_t x_128; uint8_t x_129; -x_125 = lean_ctor_get(x_118, 0); -x_126 = lean_ctor_get(x_118, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_118); -x_127 = lean_ptr_addr(x_117); -lean_dec(x_117); -x_128 = lean_ptr_addr(x_125); -x_129 = lean_usize_dec_eq(x_127, x_128); -if (x_129 == 0) +lean_object* x_127; lean_object* x_128; size_t x_129; size_t x_130; uint8_t x_131; +x_127 = lean_ctor_get(x_120, 0); +x_128 = lean_ctor_get(x_120, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_120); +x_129 = lean_ptr_addr(x_119); +lean_dec(x_119); +x_130 = lean_ptr_addr(x_127); +x_131 = lean_usize_dec_eq(x_129, x_130); +if (x_131 == 0) { -lean_object* x_130; lean_object* x_131; +lean_object* x_132; lean_object* x_133; lean_dec(x_1); -x_130 = l_Lean_Expr_proj___override(x_115, x_116, x_125); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_126); -return x_131; +x_132 = l_Lean_Expr_proj___override(x_117, x_118, x_127); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_128); +return x_133; } else { -lean_object* x_132; -lean_dec(x_125); -lean_dec(x_116); -lean_dec(x_115); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_1); -lean_ctor_set(x_132, 1, x_126); -return x_132; +lean_object* x_134; +lean_dec(x_127); +lean_dec(x_118); +lean_dec(x_117); +x_134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_1); +lean_ctor_set(x_134, 1, x_128); +return x_134; } } } else { -uint8_t x_133; +uint8_t x_135; +lean_dec(x_119); +lean_dec(x_118); lean_dec(x_117); -lean_dec(x_116); -lean_dec(x_115); lean_dec(x_1); -x_133 = !lean_is_exclusive(x_118); -if (x_133 == 0) +x_135 = !lean_is_exclusive(x_120); +if (x_135 == 0) { -return x_118; +return x_120; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_118, 0); -x_135 = lean_ctor_get(x_118, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_118); -x_136 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_136, 0, x_134); -lean_ctor_set(x_136, 1, x_135); -return x_136; +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_120, 0); +x_137 = lean_ctor_get(x_120, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_120); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); +return x_138; } } } default: { -lean_object* x_137; +lean_object* x_139; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_137 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_137, 0, x_1); -lean_ctor_set(x_137, 1, x_9); -return x_137; +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_1); +lean_ctor_set(x_139, 1, x_9); +return x_139; } } } @@ -29278,16 +29272,32 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_3); lean_dec(x_3); -x_12 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__3(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -29326,6 +29336,22 @@ lean_dec(x_4); return x_12; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__4(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -36569,7 +36595,7 @@ x_21 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_18); -x_22 = lean_alloc_ctor(9, 1, 0); +x_22 = lean_alloc_ctor(10, 1, 0); lean_ctor_set(x_22, 0, x_21); x_23 = l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2(x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_13); x_24 = lean_ctor_get(x_23, 1); @@ -43719,7 +43745,7 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1() { _start: { lean_object* x_1; @@ -43727,17 +43753,17 @@ x_1 = lean_mk_string_unchecked("ignoreUnusedAlts", 16, 16); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3() { _start: { lean_object* x_1; @@ -43745,13 +43771,13 @@ x_1 = lean_mk_string_unchecked("if true, do not generate error if an alternative return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__14; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -43760,7 +43786,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -43768,18 +43794,18 @@ x_1 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__ x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__3; x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__6; -x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1; +x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -52320,7 +52346,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -52330,27 +52356,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2; x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4() { _start: { lean_object* x_1; @@ -52358,17 +52384,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6() { _start: { lean_object* x_1; @@ -52376,37 +52402,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7; x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10() { _start: { lean_object* x_1; @@ -52414,17 +52440,17 @@ x_1 = lean_mk_string_unchecked("Match", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12() { _start: { lean_object* x_1; @@ -52432,33 +52458,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13; -x_2 = lean_unsigned_to_nat(18828u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13; +x_2 = lean_unsigned_to_nat(18831u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -55277,17 +55303,17 @@ l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__1 lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__1); l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2 = _init_l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2(); lean_mark_persistent(l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962____closed__5); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14962_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965____closed__5); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_14965_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_match_ignoreUnusedAlts = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_match_ignoreUnusedAlts); @@ -55400,35 +55426,35 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabMatch_declRange__1___clos if (builtin) {res = l___regBuiltin_Lean_Elab_Term_elabMatch_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828____closed__14); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18828_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831____closed__14); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18831_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoMatch___spec__3___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoMatch___spec__3___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 105b8fc5357a..395adc2bf1f3 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -18,14 +18,13 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____clo static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_instantiateMVarsProfiling___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___closed__1; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabMutualDef___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,7 +42,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_M LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420_(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__10; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -51,6 +50,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getAl LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_LocalContext_getFVars___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -73,7 +73,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_regis lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__3; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___closed__2; static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__3; @@ -113,11 +112,11 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__4; extern lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__8___boxed(lean_object**); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1; lean_object* l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Language_DynamicSnapshot_toTyped_x3f___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7(lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -126,6 +125,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__9___closed__2; lean_object* lean_io_promise_new(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); @@ -188,7 +188,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMu LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__7___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -197,7 +196,6 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_deprecated_oldSectionVars; @@ -211,6 +209,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_Mutua LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -226,6 +225,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_typeH LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__5; lean_object* l_Lean_Elab_shareCommonPreDefs(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3; lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -242,7 +243,8 @@ static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_M LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__12(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__15; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -260,6 +262,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDe LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendingMVarErrorMessage___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Command_elabMutualDef___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__12___lambda__1___boxed(lean_object**); @@ -309,6 +312,7 @@ lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_ uint8_t l_List_any___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); @@ -352,7 +356,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_instantiateMVarsProfiling___lambda__1(lean_ LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Elab_Command_elabMutualDef___spec__10___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5; lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; @@ -362,6 +365,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check lean_object* l_Lean_LocalDecl_index(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1; lean_object* l_Lean_Elab_Term_expandLetEqnsDecl(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -490,7 +494,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1724_(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_collectUsed___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -524,13 +527,13 @@ lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(l static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__6; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___closed__4; lean_object* l_Lean_Elab_Term_applyAttributesAt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__2; +lean_object* l_Lean_SourceInfo_getTailPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_MessageData_ofFormat(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); @@ -552,6 +555,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_E LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesForBuiltin___at_Lean_Elab_Term_elabMutualDef_go___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at_Lean_Elab_Command_elabMutualDef___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_indexOfAux___at_Lean_LocalContext_erase___spec__3(lean_object*, lean_object*, lean_object*); @@ -562,6 +566,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_El static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__14; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__18; @@ -591,13 +596,11 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0_ LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabMutualDef___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__1; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_MutualClosure_main___spec__7(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__4; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_MutualClosure_main___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -605,9 +608,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_insertReplacementForMain lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5; lean_object* l_Lean_Elab_instantiateMVarsAtPreDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__50(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__4___closed__4; @@ -626,13 +631,14 @@ static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_M LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, size_t, lean_object*); extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__1___boxed(lean_object**); extern lean_object* l_Lean_instInhabitedExpr; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4; static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__5; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -643,8 +649,11 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModif static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__4; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914_(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__8; +static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6; uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); uint8_t l_Lean_Elab_DefView_isInstance(lean_object*); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -667,12 +676,12 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_Mut LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__10___boxed(lean_object**); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__6; @@ -687,6 +696,7 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____clo LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendingMVarErrorMessage___spec__1___lambda__1___boxed(lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_utf8_prev(lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -776,7 +786,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Mutual static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__3___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__1; lean_object* l_Lean_MessageData_ofExpr(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2; lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -786,7 +795,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0_ static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___closed__2; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__11; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8; extern lean_object* l_Lean_Elab_TerminationHints_none; static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__10___closed__6; LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -806,6 +814,7 @@ lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__2(lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SourceInfo_getPos_x3f(lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -816,24 +825,24 @@ LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_E LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesForBuiltin___at_Lean_Elab_Term_elabMutualDef_go___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6; static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__2; lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_TerminationHints_rememberExtraParams(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Term_type; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6; static lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3___closed__4; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); +lean_object* l_Lean_Syntax_getTailInfo_x3f(lean_object*); uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addAutoBoundImplicits(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -896,9 +905,9 @@ static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDe LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__6(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_pushLetRecs(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -978,12 +987,13 @@ uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__12___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1; static size_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__3; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__4___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -1006,11 +1016,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_isApplicable___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__18___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2___closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__4___closed__3; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1021,7 +1030,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPe lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__15___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1100,6 +1108,7 @@ lean_object* l_Lean_Elab_Term_withLevelNames___rarg(lean_object*, lean_object*, static lean_object* l_List_mapTR_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendingMVarErrorMessage___spec__1___closed__1; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__10___closed__4; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2___closed__2; @@ -1186,7 +1195,6 @@ lean_object* lean_find_expr(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__11___boxed(lean_object**); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__10; @@ -1201,6 +1209,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMul lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_withAlwaysResolvedPromises___at_Lean_Elab_Term_elabMutualDef_go___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1; static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Term_elabMutualDef_go___spec__10___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__4___boxed(lean_object**); static lean_object* l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__3; @@ -1234,7 +1243,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_E static lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__6___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_getFVars___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_getUsedFVarsMap___rarg(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5; uint8_t l_Lean_Elab_Modifiers_isPartial(lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__11___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1260,6 +1268,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabTerminationHints___at___private_Lean_El static lean_object* l_Lean_Elab_instInhabitedDefViewElabHeader___closed__5; size_t lean_usize_land(size_t, size_t); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__8; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -15639,215 +15648,425 @@ x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_8; size_t x_9; lean_object* x_10; -x_8 = lean_box(0); -x_9 = lean_array_size(x_1); -lean_inc(x_6); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_9, x_2, x_1, x_6, x_7); -if (lean_obj_tag(x_10) == 0) +lean_object* x_11; size_t x_12; lean_object* x_13; +x_11 = lean_box(0); +x_12 = lean_array_size(x_1); +lean_inc(x_9); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__2(x_12, x_2, x_1, x_9, x_10); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_11; lean_object* x_12; size_t x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_array_size(x_11); -lean_inc(x_6); +lean_object* x_14; lean_object* x_15; lean_object* x_16; size_t x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_16 = x_13; +} else { + lean_dec_ref(x_13); + x_16 = lean_box(0); +} +x_17 = lean_array_size(x_14); +lean_inc(x_9); lean_inc(x_3); -x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(x_3, x_2, x_8, x_13, x_2, x_11, x_6, x_12); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(x_3, x_2, x_11, x_17, x_2, x_14, x_9, x_15); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); -x_18 = lean_ctor_get(x_6, 5); -lean_inc(x_18); -x_19 = 0; -x_20 = l_Lean_SourceInfo_fromRef(x_18, x_19); -lean_dec(x_18); -x_21 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__3; +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_61; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_22 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__10; +if (lean_is_exclusive(x_18)) { + lean_ctor_release(x_18, 0); + lean_ctor_release(x_18, 1); + x_21 = x_18; +} else { + lean_dec_ref(x_18); + x_21 = lean_box(0); +} +x_22 = 0; +x_23 = l_Lean_Syntax_getPos_x3f(x_4, x_22); +x_24 = l_Lean_Syntax_getTailInfo_x3f(x_5); +x_25 = lean_ctor_get(x_9, 5); +lean_inc(x_25); +x_26 = l_Lean_SourceInfo_fromRef(x_25, x_22); +lean_dec(x_25); +x_27 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__3; +lean_inc(x_26); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__10; lean_inc(x_3); -lean_inc(x_20); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_24, 2, x_3); -x_25 = lean_array_size(x_16); -x_26 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__7; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5(x_26, x_25, x_2, x_16); -x_28 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__8; -x_29 = l_Lean_Syntax_SepArray_ofElems(x_28, x_27); -lean_dec(x_27); -x_30 = l_Array_append___rarg(x_3, x_29); -lean_dec(x_29); -lean_inc(x_20); -x_31 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_31, 0, x_20); -lean_ctor_set(x_31, 1, x_23); -lean_ctor_set(x_31, 2, x_30); -x_32 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__10; -lean_inc(x_24); -lean_inc(x_20); -x_33 = l_Lean_Syntax_node1(x_20, x_32, x_24); -x_34 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__11; -lean_inc(x_20); -x_35 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_35, 0, x_20); -lean_ctor_set(x_35, 1, x_34); -x_36 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__2; -lean_inc(x_24); -x_37 = l_Lean_Syntax_node6(x_20, x_36, x_22, x_24, x_31, x_33, x_24, x_35); -if (lean_obj_tag(x_5) == 0) +lean_inc(x_26); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_26); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_3); +x_31 = lean_array_size(x_19); +x_32 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__7; +x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5(x_32, x_31, x_2, x_19); +x_34 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__8; +x_35 = l_Lean_Syntax_SepArray_ofElems(x_34, x_33); +lean_dec(x_33); +x_36 = l_Array_append___rarg(x_3, x_35); +lean_dec(x_35); +lean_inc(x_26); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_26); +lean_ctor_set(x_37, 1, x_29); +lean_ctor_set(x_37, 2, x_36); +lean_inc(x_26); +x_38 = l_Lean_Syntax_node1(x_26, x_6, x_37); +x_39 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__10; +lean_inc(x_30); +lean_inc(x_26); +x_40 = l_Lean_Syntax_node1(x_26, x_39, x_30); +x_41 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__11; +lean_inc(x_26); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_26); +lean_ctor_set(x_42, 1, x_41); +x_43 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__2; +lean_inc(x_30); +x_44 = l_Lean_Syntax_node6(x_26, x_43, x_28, x_30, x_38, x_40, x_30, x_42); +if (lean_obj_tag(x_23) == 0) { -lean_dec(x_6); -lean_ctor_set(x_14, 0, x_37); -return x_14; +lean_object* x_95; +x_95 = lean_box(2); +x_61 = x_95; +goto block_94; } else { -lean_object* x_38; lean_object* x_39; -lean_free_object(x_14); -x_38 = lean_ctor_get(x_5, 0); -lean_inc(x_38); -lean_dec(x_5); -x_39 = l_Lean_Elab_Term_expandWhereDecls(x_38, x_37, x_6, x_17); -lean_dec(x_6); -return x_39; +lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; +x_96 = lean_ctor_get(x_23, 0); +lean_inc(x_96); +lean_dec(x_23); +x_97 = lean_unsigned_to_nat(1u); +x_98 = lean_nat_add(x_96, x_97); +x_99 = 1; +x_100 = lean_alloc_ctor(1, 2, 1); +lean_ctor_set(x_100, 0, x_96); +lean_ctor_set(x_100, 1, x_98); +lean_ctor_set_uint8(x_100, sizeof(void*)*2, x_99); +x_61 = x_100; +goto block_94; +} +block_60: +{ +lean_object* x_47; +x_47 = l_Lean_SourceInfo_getTailPos_x3f(x_45, x_22); +lean_dec(x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_46); +x_48 = lean_box(2); +x_49 = l_Lean_Syntax_setInfo(x_48, x_44); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_50; +lean_dec(x_9); +if (lean_is_scalar(x_21)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_21; } +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_20); +return x_50; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_40 = lean_ctor_get(x_14, 0); -x_41 = lean_ctor_get(x_14, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_14); -x_42 = lean_ctor_get(x_6, 5); -lean_inc(x_42); -x_43 = 0; -x_44 = l_Lean_SourceInfo_fromRef(x_42, x_43); -lean_dec(x_42); -x_45 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__3; -lean_inc(x_44); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__10; -lean_inc(x_3); -lean_inc(x_44); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_44); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set(x_48, 2, x_3); -x_49 = lean_array_size(x_40); -x_50 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__7; -x_51 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5(x_50, x_49, x_2, x_40); -x_52 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__8; -x_53 = l_Lean_Syntax_SepArray_ofElems(x_52, x_51); -lean_dec(x_51); -x_54 = l_Array_append___rarg(x_3, x_53); -lean_dec(x_53); -lean_inc(x_44); -x_55 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_55, 0, x_44); -lean_ctor_set(x_55, 1, x_47); -lean_ctor_set(x_55, 2, x_54); -x_56 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__10; -lean_inc(x_48); -lean_inc(x_44); -x_57 = l_Lean_Syntax_node1(x_44, x_56, x_48); -x_58 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__11; -lean_inc(x_44); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_44); -lean_ctor_set(x_59, 1, x_58); -x_60 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__2; -lean_inc(x_48); -x_61 = l_Lean_Syntax_node6(x_44, x_60, x_46, x_48, x_55, x_57, x_48, x_59); -if (lean_obj_tag(x_5) == 0) +lean_object* x_51; lean_object* x_52; +lean_dec(x_21); +x_51 = lean_ctor_get(x_8, 0); +lean_inc(x_51); +lean_dec(x_8); +x_52 = l_Lean_Elab_Term_expandWhereDecls(x_51, x_49, x_9, x_20); +lean_dec(x_9); +return x_52; +} +} +else { -lean_object* x_62; -lean_dec(x_6); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_41); -return x_62; +lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_47, 0); +lean_inc(x_53); +lean_dec(x_47); +x_54 = 1; +x_55 = lean_alloc_ctor(1, 2, 1); +lean_ctor_set(x_55, 0, x_46); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_54); +x_56 = l_Lean_Syntax_setInfo(x_55, x_44); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_57; +lean_dec(x_9); +if (lean_is_scalar(x_21)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_21; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_20); +return x_57; } else { -lean_object* x_63; lean_object* x_64; -x_63 = lean_ctor_get(x_5, 0); -lean_inc(x_63); -lean_dec(x_5); -x_64 = l_Lean_Elab_Term_expandWhereDecls(x_63, x_61, x_6, x_41); -lean_dec(x_6); -return x_64; +lean_object* x_58; lean_object* x_59; +lean_dec(x_21); +x_58 = lean_ctor_get(x_8, 0); +lean_inc(x_58); +lean_dec(x_8); +x_59 = l_Lean_Elab_Term_expandWhereDecls(x_58, x_56, x_9, x_20); +lean_dec(x_9); +return x_59; +} } } +block_94: +{ +lean_object* x_62; lean_object* x_63; +x_62 = l_Lean_SourceInfo_getPos_x3f(x_61, x_22); +lean_dec(x_61); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_84; +x_84 = l_Lean_Syntax_getTailInfo_x3f(x_4); +if (lean_obj_tag(x_84) == 0) +{ +lean_dec(x_16); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_85; lean_object* x_86; +lean_dec(x_21); +x_85 = lean_box(2); +x_86 = l_Lean_Syntax_setInfo(x_85, x_44); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_87; +lean_dec(x_9); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_20); +return x_87; } else { -uint8_t x_65; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_65 = !lean_is_exclusive(x_14); -if (x_65 == 0) +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_8, 0); +lean_inc(x_88); +lean_dec(x_8); +x_89 = l_Lean_Elab_Term_expandWhereDecls(x_88, x_86, x_9, x_20); +lean_dec(x_9); +return x_89; +} +} +else { -return x_14; +lean_object* x_90; lean_object* x_91; +x_90 = lean_ctor_get(x_62, 0); +lean_inc(x_90); +lean_dec(x_62); +x_91 = lean_box(2); +x_45 = x_91; +x_46 = x_90; +goto block_60; +} } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_14, 0); -x_67 = lean_ctor_get(x_14, 1); +lean_object* x_92; +x_92 = lean_ctor_get(x_84, 0); +lean_inc(x_92); +lean_dec(x_84); +x_63 = x_92; +goto block_83; +} +} +else +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_24, 0); +lean_inc(x_93); +lean_dec(x_24); +x_63 = x_93; +goto block_83; +} +block_83: +{ +if (lean_obj_tag(x_63) == 0) +{ +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_63); +lean_dec(x_21); +x_64 = lean_box(2); +x_65 = l_Lean_Syntax_setInfo(x_64, x_44); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_66; +lean_dec(x_9); +if (lean_is_scalar(x_16)) { + x_66 = lean_alloc_ctor(0, 2, 0); +} else { + x_66 = x_16; +} +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_20); +return x_66; +} +else +{ +lean_object* x_67; lean_object* x_68; +lean_dec(x_16); +x_67 = lean_ctor_get(x_8, 0); lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_14); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); +lean_dec(x_8); +x_68 = l_Lean_Elab_Term_expandWhereDecls(x_67, x_65, x_9, x_20); +lean_dec(x_9); return x_68; } } +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; +lean_dec(x_16); +x_69 = lean_ctor_get(x_63, 2); +lean_inc(x_69); +lean_dec(x_63); +x_70 = lean_ctor_get(x_62, 0); +lean_inc(x_70); +lean_dec(x_62); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_69, 2); +lean_inc(x_72); +lean_dec(x_69); +x_73 = lean_string_utf8_prev(x_71, x_72); +lean_dec(x_71); +x_74 = 1; +x_75 = lean_alloc_ctor(1, 2, 1); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_72); +lean_ctor_set_uint8(x_75, sizeof(void*)*2, x_74); +x_45 = x_75; +x_46 = x_70; +goto block_60; +} } else { -uint8_t x_69; +lean_dec(x_63); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_21); +x_76 = lean_box(2); +x_77 = l_Lean_Syntax_setInfo(x_76, x_44); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_78; +lean_dec(x_9); +if (lean_is_scalar(x_16)) { + x_78 = lean_alloc_ctor(0, 2, 0); +} else { + x_78 = x_16; +} +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_20); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_16); +x_79 = lean_ctor_get(x_8, 0); +lean_inc(x_79); +lean_dec(x_8); +x_80 = l_Lean_Elab_Term_expandWhereDecls(x_79, x_77, x_9, x_20); +lean_dec(x_9); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; +lean_dec(x_16); +x_81 = lean_ctor_get(x_62, 0); +lean_inc(x_81); +lean_dec(x_62); +x_82 = lean_box(2); +x_45 = x_82; +x_46 = x_81; +goto block_60; +} +} +} +} +} +else +{ +uint8_t x_101; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); -x_69 = !lean_is_exclusive(x_10); -if (x_69 == 0) +x_101 = !lean_is_exclusive(x_18); +if (x_101 == 0) { -return x_10; +return x_18; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_10, 0); -x_71 = lean_ctor_get(x_10, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_10); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_18, 0); +x_103 = lean_ctor_get(x_18, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_18); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; +} +} +} +else +{ +uint8_t x_105; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_3); +x_105 = !lean_is_exclusive(x_13); +if (x_105 == 0) +{ +return x_13; +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_13, 0); +x_107 = lean_ctor_get(x_13, 1); +lean_inc(x_107); +lean_inc(x_106); +lean_dec(x_13); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +return x_108; } } } @@ -15875,6 +16094,26 @@ return x_5; static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__4; +x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__1; +x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; +x_4 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5() { +_start: +{ uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; x_2 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; @@ -15885,7 +16124,7 @@ lean_ctor_set(x_4, 1, x_2); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6() { _start: { lean_object* x_1; @@ -15893,14 +16132,14 @@ x_1 = lean_mk_string_unchecked("whereDecls", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5() { +static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__4; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_getBodyTerm_x3f___closed__1; x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; -x_4 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4; +x_4 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -15925,210 +16164,250 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; size_t x_14; lean_object* x_15; -x_8 = lean_unsigned_to_nat(1u); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_8 = lean_unsigned_to_nat(0u); x_9 = l_Lean_Syntax_getArg(x_1, x_8); -x_10 = l_Lean_Syntax_getArgs(x_9); -lean_dec(x_9); -x_11 = lean_array_get_size(x_10); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_nat_dec_lt(x_12, x_11); -x_14 = 0; +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4; +lean_inc(x_11); +x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); if (x_13 == 0) { -lean_object* x_61; +lean_object* x_14; lean_object* x_15; lean_dec(x_11); -lean_dec(x_10); -x_61 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_15 = x_61; -goto block_60; +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_3); +return x_15; } else { -uint8_t x_62; -x_62 = lean_nat_dec_le(x_11, x_11); -if (x_62 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; size_t x_20; lean_object* x_21; +x_16 = l_Lean_Syntax_getArg(x_11, x_8); +x_17 = l_Lean_Syntax_getArgs(x_16); +lean_dec(x_16); +x_18 = lean_array_get_size(x_17); +x_19 = lean_nat_dec_lt(x_8, x_18); +x_20 = 0; +if (x_19 == 0) { -lean_object* x_63; -lean_dec(x_11); -lean_dec(x_10); -x_63 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_15 = x_63; -goto block_60; +lean_object* x_67; +lean_dec(x_18); +lean_dec(x_17); +x_67 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_21 = x_67; +goto block_66; } else { -size_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_65 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__3; -x_66 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_10, x_14, x_64, x_65); -lean_dec(x_10); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_15 = x_67; -goto block_60; +uint8_t x_68; +x_68 = lean_nat_dec_le(x_18, x_18); +if (x_68 == 0) +{ +lean_object* x_69; +lean_dec(x_18); +lean_dec(x_17); +x_69 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_21 = x_69; +goto block_66; +} +else +{ +size_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_71 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_17, x_20, x_70, x_71); +lean_dec(x_17); +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +lean_dec(x_72); +x_21 = x_73; +goto block_66; } } -block_60: +block_66: { -size_t x_16; lean_object* x_17; -x_16 = lean_array_size(x_15); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_16, x_14, x_15); -if (lean_obj_tag(x_17) == 0) +size_t x_22; lean_object* x_23; +x_22 = lean_array_size(x_21); +x_23 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__1(x_22, x_20, x_21); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_18; lean_object* x_19; +lean_object* x_24; lean_object* x_25; +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); lean_dec(x_1); -x_18 = lean_box(1); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_3); -return x_19; +x_24 = lean_box(1); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_unsigned_to_nat(2u); -x_23 = l_Lean_Syntax_getArg(x_1, x_22); +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_23, 0); +x_28 = lean_unsigned_to_nat(2u); +x_29 = l_Lean_Syntax_getArg(x_1, x_28); lean_dec(x_1); -x_24 = l_Lean_Syntax_isNone(x_23); -if (x_24 == 0) +x_30 = l_Lean_Syntax_isNone(x_29); +if (x_30 == 0) { -uint8_t x_25; -lean_inc(x_23); -x_25 = l_Lean_Syntax_matchesNull(x_23, x_8); -if (x_25 == 0) +uint8_t x_31; +lean_inc(x_29); +x_31 = l_Lean_Syntax_matchesNull(x_29, x_10); +if (x_31 == 0) { -lean_object* x_26; lean_object* x_27; -lean_dec(x_23); -lean_free_object(x_17); -lean_dec(x_21); +lean_object* x_32; lean_object* x_33; +lean_dec(x_29); +lean_free_object(x_23); +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_26 = lean_box(1); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_3); -return x_27; +x_32 = lean_box(1); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; } else { -lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_28 = l_Lean_Syntax_getArg(x_23, x_12); -lean_dec(x_23); -x_29 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; -lean_inc(x_28); -x_30 = l_Lean_Syntax_isOfKind(x_28, x_29); -if (x_30 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = l_Lean_Syntax_getArg(x_29, x_8); +lean_dec(x_29); +x_35 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7; +lean_inc(x_34); +x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); +if (x_36 == 0) { -lean_object* x_31; lean_object* x_32; -lean_dec(x_28); -lean_free_object(x_17); -lean_dec(x_21); +lean_object* x_37; lean_object* x_38; +lean_dec(x_34); +lean_free_object(x_23); +lean_dec(x_27); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_31 = lean_box(1); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_3); -return x_32; +x_37 = lean_box(1); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_3); +return x_38; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_ctor_set(x_17, 0, x_28); -x_33 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_34 = lean_box(0); -x_35 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_21, x_14, x_33, x_34, x_17, x_2, x_3); -return x_35; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_ctor_set(x_23, 0, x_34); +x_39 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_40 = lean_box(0); +x_41 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_27, x_20, x_39, x_9, x_11, x_12, x_40, x_23, x_2, x_3); +lean_dec(x_11); +lean_dec(x_9); +return x_41; } } } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_23); -lean_free_object(x_17); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_38 = lean_box(0); -x_39 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_21, x_14, x_37, x_38, x_36, x_2, x_3); -return x_39; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_29); +lean_free_object(x_23); +x_42 = lean_box(0); +x_43 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_44 = lean_box(0); +x_45 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_27, x_20, x_43, x_9, x_11, x_12, x_44, x_42, x_2, x_3); +lean_dec(x_11); +lean_dec(x_9); +return x_45; } } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_40 = lean_ctor_get(x_17, 0); -lean_inc(x_40); -lean_dec(x_17); -x_41 = lean_unsigned_to_nat(2u); -x_42 = l_Lean_Syntax_getArg(x_1, x_41); +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_23, 0); +lean_inc(x_46); +lean_dec(x_23); +x_47 = lean_unsigned_to_nat(2u); +x_48 = l_Lean_Syntax_getArg(x_1, x_47); lean_dec(x_1); -x_43 = l_Lean_Syntax_isNone(x_42); -if (x_43 == 0) +x_49 = l_Lean_Syntax_isNone(x_48); +if (x_49 == 0) { -uint8_t x_44; -lean_inc(x_42); -x_44 = l_Lean_Syntax_matchesNull(x_42, x_8); -if (x_44 == 0) +uint8_t x_50; +lean_inc(x_48); +x_50 = l_Lean_Syntax_matchesNull(x_48, x_10); +if (x_50 == 0) { -lean_object* x_45; lean_object* x_46; -lean_dec(x_42); -lean_dec(x_40); +lean_object* x_51; lean_object* x_52; +lean_dec(x_48); +lean_dec(x_46); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_45 = lean_box(1); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_3); -return x_46; +x_51 = lean_box(1); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_3); +return x_52; } else { -lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_47 = l_Lean_Syntax_getArg(x_42, x_12); -lean_dec(x_42); -x_48 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; -lean_inc(x_47); -x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); -if (x_49 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_53 = l_Lean_Syntax_getArg(x_48, x_8); +lean_dec(x_48); +x_54 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7; +lean_inc(x_53); +x_55 = l_Lean_Syntax_isOfKind(x_53, x_54); +if (x_55 == 0) { -lean_object* x_50; lean_object* x_51; -lean_dec(x_47); -lean_dec(x_40); +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +lean_dec(x_46); +lean_dec(x_11); +lean_dec(x_9); lean_dec(x_2); -x_50 = lean_box(1); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_3); -return x_51; +x_56 = lean_box(1); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_3); +return x_57; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_47); -x_53 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_54 = lean_box(0); -x_55 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_40, x_14, x_53, x_54, x_52, x_2, x_3); -return x_55; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_53); +x_59 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_60 = lean_box(0); +x_61 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_46, x_20, x_59, x_9, x_11, x_12, x_60, x_58, x_2, x_3); +lean_dec(x_11); +lean_dec(x_9); +return x_61; } } } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_dec(x_42); -x_56 = lean_box(0); -x_57 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; -x_58 = lean_box(0); -x_59 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_40, x_14, x_57, x_58, x_56, x_2, x_3); -return x_59; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_48); +x_62 = lean_box(0); +x_63 = l_Lean_Elab_instInhabitedDefViewElabHeader___closed__1; +x_64 = lean_box(0); +x_65 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_46, x_20, x_63, x_9, x_11, x_12, x_64, x_62, x_2, x_3); +lean_dec(x_11); +lean_dec(x_9); +return x_65; +} } } } @@ -16248,15 +16527,17 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -size_t x_8; lean_object* x_9; -x_8 = lean_unbox_usize(x_2); +size_t x_11; lean_object* x_12; +x_11 = lean_unbox_usize(x_2); lean_dec(x_2); -x_9 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +x_12 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); -return x_9; +return x_12; } } static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1() { @@ -20197,7 +20478,7 @@ lean_dec(x_1); return x_12; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1() { _start: { lean_object* x_1; @@ -20205,7 +20486,7 @@ x_1 = lean_mk_string_unchecked("deprecated", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2() { _start: { lean_object* x_1; @@ -20213,17 +20494,17 @@ x_1 = lean_mk_string_unchecked("oldSectionVars", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4() { _start: { lean_object* x_1; @@ -20231,13 +20512,13 @@ x_1 = lean_mk_string_unchecked("re-enable deprecated behavior of including exact return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Lean_Elab_instantiateMVarsProfiling___closed__1; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -20246,31 +20527,31 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__4; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__6; x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1; -x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1; +x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1() { _start: { lean_object* x_1; @@ -20278,7 +20559,7 @@ x_1 = lean_mk_string_unchecked("linter", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2() { _start: { lean_object* x_1; @@ -20286,17 +20567,17 @@ x_1 = lean_mk_string_unchecked("unusedSectionVars", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4() { _start: { lean_object* x_1; @@ -20304,13 +20585,13 @@ x_1 = lean_mk_string_unchecked("enable the 'unused section variables in theorem return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; x_2 = l_Lean_Elab_instantiateMVarsProfiling___closed__1; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -20319,26 +20600,26 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__4; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__6; x_3 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders_mkTacTask___closed__1; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1; -x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1; +x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -21960,14 +22241,16 @@ return x_27; LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_Elab_Term_mkBodyInfo(x_1, x_2); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_2); +x_11 = l_Lean_Elab_Term_mkBodyInfo(x_1, x_10); +x_12 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -22136,7 +22419,7 @@ lean_dec(x_24); lean_dec(x_21); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; lean_object* x_47; uint8_t x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); @@ -22151,85 +22434,89 @@ lean_closure_set(x_30, 2, x_29); lean_inc(x_4); x_31 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__2___boxed), 9, 1); lean_closure_set(x_31, 0, x_4); -x_32 = lean_ctor_get(x_13, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_13, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_13, 2); +lean_inc(x_4); +x_32 = l_Lean_Elab_Term_mkBodyInfo(x_4, x_29); +x_33 = lean_alloc_closure((void*)(l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed), 8, 1); +lean_closure_set(x_33, 0, x_32); +x_34 = lean_ctor_get(x_13, 0); lean_inc(x_34); -x_35 = lean_ctor_get_uint8(x_13, sizeof(void*)*9); -x_36 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 1); -x_37 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 2); -x_38 = lean_ctor_get(x_13, 3); -lean_inc(x_38); -x_39 = lean_ctor_get(x_13, 4); -lean_inc(x_39); -x_40 = lean_ctor_get(x_13, 5); +x_35 = lean_ctor_get(x_13, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_13, 2); +lean_inc(x_36); +x_37 = lean_ctor_get_uint8(x_13, sizeof(void*)*9); +x_38 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 1); +x_39 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 2); +x_40 = lean_ctor_get(x_13, 3); lean_inc(x_40); -x_41 = lean_ctor_get(x_13, 6); +x_41 = lean_ctor_get(x_13, 4); lean_inc(x_41); -x_42 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 3); -x_43 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 4); -x_44 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 5); -x_45 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 6); -x_46 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 7); -x_47 = lean_ctor_get(x_13, 7); -lean_inc(x_47); -x_48 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 8); -x_49 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 9); -x_50 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_50, 0, x_32); -lean_ctor_set(x_50, 1, x_33); -lean_ctor_set(x_50, 2, x_34); -lean_ctor_set(x_50, 3, x_38); -lean_ctor_set(x_50, 4, x_39); -lean_ctor_set(x_50, 5, x_40); -lean_ctor_set(x_50, 6, x_41); -lean_ctor_set(x_50, 7, x_47); -lean_ctor_set(x_50, 8, x_5); -lean_ctor_set_uint8(x_50, sizeof(void*)*9, x_35); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 1, x_36); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 2, x_37); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 3, x_42); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 4, x_43); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 5, x_44); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 6, x_45); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 7, x_46); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 8, x_48); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 9, x_49); +x_42 = lean_ctor_get(x_13, 5); +lean_inc(x_42); +x_43 = lean_ctor_get(x_13, 6); +lean_inc(x_43); +x_44 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 3); +x_45 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 4); +x_46 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 5); +x_47 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 6); +x_48 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 7); +x_49 = lean_ctor_get(x_13, 7); +lean_inc(x_49); +x_50 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 8); +x_51 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 9); +x_52 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_52, 0, x_34); +lean_ctor_set(x_52, 1, x_35); +lean_ctor_set(x_52, 2, x_36); +lean_ctor_set(x_52, 3, x_40); +lean_ctor_set(x_52, 4, x_41); +lean_ctor_set(x_52, 5, x_42); +lean_ctor_set(x_52, 6, x_43); +lean_ctor_set(x_52, 7, x_49); +lean_ctor_set(x_52, 8, x_5); +lean_ctor_set_uint8(x_52, sizeof(void*)*9, x_37); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 1, x_38); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 2, x_39); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 3, x_44); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 4, x_45); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 5, x_46); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 6, x_47); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 7, x_48); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 8, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 9, x_51); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_51 = l_Lean_Elab_Term_withInfoContext_x27(x_4, x_30, x_31, x_50, x_14, x_15, x_16, x_17, x_18, x_27); -if (lean_obj_tag(x_51) == 0) +x_53 = l_Lean_Elab_Term_withInfoContext_x27(x_4, x_30, x_31, x_33, x_52, x_14, x_15, x_16, x_17, x_18, x_27); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_52; lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = 0; -x_55 = 1; -x_56 = 1; -x_57 = l_Lean_Meta_mkLambdaFVars(x_11, x_52, x_54, x_55, x_54, x_56, x_15, x_16, x_17, x_18, x_53); -if (lean_obj_tag(x_57) == 0) +lean_object* x_54; lean_object* x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = 0; +x_57 = 1; +x_58 = 1; +x_59 = l_Lean_Meta_mkLambdaFVars(x_11, x_54, x_56, x_57, x_56, x_58, x_15, x_16, x_17, x_18, x_55); +if (lean_obj_tag(x_59) == 0) { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_57); -if (x_58 == 0) +uint8_t x_60; +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_57, 0); -x_60 = lean_ctor_get(x_57, 1); -x_61 = lean_ctor_get(x_17, 2); -lean_inc(x_61); -x_62 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; -x_63 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_61, x_62); -lean_dec(x_61); -if (x_63 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_61 = lean_ctor_get(x_59, 0); +x_62 = lean_ctor_get(x_59, 1); +x_63 = lean_ctor_get(x_17, 2); +lean_inc(x_63); +x_64 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; +x_65 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_63, x_64); +lean_dec(x_63); +if (x_65 == 0) { lean_dec(x_18); lean_dec(x_17); @@ -22238,39 +22525,39 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } else { -uint8_t x_64; -x_64 = l_Lean_Expr_hasSorry(x_6); -if (x_64 == 0) +uint8_t x_66; +x_66 = l_Lean_Expr_hasSorry(x_6); +if (x_66 == 0) { -uint8_t x_65; -x_65 = l_Lean_Expr_hasSorry(x_59); -if (x_65 == 0) +uint8_t x_67; +x_67 = l_Lean_Expr_hasSorry(x_61); +if (x_67 == 0) { -lean_object* x_66; lean_object* x_67; -lean_free_object(x_57); -x_66 = lean_array_get_size(x_7); +lean_object* x_68; lean_object* x_69; +lean_free_object(x_59); +x_68 = lean_array_get_size(x_7); lean_inc(x_15); -x_67 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_8, x_7, x_6, x_59, x_7, x_22, x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_60); -lean_dec(x_66); -if (lean_obj_tag(x_67) == 0) +x_69 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_8, x_7, x_6, x_61, x_7, x_22, x_68, x_13, x_14, x_15, x_16, x_17, x_18, x_62); +lean_dec(x_68); +if (lean_obj_tag(x_69) == 0) { -uint8_t x_68; -x_68 = !lean_is_exclusive(x_67); -if (x_68 == 0) +uint8_t x_70; +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_67, 0); -x_70 = lean_ctor_get(x_67, 1); -x_71 = lean_array_get_size(x_69); -x_72 = lean_nat_dec_lt(x_22, x_71); -lean_dec(x_71); -if (x_72 == 0) +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_ctor_get(x_69, 1); +x_73 = lean_array_get_size(x_71); +x_74 = lean_nat_dec_lt(x_22, x_73); +lean_dec(x_73); +if (x_74 == 0) { -lean_dec(x_69); +lean_dec(x_71); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22278,86 +22565,86 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -lean_ctor_set(x_67, 0, x_59); -return x_67; +lean_ctor_set(x_69, 0, x_61); +return x_69; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -lean_free_object(x_67); -x_73 = lean_ctor_get(x_9, 0); -x_74 = l_Lean_MessageData_ofName(x_10); -x_75 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_free_object(x_69); +x_75 = lean_ctor_get(x_9, 0); +x_76 = l_Lean_MessageData_ofName(x_10); +x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -x_79 = lean_array_to_list(x_69); -x_80 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_79); -x_81 = l_Lean_MessageData_joinSep(x_79, x_80); -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_78); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_76); +x_79 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_array_to_list(x_71); +x_82 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_81); +x_83 = l_Lean_MessageData_joinSep(x_81, x_82); x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 0, x_80); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_86 = l_Lean_MessageData_joinSep(x_79, x_85); -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_85 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_88 = l_Lean_MessageData_joinSep(x_81, x_87); x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 0, x_86); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_62, x_73, x_89, x_13, x_14, x_15, x_16, x_17, x_18, x_70); +x_90 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_91 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_64, x_75, x_91, x_13, x_14, x_15, x_16, x_17, x_18, x_72); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) { -lean_object* x_92; -x_92 = lean_ctor_get(x_90, 0); -lean_dec(x_92); -lean_ctor_set(x_90, 0, x_59); -return x_90; +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); +lean_dec(x_94); +lean_ctor_set(x_92, 0, x_61); +return x_92; } else { -lean_object* x_93; lean_object* x_94; -x_93 = lean_ctor_get(x_90, 1); -lean_inc(x_93); -lean_dec(x_90); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_59); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +lean_dec(x_92); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_61); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_95 = lean_ctor_get(x_67, 0); -x_96 = lean_ctor_get(x_67, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_67); -x_97 = lean_array_get_size(x_95); -x_98 = lean_nat_dec_lt(x_22, x_97); -lean_dec(x_97); -if (x_98 == 0) +lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_97 = lean_ctor_get(x_69, 0); +x_98 = lean_ctor_get(x_69, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_69); +x_99 = lean_array_get_size(x_97); +x_100 = lean_nat_dec_lt(x_22, x_99); +lean_dec(x_99); +if (x_100 == 0) { -lean_object* x_99; -lean_dec(x_95); +lean_object* x_101; +lean_dec(x_97); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22365,75 +22652,75 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_59); -lean_ctor_set(x_99, 1, x_96); -return x_99; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_61); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_100 = lean_ctor_get(x_9, 0); -x_101 = l_Lean_MessageData_ofName(x_10); -x_102 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -x_104 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_102 = lean_ctor_get(x_9, 0); +x_103 = l_Lean_MessageData_ofName(x_10); +x_104 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_array_to_list(x_95); -x_107 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_106); -x_108 = l_Lean_MessageData_joinSep(x_106, x_107); -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_105); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_103); +x_106 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_107 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = lean_array_to_list(x_97); +x_109 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_108); +x_110 = l_Lean_MessageData_joinSep(x_108, x_109); x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 0, x_107); lean_ctor_set(x_111, 1, x_110); -x_112 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_113 = l_Lean_MessageData_joinSep(x_106, x_112); -x_114 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_112 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_115 = l_Lean_MessageData_joinSep(x_108, x_114); x_116 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 0, x_113); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_62, x_100, x_116, x_13, x_14, x_15, x_16, x_17, x_18, x_96); +x_117 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_64, x_102, x_118, x_13, x_14, x_15, x_16, x_17, x_18, x_98); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_119 = x_117; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_121 = x_119; } else { - lean_dec_ref(x_117); - x_119 = lean_box(0); + lean_dec_ref(x_119); + x_121 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_119; + x_122 = x_121; } -lean_ctor_set(x_120, 0, x_59); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_122, 0, x_61); +lean_ctor_set(x_122, 1, x_120); +return x_122; } } } else { -uint8_t x_121; -lean_dec(x_59); +uint8_t x_123; +lean_dec(x_61); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22441,23 +22728,23 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_121 = !lean_is_exclusive(x_67); -if (x_121 == 0) +x_123 = !lean_is_exclusive(x_69); +if (x_123 == 0) { -return x_67; +return x_69; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_67, 0); -x_123 = lean_ctor_get(x_67, 1); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_67); -x_124 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -return x_124; +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_69, 0); +x_125 = lean_ctor_get(x_69, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_69); +x_126 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +return x_126; } } } @@ -22470,7 +22757,7 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } } else @@ -22482,26 +22769,26 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } } } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_125 = lean_ctor_get(x_57, 0); -x_126 = lean_ctor_get(x_57, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_57); -x_127 = lean_ctor_get(x_17, 2); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_127 = lean_ctor_get(x_59, 0); +x_128 = lean_ctor_get(x_59, 1); +lean_inc(x_128); lean_inc(x_127); -x_128 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; -x_129 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_127, x_128); -lean_dec(x_127); -if (x_129 == 0) +lean_dec(x_59); +x_129 = lean_ctor_get(x_17, 2); +lean_inc(x_129); +x_130 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; +x_131 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_129, x_130); +lean_dec(x_129); +if (x_131 == 0) { -lean_object* x_130; +lean_object* x_132; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22509,48 +22796,48 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_125); -lean_ctor_set(x_130, 1, x_126); -return x_130; +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_127); +lean_ctor_set(x_132, 1, x_128); +return x_132; } else { -uint8_t x_131; -x_131 = l_Lean_Expr_hasSorry(x_6); -if (x_131 == 0) +uint8_t x_133; +x_133 = l_Lean_Expr_hasSorry(x_6); +if (x_133 == 0) { -uint8_t x_132; -x_132 = l_Lean_Expr_hasSorry(x_125); -if (x_132 == 0) +uint8_t x_134; +x_134 = l_Lean_Expr_hasSorry(x_127); +if (x_134 == 0) { -lean_object* x_133; lean_object* x_134; -x_133 = lean_array_get_size(x_7); +lean_object* x_135; lean_object* x_136; +x_135 = lean_array_get_size(x_7); lean_inc(x_15); -x_134 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_8, x_7, x_6, x_125, x_7, x_22, x_133, x_13, x_14, x_15, x_16, x_17, x_18, x_126); -lean_dec(x_133); -if (lean_obj_tag(x_134) == 0) +x_136 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__3(x_8, x_7, x_6, x_127, x_7, x_22, x_135, x_13, x_14, x_15, x_16, x_17, x_18, x_128); +lean_dec(x_135); +if (lean_obj_tag(x_136) == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_137 = x_134; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_139 = x_136; } else { - lean_dec_ref(x_134); - x_137 = lean_box(0); + lean_dec_ref(x_136); + x_139 = lean_box(0); } -x_138 = lean_array_get_size(x_135); -x_139 = lean_nat_dec_lt(x_22, x_138); -lean_dec(x_138); -if (x_139 == 0) +x_140 = lean_array_get_size(x_137); +x_141 = lean_nat_dec_lt(x_22, x_140); +lean_dec(x_140); +if (x_141 == 0) { -lean_object* x_140; -lean_dec(x_135); +lean_object* x_142; +lean_dec(x_137); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22558,79 +22845,79 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -if (lean_is_scalar(x_137)) { - x_140 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_139)) { + x_142 = lean_alloc_ctor(0, 2, 0); } else { - x_140 = x_137; + x_142 = x_139; } -lean_ctor_set(x_140, 0, x_125); -lean_ctor_set(x_140, 1, x_136); -return x_140; +lean_ctor_set(x_142, 0, x_127); +lean_ctor_set(x_142, 1, x_138); +return x_142; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -lean_dec(x_137); -x_141 = lean_ctor_get(x_9, 0); -x_142 = l_Lean_MessageData_ofName(x_10); -x_143 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_142); -x_145 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_139); +x_143 = lean_ctor_get(x_9, 0); +x_144 = l_Lean_MessageData_ofName(x_10); +x_145 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -x_147 = lean_array_to_list(x_135); -x_148 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_147); -x_149 = l_Lean_MessageData_joinSep(x_147, x_148); -x_150 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_150, 0, x_146); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +x_147 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_148 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = lean_array_to_list(x_137); +x_150 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_149); +x_151 = l_Lean_MessageData_joinSep(x_149, x_150); x_152 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 0, x_148); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_154 = l_Lean_MessageData_joinSep(x_147, x_153); -x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_152); -lean_ctor_set(x_155, 1, x_154); -x_156 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_153 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_156 = l_Lean_MessageData_joinSep(x_149, x_155); x_157 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 0, x_154); lean_ctor_set(x_157, 1, x_156); -x_158 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_128, x_141, x_157, x_13, x_14, x_15, x_16, x_17, x_18, x_136); +x_158 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_130, x_143, x_159, x_13, x_14, x_15, x_16, x_17, x_18, x_138); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_159 = lean_ctor_get(x_158, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_160 = x_158; +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_162 = x_160; } else { - lean_dec_ref(x_158); - x_160 = lean_box(0); + lean_dec_ref(x_160); + x_162 = lean_box(0); } -if (lean_is_scalar(x_160)) { - x_161 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(0, 2, 0); } else { - x_161 = x_160; + x_163 = x_162; } -lean_ctor_set(x_161, 0, x_125); -lean_ctor_set(x_161, 1, x_159); -return x_161; +lean_ctor_set(x_163, 0, x_127); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_125); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_127); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22638,31 +22925,31 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_162 = lean_ctor_get(x_134, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_134, 1); -lean_inc(x_163); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_164 = x_134; +x_164 = lean_ctor_get(x_136, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_136, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_166 = x_136; } else { - lean_dec_ref(x_134); - x_164 = lean_box(0); + lean_dec_ref(x_136); + x_166 = lean_box(0); } -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_166)) { + x_167 = lean_alloc_ctor(1, 2, 0); } else { - x_165 = x_164; + x_167 = x_166; } -lean_ctor_set(x_165, 0, x_162); -lean_ctor_set(x_165, 1, x_163); -return x_165; +lean_ctor_set(x_167, 0, x_164); +lean_ctor_set(x_167, 1, x_165); +return x_167; } } else { -lean_object* x_166; +lean_object* x_168; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22670,15 +22957,15 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_125); -lean_ctor_set(x_166, 1, x_126); -return x_166; +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_127); +lean_ctor_set(x_168, 1, x_128); +return x_168; } } else { -lean_object* x_167; +lean_object* x_169; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22686,17 +22973,17 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_125); -lean_ctor_set(x_167, 1, x_126); -return x_167; +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_127); +lean_ctor_set(x_169, 1, x_128); +return x_169; } } } } else { -uint8_t x_168; +uint8_t x_170; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22704,29 +22991,29 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_168 = !lean_is_exclusive(x_57); -if (x_168 == 0) +x_170 = !lean_is_exclusive(x_59); +if (x_170 == 0) { -return x_57; +return x_59; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_57, 0); -x_170 = lean_ctor_get(x_57, 1); -lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_57); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -return x_171; +lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_171 = lean_ctor_get(x_59, 0); +x_172 = lean_ctor_get(x_59, 1); +lean_inc(x_172); +lean_inc(x_171); +lean_dec(x_59); +x_173 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_173, 0, x_171); +lean_ctor_set(x_173, 1, x_172); +return x_173; } } } else { -uint8_t x_172; +uint8_t x_174; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22735,29 +23022,29 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_11); lean_dec(x_10); -x_172 = !lean_is_exclusive(x_51); -if (x_172 == 0) +x_174 = !lean_is_exclusive(x_53); +if (x_174 == 0) { -return x_51; +return x_53; } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_51, 0); -x_174 = lean_ctor_get(x_51, 1); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_51); -x_175 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -return x_175; +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = lean_ctor_get(x_53, 0); +x_176 = lean_ctor_get(x_53, 1); +lean_inc(x_176); +lean_inc(x_175); +lean_dec(x_53); +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set(x_177, 1, x_176); +return x_177; } } } else { -uint8_t x_176; +uint8_t x_178; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22769,23 +23056,23 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_5); lean_dec(x_4); -x_176 = !lean_is_exclusive(x_26); -if (x_176 == 0) +x_178 = !lean_is_exclusive(x_26); +if (x_178 == 0) { return x_26; } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_177 = lean_ctor_get(x_26, 0); -x_178 = lean_ctor_get(x_26, 1); -lean_inc(x_178); -lean_inc(x_177); +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_26, 0); +x_180 = lean_ctor_get(x_26, 1); +lean_inc(x_180); +lean_inc(x_179); lean_dec(x_26); -x_179 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_178); -return x_179; +x_181 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_180); +return x_181; } } } @@ -22816,7 +23103,7 @@ lean_dec(x_24); lean_dec(x_21); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; lean_object* x_47; uint8_t x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); @@ -22831,85 +23118,89 @@ lean_closure_set(x_30, 2, x_29); lean_inc(x_4); x_31 = lean_alloc_closure((void*)(l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__2___boxed), 9, 1); lean_closure_set(x_31, 0, x_4); -x_32 = lean_ctor_get(x_13, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_13, 1); -lean_inc(x_33); -x_34 = lean_ctor_get(x_13, 2); +lean_inc(x_4); +x_32 = l_Lean_Elab_Term_mkBodyInfo(x_4, x_29); +x_33 = lean_alloc_closure((void*)(l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12___rarg___boxed), 8, 1); +lean_closure_set(x_33, 0, x_32); +x_34 = lean_ctor_get(x_13, 0); lean_inc(x_34); -x_35 = lean_ctor_get_uint8(x_13, sizeof(void*)*9); -x_36 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 1); -x_37 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 2); -x_38 = lean_ctor_get(x_13, 3); -lean_inc(x_38); -x_39 = lean_ctor_get(x_13, 4); -lean_inc(x_39); -x_40 = lean_ctor_get(x_13, 5); +x_35 = lean_ctor_get(x_13, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_13, 2); +lean_inc(x_36); +x_37 = lean_ctor_get_uint8(x_13, sizeof(void*)*9); +x_38 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 1); +x_39 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 2); +x_40 = lean_ctor_get(x_13, 3); lean_inc(x_40); -x_41 = lean_ctor_get(x_13, 6); +x_41 = lean_ctor_get(x_13, 4); lean_inc(x_41); -x_42 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 3); -x_43 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 4); -x_44 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 5); -x_45 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 6); -x_46 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 7); -x_47 = lean_ctor_get(x_13, 7); -lean_inc(x_47); -x_48 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 8); -x_49 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 9); -x_50 = lean_alloc_ctor(0, 9, 10); -lean_ctor_set(x_50, 0, x_32); -lean_ctor_set(x_50, 1, x_33); -lean_ctor_set(x_50, 2, x_34); -lean_ctor_set(x_50, 3, x_38); -lean_ctor_set(x_50, 4, x_39); -lean_ctor_set(x_50, 5, x_40); -lean_ctor_set(x_50, 6, x_41); -lean_ctor_set(x_50, 7, x_47); -lean_ctor_set(x_50, 8, x_5); -lean_ctor_set_uint8(x_50, sizeof(void*)*9, x_35); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 1, x_36); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 2, x_37); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 3, x_42); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 4, x_43); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 5, x_44); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 6, x_45); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 7, x_46); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 8, x_48); -lean_ctor_set_uint8(x_50, sizeof(void*)*9 + 9, x_49); +x_42 = lean_ctor_get(x_13, 5); +lean_inc(x_42); +x_43 = lean_ctor_get(x_13, 6); +lean_inc(x_43); +x_44 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 3); +x_45 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 4); +x_46 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 5); +x_47 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 6); +x_48 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 7); +x_49 = lean_ctor_get(x_13, 7); +lean_inc(x_49); +x_50 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 8); +x_51 = lean_ctor_get_uint8(x_13, sizeof(void*)*9 + 9); +x_52 = lean_alloc_ctor(0, 9, 10); +lean_ctor_set(x_52, 0, x_34); +lean_ctor_set(x_52, 1, x_35); +lean_ctor_set(x_52, 2, x_36); +lean_ctor_set(x_52, 3, x_40); +lean_ctor_set(x_52, 4, x_41); +lean_ctor_set(x_52, 5, x_42); +lean_ctor_set(x_52, 6, x_43); +lean_ctor_set(x_52, 7, x_49); +lean_ctor_set(x_52, 8, x_5); +lean_ctor_set_uint8(x_52, sizeof(void*)*9, x_37); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 1, x_38); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 2, x_39); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 3, x_44); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 4, x_45); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 5, x_46); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 6, x_47); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 7, x_48); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 8, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*9 + 9, x_51); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_51 = l_Lean_Elab_Term_withInfoContext_x27(x_4, x_30, x_31, x_50, x_14, x_15, x_16, x_17, x_18, x_27); -if (lean_obj_tag(x_51) == 0) +x_53 = l_Lean_Elab_Term_withInfoContext_x27(x_4, x_30, x_31, x_33, x_52, x_14, x_15, x_16, x_17, x_18, x_27); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_52; lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = 0; -x_55 = 1; -x_56 = 1; -x_57 = l_Lean_Meta_mkLambdaFVars(x_11, x_52, x_54, x_55, x_54, x_56, x_15, x_16, x_17, x_18, x_53); -if (lean_obj_tag(x_57) == 0) +lean_object* x_54; lean_object* x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = 0; +x_57 = 1; +x_58 = 1; +x_59 = l_Lean_Meta_mkLambdaFVars(x_11, x_54, x_56, x_57, x_56, x_58, x_15, x_16, x_17, x_18, x_55); +if (lean_obj_tag(x_59) == 0) { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_57); -if (x_58 == 0) +uint8_t x_60; +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_57, 0); -x_60 = lean_ctor_get(x_57, 1); -x_61 = lean_ctor_get(x_17, 2); -lean_inc(x_61); -x_62 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; -x_63 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_61, x_62); -lean_dec(x_61); -if (x_63 == 0) +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_61 = lean_ctor_get(x_59, 0); +x_62 = lean_ctor_get(x_59, 1); +x_63 = lean_ctor_get(x_17, 2); +lean_inc(x_63); +x_64 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; +x_65 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_63, x_64); +lean_dec(x_63); +if (x_65 == 0) { lean_dec(x_18); lean_dec(x_17); @@ -22918,39 +23209,39 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } else { -uint8_t x_64; -x_64 = l_Lean_Expr_hasSorry(x_6); -if (x_64 == 0) +uint8_t x_66; +x_66 = l_Lean_Expr_hasSorry(x_6); +if (x_66 == 0) { -uint8_t x_65; -x_65 = l_Lean_Expr_hasSorry(x_59); -if (x_65 == 0) +uint8_t x_67; +x_67 = l_Lean_Expr_hasSorry(x_61); +if (x_67 == 0) { -lean_object* x_66; lean_object* x_67; -lean_free_object(x_57); -x_66 = lean_array_get_size(x_7); +lean_object* x_68; lean_object* x_69; +lean_free_object(x_59); +x_68 = lean_array_get_size(x_7); lean_inc(x_15); -x_67 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_8, x_7, x_6, x_59, x_7, x_22, x_66, x_13, x_14, x_15, x_16, x_17, x_18, x_60); -lean_dec(x_66); -if (lean_obj_tag(x_67) == 0) +x_69 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_8, x_7, x_6, x_61, x_7, x_22, x_68, x_13, x_14, x_15, x_16, x_17, x_18, x_62); +lean_dec(x_68); +if (lean_obj_tag(x_69) == 0) { -uint8_t x_68; -x_68 = !lean_is_exclusive(x_67); -if (x_68 == 0) +uint8_t x_70; +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_67, 0); -x_70 = lean_ctor_get(x_67, 1); -x_71 = lean_array_get_size(x_69); -x_72 = lean_nat_dec_lt(x_22, x_71); -lean_dec(x_71); -if (x_72 == 0) +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_69, 0); +x_72 = lean_ctor_get(x_69, 1); +x_73 = lean_array_get_size(x_71); +x_74 = lean_nat_dec_lt(x_22, x_73); +lean_dec(x_73); +if (x_74 == 0) { -lean_dec(x_69); +lean_dec(x_71); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -22958,86 +23249,86 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -lean_ctor_set(x_67, 0, x_59); -return x_67; +lean_ctor_set(x_69, 0, x_61); +return x_69; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; -lean_free_object(x_67); -x_73 = lean_ctor_get(x_9, 0); -x_74 = l_Lean_MessageData_ofName(x_10); -x_75 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_free_object(x_69); +x_75 = lean_ctor_get(x_9, 0); +x_76 = l_Lean_MessageData_ofName(x_10); +x_77 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -x_79 = lean_array_to_list(x_69); -x_80 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_79); -x_81 = l_Lean_MessageData_joinSep(x_79, x_80); -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_78); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_76); +x_79 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_array_to_list(x_71); +x_82 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_81); +x_83 = l_Lean_MessageData_joinSep(x_81, x_82); x_84 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 0, x_80); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_86 = l_Lean_MessageData_joinSep(x_79, x_85); -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_84); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_85 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_88 = l_Lean_MessageData_joinSep(x_81, x_87); x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 0, x_86); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_62, x_73, x_89, x_13, x_14, x_15, x_16, x_17, x_18, x_70); +x_90 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_91 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_64, x_75, x_91, x_13, x_14, x_15, x_16, x_17, x_18, x_72); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_91 = !lean_is_exclusive(x_90); -if (x_91 == 0) +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) { -lean_object* x_92; -x_92 = lean_ctor_get(x_90, 0); -lean_dec(x_92); -lean_ctor_set(x_90, 0, x_59); -return x_90; +lean_object* x_94; +x_94 = lean_ctor_get(x_92, 0); +lean_dec(x_94); +lean_ctor_set(x_92, 0, x_61); +return x_92; } else { -lean_object* x_93; lean_object* x_94; -x_93 = lean_ctor_get(x_90, 1); -lean_inc(x_93); -lean_dec(x_90); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_59); -lean_ctor_set(x_94, 1, x_93); -return x_94; +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_92, 1); +lean_inc(x_95); +lean_dec(x_92); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_61); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; -x_95 = lean_ctor_get(x_67, 0); -x_96 = lean_ctor_get(x_67, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_67); -x_97 = lean_array_get_size(x_95); -x_98 = lean_nat_dec_lt(x_22, x_97); -lean_dec(x_97); -if (x_98 == 0) +lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_97 = lean_ctor_get(x_69, 0); +x_98 = lean_ctor_get(x_69, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_69); +x_99 = lean_array_get_size(x_97); +x_100 = lean_nat_dec_lt(x_22, x_99); +lean_dec(x_99); +if (x_100 == 0) { -lean_object* x_99; -lean_dec(x_95); +lean_object* x_101; +lean_dec(x_97); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23045,75 +23336,75 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_59); -lean_ctor_set(x_99, 1, x_96); -return x_99; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_61); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_100 = lean_ctor_get(x_9, 0); -x_101 = l_Lean_MessageData_ofName(x_10); -x_102 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -x_104 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_102 = lean_ctor_get(x_9, 0); +x_103 = l_Lean_MessageData_ofName(x_10); +x_104 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_array_to_list(x_95); -x_107 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_106); -x_108 = l_Lean_MessageData_joinSep(x_106, x_107); -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_105); -lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_103); +x_106 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_107 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = lean_array_to_list(x_97); +x_109 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_108); +x_110 = l_Lean_MessageData_joinSep(x_108, x_109); x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 0, x_107); lean_ctor_set(x_111, 1, x_110); -x_112 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_113 = l_Lean_MessageData_joinSep(x_106, x_112); -x_114 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_112 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_115 = l_Lean_MessageData_joinSep(x_108, x_114); x_116 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 0, x_113); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_62, x_100, x_116, x_13, x_14, x_15, x_16, x_17, x_18, x_96); +x_117 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_64, x_102, x_118, x_13, x_14, x_15, x_16, x_17, x_18, x_98); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_118 = lean_ctor_get(x_117, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_119 = x_117; +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_121 = x_119; } else { - lean_dec_ref(x_117); - x_119 = lean_box(0); + lean_dec_ref(x_119); + x_121 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_119; + x_122 = x_121; } -lean_ctor_set(x_120, 0, x_59); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_122, 0, x_61); +lean_ctor_set(x_122, 1, x_120); +return x_122; } } } else { -uint8_t x_121; -lean_dec(x_59); +uint8_t x_123; +lean_dec(x_61); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23121,23 +23412,23 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_121 = !lean_is_exclusive(x_67); -if (x_121 == 0) +x_123 = !lean_is_exclusive(x_69); +if (x_123 == 0) { -return x_67; +return x_69; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_122 = lean_ctor_get(x_67, 0); -x_123 = lean_ctor_get(x_67, 1); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_67); -x_124 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -return x_124; +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_69, 0); +x_125 = lean_ctor_get(x_69, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_69); +x_126 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +return x_126; } } } @@ -23150,7 +23441,7 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } } else @@ -23162,26 +23453,26 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -return x_57; +return x_59; } } } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -x_125 = lean_ctor_get(x_57, 0); -x_126 = lean_ctor_get(x_57, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_57); -x_127 = lean_ctor_get(x_17, 2); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_127 = lean_ctor_get(x_59, 0); +x_128 = lean_ctor_get(x_59, 1); +lean_inc(x_128); lean_inc(x_127); -x_128 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; -x_129 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_127, x_128); -lean_dec(x_127); -if (x_129 == 0) +lean_dec(x_59); +x_129 = lean_ctor_get(x_17, 2); +lean_inc(x_129); +x_130 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__1; +x_131 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_129, x_130); +lean_dec(x_129); +if (x_131 == 0) { -lean_object* x_130; +lean_object* x_132; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23189,48 +23480,48 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_125); -lean_ctor_set(x_130, 1, x_126); -return x_130; +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_127); +lean_ctor_set(x_132, 1, x_128); +return x_132; } else { -uint8_t x_131; -x_131 = l_Lean_Expr_hasSorry(x_6); -if (x_131 == 0) +uint8_t x_133; +x_133 = l_Lean_Expr_hasSorry(x_6); +if (x_133 == 0) { -uint8_t x_132; -x_132 = l_Lean_Expr_hasSorry(x_125); -if (x_132 == 0) +uint8_t x_134; +x_134 = l_Lean_Expr_hasSorry(x_127); +if (x_134 == 0) { -lean_object* x_133; lean_object* x_134; -x_133 = lean_array_get_size(x_7); +lean_object* x_135; lean_object* x_136; +x_135 = lean_array_get_size(x_7); lean_inc(x_15); -x_134 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_8, x_7, x_6, x_125, x_7, x_22, x_133, x_13, x_14, x_15, x_16, x_17, x_18, x_126); -lean_dec(x_133); -if (lean_obj_tag(x_134) == 0) +x_136 = l_Array_filterMapM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__8(x_8, x_7, x_6, x_127, x_7, x_22, x_135, x_13, x_14, x_15, x_16, x_17, x_18, x_128); +lean_dec(x_135); +if (lean_obj_tag(x_136) == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_137 = x_134; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_139 = x_136; } else { - lean_dec_ref(x_134); - x_137 = lean_box(0); + lean_dec_ref(x_136); + x_139 = lean_box(0); } -x_138 = lean_array_get_size(x_135); -x_139 = lean_nat_dec_lt(x_22, x_138); -lean_dec(x_138); -if (x_139 == 0) +x_140 = lean_array_get_size(x_137); +x_141 = lean_nat_dec_lt(x_22, x_140); +lean_dec(x_140); +if (x_141 == 0) { -lean_object* x_140; -lean_dec(x_135); +lean_object* x_142; +lean_dec(x_137); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23238,79 +23529,79 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -if (lean_is_scalar(x_137)) { - x_140 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_139)) { + x_142 = lean_alloc_ctor(0, 2, 0); } else { - x_140 = x_137; + x_142 = x_139; } -lean_ctor_set(x_140, 0, x_125); -lean_ctor_set(x_140, 1, x_136); -return x_140; +lean_ctor_set(x_142, 0, x_127); +lean_ctor_set(x_142, 1, x_138); +return x_142; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -lean_dec(x_137); -x_141 = lean_ctor_get(x_9, 0); -x_142 = l_Lean_MessageData_ofName(x_10); -x_143 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_142); -x_145 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_139); +x_143 = lean_ctor_get(x_9, 0); +x_144 = l_Lean_MessageData_ofName(x_10); +x_145 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__3; x_146 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -x_147 = lean_array_to_list(x_135); -x_148 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; -lean_inc(x_147); -x_149 = l_Lean_MessageData_joinSep(x_147, x_148); -x_150 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_150, 0, x_146); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +x_147 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__5; +x_148 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = lean_array_to_list(x_137); +x_150 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__8; +lean_inc(x_149); +x_151 = l_Lean_MessageData_joinSep(x_149, x_150); x_152 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 0, x_148); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; -x_154 = l_Lean_MessageData_joinSep(x_147, x_153); -x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_152); -lean_ctor_set(x_155, 1, x_154); -x_156 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_153 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__10; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__13; +x_156 = l_Lean_MessageData_joinSep(x_149, x_155); x_157 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 0, x_154); lean_ctor_set(x_157, 1, x_156); -x_158 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_128, x_141, x_157, x_13, x_14, x_15, x_16, x_17, x_18, x_136); +x_158 = l_Lean_Core_withRestoreOrSaveFull___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__13___lambda__4___closed__15; +x_159 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_159, 0, x_157); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Linter_logLint___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__5(x_130, x_143, x_159, x_13, x_14, x_15, x_16, x_17, x_18, x_138); lean_dec(x_18); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_159 = lean_ctor_get(x_158, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_160 = x_158; +x_161 = lean_ctor_get(x_160, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_162 = x_160; } else { - lean_dec_ref(x_158); - x_160 = lean_box(0); + lean_dec_ref(x_160); + x_162 = lean_box(0); } -if (lean_is_scalar(x_160)) { - x_161 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(0, 2, 0); } else { - x_161 = x_160; + x_163 = x_162; } -lean_ctor_set(x_161, 0, x_125); -lean_ctor_set(x_161, 1, x_159); -return x_161; +lean_ctor_set(x_163, 0, x_127); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_125); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_127); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23318,31 +23609,31 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_162 = lean_ctor_get(x_134, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_134, 1); -lean_inc(x_163); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_164 = x_134; +x_164 = lean_ctor_get(x_136, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_136, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_166 = x_136; } else { - lean_dec_ref(x_134); - x_164 = lean_box(0); + lean_dec_ref(x_136); + x_166 = lean_box(0); } -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_166)) { + x_167 = lean_alloc_ctor(1, 2, 0); } else { - x_165 = x_164; + x_167 = x_166; } -lean_ctor_set(x_165, 0, x_162); -lean_ctor_set(x_165, 1, x_163); -return x_165; +lean_ctor_set(x_167, 0, x_164); +lean_ctor_set(x_167, 1, x_165); +return x_167; } } else { -lean_object* x_166; +lean_object* x_168; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23350,15 +23641,15 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_125); -lean_ctor_set(x_166, 1, x_126); -return x_166; +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_127); +lean_ctor_set(x_168, 1, x_128); +return x_168; } } else { -lean_object* x_167; +lean_object* x_169; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23366,17 +23657,17 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_125); -lean_ctor_set(x_167, 1, x_126); -return x_167; +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_127); +lean_ctor_set(x_169, 1, x_128); +return x_169; } } } } else { -uint8_t x_168; +uint8_t x_170; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23384,29 +23675,29 @@ lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_10); -x_168 = !lean_is_exclusive(x_57); -if (x_168 == 0) +x_170 = !lean_is_exclusive(x_59); +if (x_170 == 0) { -return x_57; +return x_59; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_57, 0); -x_170 = lean_ctor_get(x_57, 1); -lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_57); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -return x_171; +lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_171 = lean_ctor_get(x_59, 0); +x_172 = lean_ctor_get(x_59, 1); +lean_inc(x_172); +lean_inc(x_171); +lean_dec(x_59); +x_173 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_173, 0, x_171); +lean_ctor_set(x_173, 1, x_172); +return x_173; } } } else { -uint8_t x_172; +uint8_t x_174; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23415,29 +23706,29 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_11); lean_dec(x_10); -x_172 = !lean_is_exclusive(x_51); -if (x_172 == 0) +x_174 = !lean_is_exclusive(x_53); +if (x_174 == 0) { -return x_51; +return x_53; } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_51, 0); -x_174 = lean_ctor_get(x_51, 1); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_51); -x_175 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -return x_175; +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = lean_ctor_get(x_53, 0); +x_176 = lean_ctor_get(x_53, 1); +lean_inc(x_176); +lean_inc(x_175); +lean_dec(x_53); +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set(x_177, 1, x_176); +return x_177; } } } else { -uint8_t x_176; +uint8_t x_178; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -23449,23 +23740,23 @@ lean_dec(x_11); lean_dec(x_10); lean_dec(x_5); lean_dec(x_4); -x_176 = !lean_is_exclusive(x_26); -if (x_176 == 0) +x_178 = !lean_is_exclusive(x_26); +if (x_178 == 0) { return x_26; } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_177 = lean_ctor_get(x_26, 0); -x_178 = lean_ctor_get(x_26, 1); -lean_inc(x_178); -lean_inc(x_177); +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_26, 0); +x_180 = lean_ctor_get(x_26, 1); +lean_inc(x_180); +lean_inc(x_179); lean_dec(x_26); -x_179 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_178); -return x_179; +x_181 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_180); +return x_181; } } } @@ -53292,7 +53583,7 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -53302,83 +53593,83 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_MutualDef___hyg_5____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7; -x_2 = lean_unsigned_to_nat(16659u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7; +x_2 = lean_unsigned_to_nat(16914u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -53778,6 +54069,10 @@ l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__4); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5); +l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6(); +lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6); +l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__7); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1(); lean_mark_persistent(l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1); l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__2(); @@ -53838,36 +54133,36 @@ l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Ter lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__4___closed__4); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__7___lambda__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__7___lambda__1___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withHeaderSecVars_collectUsed___spec__7___lambda__1___closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381____closed__6); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6381_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628____closed__6); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6628_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_deprecated_oldSectionVars = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_deprecated_oldSectionVars); lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420____closed__6); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6420_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667____closed__6); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_MutualDef___hyg_6667_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_linter_unusedSectionVars = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_linter_unusedSectionVars); @@ -54065,23 +54360,23 @@ l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__4 = _init_l_Lean_Elab_Co lean_mark_persistent(l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__4); l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5 = _init_l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5(); lean_mark_persistent(l_Lean_Elab_Command_elabMutualDef___lambda__3___closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659____closed__8); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16659_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914____closed__8); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_MutualDef___hyg_16914_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PatternVar.c b/stage0/stdlib/Lean/Elab/PatternVar.c index 5bb4f0246589..db02e0c9bf09 100644 --- a/stage0/stdlib/Lean/Elab/PatternVar.c +++ b/stage0/stdlib/Lean/Elab/PatternVar.c @@ -52,7 +52,7 @@ static lean_object* l_Lean_Elab_Term_CollectPatternVars_instInhabitedContext___c LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___spec__13___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); @@ -66,7 +66,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVar lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); extern lean_object* l_Lean_Elab_Term_instInhabitedNamedArg; lean_object* l_Lean_Syntax_getId(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__4; uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); @@ -135,10 +135,11 @@ static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__9; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getPatternVars___lambda__1___boxed(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_getPatternsVars___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__3; static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__27; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); @@ -281,6 +282,7 @@ static lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPa lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_CollectPatternVars_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_instInhabitedState___closed__2; +static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_getPatternsVars___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected_showName___spec__1(lean_object*, lean_object*, lean_object*); @@ -482,7 +484,7 @@ lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*, uint8_t); uint8_t l_Lean_Name_isSuffixOf(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_getPatternsVars___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__21; @@ -508,7 +510,7 @@ lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect_processId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_Term_CollectPatternVars_collect___spec__2___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___spec__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -11224,252 +11226,255 @@ x_1 = lean_mk_string_unchecked("with", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { -lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; -x_19 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); -x_20 = lean_array_size(x_19); -x_21 = 0; +lean_object* x_20; size_t x_21; size_t x_22; lean_object* x_23; +x_20 = l_Lean_Syntax_TSepArray_getElems___rarg(x_1); +x_21 = lean_array_size(x_20); +x_22 = 0; +lean_inc(x_18); lean_inc(x_17); -lean_inc(x_16); -x_22 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_2, x_3, x_20, x_21, x_19, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_22) == 0) +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__3(x_2, x_3, x_21, x_22, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_16, 5); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -lean_dec(x_16); -x_26 = 0; -x_27 = l_Lean_SourceInfo_fromRef(x_25, x_26); -lean_dec(x_25); -x_28 = lean_st_ref_get(x_17, x_24); +lean_dec(x_23); +x_26 = lean_ctor_get(x_17, 5); +lean_inc(x_26); lean_dec(x_17); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_30 = x_28; +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_st_ref_get(x_18, x_25); +lean_dec(x_18); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_31 = x_29; } else { - lean_dec_ref(x_28); - x_30 = lean_box(0); + lean_dec_ref(x_29); + x_31 = lean_box(0); } -x_31 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1; -lean_inc(x_27); -x_32 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_size(x_23); -x_34 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_4, x_33, x_21, x_23); -x_35 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__4___closed__2; -x_36 = l_Lean_Syntax_SepArray_ofElems(x_35, x_34); -lean_dec(x_34); -x_37 = l_Lean_Elab_Term_CollectPatternVars_instInhabitedState___closed__1; -x_38 = l_Array_append___rarg(x_37, x_36); -lean_dec(x_36); -x_39 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; -lean_inc(x_27); -x_40 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_40, 0, x_27); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_38); -x_41 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__4; -lean_inc(x_27); -x_42 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_42, 0, x_27); -lean_ctor_set(x_42, 1, x_41); -if (lean_obj_tag(x_9) == 0) +x_32 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__1; +lean_inc(x_28); +x_33 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_33, 0, x_28); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_array_size(x_24); +x_35 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_CollectPatternVars_collect___spec__4(x_4, x_34, x_22, x_24); +x_36 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwCtorExpected___rarg___lambda__4___closed__2; +x_37 = l_Lean_Syntax_SepArray_ofElems(x_36, x_35); +lean_dec(x_35); +x_38 = l_Lean_Elab_Term_CollectPatternVars_instInhabitedState___closed__1; +x_39 = l_Array_append___rarg(x_38, x_37); +lean_dec(x_37); +x_40 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; +lean_inc(x_28); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_28); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_39); +lean_inc(x_28); +x_42 = l_Lean_Syntax_node1(x_28, x_5, x_41); +x_43 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__4; +lean_inc(x_28); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_28); +lean_ctor_set(x_44, 1, x_43); +if (lean_obj_tag(x_10) == 0) { -x_43 = x_37; -goto block_80; +x_45 = x_38; +goto block_82; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_81 = lean_ctor_get(x_9, 0); -x_82 = l_Array_append___rarg(x_37, x_81); -lean_inc(x_27); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_27); -lean_ctor_set(x_83, 1, x_39); -lean_ctor_set(x_83, 2, x_82); -x_84 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7; -lean_inc(x_27); -x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_27); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Array_mkArray2___rarg(x_83, x_85); -x_43 = x_86; -goto block_80; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_83 = lean_ctor_get(x_10, 0); +x_84 = l_Array_append___rarg(x_38, x_83); +lean_inc(x_28); +x_85 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_85, 0, x_28); +lean_ctor_set(x_85, 1, x_40); +lean_ctor_set(x_85, 2, x_84); +x_86 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__7; +lean_inc(x_28); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_28); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Array_mkArray2___rarg(x_85, x_87); +x_45 = x_88; +goto block_82; } -block_80: +block_82: { -lean_object* x_44; lean_object* x_45; -x_44 = l_Array_append___rarg(x_37, x_43); -lean_dec(x_43); -lean_inc(x_27); -x_45 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_45, 0, x_27); -lean_ctor_set(x_45, 1, x_39); -lean_ctor_set(x_45, 2, x_44); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; -lean_inc(x_27); +lean_object* x_46; lean_object* x_47; +x_46 = l_Array_append___rarg(x_38, x_45); +lean_dec(x_45); +lean_inc(x_28); x_47 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_47, 0, x_27); -lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 0, x_28); +lean_ctor_set(x_47, 1, x_40); lean_ctor_set(x_47, 2, x_46); -lean_inc(x_47); -lean_inc(x_27); -x_48 = l_Lean_Syntax_node1(x_27, x_6, x_47); -if (lean_obj_tag(x_7) == 0) +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; +lean_inc(x_28); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_28); +lean_ctor_set(x_49, 1, x_40); +lean_ctor_set(x_49, 2, x_48); +lean_inc(x_49); +lean_inc(x_28); +x_50 = l_Lean_Syntax_node1(x_28, x_7, x_49); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_49; lean_object* x_50; -x_49 = l_Lean_Syntax_node6(x_27, x_8, x_32, x_45, x_40, x_48, x_47, x_42); -if (lean_is_scalar(x_30)) { - x_50 = lean_alloc_ctor(0, 2, 0); +lean_object* x_51; lean_object* x_52; +x_51 = l_Lean_Syntax_node6(x_28, x_9, x_33, x_47, x_42, x_50, x_49, x_44); +if (lean_is_scalar(x_31)) { + x_52 = lean_alloc_ctor(0, 2, 0); } else { - x_50 = x_30; + x_52 = x_31; } -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_29); -return x_50; +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_30); +return x_52; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_47); -x_51 = lean_ctor_get(x_7, 0); -lean_inc(x_51); -lean_dec(x_7); -x_52 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; -lean_inc(x_27); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_27); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array_mkArray2___rarg(x_53, x_51); -x_55 = l_Array_append___rarg(x_37, x_54); -lean_dec(x_54); -lean_inc(x_27); -x_56 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_56, 0, x_27); -lean_ctor_set(x_56, 1, x_39); -lean_ctor_set(x_56, 2, x_55); -x_57 = l_Lean_Syntax_node6(x_27, x_8, x_32, x_45, x_40, x_48, x_56, x_42); -if (lean_is_scalar(x_30)) { - x_58 = lean_alloc_ctor(0, 2, 0); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_49); +x_53 = lean_ctor_get(x_8, 0); +lean_inc(x_53); +lean_dec(x_8); +x_54 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; +lean_inc(x_28); +x_55 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_55, 0, x_28); +lean_ctor_set(x_55, 1, x_54); +x_56 = l_Array_mkArray2___rarg(x_55, x_53); +x_57 = l_Array_append___rarg(x_38, x_56); +lean_dec(x_56); +lean_inc(x_28); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_28); +lean_ctor_set(x_58, 1, x_40); +lean_ctor_set(x_58, 2, x_57); +x_59 = l_Lean_Syntax_node6(x_28, x_9, x_33, x_47, x_42, x_50, x_58, x_44); +if (lean_is_scalar(x_31)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - x_58 = x_30; + x_60 = x_31; } -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_29); -return x_58; +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_30); +return x_60; } } else { -lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_59 = lean_ctor_get(x_5, 0); -x_60 = 1; -x_61 = l_Lean_SourceInfo_fromRef(x_59, x_60); -x_62 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__5; -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Array_mkArray1___rarg(x_63); -x_65 = l_Array_append___rarg(x_37, x_64); -lean_dec(x_64); -lean_inc(x_27); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_27); -lean_ctor_set(x_66, 1, x_39); -lean_ctor_set(x_66, 2, x_65); -lean_inc(x_27); -x_67 = l_Lean_Syntax_node1(x_27, x_6, x_66); -if (lean_obj_tag(x_7) == 0) +lean_object* x_61; uint8_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_61 = lean_ctor_get(x_6, 0); +x_62 = 1; +x_63 = l_Lean_SourceInfo_fromRef(x_61, x_62); +x_64 = l_Lean_Elab_Term_CollectPatternVars_collect_processCtorApp___closed__5; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Array_mkArray1___rarg(x_65); +x_67 = l_Array_append___rarg(x_38, x_66); +lean_dec(x_66); +lean_inc(x_28); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_28); +lean_ctor_set(x_68, 1, x_40); +lean_ctor_set(x_68, 2, x_67); +lean_inc(x_28); +x_69 = l_Lean_Syntax_node1(x_28, x_7, x_68); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; -lean_inc(x_27); -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_27); -lean_ctor_set(x_69, 1, x_39); -lean_ctor_set(x_69, 2, x_68); -x_70 = l_Lean_Syntax_node6(x_27, x_8, x_32, x_45, x_40, x_67, x_69, x_42); -if (lean_is_scalar(x_30)) { - x_71 = lean_alloc_ctor(0, 2, 0); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__5; +lean_inc(x_28); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_28); +lean_ctor_set(x_71, 1, x_40); +lean_ctor_set(x_71, 2, x_70); +x_72 = l_Lean_Syntax_node6(x_28, x_9, x_33, x_47, x_42, x_69, x_71, x_44); +if (lean_is_scalar(x_31)) { + x_73 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_30; + x_73 = x_31; } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_29); -return x_71; +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_30); +return x_73; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_72 = lean_ctor_get(x_7, 0); -lean_inc(x_72); -lean_dec(x_7); -x_73 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; -lean_inc(x_27); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_27); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_Array_mkArray2___rarg(x_74, x_72); -x_76 = l_Array_append___rarg(x_37, x_75); -lean_dec(x_75); -lean_inc(x_27); -x_77 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_77, 0, x_27); -lean_ctor_set(x_77, 1, x_39); -lean_ctor_set(x_77, 2, x_76); -x_78 = l_Lean_Syntax_node6(x_27, x_8, x_32, x_45, x_40, x_67, x_77, x_42); -if (lean_is_scalar(x_30)) { - x_79 = lean_alloc_ctor(0, 2, 0); +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_74 = lean_ctor_get(x_8, 0); +lean_inc(x_74); +lean_dec(x_8); +x_75 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__6; +lean_inc(x_28); +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_28); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Array_mkArray2___rarg(x_76, x_74); +x_78 = l_Array_append___rarg(x_38, x_77); +lean_dec(x_77); +lean_inc(x_28); +x_79 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_79, 0, x_28); +lean_ctor_set(x_79, 1, x_40); +lean_ctor_set(x_79, 2, x_78); +x_80 = l_Lean_Syntax_node6(x_28, x_9, x_33, x_47, x_42, x_69, x_79, x_44); +if (lean_is_scalar(x_31)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_79 = x_30; + x_81 = x_31; } -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_29); -return x_79; +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_30); +return x_81; } } } } else { -uint8_t x_87; +uint8_t x_89; +lean_dec(x_18); lean_dec(x_17); -lean_dec(x_16); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -x_87 = !lean_is_exclusive(x_22); -if (x_87 == 0) +lean_dec(x_5); +x_89 = !lean_is_exclusive(x_23); +if (x_89 == 0) { -return x_22; +return x_23; } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_22, 0); -x_89 = lean_ctor_get(x_22, 1); -lean_inc(x_89); -lean_inc(x_88); -lean_dec(x_22); -x_90 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_90, 0, x_88); -lean_ctor_set(x_90, 1, x_89); -return x_90; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_23, 0); +x_91 = lean_ctor_get(x_23, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_23); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } @@ -11555,117 +11560,119 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; -x_16 = l_Lean_Syntax_getArgs(x_1); -if (lean_obj_tag(x_5) == 0) +lean_object* x_17; +x_17 = l_Lean_Syntax_getArgs(x_1); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__4; -x_18 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__5; -x_19 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__6; -x_20 = lean_box(0); -x_21 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_16, x_17, x_18, x_19, x_2, x_3, x_7, x_4, x_5, x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_16); -return x_21; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__4; +x_19 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__5; +x_20 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__6; +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_17, x_18, x_19, x_20, x_2, x_3, x_4, x_8, x_5, x_6, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_17); +return x_22; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -lean_dec(x_16); -lean_dec(x_7); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_17); +lean_dec(x_8); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_22 = lean_ctor_get(x_5, 0); -x_23 = l_Lean_Syntax_TSepArray_getElems___rarg(x_22); -x_24 = lean_box(2); -x_25 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_23); -x_27 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__8; -x_28 = l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___spec__1(x_26, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); +lean_dec(x_2); +x_23 = lean_ctor_get(x_6, 0); +x_24 = l_Lean_Syntax_TSepArray_getElems___rarg(x_23); +x_25 = lean_box(2); +x_26 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2___closed__3; +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_24); +x_28 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___closed__8; +x_29 = l_Lean_throwErrorAt___at___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_processVar___spec__1(x_27, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_26); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +lean_dec(x_27); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -return x_28; +return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_unsigned_to_nat(4u); -x_17 = l_Lean_Syntax_getArg(x_1, x_16); -x_18 = l_Lean_Syntax_isNone(x_17); -if (x_18 == 0) +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_unsigned_to_nat(4u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); +x_19 = l_Lean_Syntax_isNone(x_18); +if (x_19 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_unsigned_to_nat(2u); -lean_inc(x_17); -x_20 = l_Lean_Syntax_matchesNull(x_17, x_19); -if (x_20 == 0) +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(2u); +lean_inc(x_18); +x_21 = l_Lean_Syntax_matchesNull(x_18, x_20); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_21 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_22 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -return x_21; +return x_22; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_unsigned_to_nat(1u); -x_23 = l_Lean_Syntax_getArg(x_17, x_22); -lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = lean_box(0); -x_26 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_7, x_3, x_4, x_5, x_25, x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_23 = lean_unsigned_to_nat(1u); +x_24 = l_Lean_Syntax_getArg(x_18, x_23); +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_box(0); +x_27 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_8, x_4, x_5, x_6, x_26, x_25, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_27; } } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_17); -x_27 = lean_box(0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_18); x_28 = lean_box(0); -x_29 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_7, x_3, x_4, x_5, x_28, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_29; +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_2, x_3, x_8, x_4, x_5, x_6, x_29, x_28, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_30; } } } @@ -11673,7 +11680,7 @@ static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -11689,24 +11696,41 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__3; +x_2 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__4; +x_3 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_finalize___closed__5; +x_4 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_13 = lean_unsigned_to_nat(2u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); -x_15 = lean_unsigned_to_nat(3u); -x_16 = l_Lean_Syntax_getArg(x_1, x_15); -x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__2; -lean_inc(x_16); -x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); -if (x_18 == 0) +x_15 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__2; +lean_inc(x_14); +x_16 = l_Lean_Syntax_isOfKind(x_14, x_15); +if (x_16 == 0) { -lean_object* x_19; -lean_dec(x_16); +lean_object* x_17; lean_dec(x_14); lean_dec(x_2); -x_19 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_17 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -11714,28 +11738,26 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_19; +return x_17; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Lean_Syntax_getArg(x_16, x_20); -lean_dec(x_16); -x_22 = l_Lean_Syntax_isNone(x_21); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_unsigned_to_nat(1u); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Lean_Syntax_getArg(x_14, x_18); +lean_dec(x_14); +x_20 = lean_unsigned_to_nat(3u); +x_21 = l_Lean_Syntax_getArg(x_1, x_20); +x_22 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4; lean_inc(x_21); -x_24 = l_Lean_Syntax_matchesNull(x_21, x_23); -if (x_24 == 0) +x_23 = l_Lean_Syntax_isOfKind(x_21, x_22); +if (x_23 == 0) { -lean_object* x_25; +lean_object* x_24; lean_dec(x_21); -lean_dec(x_14); +lean_dec(x_19); lean_dec(x_2); -x_25 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_24 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -11743,31 +11765,60 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_25; +return x_24; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_26 = l_Lean_Syntax_getArg(x_21, x_20); +lean_object* x_25; uint8_t x_26; +x_25 = l_Lean_Syntax_getArg(x_21, x_18); lean_dec(x_21); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_26); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_14, x_17, x_2, x_4, x_28, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_27); -lean_dec(x_14); +x_26 = l_Lean_Syntax_isNone(x_25); +if (x_26 == 0) +{ +lean_object* x_27; uint8_t x_28; +x_27 = lean_unsigned_to_nat(1u); +lean_inc(x_25); +x_28 = l_Lean_Syntax_matchesNull(x_25, x_27); +if (x_28 == 0) +{ +lean_object* x_29; +lean_dec(x_25); +lean_dec(x_19); +lean_dec(x_2); +x_29 = l___private_Lean_Elab_PatternVar_0__Lean_Elab_Term_CollectPatternVars_throwInvalidPattern___rarg(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); return x_29; } +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = l_Lean_Syntax_getArg(x_25, x_18); +lean_dec(x_25); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_19, x_15, x_22, x_2, x_4, x_32, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_31); +lean_dec(x_19); +return x_33; +} } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_21); -x_30 = lean_box(0); -x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_14, x_17, x_2, x_4, x_31, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_14); -return x_32; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_25); +x_34 = lean_box(0); +x_35 = lean_box(0); +x_36 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_19, x_15, x_22, x_2, x_4, x_35, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_19); +return x_36; +} } } } @@ -14809,42 +14860,43 @@ lean_object* x_15 = _args[14]; lean_object* x_16 = _args[15]; lean_object* x_17 = _args[16]; lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; _start: { -lean_object* x_19; -x_19 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +lean_object* x_20; +x_20 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -return x_19; +return x_20; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; -x_16 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_17; +x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -return x_16; +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { _start: { -lean_object* x_16; -x_16 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_object* x_17; +x_17 = l_Lean_Elab_Term_CollectPatternVars_collect___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -return x_16; +return x_17; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -17495,6 +17547,10 @@ l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__1 = _init_l_Le lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__1); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__2 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__2); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__3); +l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__5___closed__4); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__1 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__1); l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__2 = _init_l_Lean_Elab_Term_CollectPatternVars_collect___lambda__10___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/SetOption.c b/stage0/stdlib/Lean/Elab/SetOption.c index 5d3d2167079f..a8ba2dc6acfa 100644 --- a/stage0/stdlib/Lean/Elab/SetOption.c +++ b/stage0/stdlib/Lean/Elab/SetOption.c @@ -655,7 +655,7 @@ x_13 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_13, 0, x_1); lean_ctor_set(x_13, 1, x_2); lean_ctor_set(x_13, 2, x_12); -x_14 = lean_alloc_ctor(4, 1, 0); +x_14 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_14, 0, x_13); lean_inc(x_3); x_15 = l_Lean_Elab_pushInfoLeaf___rarg(x_3, x_4, x_14); diff --git a/stage0/stdlib/Lean/Elab/StructInst.c b/stage0/stdlib/Lean/Elab/StructInst.c index aee6b74c6e48..65110fc55143 100644 --- a/stage0/stdlib/Lean/Elab/StructInst.c +++ b/stage0/stdlib/Lean/Elab/StructInst.c @@ -45,8 +45,6 @@ static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___la static lean_object* l_Lean_Elab_Term_StructInst_formatField___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst__1___closed__4; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_StructInst_DefaultFields_collectStructNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -109,7 +107,6 @@ uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_isDefaultMissing_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__4___closed__3; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,9 +121,8 @@ static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructIns static lean_object* l_panic___at_Lean_Elab_Term_StructInst_DefaultFields_step___spec__1___closed__3; size_t lean_uint64_to_usize(uint64_t); lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMissing_x3f___spec__1___rarg___lambda__2(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2; lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); @@ -147,7 +143,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__4___boxed(lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__3___closed__5; lean_object* l_Lean_findField_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -199,6 +194,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMi LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_isDefaultMissing_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_getFieldValue_x3f___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_getFieldValue_x3f___boxed(lean_object*, lean_object*); @@ -206,13 +202,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_Struct_modifyFieldsM___rarg LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_StructInst_throwFailedToElabField___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__10___closed__5; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__10___closed__6; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_isDefaultMissing_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__7___closed__1; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__11; +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12; lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_setType(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -237,7 +234,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Stru static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_Field_toSyntax___closed__2; lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; @@ -269,6 +266,7 @@ lean_object* l_Lean_LocalDecl_index(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___closed__2; LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__7(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14; lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; lean_object* l_Lean_Meta_project_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -281,6 +279,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructI uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources_go___lambda__1___closed__4; lean_object* l_Lean_PersistentArray_set___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___closed__2; @@ -289,6 +288,7 @@ static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructIns LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_propagate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPathToBaseStructure_x3f(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -309,12 +309,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_instInhabitedField(lean_obj LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_mkDefaultValueAux_x3f___closed__3; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_StructInst_formatField___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMissing_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___closed__2; static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9; lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_markDefaultMissing(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -358,6 +359,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesi lean_object* lean_nat_to_int(lean_object*); lean_object* l_List_findSome_x3f___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getFieldIdx___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_instToFormatFieldLHS___closed__3; @@ -385,7 +387,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_FieldVal_toSyntax(lean_obje LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing_go___spec__5___rarg___lambda__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(lean_object*, lean_object*, size_t, size_t); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_components(lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -408,6 +409,7 @@ lean_object* l_Lean_Meta_reduceProjOf_x3f(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4; static lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; static lean_object* l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_collectStructNames___boxed(lean_object*, lean_object*); @@ -471,8 +473,6 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_StructInst LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_instInhabitedExplicitSourceInfo; LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg___closed__2; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__3; @@ -504,7 +504,6 @@ lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__5___boxed(lean_object*, lean_object*); @@ -558,8 +557,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_StructInst_0 static lean_object* l_List_format___at_Lean_Elab_Term_StructInst_formatStruct___spec__3___closed__5; lean_object* l_Lean_Elab_Term_mkTacticMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13; static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_formatStruct___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11; extern lean_object* l_Std_Format_defWidth; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -641,7 +642,6 @@ static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructIns LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__2; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6; LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_findDefaultMissing_x3f___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___lambda__1___boxed(lean_object*); @@ -678,6 +678,7 @@ static lean_object* l_Lean_Elab_Term_StructInst_instToStringStruct___closed__1; lean_object* lean_environment_main_module(lean_object*); static lean_object* l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__1; lean_object* l_Lean_Meta_synthInstance_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6; static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___closed__1; static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev_declRange__1___closed__4; @@ -697,6 +698,7 @@ static lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3; static lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__5; uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_instToFormatFieldLHS(lean_object*); @@ -716,7 +718,6 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructIns static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5; uint64_t l_Lean_Name_hash___override(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__3; static lean_object* l_Lean_Elab_Term_StructInst_formatStruct___closed__5; @@ -802,6 +803,7 @@ static lean_object* l_Lean_Elab_Term_StructInst_instInhabitedStruct___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Term_StructInst_Source_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4; static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing_go___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -895,10 +897,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev static lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_tryToSynthesizeDefault_loop___closed__2; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_404_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStructInstAux___spec__5___rarg___closed__2; +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10; lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1; static lean_object* l_Lean_Elab_Term_StructInst_formatStruct___closed__6; lean_object* l_Array_contains___at_Lean_registerInternalExceptionId___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_FieldVal_toSyntax___closed__1; @@ -917,6 +919,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_StructI LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_isDefaultMissing_x3f___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg___closed__5; +static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev_docString__1(lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -945,7 +948,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMis LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_mkDefaultValueAux_x3f___lambda__3(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_StructInst_formatField___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_StructInst_throwFailedToElabField___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getForallBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_propagateExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -957,11 +959,10 @@ static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___c LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__10___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_markDefaultMissing___closed__1; static lean_object* l_Lean_Elab_Term_StructInst_instToFormatFieldLHS___closed__4; -static lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_throwFailedToElabField___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_Struct_modifyFields(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_StructInst_DefaultFields_mkDefaultValueAux_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); @@ -975,6 +976,7 @@ static lean_object* l_Lean_Elab_Term_StructInst_formatStruct___closed__10; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_StructInst_Field_isSimple___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_formatField___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843_(lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_mkProjStx_x3f___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_reduce___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); @@ -1647,255 +1649,259 @@ x_1 = lean_mk_string_unchecked("with", 4, 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = l_Lean_Syntax_getArgs(x_1); -x_11 = l_Lean_Syntax_TSepArray_getElems___rarg(x_10); -lean_dec(x_10); -x_12 = lean_array_get_size(x_11); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -if (x_14 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = l_Lean_Syntax_getArgs(x_1); +x_12 = l_Lean_Syntax_TSepArray_getElems___rarg(x_11); +lean_dec(x_11); +x_13 = lean_array_get_size(x_12); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_lt(x_14, x_13); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; +lean_object* x_16; lean_object* x_17; +lean_dec(x_13); lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); +lean_dec(x_8); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_15 = lean_box(1); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_9); -return x_16; +lean_dec(x_2); +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -size_t x_17; size_t x_18; lean_object* x_19; uint8_t x_20; -x_17 = 0; -x_18 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_19 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; -x_20 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(x_19, x_11, x_17, x_18); -if (x_20 == 0) +size_t x_18; size_t x_19; lean_object* x_20; uint8_t x_21; +x_18 = 0; +x_19 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_20 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__2; +x_21 = l_Array_anyMUnsafe_any___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__1(x_20, x_12, x_18, x_19); +if (x_21 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_11); -lean_dec(x_7); +lean_object* x_22; lean_object* x_23; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_21 = lean_box(1); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_9); -return x_22; +lean_dec(x_2); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_10); +return x_23; } else { -size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_23 = lean_array_size(x_11); -x_24 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4; -x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; -x_26 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_19, x_24, x_25, x_23, x_17, x_11, x_8, x_9); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +size_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; size_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_24 = lean_array_size(x_12); +x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__4; +x_26 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__5; +x_27 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2(x_20, x_25, x_26, x_24, x_18, x_12, x_9, x_10); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_29 = x_26; +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_30 = x_27; } else { - lean_dec_ref(x_26); - x_29 = lean_box(0); + lean_dec_ref(x_27); + x_30 = lean_box(0); } -x_30 = lean_ctor_get(x_8, 5); -x_31 = 0; -x_32 = l_Lean_SourceInfo_fromRef(x_30, x_31); -x_33 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; -lean_inc(x_32); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = lean_array_size(x_27); -x_36 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; -x_37 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__3(x_36, x_35, x_17, x_27); -x_38 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8; -x_39 = l_Lean_Syntax_SepArray_ofElems(x_38, x_37); -lean_dec(x_37); -x_40 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; -x_41 = l_Array_append___rarg(x_40, x_39); -lean_dec(x_39); -x_42 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; -lean_inc(x_32); -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_32); -lean_ctor_set(x_43, 1, x_42); -lean_ctor_set(x_43, 2, x_41); -x_44 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9; -lean_inc(x_32); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_32); -lean_ctor_set(x_45, 1, x_44); -if (lean_obj_tag(x_5) == 0) +x_31 = lean_ctor_get(x_9, 5); +x_32 = 0; +x_33 = l_Lean_SourceInfo_fromRef(x_31, x_32); +x_34 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; +lean_inc(x_33); +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = lean_array_size(x_28); +x_37 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__6; +x_38 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__3(x_37, x_36, x_18, x_28); +x_39 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__8; +x_40 = l_Lean_Syntax_SepArray_ofElems(x_39, x_38); +lean_dec(x_38); +x_41 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +x_42 = l_Array_append___rarg(x_41, x_40); +lean_dec(x_40); +x_43 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +lean_inc(x_33); +x_44 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_44, 0, x_33); +lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_44, 2, x_42); +lean_inc(x_33); +x_45 = l_Lean_Syntax_node1(x_33, x_2, x_44); +x_46 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__9; +lean_inc(x_33); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_33); +lean_ctor_set(x_47, 1, x_46); +if (lean_obj_tag(x_6) == 0) { -x_46 = x_40; -goto block_83; +x_48 = x_41; +goto block_85; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_84 = lean_ctor_get(x_5, 0); -x_85 = l_Array_append___rarg(x_40, x_84); -lean_inc(x_32); -x_86 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_86, 0, x_32); -lean_ctor_set(x_86, 1, x_42); -lean_ctor_set(x_86, 2, x_85); -x_87 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__12; -lean_inc(x_32); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_32); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Array_mkArray2___rarg(x_86, x_88); -x_46 = x_89; -goto block_83; +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_86 = lean_ctor_get(x_6, 0); +x_87 = l_Array_append___rarg(x_41, x_86); +lean_inc(x_33); +x_88 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_88, 0, x_33); +lean_ctor_set(x_88, 1, x_43); +lean_ctor_set(x_88, 2, x_87); +x_89 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__12; +lean_inc(x_33); +x_90 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_90, 0, x_33); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Array_mkArray2___rarg(x_88, x_90); +x_48 = x_91; +goto block_85; } -block_83: +block_85: { -lean_object* x_47; lean_object* x_48; -x_47 = l_Array_append___rarg(x_40, x_46); -lean_dec(x_46); -lean_inc(x_32); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_32); -lean_ctor_set(x_48, 1, x_42); -lean_ctor_set(x_48, 2, x_47); -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; -lean_inc(x_32); +lean_object* x_49; lean_object* x_50; +x_49 = l_Array_append___rarg(x_41, x_48); +lean_dec(x_48); +lean_inc(x_33); x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_32); -lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 0, x_33); +lean_ctor_set(x_50, 1, x_43); lean_ctor_set(x_50, 2, x_49); -lean_inc(x_50); -lean_inc(x_32); -x_51 = l_Lean_Syntax_node1(x_32, x_3, x_50); -if (lean_obj_tag(x_7) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_52; lean_object* x_53; -x_52 = l_Lean_Syntax_node6(x_32, x_4, x_34, x_48, x_43, x_51, x_50, x_45); -if (lean_is_scalar(x_29)) { - x_53 = lean_alloc_ctor(0, 2, 0); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; +lean_inc(x_33); +x_52 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_52, 0, x_33); +lean_ctor_set(x_52, 1, x_43); +lean_ctor_set(x_52, 2, x_51); +lean_inc(x_52); +lean_inc(x_33); +x_53 = l_Lean_Syntax_node1(x_33, x_4, x_52); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_54; lean_object* x_55; +x_54 = l_Lean_Syntax_node6(x_33, x_5, x_35, x_50, x_45, x_53, x_52, x_47); +if (lean_is_scalar(x_30)) { + x_55 = lean_alloc_ctor(0, 2, 0); } else { - x_53 = x_29; + x_55 = x_30; } -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_28); -return x_53; +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_29); +return x_55; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_50); -x_54 = lean_ctor_get(x_7, 0); -lean_inc(x_54); -lean_dec(x_7); -x_55 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; -lean_inc(x_32); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_32); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Array_mkArray2___rarg(x_56, x_54); -x_58 = l_Array_append___rarg(x_40, x_57); -lean_dec(x_57); -lean_inc(x_32); -x_59 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_59, 0, x_32); -lean_ctor_set(x_59, 1, x_42); -lean_ctor_set(x_59, 2, x_58); -x_60 = l_Lean_Syntax_node6(x_32, x_4, x_34, x_48, x_43, x_51, x_59, x_45); -if (lean_is_scalar(x_29)) { - x_61 = lean_alloc_ctor(0, 2, 0); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +lean_dec(x_52); +x_56 = lean_ctor_get(x_8, 0); +lean_inc(x_56); +lean_dec(x_8); +x_57 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +lean_inc(x_33); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_33); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Array_mkArray2___rarg(x_58, x_56); +x_60 = l_Array_append___rarg(x_41, x_59); +lean_dec(x_59); +lean_inc(x_33); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_33); +lean_ctor_set(x_61, 1, x_43); +lean_ctor_set(x_61, 2, x_60); +x_62 = l_Lean_Syntax_node6(x_33, x_5, x_35, x_50, x_45, x_53, x_61, x_47); +if (lean_is_scalar(x_30)) { + x_63 = lean_alloc_ctor(0, 2, 0); } else { - x_61 = x_29; + x_63 = x_30; } -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_28); -return x_61; +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_29); +return x_63; } } else { -lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_62 = lean_ctor_get(x_2, 0); -x_63 = 1; -x_64 = l_Lean_SourceInfo_fromRef(x_62, x_63); -x_65 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11; -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Array_mkArray1___rarg(x_66); -x_68 = l_Array_append___rarg(x_40, x_67); -lean_dec(x_67); -lean_inc(x_32); -x_69 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_69, 0, x_32); -lean_ctor_set(x_69, 1, x_42); -lean_ctor_set(x_69, 2, x_68); -lean_inc(x_32); -x_70 = l_Lean_Syntax_node1(x_32, x_3, x_69); -if (lean_obj_tag(x_7) == 0) +lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_64 = lean_ctor_get(x_3, 0); +x_65 = 1; +x_66 = l_Lean_SourceInfo_fromRef(x_64, x_65); +x_67 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__11; +x_68 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +x_69 = l_Array_mkArray1___rarg(x_68); +x_70 = l_Array_append___rarg(x_41, x_69); +lean_dec(x_69); +lean_inc(x_33); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_33); +lean_ctor_set(x_71, 1, x_43); +lean_ctor_set(x_71, 2, x_70); +lean_inc(x_33); +x_72 = l_Lean_Syntax_node1(x_33, x_4, x_71); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; -lean_inc(x_32); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_32); -lean_ctor_set(x_72, 1, x_42); -lean_ctor_set(x_72, 2, x_71); -x_73 = l_Lean_Syntax_node6(x_32, x_4, x_34, x_48, x_43, x_70, x_72, x_45); -if (lean_is_scalar(x_29)) { - x_74 = lean_alloc_ctor(0, 2, 0); -} else { - x_74 = x_29; -} -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_28); -return x_74; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__10; +lean_inc(x_33); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_33); +lean_ctor_set(x_74, 1, x_43); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Lean_Syntax_node6(x_33, x_5, x_35, x_50, x_45, x_72, x_74, x_47); +if (lean_is_scalar(x_30)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_30; +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_29); +return x_76; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_75 = lean_ctor_get(x_7, 0); -lean_inc(x_75); -lean_dec(x_7); -x_76 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; -lean_inc(x_32); -x_77 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_77, 0, x_32); -lean_ctor_set(x_77, 1, x_76); -x_78 = l_Array_mkArray2___rarg(x_77, x_75); -x_79 = l_Array_append___rarg(x_40, x_78); -lean_dec(x_78); -lean_inc(x_32); -x_80 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_80, 0, x_32); -lean_ctor_set(x_80, 1, x_42); -lean_ctor_set(x_80, 2, x_79); -x_81 = l_Lean_Syntax_node6(x_32, x_4, x_34, x_48, x_43, x_70, x_80, x_45); -if (lean_is_scalar(x_29)) { - x_82 = lean_alloc_ctor(0, 2, 0); +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_77 = lean_ctor_get(x_8, 0); +lean_inc(x_77); +lean_dec(x_8); +x_78 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__11; +lean_inc(x_33); +x_79 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_79, 0, x_33); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Array_mkArray2___rarg(x_79, x_77); +x_81 = l_Array_append___rarg(x_41, x_80); +lean_dec(x_80); +lean_inc(x_33); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_33); +lean_ctor_set(x_82, 1, x_43); +lean_ctor_set(x_82, 2, x_81); +x_83 = l_Lean_Syntax_node6(x_33, x_5, x_35, x_50, x_45, x_72, x_82, x_47); +if (lean_is_scalar(x_30)) { + x_84 = lean_alloc_ctor(0, 2, 0); } else { - x_82 = x_29; + x_84 = x_30; } -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_28); -return x_82; +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_29); +return x_84; } } } @@ -1903,52 +1909,53 @@ return x_82; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_unsigned_to_nat(4u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Syntax_isNone(x_11); -if (x_12 == 0) +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_unsigned_to_nat(4u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +x_13 = l_Lean_Syntax_isNone(x_12); +if (x_13 == 0) { -lean_object* x_13; uint8_t x_14; -x_13 = lean_unsigned_to_nat(2u); -lean_inc(x_11); -x_14 = l_Lean_Syntax_matchesNull(x_11, x_13); -if (x_14 == 0) +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(2u); +lean_inc(x_12); +x_15 = l_Lean_Syntax_matchesNull(x_12, x_14); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; -lean_dec(x_11); +lean_object* x_16; lean_object* x_17; +lean_dec(x_12); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_15 = lean_box(1); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_9); -return x_16; +x_16 = lean_box(1); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +return x_17; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_unsigned_to_nat(1u); -x_18 = l_Lean_Syntax_getArg(x_11, x_17); -lean_dec(x_11); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_box(0); -x_21 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_7, x_3, x_4, x_5, x_20, x_19, x_8, x_9); -return x_21; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_Syntax_getArg(x_12, x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box(0); +x_22 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_8, x_4, x_5, x_6, x_21, x_20, x_9, x_10); +return x_22; } } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_11); -x_22 = lean_box(0); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_12); x_23 = lean_box(0); -x_24 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_7, x_3, x_4, x_5, x_23, x_22, x_8, x_9); -return x_24; +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_2, x_3, x_8, x_4, x_5, x_6, x_24, x_23, x_9, x_10); +return x_25; } } } @@ -1956,7 +1963,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbre _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -1972,77 +1979,117 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__5; +x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__6; +x_3 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__7; +x_4 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; x_7 = lean_unsigned_to_nat(2u); x_8 = l_Lean_Syntax_getArg(x_1, x_7); -x_9 = lean_unsigned_to_nat(3u); -x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_11 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; -lean_inc(x_10); -x_12 = l_Lean_Syntax_isOfKind(x_10, x_11); -if (x_12 == 0) +x_9 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +lean_inc(x_8); +x_10 = l_Lean_Syntax_isOfKind(x_8, x_9); +if (x_10 == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); +lean_object* x_11; lean_object* x_12; lean_dec(x_8); lean_dec(x_2); -x_13 = lean_box(1); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_6); -return x_14; +x_11 = lean_box(1); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_6); +return x_12; } else { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Lean_Syntax_getArg(x_10, x_15); -lean_dec(x_10); -x_17 = l_Lean_Syntax_isNone(x_16); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(1u); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_Syntax_getArg(x_8, x_13); +lean_dec(x_8); +x_15 = lean_unsigned_to_nat(3u); +x_16 = l_Lean_Syntax_getArg(x_1, x_15); +x_17 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4; lean_inc(x_16); -x_19 = l_Lean_Syntax_matchesNull(x_16, x_18); -if (x_19 == 0) +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) { -lean_object* x_20; lean_object* x_21; +lean_object* x_19; lean_object* x_20; lean_dec(x_16); -lean_dec(x_8); +lean_dec(x_14); lean_dec(x_2); -x_20 = lean_box(1); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_6); -return x_21; +x_19 = lean_box(1); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_6); +return x_20; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = l_Lean_Syntax_getArg(x_16, x_15); +lean_object* x_21; uint8_t x_22; +x_21 = l_Lean_Syntax_getArg(x_16, x_13); lean_dec(x_16); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_8, x_11, x_2, x_4, x_24, x_23, x_5, x_6); -lean_dec(x_23); -lean_dec(x_8); -return x_25; +x_22 = l_Lean_Syntax_isNone(x_21); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_unsigned_to_nat(1u); +lean_inc(x_21); +x_24 = l_Lean_Syntax_matchesNull(x_21, x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_21); +lean_dec(x_14); +lean_dec(x_2); +x_25 = lean_box(1); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_6); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = l_Lean_Syntax_getArg(x_21, x_13); +lean_dec(x_21); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_14, x_9, x_17, x_2, x_4, x_29, x_28, x_5, x_6); +lean_dec(x_28); +lean_dec(x_14); +return x_30; } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_16); -x_26 = lean_box(0); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_8, x_11, x_2, x_4, x_27, x_26, x_5, x_6); -lean_dec(x_8); -return x_28; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_21); +x_31 = lean_box(0); +x_32 = lean_box(0); +x_33 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_14, x_9, x_17, x_2, x_4, x_32, x_31, x_5, x_6); +lean_dec(x_14); +return x_33; +} } } } @@ -2160,31 +2207,31 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); +lean_object* x_11; +x_11 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); +lean_dec(x_3); lean_dec(x_1); -return x_10; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_11; +x_11 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -return x_10; +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -4814,166 +4861,168 @@ return x_61; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; x_9 = lean_box(0); x_10 = lean_unsigned_to_nat(2u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Syntax_getSepArgs(x_11); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Lean_Syntax_getArg(x_11, x_12); lean_dec(x_11); -x_13 = lean_array_get_size(x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 0) -{ -lean_object* x_16; +x_14 = l_Lean_Syntax_getSepArgs(x_13); lean_dec(x_13); -lean_dec(x_12); +x_15 = lean_array_get_size(x_14); +x_16 = lean_nat_dec_lt(x_12, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_6); lean_dec(x_2); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_8); -return x_16; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_9); +lean_ctor_set(x_17, 1, x_8); +return x_17; } else { -uint8_t x_17; -x_17 = lean_nat_dec_le(x_13, x_13); -if (x_17 == 0) +uint8_t x_18; +x_18 = lean_nat_dec_le(x_15, x_15); +if (x_18 == 0) { -lean_object* x_18; -lean_dec(x_13); -lean_dec(x_12); +lean_object* x_19; +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_6); lean_dec(x_2); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_9); -lean_ctor_set(x_18, 1, x_8); -return x_18; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_9); +lean_ctor_set(x_19, 1, x_8); +return x_19; } else { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = 0; -x_20 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_21 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3(x_12, x_19, x_20, x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_12); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); +size_t x_20; size_t x_21; lean_object* x_22; +x_20 = 0; +x_21 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_22 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3(x_14, x_20, x_21, x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_14); if (lean_obj_tag(x_22) == 0) { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_21, 0); -lean_dec(x_24); -lean_ctor_set(x_21, 0, x_9); -return x_21; +uint8_t x_24; +x_24 = !lean_is_exclusive(x_22); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_22, 0); +lean_dec(x_25); +lean_ctor_set(x_22, 0, x_9); +return x_22; } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_9); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } else { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_21); -if (x_27 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_22); +if (x_28 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_28 = lean_ctor_get(x_21, 0); -lean_dec(x_28); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; x_29 = lean_ctor_get(x_22, 0); -lean_inc(x_29); -x_30 = l_Lean_Syntax_getArg(x_29, x_14); lean_dec(x_29); -x_31 = l_Lean_Syntax_getArg(x_30, x_14); +x_30 = lean_ctor_get(x_23, 0); +lean_inc(x_30); +x_31 = l_Lean_Syntax_getArg(x_30, x_12); lean_dec(x_30); -x_32 = l_Lean_Syntax_getKind(x_31); -x_33 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; -x_34 = lean_name_eq(x_32, x_33); -lean_dec(x_32); -if (x_34 == 0) +x_32 = l_Lean_Syntax_getArg(x_31, x_12); +lean_dec(x_31); +x_33 = l_Lean_Syntax_getKind(x_32); +x_34 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; +x_35 = lean_name_eq(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) { -lean_dec(x_22); -lean_ctor_set(x_21, 0, x_9); -return x_21; +lean_dec(x_23); +lean_ctor_set(x_22, 0, x_9); +return x_22; } else { -return x_21; +return x_22; } } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_35 = lean_ctor_get(x_21, 1); -lean_inc(x_35); -lean_dec(x_21); -x_36 = lean_ctor_get(x_22, 0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_36 = lean_ctor_get(x_22, 1); lean_inc(x_36); -x_37 = l_Lean_Syntax_getArg(x_36, x_14); -lean_dec(x_36); -x_38 = l_Lean_Syntax_getArg(x_37, x_14); -lean_dec(x_37); -x_39 = l_Lean_Syntax_getKind(x_38); -x_40 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; -x_41 = lean_name_eq(x_39, x_40); -lean_dec(x_39); -if (x_41 == 0) -{ -lean_object* x_42; lean_dec(x_22); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_9); -lean_ctor_set(x_42, 1, x_35); -return x_42; -} -else +x_37 = lean_ctor_get(x_23, 0); +lean_inc(x_37); +x_38 = l_Lean_Syntax_getArg(x_37, x_12); +lean_dec(x_37); +x_39 = l_Lean_Syntax_getArg(x_38, x_12); +lean_dec(x_38); +x_40 = l_Lean_Syntax_getKind(x_39); +x_41 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; +x_42 = lean_name_eq(x_40, x_41); +lean_dec(x_40); +if (x_42 == 0) { lean_object* x_43; +lean_dec(x_23); x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_22); -lean_ctor_set(x_43, 1, x_35); +lean_ctor_set(x_43, 0, x_9); +lean_ctor_set(x_43, 1, x_36); return x_43; } +else +{ +lean_object* x_44; +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_23); +lean_ctor_set(x_44, 1, x_36); +return x_44; +} } } } else { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_21); -if (x_44 == 0) +uint8_t x_45; +x_45 = !lean_is_exclusive(x_22); +if (x_45 == 0) { -return x_21; +return x_22; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_21, 0); -x_46 = lean_ctor_get(x_21, 1); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_22, 0); +x_47 = lean_ctor_get(x_22, 1); +lean_inc(x_47); lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_21); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_dec(x_22); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } } } @@ -5820,24 +5869,24 @@ x_48 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyO x_49 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_48, x_6, x_7, x_8, x_9, x_10, x_11, x_25); if (x_36 == 0) { -lean_object* x_106; -x_106 = l_Lean_Syntax_getArg(x_33, x_15); +lean_object* x_114; +x_114 = l_Lean_Syntax_getArg(x_33, x_15); lean_dec(x_33); -x_50 = x_106; -goto block_105; +x_50 = x_114; +goto block_113; } else { x_50 = x_33; -goto block_105; +goto block_113; } -block_105: +block_113: { uint8_t x_51; x_51 = !lean_is_exclusive(x_49); if (x_51 == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; x_52 = lean_ctor_get(x_49, 0); x_53 = lean_ctor_get(x_49, 1); lean_ctor_set_tag(x_49, 1); @@ -5858,674 +5907,701 @@ x_60 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_60, 0, x_41); lean_ctor_set(x_60, 1, x_42); lean_ctor_set(x_60, 2, x_59); -x_61 = lean_unsigned_to_nat(2u); -x_62 = l_Lean_Syntax_setArg(x_47, x_61, x_60); -x_63 = lean_unbox(x_52); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_30); +x_62 = lean_array_mk(x_61); +x_63 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_41); +lean_ctor_set(x_64, 1, x_63); +lean_ctor_set(x_64, 2, x_62); +x_65 = lean_unsigned_to_nat(2u); +x_66 = l_Lean_Syntax_setArg(x_47, x_65, x_64); +x_67 = lean_unbox(x_52); lean_dec(x_52); -if (x_63 == 0) +if (x_67 == 0) { -lean_object* x_64; lean_object* x_65; -x_64 = lean_box(0); -x_65 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_62, x_48, x_4, x_2, x_64, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +lean_object* x_68; lean_object* x_69; +x_68 = lean_box(0); +x_69 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_66, x_48, x_4, x_2, x_68, x_6, x_7, x_8, x_9, x_10, x_11, x_53); lean_dec(x_14); -return x_65; +return x_69; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_inc(x_2); -x_66 = l_Lean_MessageData_ofSyntax(x_2); -x_67 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -x_68 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; -x_70 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -lean_inc(x_62); -x_71 = l_Lean_MessageData_ofSyntax(x_62); +x_70 = l_Lean_MessageData_ofSyntax(x_2); +x_71 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_67); -x_74 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_48, x_73, x_6, x_7, x_8, x_9, x_10, x_11, x_53); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); -lean_inc(x_76); -lean_dec(x_74); -x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_62, x_48, x_4, x_2, x_75, x_6, x_7, x_8, x_9, x_10, x_11, x_76); -lean_dec(x_75); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_70); +x_73 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +lean_inc(x_66); +x_75 = l_Lean_MessageData_ofSyntax(x_66); +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_71); +x_78 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_48, x_77, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); +x_81 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_66, x_48, x_4, x_2, x_79, x_6, x_7, x_8, x_9, x_10, x_11, x_80); +lean_dec(x_79); lean_dec(x_14); -return x_77; +return x_81; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_78 = lean_ctor_get(x_49, 0); -x_79 = lean_ctor_get(x_49, 1); -lean_inc(x_79); -lean_inc(x_78); +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_82 = lean_ctor_get(x_49, 0); +x_83 = lean_ctor_get(x_49, 1); +lean_inc(x_83); +lean_inc(x_82); lean_dec(x_49); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_50); -lean_ctor_set(x_80, 1, x_22); -x_81 = lean_array_mk(x_80); -x_82 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_41); -lean_ctor_set(x_83, 1, x_82); -lean_ctor_set(x_83, 2, x_81); -x_84 = l_Lean_Syntax_setArg(x_1, x_13, x_83); -x_85 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_30); -x_86 = lean_array_mk(x_85); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_50); +lean_ctor_set(x_84, 1, x_22); +x_85 = lean_array_mk(x_84); +x_86 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; x_87 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_87, 0, x_41); -lean_ctor_set(x_87, 1, x_42); -lean_ctor_set(x_87, 2, x_86); -x_88 = lean_unsigned_to_nat(2u); -x_89 = l_Lean_Syntax_setArg(x_47, x_88, x_87); -x_90 = lean_unbox(x_78); -lean_dec(x_78); -if (x_90 == 0) +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_87, 2, x_85); +x_88 = l_Lean_Syntax_setArg(x_1, x_13, x_87); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_30); +x_90 = lean_array_mk(x_89); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_41); +lean_ctor_set(x_91, 1, x_42); +lean_ctor_set(x_91, 2, x_90); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_30); +x_93 = lean_array_mk(x_92); +x_94 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_41); +lean_ctor_set(x_95, 1, x_94); +lean_ctor_set(x_95, 2, x_93); +x_96 = lean_unsigned_to_nat(2u); +x_97 = l_Lean_Syntax_setArg(x_47, x_96, x_95); +x_98 = lean_unbox(x_82); +lean_dec(x_82); +if (x_98 == 0) { -lean_object* x_91; lean_object* x_92; -x_91 = lean_box(0); -x_92 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_89, x_48, x_4, x_2, x_91, x_6, x_7, x_8, x_9, x_10, x_11, x_79); +lean_object* x_99; lean_object* x_100; +x_99 = lean_box(0); +x_100 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_97, x_48, x_4, x_2, x_99, x_6, x_7, x_8, x_9, x_10, x_11, x_83); lean_dec(x_14); -return x_92; +return x_100; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_inc(x_2); -x_93 = l_Lean_MessageData_ofSyntax(x_2); -x_94 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -x_95 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_93); -x_96 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; -x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -lean_inc(x_89); -x_98 = l_Lean_MessageData_ofSyntax(x_89); -x_99 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -x_100 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_94); -x_101 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_48, x_100, x_6, x_7, x_8, x_9, x_10, x_11, x_79); -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_89, x_48, x_4, x_2, x_102, x_6, x_7, x_8, x_9, x_10, x_11, x_103); -lean_dec(x_102); +x_101 = l_Lean_MessageData_ofSyntax(x_2); +x_102 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; +x_103 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_101); +x_104 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +lean_inc(x_97); +x_106 = l_Lean_MessageData_ofSyntax(x_97); +x_107 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_102); +x_109 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_48, x_108, x_6, x_7, x_8, x_9, x_10, x_11, x_83); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_30, x_28, x_31, x_42, x_97, x_48, x_4, x_2, x_110, x_6, x_7, x_8, x_9, x_10, x_11, x_111); +lean_dec(x_110); lean_dec(x_14); -return x_104; +return x_112; } } } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_107 = lean_ctor_get(x_22, 0); -x_108 = lean_ctor_get(x_22, 1); -lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_22); -x_109 = lean_ctor_get(x_107, 0); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_environment_main_module(x_109); -x_111 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; -x_112 = l_Lean_addMacroScope(x_110, x_111, x_21); -x_113 = lean_box(0); -x_114 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; -x_115 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_115, 0, x_20); -lean_ctor_set(x_115, 1, x_114); -lean_ctor_set(x_115, 2, x_112); -lean_ctor_set(x_115, 3, x_113); -x_116 = l_Lean_Syntax_getArg(x_16, x_13); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_115 = lean_ctor_get(x_22, 0); +x_116 = lean_ctor_get(x_22, 1); lean_inc(x_116); -x_117 = l_Lean_Syntax_getKind(x_116); -x_118 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; -x_119 = lean_name_eq(x_117, x_118); -lean_dec(x_117); -x_120 = l_Lean_Syntax_getArgs(x_16); +lean_inc(x_115); +lean_dec(x_22); +x_117 = lean_ctor_get(x_115, 0); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_environment_main_module(x_117); +x_119 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; +x_120 = l_Lean_addMacroScope(x_118, x_119, x_21); +x_121 = lean_box(0); +x_122 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; +x_123 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_123, 0, x_20); +lean_ctor_set(x_123, 1, x_122); +lean_ctor_set(x_123, 2, x_120); +lean_ctor_set(x_123, 3, x_121); +x_124 = l_Lean_Syntax_getArg(x_16, x_13); +lean_inc(x_124); +x_125 = l_Lean_Syntax_getKind(x_124); +x_126 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___spec__3___closed__2; +x_127 = lean_name_eq(x_125, x_126); +lean_dec(x_125); +x_128 = l_Lean_Syntax_getArgs(x_16); lean_dec(x_16); -x_121 = lean_array_get_size(x_120); -x_122 = l_Array_toSubarray___rarg(x_120, x_15, x_121); -x_123 = l_Array_ofSubarray___rarg(x_122); -lean_dec(x_122); -x_124 = lean_box(2); -x_125 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; -x_126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_126, 0, x_124); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_126, 2, x_123); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_126); -lean_ctor_set(x_127, 1, x_113); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_115); -lean_ctor_set(x_128, 1, x_113); -x_129 = lean_array_mk(x_128); -x_130 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_129); -lean_dec(x_129); +x_129 = lean_array_get_size(x_128); +x_130 = l_Array_toSubarray___rarg(x_128, x_15, x_129); +x_131 = l_Array_ofSubarray___rarg(x_130); +lean_dec(x_130); +x_132 = lean_box(2); +x_133 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +x_134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +lean_ctor_set(x_134, 2, x_131); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_121); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_123); +lean_ctor_set(x_136, 1, x_121); +x_137 = lean_array_mk(x_136); +x_138 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_137); +lean_dec(x_137); lean_inc(x_2); -x_131 = l_Lean_Syntax_setArg(x_2, x_15, x_130); -x_132 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; -x_133 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_132, x_6, x_7, x_8, x_9, x_10, x_11, x_108); -if (x_119 == 0) +x_139 = l_Lean_Syntax_setArg(x_2, x_15, x_138); +x_140 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; +x_141 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_140, x_6, x_7, x_8, x_9, x_10, x_11, x_116); +if (x_127 == 0) { -lean_object* x_164; -x_164 = l_Lean_Syntax_getArg(x_116, x_15); -lean_dec(x_116); -x_134 = x_164; -goto block_163; +lean_object* x_176; +x_176 = l_Lean_Syntax_getArg(x_124, x_15); +lean_dec(x_124); +x_142 = x_176; +goto block_175; } else { -x_134 = x_116; -goto block_163; +x_142 = x_124; +goto block_175; } -block_163: +block_175: { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; -x_135 = lean_ctor_get(x_133, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_133, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_133)) { - lean_ctor_release(x_133, 0); - lean_ctor_release(x_133, 1); - x_137 = x_133; +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; +x_143 = lean_ctor_get(x_141, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_141, 1); +lean_inc(x_144); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_145 = x_141; } else { - lean_dec_ref(x_133); - x_137 = lean_box(0); + lean_dec_ref(x_141); + x_145 = lean_box(0); } -if (lean_is_scalar(x_137)) { - x_138 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_145)) { + x_146 = lean_alloc_ctor(1, 2, 0); } else { - x_138 = x_137; - lean_ctor_set_tag(x_138, 1); + x_146 = x_145; + lean_ctor_set_tag(x_146, 1); } -lean_ctor_set(x_138, 0, x_134); -lean_ctor_set(x_138, 1, x_127); -x_139 = lean_array_mk(x_138); -x_140 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; -x_141 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_141, 0, x_124); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_139); -x_142 = l_Lean_Syntax_setArg(x_1, x_13, x_141); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_113); -x_144 = lean_array_mk(x_143); -x_145 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_145, 0, x_124); -lean_ctor_set(x_145, 1, x_125); -lean_ctor_set(x_145, 2, x_144); -x_146 = lean_unsigned_to_nat(2u); -x_147 = l_Lean_Syntax_setArg(x_131, x_146, x_145); -x_148 = lean_unbox(x_135); -lean_dec(x_135); -if (x_148 == 0) +lean_ctor_set(x_146, 0, x_142); +lean_ctor_set(x_146, 1, x_135); +x_147 = lean_array_mk(x_146); +x_148 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__4; +x_149 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_149, 0, x_132); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_149, 2, x_147); +x_150 = l_Lean_Syntax_setArg(x_1, x_13, x_149); +x_151 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_121); +x_152 = lean_array_mk(x_151); +x_153 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_153, 0, x_132); +lean_ctor_set(x_153, 1, x_133); +lean_ctor_set(x_153, 2, x_152); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_121); +x_155 = lean_array_mk(x_154); +x_156 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_157 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_157, 0, x_132); +lean_ctor_set(x_157, 1, x_156); +lean_ctor_set(x_157, 2, x_155); +x_158 = lean_unsigned_to_nat(2u); +x_159 = l_Lean_Syntax_setArg(x_139, x_158, x_157); +x_160 = lean_unbox(x_143); +lean_dec(x_143); +if (x_160 == 0) { -lean_object* x_149; lean_object* x_150; -x_149 = lean_box(0); -x_150 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_113, x_111, x_114, x_125, x_147, x_132, x_4, x_2, x_149, x_6, x_7, x_8, x_9, x_10, x_11, x_136); +lean_object* x_161; lean_object* x_162; +x_161 = lean_box(0); +x_162 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_121, x_119, x_122, x_133, x_159, x_140, x_4, x_2, x_161, x_6, x_7, x_8, x_9, x_10, x_11, x_144); lean_dec(x_14); -return x_150; +return x_162; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_inc(x_2); -x_151 = l_Lean_MessageData_ofSyntax(x_2); -x_152 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -x_153 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_153, 0, x_152); -lean_ctor_set(x_153, 1, x_151); -x_154 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; -x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_153); -lean_ctor_set(x_155, 1, x_154); -lean_inc(x_147); -x_156 = l_Lean_MessageData_ofSyntax(x_147); -x_157 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); -x_158 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_152); -x_159 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_132, x_158, x_6, x_7, x_8, x_9, x_10, x_11, x_136); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); -lean_inc(x_161); -lean_dec(x_159); -x_162 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_113, x_111, x_114, x_125, x_147, x_132, x_4, x_2, x_160, x_6, x_7, x_8, x_9, x_10, x_11, x_161); -lean_dec(x_160); +x_163 = l_Lean_MessageData_ofSyntax(x_2); +x_164 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; +x_165 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_163); +x_166 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__7; +x_167 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +lean_inc(x_159); +x_168 = l_Lean_MessageData_ofSyntax(x_159); +x_169 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +x_170 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_164); +x_171 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_140, x_170, x_6, x_7, x_8, x_9, x_10, x_11, x_144); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2(x_14, x_3, x_121, x_119, x_122, x_133, x_159, x_140, x_4, x_2, x_172, x_6, x_7, x_8, x_9, x_10, x_11, x_173); +lean_dec(x_172); lean_dec(x_14); -return x_162; +return x_174; } } } } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; lean_object* x_171; +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_dec(x_16); -x_165 = lean_unsigned_to_nat(2u); -x_166 = l_Lean_Syntax_getArg(x_1, x_165); +x_177 = lean_unsigned_to_nat(2u); +x_178 = l_Lean_Syntax_getArg(x_1, x_177); lean_dec(x_1); -x_167 = l_Lean_Syntax_getArg(x_14, x_13); +x_179 = l_Lean_Syntax_getArg(x_14, x_13); lean_dec(x_14); -x_168 = l_Lean_Syntax_getArg(x_167, x_15); -lean_dec(x_167); -x_169 = lean_array_get_size(x_3); -x_170 = lean_nat_dec_lt(x_13, x_169); -lean_dec(x_169); -if (x_170 == 0) +x_180 = l_Lean_Syntax_getArg(x_179, x_15); +lean_dec(x_179); +x_181 = lean_array_get_size(x_3); +x_182 = lean_nat_dec_lt(x_13, x_181); +lean_dec(x_181); +if (x_182 == 0) { -lean_object* x_326; lean_object* x_327; -x_326 = l_Lean_Elab_Term_StructInst_instInhabitedExplicitSourceInfo; -x_327 = l_outOfBounds___rarg(x_326); -x_171 = x_327; -goto block_325; +lean_object* x_338; lean_object* x_339; +x_338 = l_Lean_Elab_Term_StructInst_instInhabitedExplicitSourceInfo; +x_339 = l_outOfBounds___rarg(x_338); +x_183 = x_339; +goto block_337; } else { -lean_object* x_328; -x_328 = lean_array_fget(x_3, x_13); -x_171 = x_328; -goto block_325; +lean_object* x_340; +x_340 = lean_array_fget(x_3, x_13); +x_183 = x_340; +goto block_337; } -block_325: -{ -lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -lean_dec(x_171); -x_173 = lean_ctor_get(x_10, 5); -lean_inc(x_173); -x_174 = 0; -x_175 = l_Lean_SourceInfo_fromRef(x_173, x_174); -lean_dec(x_173); -x_176 = lean_ctor_get(x_10, 10); -lean_inc(x_176); -x_177 = lean_st_ref_get(x_11, x_12); -x_178 = !lean_is_exclusive(x_177); -if (x_178 == 0) +block_337: { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; uint8_t x_227; -x_179 = lean_ctor_get(x_177, 0); -x_180 = lean_ctor_get(x_177, 1); -x_181 = lean_ctor_get(x_179, 0); -lean_inc(x_181); -lean_dec(x_179); -x_182 = lean_environment_main_module(x_181); -x_183 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; -lean_inc(x_175); -lean_ctor_set_tag(x_177, 2); -lean_ctor_set(x_177, 1, x_183); -lean_ctor_set(x_177, 0, x_175); -x_184 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; -lean_inc(x_176); -lean_inc(x_182); -x_185 = l_Lean_addMacroScope(x_182, x_184, x_176); -x_186 = lean_box(0); -x_187 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; -lean_inc(x_175); -x_188 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_188, 0, x_175); -lean_ctor_set(x_188, 1, x_187); -lean_ctor_set(x_188, 2, x_185); -lean_ctor_set(x_188, 3, x_186); -x_189 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__4; -lean_inc(x_175); -x_190 = l_Lean_Syntax_node3(x_175, x_189, x_172, x_177, x_188); -x_191 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; -lean_inc(x_175); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_175); -lean_ctor_set(x_192, 1, x_191); -x_193 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; -lean_inc(x_176); -lean_inc(x_182); -x_194 = l_Lean_addMacroScope(x_182, x_193, x_176); -x_195 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; -lean_inc(x_175); -x_196 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_196, 0, x_175); -lean_ctor_set(x_196, 1, x_195); -lean_ctor_set(x_196, 2, x_194); -lean_ctor_set(x_196, 3, x_186); -x_197 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; -lean_inc(x_175); -x_198 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_198, 0, x_175); -lean_ctor_set(x_198, 1, x_197); -x_199 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; -lean_inc(x_175); -x_200 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_200, 0, x_175); +lean_object* x_184; lean_object* x_185; uint8_t x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +lean_dec(x_183); +x_185 = lean_ctor_get(x_10, 5); +lean_inc(x_185); +x_186 = 0; +x_187 = l_Lean_SourceInfo_fromRef(x_185, x_186); +lean_dec(x_185); +x_188 = lean_ctor_get(x_10, 10); +lean_inc(x_188); +x_189 = lean_st_ref_get(x_11, x_12); +x_190 = !lean_is_exclusive(x_189); +if (x_190 == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_191 = lean_ctor_get(x_189, 0); +x_192 = lean_ctor_get(x_189, 1); +x_193 = lean_ctor_get(x_191, 0); +lean_inc(x_193); +lean_dec(x_191); +x_194 = lean_environment_main_module(x_193); +x_195 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; +lean_inc(x_187); +lean_ctor_set_tag(x_189, 2); +lean_ctor_set(x_189, 1, x_195); +lean_ctor_set(x_189, 0, x_187); +x_196 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; +lean_inc(x_188); +lean_inc(x_194); +x_197 = l_Lean_addMacroScope(x_194, x_196, x_188); +x_198 = lean_box(0); +x_199 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; +lean_inc(x_187); +x_200 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_200, 0, x_187); lean_ctor_set(x_200, 1, x_199); -x_201 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__10; -lean_inc(x_200); -lean_inc(x_192); -lean_inc(x_175); -x_202 = l_Lean_Syntax_node5(x_175, x_201, x_192, x_196, x_198, x_168, x_200); -x_203 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__16; -lean_inc(x_175); +lean_ctor_set(x_200, 2, x_197); +lean_ctor_set(x_200, 3, x_198); +x_201 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__4; +lean_inc(x_187); +x_202 = l_Lean_Syntax_node3(x_187, x_201, x_184, x_189, x_200); +x_203 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +lean_inc(x_187); x_204 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_204, 0, x_175); +lean_ctor_set(x_204, 0, x_187); lean_ctor_set(x_204, 1, x_203); -x_205 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; -x_206 = l_Lean_addMacroScope(x_182, x_205, x_176); -x_207 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; -lean_inc(x_175); +x_205 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; +lean_inc(x_188); +lean_inc(x_194); +x_206 = l_Lean_addMacroScope(x_194, x_205, x_188); +x_207 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; +lean_inc(x_187); x_208 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_208, 0, x_175); +lean_ctor_set(x_208, 0, x_187); lean_ctor_set(x_208, 1, x_207); lean_ctor_set(x_208, 2, x_206); -lean_ctor_set(x_208, 3, x_186); -x_209 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; -lean_inc(x_175); -x_210 = l_Lean_Syntax_node1(x_175, x_209, x_208); -x_211 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; -lean_inc(x_175); -x_212 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_212, 0, x_175); -lean_ctor_set(x_212, 1, x_209); -lean_ctor_set(x_212, 2, x_211); -x_213 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__20; -lean_inc(x_175); -x_214 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_214, 0, x_175); -lean_ctor_set(x_214, 1, x_213); -x_215 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; -lean_inc(x_175); -x_216 = l_Lean_Syntax_node4(x_175, x_215, x_210, x_212, x_214, x_166); -x_217 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; -lean_inc(x_175); -x_218 = l_Lean_Syntax_node2(x_175, x_217, x_204, x_216); -x_219 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; -lean_inc(x_175); -x_220 = l_Lean_Syntax_node3(x_175, x_219, x_192, x_218, x_200); -lean_inc(x_175); -x_221 = l_Lean_Syntax_node2(x_175, x_209, x_202, x_220); -x_222 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; -x_223 = l_Lean_Syntax_node2(x_175, x_222, x_190, x_221); -x_224 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; -x_225 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_224, x_6, x_7, x_8, x_9, x_10, x_11, x_180); -x_226 = lean_ctor_get(x_225, 0); -lean_inc(x_226); -x_227 = lean_unbox(x_226); -lean_dec(x_226); -if (x_227 == 0) -{ -lean_object* x_228; lean_object* x_229; lean_object* x_230; -x_228 = lean_ctor_get(x_225, 1); -lean_inc(x_228); -lean_dec(x_225); -x_229 = lean_box(0); -x_230 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_223, x_4, x_2, x_229, x_6, x_7, x_8, x_9, x_10, x_11, x_228); -return x_230; +lean_ctor_set(x_208, 3, x_198); +x_209 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; +lean_inc(x_187); +x_210 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_210, 0, x_187); +lean_ctor_set(x_210, 1, x_209); +x_211 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; +lean_inc(x_187); +x_212 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_212, 0, x_187); +lean_ctor_set(x_212, 1, x_211); +x_213 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__10; +lean_inc(x_212); +lean_inc(x_204); +lean_inc(x_187); +x_214 = l_Lean_Syntax_node5(x_187, x_213, x_204, x_208, x_210, x_180, x_212); +x_215 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__16; +lean_inc(x_187); +x_216 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_216, 0, x_187); +lean_ctor_set(x_216, 1, x_215); +x_217 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; +x_218 = l_Lean_addMacroScope(x_194, x_217, x_188); +x_219 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; +lean_inc(x_187); +x_220 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_220, 0, x_187); +lean_ctor_set(x_220, 1, x_219); +lean_ctor_set(x_220, 2, x_218); +lean_ctor_set(x_220, 3, x_198); +x_221 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +lean_inc(x_187); +x_222 = l_Lean_Syntax_node1(x_187, x_221, x_220); +x_223 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +lean_inc(x_187); +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_187); +lean_ctor_set(x_224, 1, x_221); +lean_ctor_set(x_224, 2, x_223); +x_225 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__20; +lean_inc(x_187); +x_226 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_226, 0, x_187); +lean_ctor_set(x_226, 1, x_225); +x_227 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; +lean_inc(x_187); +x_228 = l_Lean_Syntax_node4(x_187, x_227, x_222, x_224, x_226, x_178); +x_229 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; +lean_inc(x_187); +x_230 = l_Lean_Syntax_node2(x_187, x_229, x_216, x_228); +x_231 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; +lean_inc(x_187); +x_232 = l_Lean_Syntax_node3(x_187, x_231, x_204, x_230, x_212); +lean_inc(x_187); +x_233 = l_Lean_Syntax_node2(x_187, x_221, x_214, x_232); +x_234 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; +x_235 = l_Lean_Syntax_node2(x_187, x_234, x_202, x_233); +x_236 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; +x_237 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_236, x_6, x_7, x_8, x_9, x_10, x_11, x_192); +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_unbox(x_238); +lean_dec(x_238); +if (x_239 == 0) +{ +lean_object* x_240; lean_object* x_241; lean_object* x_242; +x_240 = lean_ctor_get(x_237, 1); +lean_inc(x_240); +lean_dec(x_237); +x_241 = lean_box(0); +x_242 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_235, x_4, x_2, x_241, x_6, x_7, x_8, x_9, x_10, x_11, x_240); +return x_242; } else { -uint8_t x_231; -x_231 = !lean_is_exclusive(x_225); -if (x_231 == 0) +uint8_t x_243; +x_243 = !lean_is_exclusive(x_237); +if (x_243 == 0) { -lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; -x_232 = lean_ctor_get(x_225, 1); -x_233 = lean_ctor_get(x_225, 0); -lean_dec(x_233); -lean_inc(x_2); -x_234 = l_Lean_MessageData_ofSyntax(x_2); -x_235 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -lean_ctor_set_tag(x_225, 7); -lean_ctor_set(x_225, 1, x_234); -lean_ctor_set(x_225, 0, x_235); -x_236 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; -x_237 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_237, 0, x_225); -lean_ctor_set(x_237, 1, x_236); -lean_inc(x_223); -x_238 = l_Lean_MessageData_ofSyntax(x_223); -x_239 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_239, 0, x_237); -lean_ctor_set(x_239, 1, x_238); -x_240 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_240, 0, x_239); -lean_ctor_set(x_240, 1, x_235); -x_241 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_224, x_240, x_6, x_7, x_8, x_9, x_10, x_11, x_232); -x_242 = lean_ctor_get(x_241, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -lean_dec(x_241); -x_244 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_223, x_4, x_2, x_242, x_6, x_7, x_8, x_9, x_10, x_11, x_243); -lean_dec(x_242); -return x_244; -} -else -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_245 = lean_ctor_get(x_225, 1); -lean_inc(x_245); -lean_dec(x_225); +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_244 = lean_ctor_get(x_237, 1); +x_245 = lean_ctor_get(x_237, 0); +lean_dec(x_245); lean_inc(x_2); x_246 = l_Lean_MessageData_ofSyntax(x_2); x_247 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -x_248 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_248, 0, x_247); -lean_ctor_set(x_248, 1, x_246); -x_249 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; -x_250 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_250, 0, x_248); -lean_ctor_set(x_250, 1, x_249); -lean_inc(x_223); -x_251 = l_Lean_MessageData_ofSyntax(x_223); +lean_ctor_set_tag(x_237, 7); +lean_ctor_set(x_237, 1, x_246); +lean_ctor_set(x_237, 0, x_247); +x_248 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; +x_249 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_249, 0, x_237); +lean_ctor_set(x_249, 1, x_248); +lean_inc(x_235); +x_250 = l_Lean_MessageData_ofSyntax(x_235); +x_251 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_251, 0, x_249); +lean_ctor_set(x_251, 1, x_250); x_252 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_252, 0, x_250); -lean_ctor_set(x_252, 1, x_251); -x_253 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_253, 0, x_252); -lean_ctor_set(x_253, 1, x_247); -x_254 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_224, x_253, x_6, x_7, x_8, x_9, x_10, x_11, x_245); -x_255 = lean_ctor_get(x_254, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_247); +x_253 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_236, x_252, x_6, x_7, x_8, x_9, x_10, x_11, x_244); +x_254 = lean_ctor_get(x_253, 0); +lean_inc(x_254); +x_255 = lean_ctor_get(x_253, 1); lean_inc(x_255); -x_256 = lean_ctor_get(x_254, 1); -lean_inc(x_256); +lean_dec(x_253); +x_256 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_235, x_4, x_2, x_254, x_6, x_7, x_8, x_9, x_10, x_11, x_255); lean_dec(x_254); -x_257 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_223, x_4, x_2, x_255, x_6, x_7, x_8, x_9, x_10, x_11, x_256); -lean_dec(x_255); -return x_257; +return x_256; +} +else +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_257 = lean_ctor_get(x_237, 1); +lean_inc(x_257); +lean_dec(x_237); +lean_inc(x_2); +x_258 = l_Lean_MessageData_ofSyntax(x_2); +x_259 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; +x_260 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_260, 0, x_259); +lean_ctor_set(x_260, 1, x_258); +x_261 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; +x_262 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_262, 0, x_260); +lean_ctor_set(x_262, 1, x_261); +lean_inc(x_235); +x_263 = l_Lean_MessageData_ofSyntax(x_235); +x_264 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_264, 0, x_262); +lean_ctor_set(x_264, 1, x_263); +x_265 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_265, 0, x_264); +lean_ctor_set(x_265, 1, x_259); +x_266 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_236, x_265, x_6, x_7, x_8, x_9, x_10, x_11, x_257); +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_235, x_4, x_2, x_267, x_6, x_7, x_8, x_9, x_10, x_11, x_268); +lean_dec(x_267); +return x_269; } } } else { -lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; uint8_t x_307; -x_258 = lean_ctor_get(x_177, 0); -x_259 = lean_ctor_get(x_177, 1); -lean_inc(x_259); -lean_inc(x_258); -lean_dec(x_177); -x_260 = lean_ctor_get(x_258, 0); -lean_inc(x_260); -lean_dec(x_258); -x_261 = lean_environment_main_module(x_260); -x_262 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; -lean_inc(x_175); -x_263 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_263, 0, x_175); -lean_ctor_set(x_263, 1, x_262); -x_264 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; -lean_inc(x_176); -lean_inc(x_261); -x_265 = l_Lean_addMacroScope(x_261, x_264, x_176); -x_266 = lean_box(0); -x_267 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; -lean_inc(x_175); -x_268 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_268, 0, x_175); -lean_ctor_set(x_268, 1, x_267); -lean_ctor_set(x_268, 2, x_265); -lean_ctor_set(x_268, 3, x_266); -x_269 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__4; -lean_inc(x_175); -x_270 = l_Lean_Syntax_node3(x_175, x_269, x_172, x_263, x_268); -x_271 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; -lean_inc(x_175); -x_272 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_272, 0, x_175); -lean_ctor_set(x_272, 1, x_271); -x_273 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; -lean_inc(x_176); -lean_inc(x_261); -x_274 = l_Lean_addMacroScope(x_261, x_273, x_176); -x_275 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; -lean_inc(x_175); -x_276 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_276, 0, x_175); -lean_ctor_set(x_276, 1, x_275); -lean_ctor_set(x_276, 2, x_274); -lean_ctor_set(x_276, 3, x_266); -x_277 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; -lean_inc(x_175); -x_278 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_278, 0, x_175); -lean_ctor_set(x_278, 1, x_277); -x_279 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; -lean_inc(x_175); -x_280 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_280, 0, x_175); -lean_ctor_set(x_280, 1, x_279); -x_281 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__10; -lean_inc(x_280); +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; uint8_t x_319; +x_270 = lean_ctor_get(x_189, 0); +x_271 = lean_ctor_get(x_189, 1); +lean_inc(x_271); +lean_inc(x_270); +lean_dec(x_189); +x_272 = lean_ctor_get(x_270, 0); lean_inc(x_272); -lean_inc(x_175); -x_282 = l_Lean_Syntax_node5(x_175, x_281, x_272, x_276, x_278, x_168, x_280); -x_283 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__16; -lean_inc(x_175); +lean_dec(x_270); +x_273 = lean_environment_main_module(x_272); +x_274 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__5; +lean_inc(x_187); +x_275 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_275, 0, x_187); +lean_ctor_set(x_275, 1, x_274); +x_276 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__8; +lean_inc(x_188); +lean_inc(x_273); +x_277 = l_Lean_addMacroScope(x_273, x_276, x_188); +x_278 = lean_box(0); +x_279 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__7; +lean_inc(x_187); +x_280 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_280, 0, x_187); +lean_ctor_set(x_280, 1, x_279); +lean_ctor_set(x_280, 2, x_277); +lean_ctor_set(x_280, 3, x_278); +x_281 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__4; +lean_inc(x_187); +x_282 = l_Lean_Syntax_node3(x_187, x_281, x_184, x_275, x_280); +x_283 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__10; +lean_inc(x_187); x_284 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_284, 0, x_175); +lean_ctor_set(x_284, 0, x_187); lean_ctor_set(x_284, 1, x_283); -x_285 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; -x_286 = l_Lean_addMacroScope(x_261, x_285, x_176); -x_287 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; -lean_inc(x_175); +x_285 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__13; +lean_inc(x_188); +lean_inc(x_273); +x_286 = l_Lean_addMacroScope(x_273, x_285, x_188); +x_287 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__12; +lean_inc(x_187); x_288 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_288, 0, x_175); +lean_ctor_set(x_288, 0, x_187); lean_ctor_set(x_288, 1, x_287); lean_ctor_set(x_288, 2, x_286); -lean_ctor_set(x_288, 3, x_266); -x_289 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; -lean_inc(x_175); -x_290 = l_Lean_Syntax_node1(x_175, x_289, x_288); -x_291 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; -lean_inc(x_175); -x_292 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_292, 0, x_175); -lean_ctor_set(x_292, 1, x_289); -lean_ctor_set(x_292, 2, x_291); -x_293 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__20; -lean_inc(x_175); -x_294 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_294, 0, x_175); -lean_ctor_set(x_294, 1, x_293); -x_295 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; -lean_inc(x_175); -x_296 = l_Lean_Syntax_node4(x_175, x_295, x_290, x_292, x_294, x_166); -x_297 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; -lean_inc(x_175); -x_298 = l_Lean_Syntax_node2(x_175, x_297, x_284, x_296); -x_299 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; -lean_inc(x_175); -x_300 = l_Lean_Syntax_node3(x_175, x_299, x_272, x_298, x_280); -lean_inc(x_175); -x_301 = l_Lean_Syntax_node2(x_175, x_289, x_282, x_300); -x_302 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; -x_303 = l_Lean_Syntax_node2(x_175, x_302, x_270, x_301); -x_304 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; -x_305 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_304, x_6, x_7, x_8, x_9, x_10, x_11, x_259); -x_306 = lean_ctor_get(x_305, 0); -lean_inc(x_306); -x_307 = lean_unbox(x_306); -lean_dec(x_306); -if (x_307 == 0) +lean_ctor_set(x_288, 3, x_278); +x_289 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___spec__2___closed__5; +lean_inc(x_187); +x_290 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_290, 0, x_187); +lean_ctor_set(x_290, 1, x_289); +x_291 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__12; +lean_inc(x_187); +x_292 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_292, 0, x_187); +lean_ctor_set(x_292, 1, x_291); +x_293 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__10; +lean_inc(x_292); +lean_inc(x_284); +lean_inc(x_187); +x_294 = l_Lean_Syntax_node5(x_187, x_293, x_284, x_288, x_290, x_180, x_292); +x_295 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__16; +lean_inc(x_187); +x_296 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_296, 0, x_187); +lean_ctor_set(x_296, 1, x_295); +x_297 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__3; +x_298 = l_Lean_addMacroScope(x_273, x_297, x_188); +x_299 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__2; +lean_inc(x_187); +x_300 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_300, 0, x_187); +lean_ctor_set(x_300, 1, x_299); +lean_ctor_set(x_300, 2, x_298); +lean_ctor_set(x_300, 3, x_278); +x_301 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__3; +lean_inc(x_187); +x_302 = l_Lean_Syntax_node1(x_187, x_301, x_300); +x_303 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__1; +lean_inc(x_187); +x_304 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_304, 0, x_187); +lean_ctor_set(x_304, 1, x_301); +lean_ctor_set(x_304, 2, x_303); +x_305 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__20; +lean_inc(x_187); +x_306 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_306, 0, x_187); +lean_ctor_set(x_306, 1, x_305); +x_307 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__19; +lean_inc(x_187); +x_308 = l_Lean_Syntax_node4(x_187, x_307, x_302, x_304, x_306, x_178); +x_309 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__17; +lean_inc(x_187); +x_310 = l_Lean_Syntax_node2(x_187, x_309, x_296, x_308); +x_311 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__15; +lean_inc(x_187); +x_312 = l_Lean_Syntax_node3(x_187, x_311, x_284, x_310, x_292); +lean_inc(x_187); +x_313 = l_Lean_Syntax_node2(x_187, x_301, x_294, x_312); +x_314 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__2; +x_315 = l_Lean_Syntax_node2(x_187, x_314, x_282, x_313); +x_316 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__3___closed__5; +x_317 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_316, x_6, x_7, x_8, x_9, x_10, x_11, x_271); +x_318 = lean_ctor_get(x_317, 0); +lean_inc(x_318); +x_319 = lean_unbox(x_318); +lean_dec(x_318); +if (x_319 == 0) { -lean_object* x_308; lean_object* x_309; lean_object* x_310; -x_308 = lean_ctor_get(x_305, 1); -lean_inc(x_308); -lean_dec(x_305); -x_309 = lean_box(0); -x_310 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_303, x_4, x_2, x_309, x_6, x_7, x_8, x_9, x_10, x_11, x_308); -return x_310; +lean_object* x_320; lean_object* x_321; lean_object* x_322; +x_320 = lean_ctor_get(x_317, 1); +lean_inc(x_320); +lean_dec(x_317); +x_321 = lean_box(0); +x_322 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_315, x_4, x_2, x_321, x_6, x_7, x_8, x_9, x_10, x_11, x_320); +return x_322; } else { -lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; -x_311 = lean_ctor_get(x_305, 1); -lean_inc(x_311); -if (lean_is_exclusive(x_305)) { - lean_ctor_release(x_305, 0); - lean_ctor_release(x_305, 1); - x_312 = x_305; +lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; +x_323 = lean_ctor_get(x_317, 1); +lean_inc(x_323); +if (lean_is_exclusive(x_317)) { + lean_ctor_release(x_317, 0); + lean_ctor_release(x_317, 1); + x_324 = x_317; } else { - lean_dec_ref(x_305); - x_312 = lean_box(0); + lean_dec_ref(x_317); + x_324 = lean_box(0); } lean_inc(x_2); -x_313 = l_Lean_MessageData_ofSyntax(x_2); -x_314 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; -if (lean_is_scalar(x_312)) { - x_315 = lean_alloc_ctor(7, 2, 0); -} else { - x_315 = x_312; - lean_ctor_set_tag(x_315, 7); -} -lean_ctor_set(x_315, 0, x_314); -lean_ctor_set(x_315, 1, x_313); -x_316 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; -x_317 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_317, 0, x_315); -lean_ctor_set(x_317, 1, x_316); -lean_inc(x_303); -x_318 = l_Lean_MessageData_ofSyntax(x_303); -x_319 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_319, 0, x_317); -lean_ctor_set(x_319, 1, x_318); -x_320 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_320, 0, x_319); -lean_ctor_set(x_320, 1, x_314); -x_321 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_304, x_320, x_6, x_7, x_8, x_9, x_10, x_11, x_311); -x_322 = lean_ctor_get(x_321, 0); -lean_inc(x_322); -x_323 = lean_ctor_get(x_321, 1); -lean_inc(x_323); -lean_dec(x_321); -x_324 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_303, x_4, x_2, x_322, x_6, x_7, x_8, x_9, x_10, x_11, x_323); -lean_dec(x_322); -return x_324; +x_325 = l_Lean_MessageData_ofSyntax(x_2); +x_326 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__22; +if (lean_is_scalar(x_324)) { + x_327 = lean_alloc_ctor(7, 2, 0); +} else { + x_327 = x_324; + lean_ctor_set_tag(x_327, 7); +} +lean_ctor_set(x_327, 0, x_326); +lean_ctor_set(x_327, 1, x_325); +x_328 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__2___closed__24; +x_329 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_329, 0, x_327); +lean_ctor_set(x_329, 1, x_328); +lean_inc(x_315); +x_330 = l_Lean_MessageData_ofSyntax(x_315); +x_331 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_331, 0, x_329); +lean_ctor_set(x_331, 1, x_330); +x_332 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_332, 0, x_331); +lean_ctor_set(x_332, 1, x_326); +x_333 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_316, x_332, x_6, x_7, x_8, x_9, x_10, x_11, x_323); +x_334 = lean_ctor_get(x_333, 0); +lean_inc(x_334); +x_335 = lean_ctor_get(x_333, 1); +lean_inc(x_335); +lean_dec(x_333); +x_336 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabModifyOp___lambda__1(x_315, x_4, x_2, x_334, x_6, x_7, x_8, x_9, x_10, x_11, x_335); +lean_dec(x_334); +return x_336; } } } @@ -10681,76 +10757,79 @@ return x_70; LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_6 = lean_unsigned_to_nat(2u); x_7 = l_Lean_Syntax_getArg(x_1, x_6); -x_8 = l_Lean_Syntax_getSepArgs(x_7); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_7, x_8); lean_dec(x_7); -x_9 = lean_array_to_list(x_8); -x_10 = lean_box(0); -x_11 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView___spec__2(x_9, x_10, x_4, x_5); -if (lean_obj_tag(x_11) == 0) +x_10 = l_Lean_Syntax_getSepArgs(x_9); +lean_dec(x_9); +x_11 = lean_array_to_list(x_10); +x_12 = lean_box(0); +x_13 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkStructView___spec__2(x_11, x_12, x_4, x_5); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; -x_15 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_2); -lean_ctor_set(x_15, 2, x_14); -lean_ctor_set(x_15, 3, x_13); -lean_ctor_set(x_15, 4, x_3); -lean_ctor_set(x_11, 0, x_15); -return x_11; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +x_17 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_16); +lean_ctor_set(x_17, 3, x_15); +lean_ctor_set(x_17, 4, x_3); +lean_ctor_set(x_13, 0, x_17); +return x_13; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_11, 0); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_11); -x_18 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; -x_19 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_2); -lean_ctor_set(x_19, 2, x_18); -lean_ctor_set(x_19, 3, x_16); -lean_ctor_set(x_19, 4, x_3); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_13); +x_20 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +x_21 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_20); +lean_ctor_set(x_21, 3, x_18); +lean_ctor_set(x_21, 4, x_3); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; } } else { -uint8_t x_21; +uint8_t x_23; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_11); -if (x_21 == 0) +x_23 = !lean_is_exclusive(x_13); +if (x_23 == 0) { -return x_11; +return x_13; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_11, 0); -x_23 = lean_ctor_get(x_11, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_11); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } @@ -13194,33 +13273,35 @@ lean_inc(x_27); x_28 = !lean_is_exclusive(x_23); if (x_28 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; x_29 = lean_ctor_get(x_23, 0); x_30 = lean_ctor_get(x_23, 1); x_31 = lean_ctor_get(x_7, 1); -x_32 = l_Lean_Elab_Term_StructInst_Struct_structName(x_1); -lean_inc(x_32); -lean_inc(x_31); lean_inc(x_30); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_30); +x_33 = l_Lean_Elab_Term_StructInst_Struct_structName(x_1); +lean_inc(x_33); +lean_inc(x_31); lean_inc(x_29); -x_33 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_33, 0, x_29); -lean_ctor_set(x_33, 1, x_30); -lean_ctor_set(x_33, 2, x_31); -lean_ctor_set(x_33, 3, x_32); -x_34 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +x_34 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_31); +lean_ctor_set(x_34, 3, x_33); +x_35 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_34, 1); -x_37 = lean_ctor_get(x_34, 0); -lean_dec(x_37); -lean_inc(x_32); -x_38 = l_Lean_findField_x3f(x_2, x_32, x_30); -if (lean_obj_tag(x_38) == 0) +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_35, 1); +x_38 = lean_ctor_get(x_35, 0); +lean_dec(x_38); +lean_inc(x_33); +x_39 = l_Lean_findField_x3f(x_2, x_33, x_30); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); @@ -13228,71 +13309,71 @@ lean_dec(x_24); lean_dec(x_15); lean_dec(x_14); lean_dec(x_4); -x_39 = l_Lean_MessageData_ofName(x_30); -x_40 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_39); -lean_ctor_set(x_34, 0, x_40); -x_41 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; +x_40 = l_Lean_MessageData_ofName(x_30); +x_41 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_40); +lean_ctor_set(x_35, 0, x_41); +x_42 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_41); -lean_ctor_set(x_23, 0, x_34); -x_42 = 0; -x_43 = l_Lean_MessageData_ofConstName(x_32, x_42); -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_23); -lean_ctor_set(x_44, 1, x_43); +lean_ctor_set(x_23, 1, x_42); +lean_ctor_set(x_23, 0, x_35); +x_43 = 0; +x_44 = l_Lean_MessageData_ofConstName(x_33, x_43); x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_40); -x_46 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +lean_ctor_set(x_45, 0, x_23); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_41); +x_47 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_46, x_5, x_6, x_7, x_8, x_9, x_10, x_37); lean_dec(x_29); -x_47 = !lean_is_exclusive(x_46); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_47); +if (x_48 == 0) { -return x_46; +return x_47; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_46, 0); -x_49 = lean_ctor_get(x_46, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_47, 0); +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_46); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_dec(x_47); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } else { -lean_object* x_51; uint8_t x_52; -x_51 = lean_ctor_get(x_38, 0); -lean_inc(x_51); -lean_dec(x_38); -x_52 = lean_name_eq(x_51, x_32); -if (x_52 == 0) -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_14); +lean_object* x_52; uint8_t x_53; +x_52 = lean_ctor_get(x_39, 0); +lean_inc(x_52); +lean_dec(x_39); +x_53 = lean_name_eq(x_52, x_33); if (x_53 == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_54 = lean_ctor_get(x_14, 3); -lean_dec(x_54); -x_55 = lean_ctor_get(x_14, 2); +uint8_t x_54; +x_54 = !lean_is_exclusive(x_14); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_55 = lean_ctor_get(x_14, 3); lean_dec(x_55); -x_56 = lean_ctor_get(x_14, 1); +x_56 = lean_ctor_get(x_14, 2); lean_dec(x_56); -x_57 = lean_ctor_get(x_14, 0); +x_57 = lean_ctor_get(x_14, 1); lean_dec(x_57); -x_58 = l_Lean_getPathToBaseStructure_x3f(x_2, x_51, x_32); -lean_dec(x_51); -if (lean_obj_tag(x_58) == 0) +x_58 = lean_ctor_get(x_14, 0); +lean_dec(x_58); +x_59 = l_Lean_getPathToBaseStructure_x3f(x_2, x_52, x_33); +lean_dec(x_52); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_free_object(x_14); lean_dec(x_27); lean_dec(x_26); @@ -13300,204 +13381,204 @@ lean_dec(x_25); lean_dec(x_24); lean_dec(x_15); lean_dec(x_4); -x_59 = l_Lean_MessageData_ofName(x_30); -x_60 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_59); -lean_ctor_set(x_34, 0, x_60); -x_61 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; +x_60 = l_Lean_MessageData_ofName(x_30); +x_61 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_60); +lean_ctor_set(x_35, 0, x_61); +x_62 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_61); -lean_ctor_set(x_23, 0, x_34); -x_62 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +lean_ctor_set(x_23, 1, x_62); +lean_ctor_set(x_23, 0, x_35); +x_63 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_37); lean_dec(x_29); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) { -return x_62; +return x_63; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_62, 0); -x_65 = lean_ctor_get(x_62, 1); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 0); +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); lean_inc(x_65); -lean_inc(x_64); -lean_dec(x_62); -x_66 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_dec(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; } } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -lean_free_object(x_34); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +lean_free_object(x_35); lean_free_object(x_23); lean_dec(x_30); -x_67 = lean_ctor_get(x_58, 0); -lean_inc(x_67); -lean_dec(x_58); -x_68 = lean_box(0); -x_69 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_67, x_68); +x_68 = lean_ctor_get(x_59, 0); +lean_inc(x_68); +lean_dec(x_59); +x_69 = lean_box(0); +x_70 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_68, x_69); lean_inc(x_15); -x_70 = l_List_appendTR___rarg(x_69, x_15); -x_71 = !lean_is_exclusive(x_15); -if (x_71 == 0) +x_71 = l_List_appendTR___rarg(x_70, x_15); +x_72 = !lean_is_exclusive(x_15); +if (x_72 == 0) { -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_15, 1); -lean_dec(x_72); -x_73 = lean_ctor_get(x_15, 0); +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_15, 1); lean_dec(x_73); -lean_ctor_set(x_14, 1, x_70); +x_74 = lean_ctor_get(x_15, 0); +lean_dec(x_74); +lean_ctor_set(x_14, 1, x_71); lean_ctor_set(x_15, 1, x_4); lean_ctor_set(x_15, 0, x_14); x_3 = x_24; x_4 = x_15; -x_11 = x_36; +x_11 = x_37; goto _start; } else { -lean_object* x_75; +lean_object* x_76; lean_dec(x_15); -lean_ctor_set(x_14, 1, x_70); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_14); -lean_ctor_set(x_75, 1, x_4); +lean_ctor_set(x_14, 1, x_71); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_14); +lean_ctor_set(x_76, 1, x_4); x_3 = x_24; -x_4 = x_75; -x_11 = x_36; +x_4 = x_76; +x_11 = x_37; goto _start; } } } else { -lean_object* x_77; +lean_object* x_78; lean_dec(x_14); -x_77 = l_Lean_getPathToBaseStructure_x3f(x_2, x_51, x_32); -lean_dec(x_51); -if (lean_obj_tag(x_77) == 0) +x_78 = l_Lean_getPathToBaseStructure_x3f(x_2, x_52, x_33); +lean_dec(x_52); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); lean_dec(x_15); lean_dec(x_4); -x_78 = l_Lean_MessageData_ofName(x_30); -x_79 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; -lean_ctor_set_tag(x_34, 7); -lean_ctor_set(x_34, 1, x_78); -lean_ctor_set(x_34, 0, x_79); -x_80 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; +x_79 = l_Lean_MessageData_ofName(x_30); +x_80 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; +lean_ctor_set_tag(x_35, 7); +lean_ctor_set(x_35, 1, x_79); +lean_ctor_set(x_35, 0, x_80); +x_81 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_80); -lean_ctor_set(x_23, 0, x_34); -x_81 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_36); +lean_ctor_set(x_23, 1, x_81); +lean_ctor_set(x_23, 0, x_35); +x_82 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_37); lean_dec(x_29); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); +x_83 = lean_ctor_get(x_82, 0); lean_inc(x_83); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_84 = x_81; +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_85 = x_82; } else { - lean_dec_ref(x_81); - x_84 = lean_box(0); + lean_dec_ref(x_82); + x_85 = lean_box(0); } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_84; + x_86 = x_85; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -lean_free_object(x_34); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_free_object(x_35); lean_free_object(x_23); lean_dec(x_30); -x_86 = lean_ctor_get(x_77, 0); -lean_inc(x_86); -lean_dec(x_77); -x_87 = lean_box(0); -x_88 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_86, x_87); +x_87 = lean_ctor_get(x_78, 0); +lean_inc(x_87); +lean_dec(x_78); +x_88 = lean_box(0); +x_89 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_87, x_88); lean_inc(x_15); -x_89 = l_List_appendTR___rarg(x_88, x_15); +x_90 = l_List_appendTR___rarg(x_89, x_15); if (lean_is_exclusive(x_15)) { lean_ctor_release(x_15, 0); lean_ctor_release(x_15, 1); - x_90 = x_15; + x_91 = x_15; } else { lean_dec_ref(x_15); - x_90 = lean_box(0); -} -x_91 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_91, 0, x_25); -lean_ctor_set(x_91, 1, x_89); -lean_ctor_set(x_91, 2, x_26); -lean_ctor_set(x_91, 3, x_27); -if (lean_is_scalar(x_90)) { - x_92 = lean_alloc_ctor(1, 2, 0); + x_91 = lean_box(0); +} +x_92 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_92, 0, x_25); +lean_ctor_set(x_92, 1, x_90); +lean_ctor_set(x_92, 2, x_26); +lean_ctor_set(x_92, 3, x_27); +if (lean_is_scalar(x_91)) { + x_93 = lean_alloc_ctor(1, 2, 0); } else { - x_92 = x_90; + x_93 = x_91; } -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_4); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_4); x_3 = x_24; -x_4 = x_92; -x_11 = x_36; +x_4 = x_93; +x_11 = x_37; goto _start; } } } else { -uint8_t x_94; -lean_dec(x_51); -lean_free_object(x_34); -lean_dec(x_32); +uint8_t x_95; +lean_dec(x_52); +lean_free_object(x_35); +lean_dec(x_33); lean_free_object(x_23); lean_dec(x_30); lean_dec(x_29); lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); -x_94 = !lean_is_exclusive(x_15); -if (x_94 == 0) +x_95 = !lean_is_exclusive(x_15); +if (x_95 == 0) { -lean_object* x_95; lean_object* x_96; -x_95 = lean_ctor_get(x_15, 1); -lean_dec(x_95); -x_96 = lean_ctor_get(x_15, 0); +lean_object* x_96; lean_object* x_97; +x_96 = lean_ctor_get(x_15, 1); lean_dec(x_96); +x_97 = lean_ctor_get(x_15, 0); +lean_dec(x_97); lean_ctor_set(x_15, 1, x_4); lean_ctor_set(x_15, 0, x_14); x_3 = x_24; x_4 = x_15; -x_11 = x_36; +x_11 = x_37; goto _start; } else { -lean_object* x_98; +lean_object* x_99; lean_dec(x_15); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_14); -lean_ctor_set(x_98, 1, x_4); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_14); +lean_ctor_set(x_99, 1, x_4); x_3 = x_24; -x_4 = x_98; -x_11 = x_36; +x_4 = x_99; +x_11 = x_37; goto _start; } } @@ -13505,15 +13586,15 @@ goto _start; } else { -lean_object* x_100; lean_object* x_101; -x_100 = lean_ctor_get(x_34, 1); -lean_inc(x_100); -lean_dec(x_34); -lean_inc(x_32); -x_101 = l_Lean_findField_x3f(x_2, x_32, x_30); -if (lean_obj_tag(x_101) == 0) +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_35, 1); +lean_inc(x_101); +lean_dec(x_35); +lean_inc(x_33); +x_102 = l_Lean_findField_x3f(x_2, x_33, x_30); +if (lean_obj_tag(x_102) == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); @@ -13521,157 +13602,157 @@ lean_dec(x_24); lean_dec(x_15); lean_dec(x_14); lean_dec(x_4); -x_102 = l_Lean_MessageData_ofName(x_30); -x_103 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; -x_104 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_102); -x_105 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; +x_103 = l_Lean_MessageData_ofName(x_30); +x_104 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; +x_105 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_103); +x_106 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_105); -lean_ctor_set(x_23, 0, x_104); -x_106 = 0; -x_107 = l_Lean_MessageData_ofConstName(x_32, x_106); -x_108 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_108, 0, x_23); -lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_23, 1, x_106); +lean_ctor_set(x_23, 0, x_105); +x_107 = 0; +x_108 = l_Lean_MessageData_ofConstName(x_33, x_107); x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_103); -x_110 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_109, x_5, x_6, x_7, x_8, x_9, x_10, x_100); +lean_ctor_set(x_109, 0, x_23); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_104); +x_111 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_110, x_5, x_6, x_7, x_8, x_9, x_10, x_101); lean_dec(x_29); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_110, 1); +x_112 = lean_ctor_get(x_111, 0); lean_inc(x_112); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_113 = x_110; +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - lean_dec_ref(x_110); - x_113 = lean_box(0); + lean_dec_ref(x_111); + x_114 = lean_box(0); } -if (lean_is_scalar(x_113)) { - x_114 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); } else { - x_114 = x_113; + x_115 = x_114; } -lean_ctor_set(x_114, 0, x_111); -lean_ctor_set(x_114, 1, x_112); -return x_114; +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; } else { -lean_object* x_115; uint8_t x_116; -x_115 = lean_ctor_get(x_101, 0); -lean_inc(x_115); -lean_dec(x_101); -x_116 = lean_name_eq(x_115, x_32); -if (x_116 == 0) +lean_object* x_116; uint8_t x_117; +x_116 = lean_ctor_get(x_102, 0); +lean_inc(x_116); +lean_dec(x_102); +x_117 = lean_name_eq(x_116, x_33); +if (x_117 == 0) { -lean_object* x_117; lean_object* x_118; +lean_object* x_118; lean_object* x_119; if (lean_is_exclusive(x_14)) { lean_ctor_release(x_14, 0); lean_ctor_release(x_14, 1); lean_ctor_release(x_14, 2); lean_ctor_release(x_14, 3); - x_117 = x_14; + x_118 = x_14; } else { lean_dec_ref(x_14); - x_117 = lean_box(0); + x_118 = lean_box(0); } -x_118 = l_Lean_getPathToBaseStructure_x3f(x_2, x_115, x_32); -lean_dec(x_115); -if (lean_obj_tag(x_118) == 0) +x_119 = l_Lean_getPathToBaseStructure_x3f(x_2, x_116, x_33); +lean_dec(x_116); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; -lean_dec(x_117); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_118); lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); lean_dec(x_15); lean_dec(x_4); -x_119 = l_Lean_MessageData_ofName(x_30); -x_120 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; -x_121 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_119); -x_122 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; +x_120 = l_Lean_MessageData_ofName(x_30); +x_121 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; +x_122 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_120); +x_123 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; lean_ctor_set_tag(x_23, 7); -lean_ctor_set(x_23, 1, x_122); -lean_ctor_set(x_23, 0, x_121); -x_123 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_100); +lean_ctor_set(x_23, 1, x_123); +lean_ctor_set(x_23, 0, x_122); +x_124 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_29, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_101); lean_dec(x_29); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); +x_125 = lean_ctor_get(x_124, 0); lean_inc(x_125); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_126 = x_123; +x_126 = lean_ctor_get(x_124, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_127 = x_124; } else { - lean_dec_ref(x_123); - x_126 = lean_box(0); + lean_dec_ref(x_124); + x_127 = lean_box(0); } -if (lean_is_scalar(x_126)) { - x_127 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(1, 2, 0); } else { - x_127 = x_126; + x_128 = x_127; } -lean_ctor_set(x_127, 0, x_124); -lean_ctor_set(x_127, 1, x_125); -return x_127; +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +return x_128; } else { -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_free_object(x_23); lean_dec(x_30); -x_128 = lean_ctor_get(x_118, 0); -lean_inc(x_128); -lean_dec(x_118); -x_129 = lean_box(0); -x_130 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_128, x_129); +x_129 = lean_ctor_get(x_119, 0); +lean_inc(x_129); +lean_dec(x_119); +x_130 = lean_box(0); +x_131 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_29, x_129, x_130); lean_inc(x_15); -x_131 = l_List_appendTR___rarg(x_130, x_15); +x_132 = l_List_appendTR___rarg(x_131, x_15); if (lean_is_exclusive(x_15)) { lean_ctor_release(x_15, 0); lean_ctor_release(x_15, 1); - x_132 = x_15; + x_133 = x_15; } else { lean_dec_ref(x_15); - x_132 = lean_box(0); + x_133 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_133 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_118)) { + x_134 = lean_alloc_ctor(0, 4, 0); } else { - x_133 = x_117; + x_134 = x_118; } -lean_ctor_set(x_133, 0, x_25); -lean_ctor_set(x_133, 1, x_131); -lean_ctor_set(x_133, 2, x_26); -lean_ctor_set(x_133, 3, x_27); -if (lean_is_scalar(x_132)) { - x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_25); +lean_ctor_set(x_134, 1, x_132); +lean_ctor_set(x_134, 2, x_26); +lean_ctor_set(x_134, 3, x_27); +if (lean_is_scalar(x_133)) { + x_135 = lean_alloc_ctor(1, 2, 0); } else { - x_134 = x_132; + x_135 = x_133; } -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_4); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_4); x_3 = x_24; -x_4 = x_134; -x_11 = x_100; +x_4 = x_135; +x_11 = x_101; goto _start; } } else { -lean_object* x_136; lean_object* x_137; -lean_dec(x_115); -lean_dec(x_32); +lean_object* x_137; lean_object* x_138; +lean_dec(x_116); +lean_dec(x_33); lean_free_object(x_23); lean_dec(x_30); lean_dec(x_29); @@ -13681,21 +13762,21 @@ lean_dec(x_25); if (lean_is_exclusive(x_15)) { lean_ctor_release(x_15, 0); lean_ctor_release(x_15, 1); - x_136 = x_15; + x_137 = x_15; } else { lean_dec_ref(x_15); - x_136 = lean_box(0); + x_137 = lean_box(0); } -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); } else { - x_137 = x_136; + x_138 = x_137; } -lean_ctor_set(x_137, 0, x_14); -lean_ctor_set(x_137, 1, x_4); +lean_ctor_set(x_138, 0, x_14); +lean_ctor_set(x_138, 1, x_4); x_3 = x_24; -x_4 = x_137; -x_11 = x_100; +x_4 = x_138; +x_11 = x_101; goto _start; } } @@ -13703,39 +13784,41 @@ goto _start; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_139 = lean_ctor_get(x_23, 0); -x_140 = lean_ctor_get(x_23, 1); +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_140 = lean_ctor_get(x_23, 0); +x_141 = lean_ctor_get(x_23, 1); +lean_inc(x_141); lean_inc(x_140); -lean_inc(x_139); lean_dec(x_23); -x_141 = lean_ctor_get(x_7, 1); -x_142 = l_Lean_Elab_Term_StructInst_Struct_structName(x_1); -lean_inc(x_142); +x_142 = lean_ctor_get(x_7, 1); lean_inc(x_141); +x_143 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_143, 0, x_141); +x_144 = l_Lean_Elab_Term_StructInst_Struct_structName(x_1); +lean_inc(x_144); +lean_inc(x_142); lean_inc(x_140); -lean_inc(x_139); -x_143 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_143, 0, x_139); -lean_ctor_set(x_143, 1, x_140); -lean_ctor_set(x_143, 2, x_141); -lean_ctor_set(x_143, 3, x_142); -x_144 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_143, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_145 = lean_ctor_get(x_144, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_146 = x_144; +x_145 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_145, 0, x_140); +lean_ctor_set(x_145, 1, x_143); +lean_ctor_set(x_145, 2, x_142); +lean_ctor_set(x_145, 3, x_144); +x_146 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_145, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_147 = lean_ctor_get(x_146, 1); +lean_inc(x_147); +if (lean_is_exclusive(x_146)) { + lean_ctor_release(x_146, 0); + lean_ctor_release(x_146, 1); + x_148 = x_146; } else { - lean_dec_ref(x_144); - x_146 = lean_box(0); + lean_dec_ref(x_146); + x_148 = lean_box(0); } -lean_inc(x_142); -x_147 = l_Lean_findField_x3f(x_2, x_142, x_140); -if (lean_obj_tag(x_147) == 0) +lean_inc(x_144); +x_149 = l_Lean_findField_x3f(x_2, x_144, x_141); +if (lean_obj_tag(x_149) == 0) { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); @@ -13743,191 +13826,191 @@ lean_dec(x_24); lean_dec(x_15); lean_dec(x_14); lean_dec(x_4); -x_148 = l_Lean_MessageData_ofName(x_140); -x_149 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; -if (lean_is_scalar(x_146)) { - x_150 = lean_alloc_ctor(7, 2, 0); -} else { - x_150 = x_146; - lean_ctor_set_tag(x_150, 7); -} -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_148); -x_151 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; -x_152 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -x_153 = 0; -x_154 = l_Lean_MessageData_ofConstName(x_142, x_153); -x_155 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_155, 0, x_152); -lean_ctor_set(x_155, 1, x_154); -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_149); -x_157 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_139, x_156, x_5, x_6, x_7, x_8, x_9, x_10, x_145); -lean_dec(x_139); -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_157, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_160 = x_157; +x_150 = l_Lean_MessageData_ofName(x_141); +x_151 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__2; +if (lean_is_scalar(x_148)) { + x_152 = lean_alloc_ctor(7, 2, 0); +} else { + x_152 = x_148; + lean_ctor_set_tag(x_152, 7); +} +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_150); +x_153 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__4; +x_154 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = 0; +x_156 = l_Lean_MessageData_ofConstName(x_144, x_155); +x_157 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_157, 0, x_154); +lean_ctor_set(x_157, 1, x_156); +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_151); +x_159 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_140, x_158, x_5, x_6, x_7, x_8, x_9, x_10, x_147); +lean_dec(x_140); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_162 = x_159; } else { - lean_dec_ref(x_157); - x_160 = lean_box(0); + lean_dec_ref(x_159); + x_162 = lean_box(0); } -if (lean_is_scalar(x_160)) { - x_161 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); } else { - x_161 = x_160; + x_163 = x_162; } -lean_ctor_set(x_161, 0, x_158); -lean_ctor_set(x_161, 1, x_159); -return x_161; +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } else { -lean_object* x_162; uint8_t x_163; -x_162 = lean_ctor_get(x_147, 0); -lean_inc(x_162); -lean_dec(x_147); -x_163 = lean_name_eq(x_162, x_142); -if (x_163 == 0) +lean_object* x_164; uint8_t x_165; +x_164 = lean_ctor_get(x_149, 0); +lean_inc(x_164); +lean_dec(x_149); +x_165 = lean_name_eq(x_164, x_144); +if (x_165 == 0) { -lean_object* x_164; lean_object* x_165; +lean_object* x_166; lean_object* x_167; if (lean_is_exclusive(x_14)) { lean_ctor_release(x_14, 0); lean_ctor_release(x_14, 1); lean_ctor_release(x_14, 2); lean_ctor_release(x_14, 3); - x_164 = x_14; + x_166 = x_14; } else { lean_dec_ref(x_14); - x_164 = lean_box(0); + x_166 = lean_box(0); } -x_165 = l_Lean_getPathToBaseStructure_x3f(x_2, x_162, x_142); -lean_dec(x_162); -if (lean_obj_tag(x_165) == 0) -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_167 = l_Lean_getPathToBaseStructure_x3f(x_2, x_164, x_144); lean_dec(x_164); +if (lean_obj_tag(x_167) == 0) +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_166); lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); lean_dec(x_24); lean_dec(x_15); lean_dec(x_4); -x_166 = l_Lean_MessageData_ofName(x_140); -x_167 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; -if (lean_is_scalar(x_146)) { - x_168 = lean_alloc_ctor(7, 2, 0); +x_168 = l_Lean_MessageData_ofName(x_141); +x_169 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__6; +if (lean_is_scalar(x_148)) { + x_170 = lean_alloc_ctor(7, 2, 0); } else { - x_168 = x_146; - lean_ctor_set_tag(x_168, 7); + x_170 = x_148; + lean_ctor_set_tag(x_170, 7); } -lean_ctor_set(x_168, 0, x_167); -lean_ctor_set(x_168, 1, x_166); -x_169 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; -x_170 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_170, 0, x_168); -lean_ctor_set(x_170, 1, x_169); -x_171 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_139, x_170, x_5, x_6, x_7, x_8, x_9, x_10, x_145); -lean_dec(x_139); -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_174 = x_171; +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_168); +x_171 = l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; +x_172 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_173 = l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNumLitFields___spec__1(x_140, x_172, x_5, x_6, x_7, x_8, x_9, x_10, x_147); +lean_dec(x_140); +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_176 = x_173; } else { - lean_dec_ref(x_171); - x_174 = lean_box(0); + lean_dec_ref(x_173); + x_176 = lean_box(0); } -if (lean_is_scalar(x_174)) { - x_175 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_176)) { + x_177 = lean_alloc_ctor(1, 2, 0); } else { - x_175 = x_174; + x_177 = x_176; } -lean_ctor_set(x_175, 0, x_172); -lean_ctor_set(x_175, 1, x_173); -return x_175; +lean_ctor_set(x_177, 0, x_174); +lean_ctor_set(x_177, 1, x_175); +return x_177; } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -lean_dec(x_146); -lean_dec(x_140); -x_176 = lean_ctor_get(x_165, 0); -lean_inc(x_176); -lean_dec(x_165); -x_177 = lean_box(0); -x_178 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_139, x_176, x_177); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_148); +lean_dec(x_141); +x_178 = lean_ctor_get(x_167, 0); +lean_inc(x_178); +lean_dec(x_167); +x_179 = lean_box(0); +x_180 = l_List_mapTR_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__2(x_140, x_178, x_179); lean_inc(x_15); -x_179 = l_List_appendTR___rarg(x_178, x_15); +x_181 = l_List_appendTR___rarg(x_180, x_15); if (lean_is_exclusive(x_15)) { lean_ctor_release(x_15, 0); lean_ctor_release(x_15, 1); - x_180 = x_15; + x_182 = x_15; } else { lean_dec_ref(x_15); - x_180 = lean_box(0); + x_182 = lean_box(0); } -if (lean_is_scalar(x_164)) { - x_181 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_166)) { + x_183 = lean_alloc_ctor(0, 4, 0); } else { - x_181 = x_164; + x_183 = x_166; } -lean_ctor_set(x_181, 0, x_25); -lean_ctor_set(x_181, 1, x_179); -lean_ctor_set(x_181, 2, x_26); -lean_ctor_set(x_181, 3, x_27); -if (lean_is_scalar(x_180)) { - x_182 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_183, 0, x_25); +lean_ctor_set(x_183, 1, x_181); +lean_ctor_set(x_183, 2, x_26); +lean_ctor_set(x_183, 3, x_27); +if (lean_is_scalar(x_182)) { + x_184 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_180; + x_184 = x_182; } -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_4); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_4); x_3 = x_24; -x_4 = x_182; -x_11 = x_145; +x_4 = x_184; +x_11 = x_147; goto _start; } } else { -lean_object* x_184; lean_object* x_185; -lean_dec(x_162); -lean_dec(x_146); -lean_dec(x_142); +lean_object* x_186; lean_object* x_187; +lean_dec(x_164); +lean_dec(x_148); +lean_dec(x_144); +lean_dec(x_141); lean_dec(x_140); -lean_dec(x_139); lean_dec(x_27); lean_dec(x_26); lean_dec(x_25); if (lean_is_exclusive(x_15)) { lean_ctor_release(x_15, 0); lean_ctor_release(x_15, 1); - x_184 = x_15; + x_186 = x_15; } else { lean_dec_ref(x_15); - x_184 = lean_box(0); + x_186 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_187 = x_186; } -lean_ctor_set(x_185, 0, x_14); -lean_ctor_set(x_185, 1, x_4); +lean_ctor_set(x_187, 0, x_14); +lean_ctor_set(x_187, 1, x_4); x_3 = x_24; -x_4 = x_185; -x_11 = x_145; +x_4 = x_187; +x_11 = x_147; goto _start; } } @@ -13935,37 +14018,37 @@ goto _start; } else { -uint8_t x_187; +uint8_t x_189; lean_dec(x_23); -x_187 = !lean_is_exclusive(x_15); -if (x_187 == 0) +x_189 = !lean_is_exclusive(x_15); +if (x_189 == 0) { -lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_188 = lean_ctor_get(x_15, 1); -lean_dec(x_188); -x_189 = lean_ctor_get(x_15, 0); -lean_dec(x_189); -x_190 = lean_ctor_get(x_3, 1); -lean_inc(x_190); +lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_190 = lean_ctor_get(x_15, 1); +lean_dec(x_190); +x_191 = lean_ctor_get(x_15, 0); +lean_dec(x_191); +x_192 = lean_ctor_get(x_3, 1); +lean_inc(x_192); lean_dec(x_3); lean_ctor_set(x_15, 1, x_4); lean_ctor_set(x_15, 0, x_14); -x_3 = x_190; +x_3 = x_192; x_4 = x_15; goto _start; } else { -lean_object* x_192; lean_object* x_193; +lean_object* x_194; lean_object* x_195; lean_dec(x_15); -x_192 = lean_ctor_get(x_3, 1); -lean_inc(x_192); +x_194 = lean_ctor_get(x_3, 1); +lean_inc(x_194); lean_dec(x_3); -x_193 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_193, 0, x_14); -lean_ctor_set(x_193, 1, x_4); -x_3 = x_192; -x_4 = x_193; +x_195 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_195, 0, x_14); +lean_ctor_set(x_195, 1, x_4); +x_3 = x_194; +x_4 = x_195; goto _start; } } @@ -16773,7 +16856,7 @@ lean_inc(x_18); x_49 = l_Lean_isSubobjectField_x3f(x_2, x_48, x_18); if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; x_50 = lean_unsigned_to_nat(4u); x_51 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; lean_inc(x_3); @@ -16791,179 +16874,188 @@ x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_58); -x_62 = lean_unsigned_to_nat(2u); -x_63 = l_Lean_Syntax_setArg(x_52, x_62, x_61); -x_64 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_array_get_size(x_65); -x_67 = lean_unsigned_to_nat(0u); -x_68 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_65, x_67, x_66, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_66); -lean_dec(x_65); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_22); +x_63 = lean_array_mk(x_62); +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_63); +x_66 = lean_unsigned_to_nat(2u); +x_67 = l_Lean_Syntax_setArg(x_52, x_66, x_65); +x_68 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Array_isEmpty___rarg(x_69); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_64, 1); -lean_inc(x_72); -lean_dec(x_64); -x_73 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_69); +x_70 = lean_array_get_size(x_69); +x_71 = lean_unsigned_to_nat(0u); +x_72 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_69, x_71, x_70, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_70); lean_dec(x_69); -x_74 = lean_unsigned_to_nat(1u); -x_75 = l_Lean_Syntax_setArg(x_63, x_74, x_73); -if (lean_obj_tag(x_72) == 0) +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Array_isEmpty___rarg(x_73); +if (x_75 == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_unsigned_to_nat(3u); -x_77 = l_Lean_Syntax_setArg(x_75, x_76, x_51); -x_26 = x_77; -x_27 = x_70; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_68, 1); +lean_inc(x_76); +lean_dec(x_68); +x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_73); +lean_dec(x_73); +x_78 = lean_unsigned_to_nat(1u); +x_79 = l_Lean_Syntax_setArg(x_67, x_78, x_77); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_unsigned_to_nat(3u); +x_81 = l_Lean_Syntax_setArg(x_79, x_80, x_51); +x_26 = x_81; +x_27 = x_74; goto block_47; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_72, 0); -lean_inc(x_78); -lean_dec(x_72); -x_79 = lean_unsigned_to_nat(3u); -x_80 = l_Lean_Syntax_setArg(x_75, x_79, x_78); -x_26 = x_80; -x_27 = x_70; +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_76, 0); +lean_inc(x_82); +lean_dec(x_76); +x_83 = lean_unsigned_to_nat(3u); +x_84 = l_Lean_Syntax_setArg(x_79, x_83, x_82); +x_26 = x_84; +x_27 = x_74; goto block_47; } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_69); -x_81 = lean_ctor_get(x_64, 1); -lean_inc(x_81); -lean_dec(x_64); -x_82 = lean_unsigned_to_nat(1u); -x_83 = l_Lean_Syntax_setArg(x_63, x_82, x_51); -if (lean_obj_tag(x_81) == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_73); +x_85 = lean_ctor_get(x_68, 1); +lean_inc(x_85); +lean_dec(x_68); +x_86 = lean_unsigned_to_nat(1u); +x_87 = l_Lean_Syntax_setArg(x_67, x_86, x_51); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_84; lean_object* x_85; -x_84 = lean_unsigned_to_nat(3u); -x_85 = l_Lean_Syntax_setArg(x_83, x_84, x_51); -x_26 = x_85; -x_27 = x_70; +lean_object* x_88; lean_object* x_89; +x_88 = lean_unsigned_to_nat(3u); +x_89 = l_Lean_Syntax_setArg(x_87, x_88, x_51); +x_26 = x_89; +x_27 = x_74; goto block_47; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_81, 0); -lean_inc(x_86); -lean_dec(x_81); -x_87 = lean_unsigned_to_nat(3u); -x_88 = l_Lean_Syntax_setArg(x_83, x_87, x_86); -x_26 = x_88; -x_27 = x_70; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_85, 0); +lean_inc(x_90); +lean_dec(x_85); +x_91 = lean_unsigned_to_nat(3u); +x_92 = l_Lean_Syntax_setArg(x_87, x_91, x_90); +x_26 = x_92; +x_27 = x_74; goto block_47; } } } else { -uint8_t x_89; +uint8_t x_93; lean_dec(x_20); lean_dec(x_18); lean_dec(x_17); -x_89 = !lean_is_exclusive(x_49); -if (x_89 == 0) +x_93 = !lean_is_exclusive(x_49); +if (x_93 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = lean_ctor_get(x_49, 0); -x_91 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_92 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_94 = lean_ctor_get(x_49, 0); +x_95 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_96 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_93 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_93, 0, x_3); -lean_ctor_set(x_93, 1, x_90); -lean_ctor_set(x_93, 2, x_92); -lean_ctor_set(x_93, 3, x_23); -lean_ctor_set(x_93, 4, x_91); +x_97 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_97, 0, x_3); +lean_ctor_set(x_97, 1, x_94); +lean_ctor_set(x_97, 2, x_96); +lean_ctor_set(x_97, 3, x_23); +lean_ctor_set(x_97, 4, x_95); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_94 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_93, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = !lean_is_exclusive(x_25); -if (x_97 == 0) +x_98 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_97, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_98 = lean_ctor_get(x_25, 1); -x_99 = lean_ctor_get(x_25, 2); -lean_dec(x_99); -x_100 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_101 = l_List_head_x21___rarg(x_100, x_98); +lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); lean_dec(x_98); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_22); -lean_ctor_set(x_49, 0, x_95); +x_101 = !lean_is_exclusive(x_25); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_102 = lean_ctor_get(x_25, 1); +x_103 = lean_ctor_get(x_25, 2); +lean_dec(x_103); +x_104 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_105 = l_List_head_x21___rarg(x_104, x_102); +lean_dec(x_102); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_22); +lean_ctor_set(x_49, 0, x_99); lean_ctor_set(x_25, 2, x_49); -lean_ctor_set(x_25, 1, x_102); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_25); -lean_ctor_set(x_103, 1, x_5); +lean_ctor_set(x_25, 1, x_106); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_25); +lean_ctor_set(x_107, 1, x_5); x_4 = x_16; -x_5 = x_103; -x_12 = x_96; +x_5 = x_107; +x_12 = x_100; goto _start; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_105 = lean_ctor_get(x_25, 0); -x_106 = lean_ctor_get(x_25, 1); -x_107 = lean_ctor_get(x_25, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_25, 0); +x_110 = lean_ctor_get(x_25, 1); +x_111 = lean_ctor_get(x_25, 3); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_25); -x_108 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_109 = l_List_head_x21___rarg(x_108, x_106); -lean_dec(x_106); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_22); -lean_ctor_set(x_49, 0, x_95); -x_111 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_49); -lean_ctor_set(x_111, 3, x_107); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_5); +x_112 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_113 = l_List_head_x21___rarg(x_112, x_110); +lean_dec(x_110); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_22); +lean_ctor_set(x_49, 0, x_99); +x_115 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_49); +lean_ctor_set(x_115, 3, x_111); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_5); x_4 = x_16; -x_5 = x_112; -x_12 = x_96; +x_5 = x_116; +x_12 = x_100; goto _start; } } else { -uint8_t x_114; +uint8_t x_118; lean_free_object(x_49); lean_dec(x_25); lean_dec(x_16); @@ -16975,100 +17067,100 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_114 = !lean_is_exclusive(x_94); -if (x_114 == 0) +x_118 = !lean_is_exclusive(x_98); +if (x_118 == 0) { -return x_94; +return x_98; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_94, 0); -x_116 = lean_ctor_get(x_94, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_94); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_98, 0); +x_120 = lean_ctor_get(x_98, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_98); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_118 = lean_ctor_get(x_49, 0); -lean_inc(x_118); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_122 = lean_ctor_get(x_49, 0); +lean_inc(x_122); lean_dec(x_49); -x_119 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_120 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +x_123 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_124 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_121 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_121, 0, x_3); -lean_ctor_set(x_121, 1, x_118); -lean_ctor_set(x_121, 2, x_120); -lean_ctor_set(x_121, 3, x_23); -lean_ctor_set(x_121, 4, x_119); +x_125 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_125, 0, x_3); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_124); +lean_ctor_set(x_125, 3, x_23); +lean_ctor_set(x_125, 4, x_123); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_122 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_121, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_122) == 0) +x_126 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_125, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_126) == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_ctor_get(x_25, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_25, 1); -lean_inc(x_126); -x_127 = lean_ctor_get(x_25, 3); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_ctor_get(x_25, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_25, 1); +lean_inc(x_130); +x_131 = lean_ctor_get(x_25, 3); +lean_inc(x_131); if (lean_is_exclusive(x_25)) { lean_ctor_release(x_25, 0); lean_ctor_release(x_25, 1); lean_ctor_release(x_25, 2); lean_ctor_release(x_25, 3); - x_128 = x_25; + x_132 = x_25; } else { lean_dec_ref(x_25); - x_128 = lean_box(0); + x_132 = lean_box(0); } -x_129 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_130 = l_List_head_x21___rarg(x_129, x_126); -lean_dec(x_126); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_22); -x_132 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_132, 0, x_123); -if (lean_is_scalar(x_128)) { - x_133 = lean_alloc_ctor(0, 4, 0); -} else { - x_133 = x_128; -} -lean_ctor_set(x_133, 0, x_125); -lean_ctor_set(x_133, 1, x_131); -lean_ctor_set(x_133, 2, x_132); -lean_ctor_set(x_133, 3, x_127); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_5); +x_133 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_134 = l_List_head_x21___rarg(x_133, x_130); +lean_dec(x_130); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_22); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_127); +if (lean_is_scalar(x_132)) { + x_137 = lean_alloc_ctor(0, 4, 0); +} else { + x_137 = x_132; +} +lean_ctor_set(x_137, 0, x_129); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_136); +lean_ctor_set(x_137, 3, x_131); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_5); x_4 = x_16; -x_5 = x_134; -x_12 = x_124; +x_5 = x_138; +x_12 = x_128; goto _start; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_dec(x_25); lean_dec(x_16); lean_dec(x_11); @@ -17079,26 +17171,26 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_136 = lean_ctor_get(x_122, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_122, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_138 = x_122; +x_140 = lean_ctor_get(x_126, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_126, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_142 = x_126; } else { - lean_dec_ref(x_122); - x_138 = lean_box(0); + lean_dec_ref(x_126); + x_142 = lean_box(0); } -if (lean_is_scalar(x_138)) { - x_139 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(1, 2, 0); } else { - x_139 = x_138; + x_143 = x_142; } -lean_ctor_set(x_139, 0, x_136); -lean_ctor_set(x_139, 1, x_137); -return x_139; +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_141); +return x_143; } } } @@ -17183,22 +17275,22 @@ goto _start; } else { -lean_object* x_140; lean_object* x_141; +lean_object* x_144; lean_object* x_145; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -x_140 = lean_ctor_get(x_21, 0); -lean_inc(x_140); +x_144 = lean_ctor_get(x_21, 0); +lean_inc(x_144); lean_dec(x_21); if (lean_is_scalar(x_17)) { - x_141 = lean_alloc_ctor(1, 2, 0); + x_145 = lean_alloc_ctor(1, 2, 0); } else { - x_141 = x_17; + x_145 = x_17; } -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_5); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_5); x_4 = x_16; -x_5 = x_141; +x_5 = x_145; goto _start; } } @@ -17265,7 +17357,7 @@ lean_inc(x_18); x_49 = l_Lean_isSubobjectField_x3f(x_2, x_48, x_18); if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; x_50 = lean_unsigned_to_nat(4u); x_51 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; lean_inc(x_3); @@ -17283,179 +17375,188 @@ x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_58); -x_62 = lean_unsigned_to_nat(2u); -x_63 = l_Lean_Syntax_setArg(x_52, x_62, x_61); -x_64 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_array_get_size(x_65); -x_67 = lean_unsigned_to_nat(0u); -x_68 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_65, x_67, x_66, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_66); -lean_dec(x_65); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_22); +x_63 = lean_array_mk(x_62); +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_63); +x_66 = lean_unsigned_to_nat(2u); +x_67 = l_Lean_Syntax_setArg(x_52, x_66, x_65); +x_68 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Array_isEmpty___rarg(x_69); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_64, 1); -lean_inc(x_72); -lean_dec(x_64); -x_73 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_69); +x_70 = lean_array_get_size(x_69); +x_71 = lean_unsigned_to_nat(0u); +x_72 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_69, x_71, x_70, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_70); lean_dec(x_69); -x_74 = lean_unsigned_to_nat(1u); -x_75 = l_Lean_Syntax_setArg(x_63, x_74, x_73); -if (lean_obj_tag(x_72) == 0) +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Array_isEmpty___rarg(x_73); +if (x_75 == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_unsigned_to_nat(3u); -x_77 = l_Lean_Syntax_setArg(x_75, x_76, x_51); -x_26 = x_77; -x_27 = x_70; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_68, 1); +lean_inc(x_76); +lean_dec(x_68); +x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_73); +lean_dec(x_73); +x_78 = lean_unsigned_to_nat(1u); +x_79 = l_Lean_Syntax_setArg(x_67, x_78, x_77); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_unsigned_to_nat(3u); +x_81 = l_Lean_Syntax_setArg(x_79, x_80, x_51); +x_26 = x_81; +x_27 = x_74; goto block_47; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_72, 0); -lean_inc(x_78); -lean_dec(x_72); -x_79 = lean_unsigned_to_nat(3u); -x_80 = l_Lean_Syntax_setArg(x_75, x_79, x_78); -x_26 = x_80; -x_27 = x_70; +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_76, 0); +lean_inc(x_82); +lean_dec(x_76); +x_83 = lean_unsigned_to_nat(3u); +x_84 = l_Lean_Syntax_setArg(x_79, x_83, x_82); +x_26 = x_84; +x_27 = x_74; goto block_47; } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_69); -x_81 = lean_ctor_get(x_64, 1); -lean_inc(x_81); -lean_dec(x_64); -x_82 = lean_unsigned_to_nat(1u); -x_83 = l_Lean_Syntax_setArg(x_63, x_82, x_51); -if (lean_obj_tag(x_81) == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_73); +x_85 = lean_ctor_get(x_68, 1); +lean_inc(x_85); +lean_dec(x_68); +x_86 = lean_unsigned_to_nat(1u); +x_87 = l_Lean_Syntax_setArg(x_67, x_86, x_51); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_84; lean_object* x_85; -x_84 = lean_unsigned_to_nat(3u); -x_85 = l_Lean_Syntax_setArg(x_83, x_84, x_51); -x_26 = x_85; -x_27 = x_70; +lean_object* x_88; lean_object* x_89; +x_88 = lean_unsigned_to_nat(3u); +x_89 = l_Lean_Syntax_setArg(x_87, x_88, x_51); +x_26 = x_89; +x_27 = x_74; goto block_47; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_81, 0); -lean_inc(x_86); -lean_dec(x_81); -x_87 = lean_unsigned_to_nat(3u); -x_88 = l_Lean_Syntax_setArg(x_83, x_87, x_86); -x_26 = x_88; -x_27 = x_70; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_85, 0); +lean_inc(x_90); +lean_dec(x_85); +x_91 = lean_unsigned_to_nat(3u); +x_92 = l_Lean_Syntax_setArg(x_87, x_91, x_90); +x_26 = x_92; +x_27 = x_74; goto block_47; } } } else { -uint8_t x_89; +uint8_t x_93; lean_dec(x_20); lean_dec(x_18); lean_dec(x_17); -x_89 = !lean_is_exclusive(x_49); -if (x_89 == 0) +x_93 = !lean_is_exclusive(x_49); +if (x_93 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = lean_ctor_get(x_49, 0); -x_91 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_92 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_94 = lean_ctor_get(x_49, 0); +x_95 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_96 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_93 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_93, 0, x_3); -lean_ctor_set(x_93, 1, x_90); -lean_ctor_set(x_93, 2, x_92); -lean_ctor_set(x_93, 3, x_23); -lean_ctor_set(x_93, 4, x_91); +x_97 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_97, 0, x_3); +lean_ctor_set(x_97, 1, x_94); +lean_ctor_set(x_97, 2, x_96); +lean_ctor_set(x_97, 3, x_23); +lean_ctor_set(x_97, 4, x_95); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_94 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_93, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = !lean_is_exclusive(x_25); -if (x_97 == 0) +x_98 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_97, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_98 = lean_ctor_get(x_25, 1); -x_99 = lean_ctor_get(x_25, 2); -lean_dec(x_99); -x_100 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_101 = l_List_head_x21___rarg(x_100, x_98); +lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); lean_dec(x_98); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_22); -lean_ctor_set(x_49, 0, x_95); +x_101 = !lean_is_exclusive(x_25); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_102 = lean_ctor_get(x_25, 1); +x_103 = lean_ctor_get(x_25, 2); +lean_dec(x_103); +x_104 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_105 = l_List_head_x21___rarg(x_104, x_102); +lean_dec(x_102); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_22); +lean_ctor_set(x_49, 0, x_99); lean_ctor_set(x_25, 2, x_49); -lean_ctor_set(x_25, 1, x_102); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_25); -lean_ctor_set(x_103, 1, x_5); +lean_ctor_set(x_25, 1, x_106); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_25); +lean_ctor_set(x_107, 1, x_5); x_4 = x_16; -x_5 = x_103; -x_12 = x_96; +x_5 = x_107; +x_12 = x_100; goto _start; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_105 = lean_ctor_get(x_25, 0); -x_106 = lean_ctor_get(x_25, 1); -x_107 = lean_ctor_get(x_25, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_25, 0); +x_110 = lean_ctor_get(x_25, 1); +x_111 = lean_ctor_get(x_25, 3); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_25); -x_108 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_109 = l_List_head_x21___rarg(x_108, x_106); -lean_dec(x_106); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_22); -lean_ctor_set(x_49, 0, x_95); -x_111 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_49); -lean_ctor_set(x_111, 3, x_107); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_5); +x_112 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_113 = l_List_head_x21___rarg(x_112, x_110); +lean_dec(x_110); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_22); +lean_ctor_set(x_49, 0, x_99); +x_115 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_49); +lean_ctor_set(x_115, 3, x_111); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_5); x_4 = x_16; -x_5 = x_112; -x_12 = x_96; +x_5 = x_116; +x_12 = x_100; goto _start; } } else { -uint8_t x_114; +uint8_t x_118; lean_free_object(x_49); lean_dec(x_25); lean_dec(x_16); @@ -17467,100 +17568,100 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_114 = !lean_is_exclusive(x_94); -if (x_114 == 0) +x_118 = !lean_is_exclusive(x_98); +if (x_118 == 0) { -return x_94; +return x_98; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_94, 0); -x_116 = lean_ctor_get(x_94, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_94); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_98, 0); +x_120 = lean_ctor_get(x_98, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_98); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_118 = lean_ctor_get(x_49, 0); -lean_inc(x_118); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_122 = lean_ctor_get(x_49, 0); +lean_inc(x_122); lean_dec(x_49); -x_119 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_120 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +x_123 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_124 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_121 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_121, 0, x_3); -lean_ctor_set(x_121, 1, x_118); -lean_ctor_set(x_121, 2, x_120); -lean_ctor_set(x_121, 3, x_23); -lean_ctor_set(x_121, 4, x_119); +x_125 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_125, 0, x_3); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_124); +lean_ctor_set(x_125, 3, x_23); +lean_ctor_set(x_125, 4, x_123); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_122 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_121, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_122) == 0) +x_126 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_125, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_126) == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_ctor_get(x_25, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_25, 1); -lean_inc(x_126); -x_127 = lean_ctor_get(x_25, 3); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_ctor_get(x_25, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_25, 1); +lean_inc(x_130); +x_131 = lean_ctor_get(x_25, 3); +lean_inc(x_131); if (lean_is_exclusive(x_25)) { lean_ctor_release(x_25, 0); lean_ctor_release(x_25, 1); lean_ctor_release(x_25, 2); lean_ctor_release(x_25, 3); - x_128 = x_25; + x_132 = x_25; } else { lean_dec_ref(x_25); - x_128 = lean_box(0); + x_132 = lean_box(0); } -x_129 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_130 = l_List_head_x21___rarg(x_129, x_126); -lean_dec(x_126); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_22); -x_132 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_132, 0, x_123); -if (lean_is_scalar(x_128)) { - x_133 = lean_alloc_ctor(0, 4, 0); -} else { - x_133 = x_128; -} -lean_ctor_set(x_133, 0, x_125); -lean_ctor_set(x_133, 1, x_131); -lean_ctor_set(x_133, 2, x_132); -lean_ctor_set(x_133, 3, x_127); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_5); +x_133 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_134 = l_List_head_x21___rarg(x_133, x_130); +lean_dec(x_130); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_22); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_127); +if (lean_is_scalar(x_132)) { + x_137 = lean_alloc_ctor(0, 4, 0); +} else { + x_137 = x_132; +} +lean_ctor_set(x_137, 0, x_129); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_136); +lean_ctor_set(x_137, 3, x_131); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_5); x_4 = x_16; -x_5 = x_134; -x_12 = x_124; +x_5 = x_138; +x_12 = x_128; goto _start; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_dec(x_25); lean_dec(x_16); lean_dec(x_11); @@ -17571,26 +17672,26 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_136 = lean_ctor_get(x_122, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_122, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_138 = x_122; +x_140 = lean_ctor_get(x_126, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_126, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_142 = x_126; } else { - lean_dec_ref(x_122); - x_138 = lean_box(0); + lean_dec_ref(x_126); + x_142 = lean_box(0); } -if (lean_is_scalar(x_138)) { - x_139 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(1, 2, 0); } else { - x_139 = x_138; + x_143 = x_142; } -lean_ctor_set(x_139, 0, x_136); -lean_ctor_set(x_139, 1, x_137); -return x_139; +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_141); +return x_143; } } } @@ -17675,22 +17776,22 @@ goto _start; } else { -lean_object* x_140; lean_object* x_141; +lean_object* x_144; lean_object* x_145; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -x_140 = lean_ctor_get(x_21, 0); -lean_inc(x_140); +x_144 = lean_ctor_get(x_21, 0); +lean_inc(x_144); lean_dec(x_21); if (lean_is_scalar(x_17)) { - x_141 = lean_alloc_ctor(1, 2, 0); + x_145 = lean_alloc_ctor(1, 2, 0); } else { - x_141 = x_17; + x_145 = x_17; } -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_5); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_5); x_4 = x_16; -x_5 = x_141; +x_5 = x_145; goto _start; } } @@ -17780,7 +17881,7 @@ lean_inc(x_18); x_49 = l_Lean_isSubobjectField_x3f(x_2, x_48, x_18); if (lean_obj_tag(x_49) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; x_50 = lean_unsigned_to_nat(4u); x_51 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__4; lean_inc(x_3); @@ -17798,179 +17899,188 @@ x_61 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); lean_ctor_set(x_61, 2, x_58); -x_62 = lean_unsigned_to_nat(2u); -x_63 = l_Lean_Syntax_setArg(x_52, x_62, x_61); -x_64 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -x_66 = lean_array_get_size(x_65); -x_67 = lean_unsigned_to_nat(0u); -x_68 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_65, x_67, x_66, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_66); -lean_dec(x_65); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_22); +x_63 = lean_array_mk(x_62); +x_64 = l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2; +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_59); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_63); +x_66 = lean_unsigned_to_nat(2u); +x_67 = l_Lean_Syntax_setArg(x_52, x_66, x_65); +x_68 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -x_70 = lean_ctor_get(x_68, 1); -lean_inc(x_70); -lean_dec(x_68); -x_71 = l_Array_isEmpty___rarg(x_69); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_64, 1); -lean_inc(x_72); -lean_dec(x_64); -x_73 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_69); +x_70 = lean_array_get_size(x_69); +x_71 = lean_unsigned_to_nat(0u); +x_72 = l_Array_filterMapM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_groupFields___spec__3(x_18, x_69, x_71, x_70, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_70); lean_dec(x_69); -x_74 = lean_unsigned_to_nat(1u); -x_75 = l_Lean_Syntax_setArg(x_63, x_74, x_73); -if (lean_obj_tag(x_72) == 0) +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Array_isEmpty___rarg(x_73); +if (x_75 == 0) { -lean_object* x_76; lean_object* x_77; -x_76 = lean_unsigned_to_nat(3u); -x_77 = l_Lean_Syntax_setArg(x_75, x_76, x_51); -x_26 = x_77; -x_27 = x_70; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_68, 1); +lean_inc(x_76); +lean_dec(x_68); +x_77 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkSourcesWithSyntax(x_73); +lean_dec(x_73); +x_78 = lean_unsigned_to_nat(1u); +x_79 = l_Lean_Syntax_setArg(x_67, x_78, x_77); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_unsigned_to_nat(3u); +x_81 = l_Lean_Syntax_setArg(x_79, x_80, x_51); +x_26 = x_81; +x_27 = x_74; goto block_47; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_72, 0); -lean_inc(x_78); -lean_dec(x_72); -x_79 = lean_unsigned_to_nat(3u); -x_80 = l_Lean_Syntax_setArg(x_75, x_79, x_78); -x_26 = x_80; -x_27 = x_70; +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_76, 0); +lean_inc(x_82); +lean_dec(x_76); +x_83 = lean_unsigned_to_nat(3u); +x_84 = l_Lean_Syntax_setArg(x_79, x_83, x_82); +x_26 = x_84; +x_27 = x_74; goto block_47; } } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_69); -x_81 = lean_ctor_get(x_64, 1); -lean_inc(x_81); -lean_dec(x_64); -x_82 = lean_unsigned_to_nat(1u); -x_83 = l_Lean_Syntax_setArg(x_63, x_82, x_51); -if (lean_obj_tag(x_81) == 0) +lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_73); +x_85 = lean_ctor_get(x_68, 1); +lean_inc(x_85); +lean_dec(x_68); +x_86 = lean_unsigned_to_nat(1u); +x_87 = l_Lean_Syntax_setArg(x_67, x_86, x_51); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_84; lean_object* x_85; -x_84 = lean_unsigned_to_nat(3u); -x_85 = l_Lean_Syntax_setArg(x_83, x_84, x_51); -x_26 = x_85; -x_27 = x_70; +lean_object* x_88; lean_object* x_89; +x_88 = lean_unsigned_to_nat(3u); +x_89 = l_Lean_Syntax_setArg(x_87, x_88, x_51); +x_26 = x_89; +x_27 = x_74; goto block_47; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_81, 0); -lean_inc(x_86); -lean_dec(x_81); -x_87 = lean_unsigned_to_nat(3u); -x_88 = l_Lean_Syntax_setArg(x_83, x_87, x_86); -x_26 = x_88; -x_27 = x_70; +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_85, 0); +lean_inc(x_90); +lean_dec(x_85); +x_91 = lean_unsigned_to_nat(3u); +x_92 = l_Lean_Syntax_setArg(x_87, x_91, x_90); +x_26 = x_92; +x_27 = x_74; goto block_47; } } } else { -uint8_t x_89; +uint8_t x_93; lean_dec(x_20); lean_dec(x_18); lean_dec(x_17); -x_89 = !lean_is_exclusive(x_49); -if (x_89 == 0) +x_93 = !lean_is_exclusive(x_49); +if (x_93 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_90 = lean_ctor_get(x_49, 0); -x_91 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_92 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_94 = lean_ctor_get(x_49, 0); +x_95 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_96 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_93 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_93, 0, x_3); -lean_ctor_set(x_93, 1, x_90); -lean_ctor_set(x_93, 2, x_92); -lean_ctor_set(x_93, 3, x_23); -lean_ctor_set(x_93, 4, x_91); +x_97 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_97, 0, x_3); +lean_ctor_set(x_97, 1, x_94); +lean_ctor_set(x_97, 2, x_96); +lean_ctor_set(x_97, 3, x_23); +lean_ctor_set(x_97, 4, x_95); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_94 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_93, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_94) == 0) -{ -lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = !lean_is_exclusive(x_25); -if (x_97 == 0) +x_98 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_97, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_98 = lean_ctor_get(x_25, 1); -x_99 = lean_ctor_get(x_25, 2); -lean_dec(x_99); -x_100 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_101 = l_List_head_x21___rarg(x_100, x_98); +lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); lean_dec(x_98); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_22); -lean_ctor_set(x_49, 0, x_95); +x_101 = !lean_is_exclusive(x_25); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_102 = lean_ctor_get(x_25, 1); +x_103 = lean_ctor_get(x_25, 2); +lean_dec(x_103); +x_104 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_105 = l_List_head_x21___rarg(x_104, x_102); +lean_dec(x_102); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_22); +lean_ctor_set(x_49, 0, x_99); lean_ctor_set(x_25, 2, x_49); -lean_ctor_set(x_25, 1, x_102); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_25); -lean_ctor_set(x_103, 1, x_5); +lean_ctor_set(x_25, 1, x_106); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_25); +lean_ctor_set(x_107, 1, x_5); x_4 = x_16; -x_5 = x_103; -x_12 = x_96; +x_5 = x_107; +x_12 = x_100; goto _start; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_105 = lean_ctor_get(x_25, 0); -x_106 = lean_ctor_get(x_25, 1); -x_107 = lean_ctor_get(x_25, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_25, 0); +x_110 = lean_ctor_get(x_25, 1); +x_111 = lean_ctor_get(x_25, 3); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_25); -x_108 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_109 = l_List_head_x21___rarg(x_108, x_106); -lean_dec(x_106); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_22); -lean_ctor_set(x_49, 0, x_95); -x_111 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_111, 0, x_105); -lean_ctor_set(x_111, 1, x_110); -lean_ctor_set(x_111, 2, x_49); -lean_ctor_set(x_111, 3, x_107); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_5); +x_112 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_113 = l_List_head_x21___rarg(x_112, x_110); +lean_dec(x_110); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_22); +lean_ctor_set(x_49, 0, x_99); +x_115 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_115, 2, x_49); +lean_ctor_set(x_115, 3, x_111); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_5); x_4 = x_16; -x_5 = x_112; -x_12 = x_96; +x_5 = x_116; +x_12 = x_100; goto _start; } } else { -uint8_t x_114; +uint8_t x_118; lean_free_object(x_49); lean_dec(x_25); lean_dec(x_16); @@ -17982,100 +18092,100 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_114 = !lean_is_exclusive(x_94); -if (x_114 == 0) +x_118 = !lean_is_exclusive(x_98); +if (x_118 == 0) { -return x_94; +return x_98; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_94, 0); -x_116 = lean_ctor_get(x_94, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_94); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_98, 0); +x_120 = lean_ctor_get(x_98, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_98); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_118 = lean_ctor_get(x_49, 0); -lean_inc(x_118); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_122 = lean_ctor_get(x_49, 0); +lean_inc(x_122); lean_dec(x_49); -x_119 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); -x_120 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; +x_123 = l_Lean_Elab_Term_StructInst_Struct_source(x_1); +x_124 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandNonAtomicExplicitSources___lambda__1___closed__1; lean_inc(x_3); -x_121 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_121, 0, x_3); -lean_ctor_set(x_121, 1, x_118); -lean_ctor_set(x_121, 2, x_120); -lean_ctor_set(x_121, 3, x_23); -lean_ctor_set(x_121, 4, x_119); +x_125 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_125, 0, x_3); +lean_ctor_set(x_125, 1, x_122); +lean_ctor_set(x_125, 2, x_124); +lean_ctor_set(x_125, 3, x_23); +lean_ctor_set(x_125, 4, x_123); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_122 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_121, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_122) == 0) +x_126 = l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandStruct(x_125, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_126) == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_ctor_get(x_25, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_25, 1); -lean_inc(x_126); -x_127 = lean_ctor_get(x_25, 3); +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_127 = lean_ctor_get(x_126, 0); lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_ctor_get(x_25, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_25, 1); +lean_inc(x_130); +x_131 = lean_ctor_get(x_25, 3); +lean_inc(x_131); if (lean_is_exclusive(x_25)) { lean_ctor_release(x_25, 0); lean_ctor_release(x_25, 1); lean_ctor_release(x_25, 2); lean_ctor_release(x_25, 3); - x_128 = x_25; + x_132 = x_25; } else { lean_dec_ref(x_25); - x_128 = lean_box(0); + x_132 = lean_box(0); } -x_129 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; -x_130 = l_List_head_x21___rarg(x_129, x_126); -lean_dec(x_126); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_22); -x_132 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_132, 0, x_123); -if (lean_is_scalar(x_128)) { - x_133 = lean_alloc_ctor(0, 4, 0); -} else { - x_133 = x_128; -} -lean_ctor_set(x_133, 0, x_125); -lean_ctor_set(x_133, 1, x_131); -lean_ctor_set(x_133, 2, x_132); -lean_ctor_set(x_133, 3, x_127); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_5); +x_133 = l_Lean_Elab_Term_StructInst_instInhabitedFieldLHS; +x_134 = l_List_head_x21___rarg(x_133, x_130); +lean_dec(x_130); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_22); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_127); +if (lean_is_scalar(x_132)) { + x_137 = lean_alloc_ctor(0, 4, 0); +} else { + x_137 = x_132; +} +lean_ctor_set(x_137, 0, x_129); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_136); +lean_ctor_set(x_137, 3, x_131); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_5); x_4 = x_16; -x_5 = x_134; -x_12 = x_124; +x_5 = x_138; +x_12 = x_128; goto _start; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_dec(x_25); lean_dec(x_16); lean_dec(x_11); @@ -18086,26 +18196,26 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_136 = lean_ctor_get(x_122, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_122, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_138 = x_122; +x_140 = lean_ctor_get(x_126, 0); +lean_inc(x_140); +x_141 = lean_ctor_get(x_126, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_142 = x_126; } else { - lean_dec_ref(x_122); - x_138 = lean_box(0); + lean_dec_ref(x_126); + x_142 = lean_box(0); } -if (lean_is_scalar(x_138)) { - x_139 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_142)) { + x_143 = lean_alloc_ctor(1, 2, 0); } else { - x_139 = x_138; + x_143 = x_142; } -lean_ctor_set(x_139, 0, x_136); -lean_ctor_set(x_139, 1, x_137); -return x_139; +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_141); +return x_143; } } } @@ -18190,22 +18300,22 @@ goto _start; } else { -lean_object* x_140; lean_object* x_141; +lean_object* x_144; lean_object* x_145; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); -x_140 = lean_ctor_get(x_21, 0); -lean_inc(x_140); +x_144 = lean_ctor_get(x_21, 0); +lean_inc(x_144); lean_dec(x_21); if (lean_is_scalar(x_17)) { - x_141 = lean_alloc_ctor(1, 2, 0); + x_145 = lean_alloc_ctor(1, 2, 0); } else { - x_141 = x_17; + x_145 = x_17; } -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_5); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_5); x_4 = x_16; -x_5 = x_141; +x_5 = x_145; goto _start; } } @@ -21857,7 +21967,7 @@ lean_ctor_set(x_32, 1, x_3); lean_ctor_set(x_32, 2, x_30); lean_ctor_set(x_32, 3, x_28); lean_ctor_set(x_32, 4, x_5); -x_33 = lean_alloc_ctor(5, 1, 0); +x_33 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_33, 0, x_32); x_34 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_35 = lean_alloc_ctor(1, 2, 0); @@ -22071,7 +22181,7 @@ lean_ctor_set(x_85, 1, x_3); lean_ctor_set(x_85, 2, x_83); lean_ctor_set(x_85, 3, x_79); lean_ctor_set(x_85, 4, x_5); -x_86 = lean_alloc_ctor(5, 1, 0); +x_86 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_86, 0, x_85); x_87 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_88 = lean_alloc_ctor(1, 2, 0); @@ -22263,7 +22373,7 @@ lean_ctor_set(x_126, 1, x_3); lean_ctor_set(x_126, 2, x_124); lean_ctor_set(x_126, 3, x_120); lean_ctor_set(x_126, 4, x_5); -x_127 = lean_alloc_ctor(5, 1, 0); +x_127 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_127, 0, x_126); x_128 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_129 = lean_alloc_ctor(1, 2, 0); @@ -22363,7 +22473,7 @@ lean_ctor_set(x_155, 1, x_3); lean_ctor_set(x_155, 2, x_153); lean_ctor_set(x_155, 3, x_149); lean_ctor_set(x_155, 4, x_5); -x_156 = lean_alloc_ctor(5, 1, 0); +x_156 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_156, 0, x_155); x_157 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_158 = lean_alloc_ctor(1, 2, 0); @@ -22543,7 +22653,7 @@ lean_ctor_set(x_195, 1, x_3); lean_ctor_set(x_195, 2, x_193); lean_ctor_set(x_195, 3, x_188); lean_ctor_set(x_195, 4, x_5); -x_196 = lean_alloc_ctor(5, 1, 0); +x_196 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_196, 0, x_195); x_197 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_198 = lean_alloc_ctor(1, 2, 0); @@ -22716,7 +22826,7 @@ lean_ctor_set(x_228, 1, x_3); lean_ctor_set(x_228, 2, x_226); lean_ctor_set(x_228, 3, x_220); lean_ctor_set(x_228, 4, x_5); -x_229 = lean_alloc_ctor(5, 1, 0); +x_229 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_229, 0, x_228); x_230 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_231 = lean_alloc_ctor(1, 2, 0); @@ -22909,7 +23019,7 @@ lean_ctor_set(x_281, 1, x_3); lean_ctor_set(x_281, 2, x_279); lean_ctor_set(x_281, 3, x_278); lean_ctor_set(x_281, 4, x_5); -x_282 = lean_alloc_ctor(5, 1, 0); +x_282 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_282, 0, x_281); x_283 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; lean_ctor_set_tag(x_274, 1); @@ -23009,7 +23119,7 @@ lean_ctor_set(x_310, 1, x_3); lean_ctor_set(x_310, 2, x_308); lean_ctor_set(x_310, 3, x_307); lean_ctor_set(x_310, 4, x_5); -x_311 = lean_alloc_ctor(5, 1, 0); +x_311 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_311, 0, x_310); x_312 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_313 = lean_alloc_ctor(1, 2, 0); @@ -23140,7 +23250,7 @@ lean_ctor_set(x_354, 1, x_3); lean_ctor_set(x_354, 2, x_352); lean_ctor_set(x_354, 3, x_348); lean_ctor_set(x_354, 4, x_5); -x_355 = lean_alloc_ctor(5, 1, 0); +x_355 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_355, 0, x_354); x_356 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; lean_ctor_set_tag(x_346, 1); @@ -23241,7 +23351,7 @@ lean_ctor_set(x_384, 1, x_3); lean_ctor_set(x_384, 2, x_382); lean_ctor_set(x_384, 3, x_378); lean_ctor_set(x_384, 4, x_5); -x_385 = lean_alloc_ctor(5, 1, 0); +x_385 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_385, 0, x_384); x_386 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_387 = lean_alloc_ctor(1, 2, 0); @@ -23451,7 +23561,7 @@ lean_ctor_set(x_444, 1, x_3); lean_ctor_set(x_444, 2, x_442); lean_ctor_set(x_444, 3, x_436); lean_ctor_set(x_444, 4, x_5); -lean_ctor_set_tag(x_404, 5); +lean_ctor_set_tag(x_404, 6); lean_ctor_set(x_404, 0, x_444); x_445 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_446 = lean_alloc_ctor(1, 2, 0); @@ -23685,7 +23795,7 @@ lean_ctor_set(x_499, 1, x_3); lean_ctor_set(x_499, 2, x_497); lean_ctor_set(x_499, 3, x_491); lean_ctor_set(x_499, 4, x_5); -lean_ctor_set_tag(x_404, 5); +lean_ctor_set_tag(x_404, 6); lean_ctor_set(x_404, 0, x_499); x_500 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_501 = lean_alloc_ctor(1, 2, 0); @@ -23921,7 +24031,7 @@ lean_ctor_set(x_552, 1, x_3); lean_ctor_set(x_552, 2, x_550); lean_ctor_set(x_552, 3, x_544); lean_ctor_set(x_552, 4, x_5); -x_553 = lean_alloc_ctor(5, 1, 0); +x_553 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_553, 0, x_552); x_554 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_555 = lean_alloc_ctor(1, 2, 0); @@ -24141,7 +24251,7 @@ lean_ctor_set(x_603, 1, x_3); lean_ctor_set(x_603, 2, x_601); lean_ctor_set(x_603, 3, x_600); lean_ctor_set(x_603, 4, x_5); -x_604 = lean_alloc_ctor(5, 1, 0); +x_604 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_604, 0, x_603); x_605 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; lean_ctor_set_tag(x_596, 1); @@ -24241,7 +24351,7 @@ lean_ctor_set(x_632, 1, x_3); lean_ctor_set(x_632, 2, x_630); lean_ctor_set(x_632, 3, x_629); lean_ctor_set(x_632, 4, x_5); -x_633 = lean_alloc_ctor(5, 1, 0); +x_633 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_633, 0, x_632); x_634 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_635 = lean_alloc_ctor(1, 2, 0); @@ -24371,7 +24481,7 @@ lean_ctor_set(x_675, 1, x_3); lean_ctor_set(x_675, 2, x_673); lean_ctor_set(x_675, 3, x_669); lean_ctor_set(x_675, 4, x_5); -x_676 = lean_alloc_ctor(5, 1, 0); +x_676 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_676, 0, x_675); x_677 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; lean_ctor_set_tag(x_667, 1); @@ -24472,7 +24582,7 @@ lean_ctor_set(x_705, 1, x_3); lean_ctor_set(x_705, 2, x_703); lean_ctor_set(x_705, 3, x_699); lean_ctor_set(x_705, 4, x_5); -x_706 = lean_alloc_ctor(5, 1, 0); +x_706 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_706, 0, x_705); x_707 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; x_708 = lean_alloc_ctor(1, 2, 0); @@ -24686,10 +24796,10 @@ lean_ctor_set(x_762, 2, x_760); lean_ctor_set(x_762, 3, x_754); lean_ctor_set(x_762, 4, x_5); if (lean_is_scalar(x_731)) { - x_763 = lean_alloc_ctor(5, 1, 0); + x_763 = lean_alloc_ctor(6, 1, 0); } else { x_763 = x_731; - lean_ctor_set_tag(x_763, 5); + lean_ctor_set_tag(x_763, 6); } lean_ctor_set(x_763, 0, x_762); x_764 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; @@ -24917,7 +25027,7 @@ lean_ctor_set(x_815, 1, x_3); lean_ctor_set(x_815, 2, x_813); lean_ctor_set(x_815, 3, x_812); lean_ctor_set(x_815, 4, x_5); -x_816 = lean_alloc_ctor(5, 1, 0); +x_816 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_816, 0, x_815); x_817 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; if (lean_is_scalar(x_811)) { @@ -25058,7 +25168,7 @@ lean_ctor_set(x_859, 1, x_3); lean_ctor_set(x_859, 2, x_857); lean_ctor_set(x_859, 3, x_852); lean_ctor_set(x_859, 4, x_5); -x_860 = lean_alloc_ctor(5, 1, 0); +x_860 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_860, 0, x_859); x_861 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___lambda__1___closed__3; if (lean_is_scalar(x_854)) { @@ -28674,7 +28784,7 @@ static lean_object* _init_l_Lean_Elab_Term_StructInst_DefaultFields_getFieldName lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__2___closed__1; x_2 = l_Lean_Elab_Term_StructInst_DefaultFields_getFieldName___closed__1; -x_3 = lean_unsigned_to_nat(781u); +x_3 = lean_unsigned_to_nat(783u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -32915,7 +33025,7 @@ static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Elab_Term_StructInst_D lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__2___closed__1; x_2 = l_List_forIn_x27_loop___at_Lean_Elab_Term_StructInst_DefaultFields_step___spec__3___closed__1; -x_3 = lean_unsigned_to_nat(917u); +x_3 = lean_unsigned_to_nat(919u); x_4 = lean_unsigned_to_nat(25u); x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_getStructSource___spec__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -38186,7 +38296,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -38196,37 +38306,37 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType__1___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2; x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5() { _start: { lean_object* x_1; @@ -38234,17 +38344,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4; -x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4; +x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7() { _start: { lean_object* x_1; @@ -38252,47 +38362,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6; -x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6; +x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8; x_2 = l_Lean_Elab_Term_StructInst_expandStructInstExpectedType___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType__1___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10; x_2 = l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstExpectedType__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12() { _start: { lean_object* x_1; @@ -38300,33 +38410,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11; -x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12; +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11; +x_2 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13; -x_2 = lean_unsigned_to_nat(13700u); +x_1 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13; +x_2 = lean_unsigned_to_nat(13843u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_List_foldlM___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct___spec__5___closed__3; x_3 = 0; -x_4 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14; +x_4 = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -38489,6 +38599,10 @@ l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1 lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__1); l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__2); +l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__3); +l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4 = _init_l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__3___closed__4); l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__1); l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev__1___closed__2(); @@ -38993,35 +39107,35 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst_dec if (builtin) {res = l___regBuiltin_Lean_Elab_Term_StructInst_elabStructInst_declRange__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__1); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__2); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__3); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__4); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__5); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__6); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__7); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__8); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__9); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__10); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__11); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__12); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__13); -l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700____closed__14); -if (builtin) {res = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13700_(lean_io_mk_world()); +}l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__1); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__2); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__3); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__4); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__5); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__6); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__7); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__8); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__9); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__10); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__11); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__12); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__13); +l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14 = _init_l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843____closed__14); +if (builtin) {res = l_Lean_Elab_Term_StructInst_initFn____x40_Lean_Elab_StructInst___hyg_13843_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index 3a4a7e97d878..efcf1bdfc825 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -16636,7 +16636,7 @@ lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = lean_alloc_ctor(10, 1, 0); +x_30 = lean_alloc_ctor(11, 1, 0); lean_ctor_set(x_30, 0, x_2); x_31 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(x_30, x_14, x_15, x_16, x_17, x_18, x_19, x_29); x_32 = lean_ctor_get(x_31, 1); diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index 744a5deddf5c..856d005c770c 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -43,7 +43,6 @@ uint32_t lean_string_utf8_get(lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__11; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__7___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_Command_isLocalAttrKind(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__4; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__4; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__9___closed__1; @@ -55,6 +54,7 @@ static lean_object* l_panic___at_Lean_Elab_Term_toParserDescr_processAlias___spe LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkNameFromParserSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_withNotFirst(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__39; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___closed__9; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___closed__7; @@ -62,6 +62,7 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__6; lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandOptPrecedence(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getQuotContent(lean_object*); @@ -86,6 +87,7 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__ static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__4; static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__2; uint8_t l_Lean_Exception_isInterrupt(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7; lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___boxed(lean_object*, lean_object*); @@ -100,13 +102,13 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec__ lean_object* l_Lean_TSyntax_getString(lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__71; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__30; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Term_toParserDescr_isValidAtom(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_isAtomLikeSyntax(lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_markAsTrailingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); @@ -141,10 +143,10 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_resolveSyntaxKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445_(lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__13; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -155,11 +157,13 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange__1___closed__5; static lean_object* l_Lean_Elab_Term_elabParserName_x3f___closed__3; lean_object* l_Lean_Parser_getParserAliasInfo(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__19; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; +static lean_object* l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1; LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_Elab_Command_elabDeclareSyntaxCat___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__5; @@ -170,6 +174,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev(lean_object*, lean static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange__1___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__10; static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__5; @@ -230,7 +235,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___c lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabParserName___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__16; @@ -302,6 +306,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_next(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Command_commandElabAttribute; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__2; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1; @@ -309,6 +314,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2 static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__41; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1; lean_object* l_Char_toUpper(uint32_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__58; @@ -360,7 +366,6 @@ lean_object* l_Lean_Name_getPrefix(lean_object*); static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__3; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__16; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__18; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange__1___closed__2; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__3; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__4; @@ -389,6 +394,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_toParserDesc lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__18; @@ -400,7 +406,6 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__10; lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__20; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__7___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_toParserDescr_process___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces___at_Lean_Elab_Command_resolveSyntaxKind___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -430,6 +435,7 @@ static lean_object* l_Lean_Elab_Term_addCategoryInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__13; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___closed__1; lean_object* l_Lean_Parser_ensureConstantParserAlias(lean_object*, lean_object*); @@ -452,10 +458,10 @@ LEAN_EXPORT lean_object* l_String_mapAux___at_Lean_Elab_Command_mkNameFromParser static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_toParserDescr_process___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419_(lean_object*); static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__14; static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__2; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +uint8_t l_String_anyAux___at_addParenHeuristic___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__2(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange__1___closed__7; @@ -492,7 +498,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_resolveSyntaxKind(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__14; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1(lean_object*); @@ -514,7 +519,6 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn uint8_t l_Lean_Parser_isParserCategory(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_Command_resolveSyntaxKind___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange__1(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12; static lean_object* l_Lean_Elab_Term_elabParserName___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__19; @@ -530,6 +534,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1( static lean_object* l_Lean_Elab_Term_checkLeftRec___closed__3; static lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___closed__2; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Term_markAsTrailingParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabParserName___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__5; @@ -580,6 +585,7 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__20; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabSyntax___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); @@ -590,7 +596,6 @@ static lean_object* l_Lean_Elab_Command_inferMacroRulesAltKind___closed__1; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__32; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -626,7 +631,6 @@ static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyn static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange__1___closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__57; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__12; static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__11; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_mkNameFromParserSyntax_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -656,6 +660,7 @@ static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__1; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabSyntax___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14; static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange__1(lean_object*); @@ -684,10 +689,8 @@ extern lean_object* l_Lean_firstFrontendMacroScope; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSeq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax_declRange__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_expandNoKindMacroRulesAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__5; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__3; @@ -718,7 +721,6 @@ lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__1; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__11; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__54; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_checkLeftRec___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -727,7 +729,6 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__2___rarg___closed__2; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3___closed__1; static lean_object* l_Lean_Elab_Command_resolveSyntaxKind___closed__2; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; @@ -747,6 +748,7 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___cl static double l_Lean_addTrace___at_Lean_Elab_Term_checkLeftRec___spec__3___closed__1; static lean_object* l_Lean_Elab_Term_toParserDescr_processParserCategory___lambda__1___closed__6; uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__6; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__1; @@ -772,6 +774,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntax__1(lean_obj lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__4; static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__13; +static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21; static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__10; static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__3___closed__2; @@ -787,7 +790,6 @@ static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__ static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__13; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_checkLeftRec___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__15; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_checkLeftRec___spec__7___closed__6; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__8; @@ -798,6 +800,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_toParserDescr_pro LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__12; lean_object* l_String_toSubstring_x27(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8; extern lean_object* l_Lean_Parser_leadPrec; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -811,6 +814,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_resolveSyntaxK static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___closed__2; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__11(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; @@ -5918,230 +5922,310 @@ lean_dec(x_1); return x_11; } } -LEAN_EXPORT uint8_t l_Lean_Elab_Term_toParserDescr_isValidAtom(lean_object* x_1) { +static lean_object* _init_l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1() { _start: { -lean_object* x_2; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_string_utf8_byte_size(x_1); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_eq(x_13, x_14); -if (x_15 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("''", 2, 2); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Elab_Term_toParserDescr_isValidAtom(lean_object* x_1) { +_start: { -uint32_t x_16; uint32_t x_17; uint8_t x_18; -x_16 = lean_string_utf8_get(x_1, x_14); -x_17 = 39; -x_18 = lean_uint32_dec_eq(x_16, x_17); -if (x_18 == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_13; lean_object* x_23; lean_object* x_64; uint8_t x_65; +x_2 = lean_string_utf8_byte_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Substring_takeWhileAux___at_Substring_trimLeft___spec__1(x_1, x_2, x_3); +x_5 = l_Substring_takeRightWhileAux___at_Substring_trimRight___spec__1(x_1, x_4, x_2); +x_6 = lean_string_utf8_extract(x_1, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +x_64 = lean_string_utf8_byte_size(x_6); +x_65 = lean_nat_dec_eq(x_64, x_3); +lean_dec(x_64); +if (x_65 == 0) { -uint32_t x_19; uint8_t x_20; -x_19 = 34; -x_20 = lean_uint32_dec_eq(x_16, x_19); -if (x_20 == 0) +uint32_t x_66; uint32_t x_67; uint8_t x_68; +x_66 = lean_string_utf8_get(x_6, x_3); +x_67 = 39; +x_68 = lean_uint32_dec_eq(x_66, x_67); +if (x_68 == 0) { -uint32_t x_21; uint8_t x_22; -x_21 = 96; -x_22 = lean_uint32_dec_eq(x_16, x_21); -if (x_22 == 0) +uint32_t x_69; uint8_t x_70; +x_69 = 34; +x_70 = lean_uint32_dec_eq(x_66, x_69); +if (x_70 == 0) { -lean_object* x_23; -lean_dec(x_13); -x_23 = lean_box(0); -x_2 = x_23; -goto block_12; +lean_object* x_71; +x_71 = lean_box(0); +x_23 = x_71; +goto block_63; } else { -lean_object* x_24; uint8_t x_25; -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_dec_eq(x_13, x_24); -lean_dec(x_13); -if (x_25 == 0) -{ -uint32_t x_26; lean_object* x_27; uint32_t x_38; uint8_t x_39; -x_26 = lean_string_utf8_get(x_1, x_24); -x_38 = 65; -x_39 = lean_uint32_dec_le(x_38, x_26); -if (x_39 == 0) +uint8_t x_72; +lean_dec(x_6); +x_72 = 0; +return x_72; +} +} +else { -uint32_t x_40; uint8_t x_41; -x_40 = 97; -x_41 = lean_uint32_dec_le(x_40, x_26); -if (x_41 == 0) +lean_object* x_73; uint8_t x_74; +x_73 = l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1; +x_74 = lean_string_dec_eq(x_6, x_73); +if (x_74 == 0) { -lean_object* x_42; -x_42 = lean_box(0); -x_27 = x_42; -goto block_37; +uint8_t x_75; +lean_dec(x_6); +x_75 = 0; +return x_75; } else { -uint32_t x_43; uint8_t x_44; -x_43 = 122; -x_44 = lean_uint32_dec_le(x_26, x_43); -if (x_44 == 0) +uint32_t x_76; uint8_t x_77; +x_76 = 34; +x_77 = lean_uint32_dec_eq(x_66, x_76); +if (x_77 == 0) { -lean_object* x_45; -x_45 = lean_box(0); -x_27 = x_45; -goto block_37; +lean_object* x_78; +x_78 = lean_box(0); +x_23 = x_78; +goto block_63; } else { -uint8_t x_46; -x_46 = 0; -return x_46; +uint8_t x_79; +lean_dec(x_6); +x_79 = 0; +return x_79; +} } } } else { -uint32_t x_47; uint8_t x_48; -x_47 = 90; -x_48 = lean_uint32_dec_le(x_26, x_47); -if (x_48 == 0) +uint8_t x_80; +lean_dec(x_6); +x_80 = 0; +return x_80; +} +block_12: { -uint32_t x_49; uint8_t x_50; -x_49 = 97; -x_50 = lean_uint32_dec_le(x_49, x_26); -if (x_50 == 0) +lean_object* x_8; uint8_t x_9; +lean_dec(x_7); +x_8 = lean_string_utf8_byte_size(x_6); +x_9 = l_String_anyAux___at_addParenHeuristic___spec__1(x_6, x_8, x_3); +lean_dec(x_8); +lean_dec(x_6); +if (x_9 == 0) { -lean_object* x_51; -x_51 = lean_box(0); -x_27 = x_51; -goto block_37; +uint8_t x_10; +x_10 = 1; +return x_10; } else { -uint32_t x_52; uint8_t x_53; -x_52 = 122; -x_53 = lean_uint32_dec_le(x_26, x_52); -if (x_53 == 0) +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +block_22: { -lean_object* x_54; -x_54 = lean_box(0); -x_27 = x_54; -goto block_37; +uint32_t x_14; uint32_t x_15; uint8_t x_16; +lean_dec(x_13); +x_14 = lean_string_utf8_get(x_6, x_3); +x_15 = 48; +x_16 = lean_uint32_dec_le(x_15, x_14); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_box(0); +x_7 = x_17; +goto block_12; } else { -uint8_t x_55; -x_55 = 0; -return x_55; -} -} +uint32_t x_18; uint8_t x_19; +x_18 = 57; +x_19 = lean_uint32_dec_le(x_14, x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_box(0); +x_7 = x_20; +goto block_12; } else { -uint8_t x_56; -x_56 = 0; -return x_56; +uint8_t x_21; +lean_dec(x_6); +x_21 = 0; +return x_21; } } -block_37: +} +block_63: { -uint32_t x_28; uint8_t x_29; -lean_dec(x_27); -x_28 = 95; -x_29 = lean_uint32_dec_eq(x_26, x_28); -if (x_29 == 0) +uint32_t x_24; uint32_t x_25; uint8_t x_26; +lean_dec(x_23); +x_24 = lean_string_utf8_get(x_6, x_3); +x_25 = 96; +x_26 = lean_uint32_dec_eq(x_24, x_25); +if (x_26 == 0) { -uint8_t x_30; -x_30 = l_Lean_isLetterLike(x_26); +lean_object* x_27; +x_27 = lean_box(0); +x_13 = x_27; +goto block_22; +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_string_utf8_byte_size(x_6); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_dec_eq(x_28, x_29); +lean_dec(x_28); if (x_30 == 0) { -uint32_t x_31; uint8_t x_32; -x_31 = l_Lean_idBeginEscape; -x_32 = lean_uint32_dec_eq(x_26, x_31); -if (x_32 == 0) +uint32_t x_31; lean_object* x_32; uint32_t x_43; uint8_t x_44; +x_31 = lean_string_utf8_get(x_6, x_29); +x_43 = 65; +x_44 = lean_uint32_dec_le(x_43, x_31); +if (x_44 == 0) { -lean_object* x_33; -x_33 = lean_box(0); -x_2 = x_33; -goto block_12; +uint32_t x_45; uint8_t x_46; +x_45 = 97; +x_46 = lean_uint32_dec_le(x_45, x_31); +if (x_46 == 0) +{ +lean_object* x_47; +x_47 = lean_box(0); +x_32 = x_47; +goto block_42; } else { -uint8_t x_34; -x_34 = 0; -return x_34; -} +uint32_t x_48; uint8_t x_49; +x_48 = 122; +x_49 = lean_uint32_dec_le(x_31, x_48); +if (x_49 == 0) +{ +lean_object* x_50; +x_50 = lean_box(0); +x_32 = x_50; +goto block_42; } else { -uint8_t x_35; -x_35 = 0; -return x_35; +uint8_t x_51; +lean_dec(x_6); +x_51 = 0; +return x_51; +} } } else { -uint8_t x_36; -x_36 = 0; -return x_36; -} +uint32_t x_52; uint8_t x_53; +x_52 = 90; +x_53 = lean_uint32_dec_le(x_31, x_52); +if (x_53 == 0) +{ +uint32_t x_54; uint8_t x_55; +x_54 = 97; +x_55 = lean_uint32_dec_le(x_54, x_31); +if (x_55 == 0) +{ +lean_object* x_56; +x_56 = lean_box(0); +x_32 = x_56; +goto block_42; } +else +{ +uint32_t x_57; uint8_t x_58; +x_57 = 122; +x_58 = lean_uint32_dec_le(x_31, x_57); +if (x_58 == 0) +{ +lean_object* x_59; +x_59 = lean_box(0); +x_32 = x_59; +goto block_42; } else { -uint8_t x_57; -x_57 = 0; -return x_57; +uint8_t x_60; +lean_dec(x_6); +x_60 = 0; +return x_60; } } } else { -uint8_t x_58; -lean_dec(x_13); -x_58 = 0; -return x_58; +uint8_t x_61; +lean_dec(x_6); +x_61 = 0; +return x_61; } } +block_42: +{ +uint32_t x_33; uint8_t x_34; +lean_dec(x_32); +x_33 = 95; +x_34 = lean_uint32_dec_eq(x_31, x_33); +if (x_34 == 0) +{ +uint8_t x_35; +x_35 = l_Lean_isLetterLike(x_31); +if (x_35 == 0) +{ +uint32_t x_36; uint8_t x_37; +x_36 = l_Lean_idBeginEscape; +x_37 = lean_uint32_dec_eq(x_31, x_36); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = lean_box(0); +x_13 = x_38; +goto block_22; +} else { -uint8_t x_59; -lean_dec(x_13); -x_59 = 0; -return x_59; +uint8_t x_39; +lean_dec(x_6); +x_39 = 0; +return x_39; } } else { -uint8_t x_60; -lean_dec(x_13); -x_60 = 0; -return x_60; +uint8_t x_40; +lean_dec(x_6); +x_40 = 0; +return x_40; } -block_12: -{ -lean_object* x_3; uint32_t x_4; uint32_t x_5; uint8_t x_6; -lean_dec(x_2); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_string_utf8_get(x_1, x_3); -x_5 = 48; -x_6 = lean_uint32_dec_le(x_5, x_4); -if (x_6 == 0) -{ -uint8_t x_7; -x_7 = 1; -return x_7; } else { -uint32_t x_8; uint8_t x_9; -x_8 = 57; -x_9 = lean_uint32_dec_le(x_4, x_8); -if (x_9 == 0) -{ -uint8_t x_10; -x_10 = 1; -return x_10; +uint8_t x_41; +lean_dec(x_6); +x_41 = 0; +return x_41; +} +} } else { -uint8_t x_11; -x_11 = 0; -return x_11; +uint8_t x_62; +lean_dec(x_6); +x_62 = 0; +return x_62; } } } @@ -15626,7 +15710,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__16( _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -15645,13 +15729,33 @@ return x_5; static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; +x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; +x_3 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__3; +x_4 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19() { +static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21() { _start: { lean_object* x_1; @@ -15659,14 +15763,14 @@ x_1 = lean_mk_string_unchecked("catBehaviorBoth", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20() { +static lean_object* _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__2; x_3 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; -x_4 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +x_4 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -15694,68 +15798,68 @@ x_17 = l_Lean_Name_append(x_16, x_10); x_18 = lean_st_ref_get(x_3, x_4); if (lean_obj_tag(x_7) == 0) { -lean_object* x_564; -x_564 = lean_box(0); -x_19 = x_564; -goto block_563; +lean_object* x_576; +x_576 = lean_box(0); +x_19 = x_576; +goto block_575; } else { -uint8_t x_565; -x_565 = !lean_is_exclusive(x_7); -if (x_565 == 0) +uint8_t x_577; +x_577 = !lean_is_exclusive(x_7); +if (x_577 == 0) { x_19 = x_7; -goto block_563; +goto block_575; } else { -lean_object* x_566; lean_object* x_567; -x_566 = lean_ctor_get(x_7, 0); -lean_inc(x_566); +lean_object* x_578; lean_object* x_579; +x_578 = lean_ctor_get(x_7, 0); +lean_inc(x_578); lean_dec(x_7); -x_567 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_567, 0, x_566); -x_19 = x_567; -goto block_563; +x_579 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_579, 0, x_578); +x_19 = x_579; +goto block_575; } } -block_563: +block_575: { uint8_t x_20; if (x_13 == 0) { -lean_object* x_556; lean_object* x_557; lean_object* x_558; uint8_t x_559; -x_556 = l_Lean_Syntax_getArg(x_12, x_11); +lean_object* x_568; lean_object* x_569; lean_object* x_570; uint8_t x_571; +x_568 = l_Lean_Syntax_getArg(x_12, x_11); lean_dec(x_12); -x_557 = l_Lean_Syntax_getKind(x_556); -x_558 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; -x_559 = lean_name_eq(x_557, x_558); -lean_dec(x_557); -if (x_559 == 0) +x_569 = l_Lean_Syntax_getKind(x_568); +x_570 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22; +x_571 = lean_name_eq(x_569, x_570); +lean_dec(x_569); +if (x_571 == 0) { -uint8_t x_560; -x_560 = 1; -x_20 = x_560; -goto block_555; +uint8_t x_572; +x_572 = 1; +x_20 = x_572; +goto block_567; } else { -uint8_t x_561; -x_561 = 2; -x_20 = x_561; -goto block_555; +uint8_t x_573; +x_573 = 2; +x_20 = x_573; +goto block_567; } } else { -uint8_t x_562; +uint8_t x_574; lean_dec(x_12); -x_562 = 0; -x_20 = x_562; -goto block_555; +x_574 = 0; +x_20 = x_574; +goto block_567; } -block_555: +block_567: { uint8_t x_21; x_21 = !lean_is_exclusive(x_18); @@ -15807,7 +15911,7 @@ x_42 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_41); x_43 = !lean_is_exclusive(x_42); if (x_43 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; x_44 = lean_ctor_get(x_42, 0); x_45 = lean_ctor_get(x_42, 1); x_46 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; @@ -15885,1237 +15989,1228 @@ x_75 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; lean_inc(x_48); lean_inc(x_37); x_76 = l_Lean_Syntax_node1(x_37, x_75, x_48); -x_77 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_48, 3); +x_77 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_48); lean_inc(x_37); -x_78 = l_Lean_Syntax_node6(x_37, x_77, x_18, x_48, x_48, x_76, x_48, x_72); -x_79 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_78 = l_Lean_Syntax_node1(x_37, x_77, x_48); +x_79 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_48, 2); +lean_inc(x_37); +x_80 = l_Lean_Syntax_node6(x_37, x_79, x_18, x_48, x_76, x_78, x_48, x_72); +x_81 = l_Lean_Elab_Term_toParserDescr_process___closed__2; lean_inc(x_37); -x_80 = l_Lean_Syntax_node2(x_37, x_79, x_74, x_78); -x_81 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +x_82 = l_Lean_Syntax_node2(x_37, x_81, x_74, x_80); +x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; lean_inc_n(x_48, 2); lean_inc(x_37); -x_82 = l_Lean_Syntax_node2(x_37, x_81, x_48, x_48); -x_83 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +x_84 = l_Lean_Syntax_node2(x_37, x_83, x_48, x_48); +x_85 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; lean_inc(x_48); lean_inc(x_37); -x_84 = l_Lean_Syntax_node4(x_37, x_83, x_28, x_80, x_82, x_48); -x_85 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +x_86 = l_Lean_Syntax_node4(x_37, x_85, x_28, x_82, x_84, x_48); +x_87 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; lean_inc(x_48); lean_inc(x_37); -x_86 = l_Lean_Syntax_node5(x_37, x_85, x_42, x_58, x_68, x_84, x_48); +x_88 = l_Lean_Syntax_node5(x_37, x_87, x_42, x_58, x_68, x_86, x_48); if (lean_obj_tag(x_19) == 0) { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_87 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_89 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; lean_inc(x_37); -x_88 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_88, 0, x_37); -lean_ctor_set(x_88, 1, x_46); -lean_ctor_set(x_88, 2, x_87); -x_89 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_90 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_90, 0, x_37); +lean_ctor_set(x_90, 1, x_46); +lean_ctor_set(x_90, 2, x_89); +x_91 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc_n(x_48, 4); lean_inc(x_37); -x_90 = l_Lean_Syntax_node6(x_37, x_89, x_88, x_48, x_48, x_48, x_48, x_48); -x_91 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_92 = l_Lean_Syntax_node2(x_37, x_91, x_90, x_86); +x_92 = l_Lean_Syntax_node6(x_37, x_91, x_90, x_48, x_48, x_48, x_48, x_48); +x_93 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_94 = l_Lean_Syntax_node2(x_37, x_93, x_92, x_88); lean_inc(x_3); lean_inc(x_2); -x_93 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_45); -if (lean_obj_tag(x_93) == 0) +x_95 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_45); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_94; lean_object* x_95; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec(x_93); -x_95 = l_Lean_Elab_Command_elabCommand(x_92, x_2, x_3, x_94); -return x_95; +lean_object* x_96; lean_object* x_97; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_97 = l_Lean_Elab_Command_elabCommand(x_94, x_2, x_3, x_96); +return x_97; } else { -uint8_t x_96; -lean_dec(x_92); +uint8_t x_98; +lean_dec(x_94); lean_dec(x_3); lean_dec(x_2); -x_96 = !lean_is_exclusive(x_93); -if (x_96 == 0) +x_98 = !lean_is_exclusive(x_95); +if (x_98 == 0) { -return x_93; +return x_95; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_93, 0); -x_98 = lean_ctor_get(x_93, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_93); -x_99 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -return x_99; +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_95, 0); +x_100 = lean_ctor_get(x_95, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_95); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; } } } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_100 = lean_ctor_get(x_19, 0); -lean_inc(x_100); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_102 = lean_ctor_get(x_19, 0); +lean_inc(x_102); lean_dec(x_19); -x_101 = l_Array_mkArray1___rarg(x_100); -x_102 = l_Array_append___rarg(x_47, x_101); -lean_dec(x_101); +x_103 = l_Array_mkArray1___rarg(x_102); +x_104 = l_Array_append___rarg(x_47, x_103); +lean_dec(x_103); lean_inc(x_37); -x_103 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_103, 0, x_37); -lean_ctor_set(x_103, 1, x_46); -lean_ctor_set(x_103, 2, x_102); -x_104 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +x_105 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_105, 0, x_37); +lean_ctor_set(x_105, 1, x_46); +lean_ctor_set(x_105, 2, x_104); +x_106 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; lean_inc_n(x_48, 4); lean_inc(x_37); -x_105 = l_Lean_Syntax_node6(x_37, x_104, x_103, x_48, x_48, x_48, x_48, x_48); -x_106 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_107 = l_Lean_Syntax_node2(x_37, x_106, x_105, x_86); +x_107 = l_Lean_Syntax_node6(x_37, x_106, x_105, x_48, x_48, x_48, x_48, x_48); +x_108 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_109 = l_Lean_Syntax_node2(x_37, x_108, x_107, x_88); lean_inc(x_3); lean_inc(x_2); -x_108 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_45); -if (lean_obj_tag(x_108) == 0) +x_110 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_45); +if (lean_obj_tag(x_110) == 0) { -lean_object* x_109; lean_object* x_110; -x_109 = lean_ctor_get(x_108, 1); -lean_inc(x_109); -lean_dec(x_108); -x_110 = l_Lean_Elab_Command_elabCommand(x_107, x_2, x_3, x_109); -return x_110; +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +lean_dec(x_110); +x_112 = l_Lean_Elab_Command_elabCommand(x_109, x_2, x_3, x_111); +return x_112; } else { -uint8_t x_111; -lean_dec(x_107); +uint8_t x_113; +lean_dec(x_109); lean_dec(x_3); lean_dec(x_2); -x_111 = !lean_is_exclusive(x_108); -if (x_111 == 0) +x_113 = !lean_is_exclusive(x_110); +if (x_113 == 0) { -return x_108; +return x_110; } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_112 = lean_ctor_get(x_108, 0); -x_113 = lean_ctor_get(x_108, 1); -lean_inc(x_113); -lean_inc(x_112); -lean_dec(x_108); -x_114 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_114, 0, x_112); -lean_ctor_set(x_114, 1, x_113); -return x_114; +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_110, 0); +x_115 = lean_ctor_get(x_110, 1); +lean_inc(x_115); +lean_inc(x_114); +lean_dec(x_110); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +return x_116; } } } } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_115 = lean_ctor_get(x_42, 0); -x_116 = lean_ctor_get(x_42, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_42); -x_117 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; -x_118 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; -lean_inc(x_37); -x_119 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_119, 0, x_37); -lean_ctor_set(x_119, 1, x_117); -lean_ctor_set(x_119, 2, x_118); -x_120 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_117 = lean_ctor_get(x_42, 0); +x_118 = lean_ctor_get(x_42, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_42); +x_119 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_120 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; lean_inc(x_37); -x_121 = lean_alloc_ctor(2, 2, 0); +x_121 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_121, 0, x_37); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; -x_123 = l_Lean_Name_append(x_122, x_17); -x_124 = 1; -x_125 = l_Lean_mkIdentFrom(x_9, x_123, x_124); +lean_ctor_set(x_121, 1, x_119); +lean_ctor_set(x_121, 2, x_120); +x_122 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +lean_inc(x_37); +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_37); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; +x_125 = l_Lean_Name_append(x_124, x_17); +x_126 = 1; +x_127 = l_Lean_mkIdentFrom(x_9, x_125, x_126); lean_dec(x_9); -x_126 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_128 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; lean_ctor_set_tag(x_38, 1); -lean_ctor_set(x_38, 1, x_126); -lean_ctor_set(x_38, 0, x_125); -x_127 = lean_array_mk(x_38); -x_128 = lean_box(2); -x_129 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; -x_130 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_130, 2, x_127); -x_131 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +lean_ctor_set(x_38, 1, x_128); +lean_ctor_set(x_38, 0, x_127); +x_129 = lean_array_mk(x_38); +x_130 = lean_box(2); +x_131 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_132 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_132, 2, x_129); +x_133 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; lean_inc(x_37); lean_ctor_set_tag(x_32, 2); -lean_ctor_set(x_32, 1, x_131); +lean_ctor_set(x_32, 1, x_133); lean_ctor_set(x_32, 0, x_37); -x_132 = l_Lean_addMacroScope(x_115, x_16, x_40); -x_133 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; -x_134 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; +x_134 = l_Lean_addMacroScope(x_117, x_16, x_40); +x_135 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; +x_136 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; lean_inc(x_37); -x_135 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_135, 0, x_37); -lean_ctor_set(x_135, 1, x_133); -lean_ctor_set(x_135, 2, x_132); -lean_ctor_set(x_135, 3, x_134); -x_136 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +x_137 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_137, 0, x_37); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_134); +lean_ctor_set(x_137, 3, x_136); +x_138 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_37); -x_137 = l_Lean_Syntax_node2(x_37, x_136, x_32, x_135); +x_139 = l_Lean_Syntax_node2(x_37, x_138, x_32, x_137); lean_inc(x_37); -x_138 = l_Lean_Syntax_node1(x_37, x_117, x_137); -x_139 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -lean_inc(x_119); +x_140 = l_Lean_Syntax_node1(x_37, x_119, x_139); +x_141 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_121); lean_inc(x_37); -x_140 = l_Lean_Syntax_node2(x_37, x_139, x_119, x_138); -x_141 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +x_142 = l_Lean_Syntax_node2(x_37, x_141, x_121, x_140); +x_143 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; lean_inc(x_37); lean_ctor_set_tag(x_28, 2); -lean_ctor_set(x_28, 1, x_141); +lean_ctor_set(x_28, 1, x_143); lean_ctor_set(x_28, 0, x_37); -x_142 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; +x_144 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; lean_inc(x_37); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_142); +lean_ctor_set(x_18, 1, x_144); lean_ctor_set(x_18, 0, x_37); -x_143 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; +x_145 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; lean_inc(x_37); -x_144 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_144, 0, x_37); -lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; -lean_inc(x_144); +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_37); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +lean_inc(x_146); lean_inc(x_18); lean_inc(x_37); -x_146 = l_Lean_Syntax_node2(x_37, x_145, x_18, x_144); -x_147 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; -lean_inc(x_119); +x_148 = l_Lean_Syntax_node2(x_37, x_147, x_18, x_146); +x_149 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; +lean_inc(x_121); lean_inc(x_37); -x_148 = l_Lean_Syntax_node1(x_37, x_147, x_119); -x_149 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_119, 3); +x_150 = l_Lean_Syntax_node1(x_37, x_149, x_121); +x_151 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_121); lean_inc(x_37); -x_150 = l_Lean_Syntax_node6(x_37, x_149, x_18, x_119, x_119, x_148, x_119, x_144); -x_151 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_152 = l_Lean_Syntax_node1(x_37, x_151, x_121); +x_153 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_121, 2); lean_inc(x_37); -x_152 = l_Lean_Syntax_node2(x_37, x_151, x_146, x_150); -x_153 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; -lean_inc_n(x_119, 2); +x_154 = l_Lean_Syntax_node6(x_37, x_153, x_18, x_121, x_150, x_152, x_121, x_146); +x_155 = l_Lean_Elab_Term_toParserDescr_process___closed__2; lean_inc(x_37); -x_154 = l_Lean_Syntax_node2(x_37, x_153, x_119, x_119); -x_155 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -lean_inc(x_119); +x_156 = l_Lean_Syntax_node2(x_37, x_155, x_148, x_154); +x_157 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +lean_inc_n(x_121, 2); lean_inc(x_37); -x_156 = l_Lean_Syntax_node4(x_37, x_155, x_28, x_152, x_154, x_119); -x_157 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; -lean_inc(x_119); +x_158 = l_Lean_Syntax_node2(x_37, x_157, x_121, x_121); +x_159 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +lean_inc(x_121); +lean_inc(x_37); +x_160 = l_Lean_Syntax_node4(x_37, x_159, x_28, x_156, x_158, x_121); +x_161 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +lean_inc(x_121); lean_inc(x_37); -x_158 = l_Lean_Syntax_node5(x_37, x_157, x_121, x_130, x_140, x_156, x_119); +x_162 = l_Lean_Syntax_node5(x_37, x_161, x_123, x_132, x_142, x_160, x_121); if (lean_obj_tag(x_19) == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_159 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_163 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; lean_inc(x_37); -x_160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_160, 0, x_37); -lean_ctor_set(x_160, 1, x_117); -lean_ctor_set(x_160, 2, x_159); -x_161 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_119, 4); +x_164 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_164, 0, x_37); +lean_ctor_set(x_164, 1, x_119); +lean_ctor_set(x_164, 2, x_163); +x_165 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_121, 4); lean_inc(x_37); -x_162 = l_Lean_Syntax_node6(x_37, x_161, x_160, x_119, x_119, x_119, x_119, x_119); -x_163 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_164 = l_Lean_Syntax_node2(x_37, x_163, x_162, x_158); +x_166 = l_Lean_Syntax_node6(x_37, x_165, x_164, x_121, x_121, x_121, x_121, x_121); +x_167 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_168 = l_Lean_Syntax_node2(x_37, x_167, x_166, x_162); lean_inc(x_3); lean_inc(x_2); -x_165 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_116); -if (lean_obj_tag(x_165) == 0) +x_169 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_118); +if (lean_obj_tag(x_169) == 0) { -lean_object* x_166; lean_object* x_167; -x_166 = lean_ctor_get(x_165, 1); -lean_inc(x_166); -lean_dec(x_165); -x_167 = l_Lean_Elab_Command_elabCommand(x_164, x_2, x_3, x_166); -return x_167; +lean_object* x_170; lean_object* x_171; +x_170 = lean_ctor_get(x_169, 1); +lean_inc(x_170); +lean_dec(x_169); +x_171 = l_Lean_Elab_Command_elabCommand(x_168, x_2, x_3, x_170); +return x_171; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -lean_dec(x_164); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_168); lean_dec(x_3); lean_dec(x_2); -x_168 = lean_ctor_get(x_165, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_165, 1); -lean_inc(x_169); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - lean_ctor_release(x_165, 1); - x_170 = x_165; +x_172 = lean_ctor_get(x_169, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_169, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_174 = x_169; } else { - lean_dec_ref(x_165); - x_170 = lean_box(0); + lean_dec_ref(x_169); + x_174 = lean_box(0); } -if (lean_is_scalar(x_170)) { - x_171 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_174)) { + x_175 = lean_alloc_ctor(1, 2, 0); } else { - x_171 = x_170; + x_175 = x_174; } -lean_ctor_set(x_171, 0, x_168); -lean_ctor_set(x_171, 1, x_169); -return x_171; +lean_ctor_set(x_175, 0, x_172); +lean_ctor_set(x_175, 1, x_173); +return x_175; } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_172 = lean_ctor_get(x_19, 0); -lean_inc(x_172); +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_176 = lean_ctor_get(x_19, 0); +lean_inc(x_176); lean_dec(x_19); -x_173 = l_Array_mkArray1___rarg(x_172); -x_174 = l_Array_append___rarg(x_118, x_173); -lean_dec(x_173); +x_177 = l_Array_mkArray1___rarg(x_176); +x_178 = l_Array_append___rarg(x_120, x_177); +lean_dec(x_177); lean_inc(x_37); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_37); -lean_ctor_set(x_175, 1, x_117); -lean_ctor_set(x_175, 2, x_174); -x_176 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_119, 4); +x_179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_179, 0, x_37); +lean_ctor_set(x_179, 1, x_119); +lean_ctor_set(x_179, 2, x_178); +x_180 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_121, 4); lean_inc(x_37); -x_177 = l_Lean_Syntax_node6(x_37, x_176, x_175, x_119, x_119, x_119, x_119, x_119); -x_178 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_179 = l_Lean_Syntax_node2(x_37, x_178, x_177, x_158); +x_181 = l_Lean_Syntax_node6(x_37, x_180, x_179, x_121, x_121, x_121, x_121, x_121); +x_182 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_183 = l_Lean_Syntax_node2(x_37, x_182, x_181, x_162); lean_inc(x_3); lean_inc(x_2); -x_180 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_116); -if (lean_obj_tag(x_180) == 0) +x_184 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_118); +if (lean_obj_tag(x_184) == 0) { -lean_object* x_181; lean_object* x_182; -x_181 = lean_ctor_get(x_180, 1); -lean_inc(x_181); -lean_dec(x_180); -x_182 = l_Lean_Elab_Command_elabCommand(x_179, x_2, x_3, x_181); -return x_182; +lean_object* x_185; lean_object* x_186; +x_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +lean_dec(x_184); +x_186 = l_Lean_Elab_Command_elabCommand(x_183, x_2, x_3, x_185); +return x_186; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_179); +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_183); lean_dec(x_3); lean_dec(x_2); -x_183 = lean_ctor_get(x_180, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_180, 1); -lean_inc(x_184); -if (lean_is_exclusive(x_180)) { - lean_ctor_release(x_180, 0); - lean_ctor_release(x_180, 1); - x_185 = x_180; +x_187 = lean_ctor_get(x_184, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_184, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_189 = x_184; } else { - lean_dec_ref(x_180); - x_185 = lean_box(0); + lean_dec_ref(x_184); + x_189 = lean_box(0); } -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_189)) { + x_190 = lean_alloc_ctor(1, 2, 0); } else { - x_186 = x_185; + x_190 = x_189; } -lean_ctor_set(x_186, 0, x_183); -lean_ctor_set(x_186, 1, x_184); -return x_186; +lean_ctor_set(x_190, 0, x_187); +lean_ctor_set(x_190, 1, x_188); +return x_190; } } } } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_187 = lean_ctor_get(x_38, 0); -x_188 = lean_ctor_get(x_38, 1); -lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_38); -x_189 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_188); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_191 = lean_ctor_get(x_38, 0); +x_192 = lean_ctor_get(x_38, 1); +lean_inc(x_192); lean_inc(x_191); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_192 = x_189; +lean_dec(x_38); +x_193 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_192); +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_193, 1); +lean_inc(x_195); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_196 = x_193; } else { - lean_dec_ref(x_189); - x_192 = lean_box(0); + lean_dec_ref(x_193); + x_196 = lean_box(0); } -x_193 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; -x_194 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; +x_197 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_198 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; lean_inc(x_37); -x_195 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_195, 0, x_37); -lean_ctor_set(x_195, 1, x_193); -lean_ctor_set(x_195, 2, x_194); -x_196 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_37); +lean_ctor_set(x_199, 1, x_197); +lean_ctor_set(x_199, 2, x_198); +x_200 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; lean_inc(x_37); -if (lean_is_scalar(x_192)) { - x_197 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_196)) { + x_201 = lean_alloc_ctor(2, 2, 0); } else { - x_197 = x_192; - lean_ctor_set_tag(x_197, 2); -} -lean_ctor_set(x_197, 0, x_37); -lean_ctor_set(x_197, 1, x_196); -x_198 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; -x_199 = l_Lean_Name_append(x_198, x_17); -x_200 = 1; -x_201 = l_Lean_mkIdentFrom(x_9, x_199, x_200); + x_201 = x_196; + lean_ctor_set_tag(x_201, 2); +} +lean_ctor_set(x_201, 0, x_37); +lean_ctor_set(x_201, 1, x_200); +x_202 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; +x_203 = l_Lean_Name_append(x_202, x_17); +x_204 = 1; +x_205 = l_Lean_mkIdentFrom(x_9, x_203, x_204); lean_dec(x_9); -x_202 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -x_203 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_203, 0, x_201); -lean_ctor_set(x_203, 1, x_202); -x_204 = lean_array_mk(x_203); -x_205 = lean_box(2); -x_206 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; -x_207 = lean_alloc_ctor(1, 3, 0); +x_206 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +x_207 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_207, 0, x_205); lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_204); -x_208 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +x_208 = lean_array_mk(x_207); +x_209 = lean_box(2); +x_210 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_211 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +lean_ctor_set(x_211, 2, x_208); +x_212 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; lean_inc(x_37); lean_ctor_set_tag(x_32, 2); -lean_ctor_set(x_32, 1, x_208); +lean_ctor_set(x_32, 1, x_212); lean_ctor_set(x_32, 0, x_37); -x_209 = l_Lean_addMacroScope(x_190, x_16, x_187); -x_210 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; -x_211 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; +x_213 = l_Lean_addMacroScope(x_194, x_16, x_191); +x_214 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; +x_215 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; lean_inc(x_37); -x_212 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_212, 0, x_37); -lean_ctor_set(x_212, 1, x_210); -lean_ctor_set(x_212, 2, x_209); -lean_ctor_set(x_212, 3, x_211); -x_213 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +x_216 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_216, 0, x_37); +lean_ctor_set(x_216, 1, x_214); +lean_ctor_set(x_216, 2, x_213); +lean_ctor_set(x_216, 3, x_215); +x_217 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; lean_inc(x_37); -x_214 = l_Lean_Syntax_node2(x_37, x_213, x_32, x_212); +x_218 = l_Lean_Syntax_node2(x_37, x_217, x_32, x_216); lean_inc(x_37); -x_215 = l_Lean_Syntax_node1(x_37, x_193, x_214); -x_216 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -lean_inc(x_195); +x_219 = l_Lean_Syntax_node1(x_37, x_197, x_218); +x_220 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_199); lean_inc(x_37); -x_217 = l_Lean_Syntax_node2(x_37, x_216, x_195, x_215); -x_218 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +x_221 = l_Lean_Syntax_node2(x_37, x_220, x_199, x_219); +x_222 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; lean_inc(x_37); lean_ctor_set_tag(x_28, 2); -lean_ctor_set(x_28, 1, x_218); +lean_ctor_set(x_28, 1, x_222); lean_ctor_set(x_28, 0, x_37); -x_219 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; +x_223 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; lean_inc(x_37); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_219); +lean_ctor_set(x_18, 1, x_223); lean_ctor_set(x_18, 0, x_37); -x_220 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; +x_224 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; lean_inc(x_37); -x_221 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_221, 0, x_37); -lean_ctor_set(x_221, 1, x_220); -x_222 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; -lean_inc(x_221); +x_225 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_225, 0, x_37); +lean_ctor_set(x_225, 1, x_224); +x_226 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +lean_inc(x_225); lean_inc(x_18); lean_inc(x_37); -x_223 = l_Lean_Syntax_node2(x_37, x_222, x_18, x_221); -x_224 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; -lean_inc(x_195); +x_227 = l_Lean_Syntax_node2(x_37, x_226, x_18, x_225); +x_228 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; +lean_inc(x_199); lean_inc(x_37); -x_225 = l_Lean_Syntax_node1(x_37, x_224, x_195); -x_226 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_195, 3); +x_229 = l_Lean_Syntax_node1(x_37, x_228, x_199); +x_230 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_199); lean_inc(x_37); -x_227 = l_Lean_Syntax_node6(x_37, x_226, x_18, x_195, x_195, x_225, x_195, x_221); -x_228 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +x_231 = l_Lean_Syntax_node1(x_37, x_230, x_199); +x_232 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_199, 2); lean_inc(x_37); -x_229 = l_Lean_Syntax_node2(x_37, x_228, x_223, x_227); -x_230 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; -lean_inc_n(x_195, 2); +x_233 = l_Lean_Syntax_node6(x_37, x_232, x_18, x_199, x_229, x_231, x_199, x_225); +x_234 = l_Lean_Elab_Term_toParserDescr_process___closed__2; lean_inc(x_37); -x_231 = l_Lean_Syntax_node2(x_37, x_230, x_195, x_195); -x_232 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -lean_inc(x_195); +x_235 = l_Lean_Syntax_node2(x_37, x_234, x_227, x_233); +x_236 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +lean_inc_n(x_199, 2); lean_inc(x_37); -x_233 = l_Lean_Syntax_node4(x_37, x_232, x_28, x_229, x_231, x_195); -x_234 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; -lean_inc(x_195); +x_237 = l_Lean_Syntax_node2(x_37, x_236, x_199, x_199); +x_238 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +lean_inc(x_199); +lean_inc(x_37); +x_239 = l_Lean_Syntax_node4(x_37, x_238, x_28, x_235, x_237, x_199); +x_240 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +lean_inc(x_199); lean_inc(x_37); -x_235 = l_Lean_Syntax_node5(x_37, x_234, x_197, x_207, x_217, x_233, x_195); +x_241 = l_Lean_Syntax_node5(x_37, x_240, x_201, x_211, x_221, x_239, x_199); if (lean_obj_tag(x_19) == 0) { -lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; -x_236 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_242 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; lean_inc(x_37); -x_237 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_237, 0, x_37); -lean_ctor_set(x_237, 1, x_193); -lean_ctor_set(x_237, 2, x_236); -x_238 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_195, 4); +x_243 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_243, 0, x_37); +lean_ctor_set(x_243, 1, x_197); +lean_ctor_set(x_243, 2, x_242); +x_244 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_199, 4); lean_inc(x_37); -x_239 = l_Lean_Syntax_node6(x_37, x_238, x_237, x_195, x_195, x_195, x_195, x_195); -x_240 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_241 = l_Lean_Syntax_node2(x_37, x_240, x_239, x_235); +x_245 = l_Lean_Syntax_node6(x_37, x_244, x_243, x_199, x_199, x_199, x_199, x_199); +x_246 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_247 = l_Lean_Syntax_node2(x_37, x_246, x_245, x_241); lean_inc(x_3); lean_inc(x_2); -x_242 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_191); -if (lean_obj_tag(x_242) == 0) +x_248 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_195); +if (lean_obj_tag(x_248) == 0) { -lean_object* x_243; lean_object* x_244; -x_243 = lean_ctor_get(x_242, 1); -lean_inc(x_243); -lean_dec(x_242); -x_244 = l_Lean_Elab_Command_elabCommand(x_241, x_2, x_3, x_243); -return x_244; +lean_object* x_249; lean_object* x_250; +x_249 = lean_ctor_get(x_248, 1); +lean_inc(x_249); +lean_dec(x_248); +x_250 = l_Lean_Elab_Command_elabCommand(x_247, x_2, x_3, x_249); +return x_250; } else { -lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; -lean_dec(x_241); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +lean_dec(x_247); lean_dec(x_3); lean_dec(x_2); -x_245 = lean_ctor_get(x_242, 0); -lean_inc(x_245); -x_246 = lean_ctor_get(x_242, 1); -lean_inc(x_246); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - lean_ctor_release(x_242, 1); - x_247 = x_242; +x_251 = lean_ctor_get(x_248, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_248, 1); +lean_inc(x_252); +if (lean_is_exclusive(x_248)) { + lean_ctor_release(x_248, 0); + lean_ctor_release(x_248, 1); + x_253 = x_248; } else { - lean_dec_ref(x_242); - x_247 = lean_box(0); + lean_dec_ref(x_248); + x_253 = lean_box(0); } -if (lean_is_scalar(x_247)) { - x_248 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_253)) { + x_254 = lean_alloc_ctor(1, 2, 0); } else { - x_248 = x_247; + x_254 = x_253; } -lean_ctor_set(x_248, 0, x_245); -lean_ctor_set(x_248, 1, x_246); -return x_248; +lean_ctor_set(x_254, 0, x_251); +lean_ctor_set(x_254, 1, x_252); +return x_254; } } else { -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; -x_249 = lean_ctor_get(x_19, 0); -lean_inc(x_249); +lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_255 = lean_ctor_get(x_19, 0); +lean_inc(x_255); lean_dec(x_19); -x_250 = l_Array_mkArray1___rarg(x_249); -x_251 = l_Array_append___rarg(x_194, x_250); -lean_dec(x_250); +x_256 = l_Array_mkArray1___rarg(x_255); +x_257 = l_Array_append___rarg(x_198, x_256); +lean_dec(x_256); lean_inc(x_37); -x_252 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_252, 0, x_37); -lean_ctor_set(x_252, 1, x_193); -lean_ctor_set(x_252, 2, x_251); -x_253 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_195, 4); +x_258 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_258, 0, x_37); +lean_ctor_set(x_258, 1, x_197); +lean_ctor_set(x_258, 2, x_257); +x_259 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_199, 4); lean_inc(x_37); -x_254 = l_Lean_Syntax_node6(x_37, x_253, x_252, x_195, x_195, x_195, x_195, x_195); -x_255 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_256 = l_Lean_Syntax_node2(x_37, x_255, x_254, x_235); +x_260 = l_Lean_Syntax_node6(x_37, x_259, x_258, x_199, x_199, x_199, x_199, x_199); +x_261 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_262 = l_Lean_Syntax_node2(x_37, x_261, x_260, x_241); lean_inc(x_3); lean_inc(x_2); -x_257 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_191); -if (lean_obj_tag(x_257) == 0) +x_263 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_195); +if (lean_obj_tag(x_263) == 0) { -lean_object* x_258; lean_object* x_259; -x_258 = lean_ctor_get(x_257, 1); -lean_inc(x_258); -lean_dec(x_257); -x_259 = l_Lean_Elab_Command_elabCommand(x_256, x_2, x_3, x_258); -return x_259; +lean_object* x_264; lean_object* x_265; +x_264 = lean_ctor_get(x_263, 1); +lean_inc(x_264); +lean_dec(x_263); +x_265 = l_Lean_Elab_Command_elabCommand(x_262, x_2, x_3, x_264); +return x_265; } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; -lean_dec(x_256); +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_262); lean_dec(x_3); lean_dec(x_2); -x_260 = lean_ctor_get(x_257, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_257, 1); -lean_inc(x_261); -if (lean_is_exclusive(x_257)) { - lean_ctor_release(x_257, 0); - lean_ctor_release(x_257, 1); - x_262 = x_257; +x_266 = lean_ctor_get(x_263, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_263, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_268 = x_263; } else { - lean_dec_ref(x_257); - x_262 = lean_box(0); + lean_dec_ref(x_263); + x_268 = lean_box(0); } -if (lean_is_scalar(x_262)) { - x_263 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(1, 2, 0); } else { - x_263 = x_262; + x_269 = x_268; } -lean_ctor_set(x_263, 0, x_260); -lean_ctor_set(x_263, 1, x_261); -return x_263; +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set(x_269, 1, x_267); +return x_269; } } } } else { -lean_object* x_264; lean_object* x_265; uint8_t x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; -x_264 = lean_ctor_get(x_32, 0); -x_265 = lean_ctor_get(x_32, 1); -lean_inc(x_265); -lean_inc(x_264); -lean_dec(x_32); -x_266 = 0; -x_267 = l_Lean_SourceInfo_fromRef(x_264, x_266); -lean_dec(x_264); -x_268 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_265); -x_269 = lean_ctor_get(x_268, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_268, 1); +lean_object* x_270; lean_object* x_271; uint8_t x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; +x_270 = lean_ctor_get(x_32, 0); +x_271 = lean_ctor_get(x_32, 1); +lean_inc(x_271); lean_inc(x_270); -if (lean_is_exclusive(x_268)) { - lean_ctor_release(x_268, 0); - lean_ctor_release(x_268, 1); - x_271 = x_268; +lean_dec(x_32); +x_272 = 0; +x_273 = l_Lean_SourceInfo_fromRef(x_270, x_272); +lean_dec(x_270); +x_274 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_271); +x_275 = lean_ctor_get(x_274, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_274, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_274)) { + lean_ctor_release(x_274, 0); + lean_ctor_release(x_274, 1); + x_277 = x_274; } else { - lean_dec_ref(x_268); - x_271 = lean_box(0); + lean_dec_ref(x_274); + x_277 = lean_box(0); } -x_272 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_270); -x_273 = lean_ctor_get(x_272, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_272, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_272)) { - lean_ctor_release(x_272, 0); - lean_ctor_release(x_272, 1); - x_275 = x_272; +x_278 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_276); +x_279 = lean_ctor_get(x_278, 0); +lean_inc(x_279); +x_280 = lean_ctor_get(x_278, 1); +lean_inc(x_280); +if (lean_is_exclusive(x_278)) { + lean_ctor_release(x_278, 0); + lean_ctor_release(x_278, 1); + x_281 = x_278; } else { - lean_dec_ref(x_272); - x_275 = lean_box(0); + lean_dec_ref(x_278); + x_281 = lean_box(0); } -x_276 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; -x_277 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; -lean_inc(x_267); -x_278 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_278, 0, x_267); -lean_ctor_set(x_278, 1, x_276); -lean_ctor_set(x_278, 2, x_277); -x_279 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; -lean_inc(x_267); -if (lean_is_scalar(x_275)) { - x_280 = lean_alloc_ctor(2, 2, 0); +x_282 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_283 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; +lean_inc(x_273); +x_284 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_284, 0, x_273); +lean_ctor_set(x_284, 1, x_282); +lean_ctor_set(x_284, 2, x_283); +x_285 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +lean_inc(x_273); +if (lean_is_scalar(x_281)) { + x_286 = lean_alloc_ctor(2, 2, 0); } else { - x_280 = x_275; - lean_ctor_set_tag(x_280, 2); -} -lean_ctor_set(x_280, 0, x_267); -lean_ctor_set(x_280, 1, x_279); -x_281 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; -x_282 = l_Lean_Name_append(x_281, x_17); -x_283 = 1; -x_284 = l_Lean_mkIdentFrom(x_9, x_282, x_283); + x_286 = x_281; + lean_ctor_set_tag(x_286, 2); +} +lean_ctor_set(x_286, 0, x_273); +lean_ctor_set(x_286, 1, x_285); +x_287 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; +x_288 = l_Lean_Name_append(x_287, x_17); +x_289 = 1; +x_290 = l_Lean_mkIdentFrom(x_9, x_288, x_289); lean_dec(x_9); -x_285 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -if (lean_is_scalar(x_271)) { - x_286 = lean_alloc_ctor(1, 2, 0); +x_291 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +if (lean_is_scalar(x_277)) { + x_292 = lean_alloc_ctor(1, 2, 0); } else { - x_286 = x_271; - lean_ctor_set_tag(x_286, 1); + x_292 = x_277; + lean_ctor_set_tag(x_292, 1); } -lean_ctor_set(x_286, 0, x_284); -lean_ctor_set(x_286, 1, x_285); -x_287 = lean_array_mk(x_286); -x_288 = lean_box(2); -x_289 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_288); -lean_ctor_set(x_290, 1, x_289); -lean_ctor_set(x_290, 2, x_287); -x_291 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; -lean_inc(x_267); -x_292 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_292, 0, x_267); +lean_ctor_set(x_292, 0, x_290); lean_ctor_set(x_292, 1, x_291); -x_293 = l_Lean_addMacroScope(x_273, x_16, x_269); -x_294 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; -x_295 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; -lean_inc(x_267); -x_296 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_296, 0, x_267); -lean_ctor_set(x_296, 1, x_294); +x_293 = lean_array_mk(x_292); +x_294 = lean_box(2); +x_295 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_296 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_296, 0, x_294); +lean_ctor_set(x_296, 1, x_295); lean_ctor_set(x_296, 2, x_293); -lean_ctor_set(x_296, 3, x_295); -x_297 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -lean_inc(x_267); -x_298 = l_Lean_Syntax_node2(x_267, x_297, x_292, x_296); -lean_inc(x_267); -x_299 = l_Lean_Syntax_node1(x_267, x_276, x_298); -x_300 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -lean_inc(x_278); -lean_inc(x_267); -x_301 = l_Lean_Syntax_node2(x_267, x_300, x_278, x_299); -x_302 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -lean_inc(x_267); +x_297 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +lean_inc(x_273); +x_298 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_298, 0, x_273); +lean_ctor_set(x_298, 1, x_297); +x_299 = l_Lean_addMacroScope(x_279, x_16, x_275); +x_300 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; +x_301 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; +lean_inc(x_273); +x_302 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_302, 0, x_273); +lean_ctor_set(x_302, 1, x_300); +lean_ctor_set(x_302, 2, x_299); +lean_ctor_set(x_302, 3, x_301); +x_303 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_273); +x_304 = l_Lean_Syntax_node2(x_273, x_303, x_298, x_302); +lean_inc(x_273); +x_305 = l_Lean_Syntax_node1(x_273, x_282, x_304); +x_306 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_284); +lean_inc(x_273); +x_307 = l_Lean_Syntax_node2(x_273, x_306, x_284, x_305); +x_308 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +lean_inc(x_273); lean_ctor_set_tag(x_28, 2); -lean_ctor_set(x_28, 1, x_302); -lean_ctor_set(x_28, 0, x_267); -x_303 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; -lean_inc(x_267); +lean_ctor_set(x_28, 1, x_308); +lean_ctor_set(x_28, 0, x_273); +x_309 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; +lean_inc(x_273); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_303); -lean_ctor_set(x_18, 0, x_267); -x_304 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; -lean_inc(x_267); -x_305 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_305, 0, x_267); -lean_ctor_set(x_305, 1, x_304); -x_306 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; -lean_inc(x_305); +lean_ctor_set(x_18, 1, x_309); +lean_ctor_set(x_18, 0, x_273); +x_310 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; +lean_inc(x_273); +x_311 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_311, 0, x_273); +lean_ctor_set(x_311, 1, x_310); +x_312 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +lean_inc(x_311); lean_inc(x_18); -lean_inc(x_267); -x_307 = l_Lean_Syntax_node2(x_267, x_306, x_18, x_305); -x_308 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; -lean_inc(x_278); -lean_inc(x_267); -x_309 = l_Lean_Syntax_node1(x_267, x_308, x_278); -x_310 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_278, 3); -lean_inc(x_267); -x_311 = l_Lean_Syntax_node6(x_267, x_310, x_18, x_278, x_278, x_309, x_278, x_305); -x_312 = l_Lean_Elab_Term_toParserDescr_process___closed__2; -lean_inc(x_267); -x_313 = l_Lean_Syntax_node2(x_267, x_312, x_307, x_311); -x_314 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; -lean_inc_n(x_278, 2); -lean_inc(x_267); -x_315 = l_Lean_Syntax_node2(x_267, x_314, x_278, x_278); -x_316 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -lean_inc(x_278); -lean_inc(x_267); -x_317 = l_Lean_Syntax_node4(x_267, x_316, x_28, x_313, x_315, x_278); -x_318 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; -lean_inc(x_278); -lean_inc(x_267); -x_319 = l_Lean_Syntax_node5(x_267, x_318, x_280, x_290, x_301, x_317, x_278); +lean_inc(x_273); +x_313 = l_Lean_Syntax_node2(x_273, x_312, x_18, x_311); +x_314 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; +lean_inc(x_284); +lean_inc(x_273); +x_315 = l_Lean_Syntax_node1(x_273, x_314, x_284); +x_316 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_284); +lean_inc(x_273); +x_317 = l_Lean_Syntax_node1(x_273, x_316, x_284); +x_318 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_284, 2); +lean_inc(x_273); +x_319 = l_Lean_Syntax_node6(x_273, x_318, x_18, x_284, x_315, x_317, x_284, x_311); +x_320 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +lean_inc(x_273); +x_321 = l_Lean_Syntax_node2(x_273, x_320, x_313, x_319); +x_322 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +lean_inc_n(x_284, 2); +lean_inc(x_273); +x_323 = l_Lean_Syntax_node2(x_273, x_322, x_284, x_284); +x_324 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +lean_inc(x_284); +lean_inc(x_273); +x_325 = l_Lean_Syntax_node4(x_273, x_324, x_28, x_321, x_323, x_284); +x_326 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +lean_inc(x_284); +lean_inc(x_273); +x_327 = l_Lean_Syntax_node5(x_273, x_326, x_286, x_296, x_307, x_325, x_284); if (lean_obj_tag(x_19) == 0) { -lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; -x_320 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; -lean_inc(x_267); -x_321 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_321, 0, x_267); -lean_ctor_set(x_321, 1, x_276); -lean_ctor_set(x_321, 2, x_320); -x_322 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_278, 4); -lean_inc(x_267); -x_323 = l_Lean_Syntax_node6(x_267, x_322, x_321, x_278, x_278, x_278, x_278, x_278); -x_324 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_325 = l_Lean_Syntax_node2(x_267, x_324, x_323, x_319); +lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_328 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; +lean_inc(x_273); +x_329 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_329, 0, x_273); +lean_ctor_set(x_329, 1, x_282); +lean_ctor_set(x_329, 2, x_328); +x_330 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_284, 4); +lean_inc(x_273); +x_331 = l_Lean_Syntax_node6(x_273, x_330, x_329, x_284, x_284, x_284, x_284, x_284); +x_332 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_333 = l_Lean_Syntax_node2(x_273, x_332, x_331, x_327); lean_inc(x_3); lean_inc(x_2); -x_326 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_274); -if (lean_obj_tag(x_326) == 0) +x_334 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_280); +if (lean_obj_tag(x_334) == 0) { -lean_object* x_327; lean_object* x_328; -x_327 = lean_ctor_get(x_326, 1); -lean_inc(x_327); -lean_dec(x_326); -x_328 = l_Lean_Elab_Command_elabCommand(x_325, x_2, x_3, x_327); -return x_328; +lean_object* x_335; lean_object* x_336; +x_335 = lean_ctor_get(x_334, 1); +lean_inc(x_335); +lean_dec(x_334); +x_336 = l_Lean_Elab_Command_elabCommand(x_333, x_2, x_3, x_335); +return x_336; } else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; -lean_dec(x_325); +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; +lean_dec(x_333); lean_dec(x_3); lean_dec(x_2); -x_329 = lean_ctor_get(x_326, 0); -lean_inc(x_329); -x_330 = lean_ctor_get(x_326, 1); -lean_inc(x_330); -if (lean_is_exclusive(x_326)) { - lean_ctor_release(x_326, 0); - lean_ctor_release(x_326, 1); - x_331 = x_326; +x_337 = lean_ctor_get(x_334, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_334, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_334)) { + lean_ctor_release(x_334, 0); + lean_ctor_release(x_334, 1); + x_339 = x_334; } else { - lean_dec_ref(x_326); - x_331 = lean_box(0); + lean_dec_ref(x_334); + x_339 = lean_box(0); } -if (lean_is_scalar(x_331)) { - x_332 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_339)) { + x_340 = lean_alloc_ctor(1, 2, 0); } else { - x_332 = x_331; + x_340 = x_339; } -lean_ctor_set(x_332, 0, x_329); -lean_ctor_set(x_332, 1, x_330); -return x_332; +lean_ctor_set(x_340, 0, x_337); +lean_ctor_set(x_340, 1, x_338); +return x_340; } } else { -lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; -x_333 = lean_ctor_get(x_19, 0); -lean_inc(x_333); +lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; +x_341 = lean_ctor_get(x_19, 0); +lean_inc(x_341); lean_dec(x_19); -x_334 = l_Array_mkArray1___rarg(x_333); -x_335 = l_Array_append___rarg(x_277, x_334); -lean_dec(x_334); -lean_inc(x_267); -x_336 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_336, 0, x_267); -lean_ctor_set(x_336, 1, x_276); -lean_ctor_set(x_336, 2, x_335); -x_337 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_278, 4); -lean_inc(x_267); -x_338 = l_Lean_Syntax_node6(x_267, x_337, x_336, x_278, x_278, x_278, x_278, x_278); -x_339 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_340 = l_Lean_Syntax_node2(x_267, x_339, x_338, x_319); +x_342 = l_Array_mkArray1___rarg(x_341); +x_343 = l_Array_append___rarg(x_283, x_342); +lean_dec(x_342); +lean_inc(x_273); +x_344 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_344, 0, x_273); +lean_ctor_set(x_344, 1, x_282); +lean_ctor_set(x_344, 2, x_343); +x_345 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_284, 4); +lean_inc(x_273); +x_346 = l_Lean_Syntax_node6(x_273, x_345, x_344, x_284, x_284, x_284, x_284, x_284); +x_347 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_348 = l_Lean_Syntax_node2(x_273, x_347, x_346, x_327); lean_inc(x_3); lean_inc(x_2); -x_341 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_274); -if (lean_obj_tag(x_341) == 0) +x_349 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_280); +if (lean_obj_tag(x_349) == 0) { -lean_object* x_342; lean_object* x_343; -x_342 = lean_ctor_get(x_341, 1); -lean_inc(x_342); -lean_dec(x_341); -x_343 = l_Lean_Elab_Command_elabCommand(x_340, x_2, x_3, x_342); -return x_343; +lean_object* x_350; lean_object* x_351; +x_350 = lean_ctor_get(x_349, 1); +lean_inc(x_350); +lean_dec(x_349); +x_351 = l_Lean_Elab_Command_elabCommand(x_348, x_2, x_3, x_350); +return x_351; } else { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; -lean_dec(x_340); +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_348); lean_dec(x_3); lean_dec(x_2); -x_344 = lean_ctor_get(x_341, 0); -lean_inc(x_344); -x_345 = lean_ctor_get(x_341, 1); -lean_inc(x_345); -if (lean_is_exclusive(x_341)) { - lean_ctor_release(x_341, 0); - lean_ctor_release(x_341, 1); - x_346 = x_341; +x_352 = lean_ctor_get(x_349, 0); +lean_inc(x_352); +x_353 = lean_ctor_get(x_349, 1); +lean_inc(x_353); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + x_354 = x_349; } else { - lean_dec_ref(x_341); - x_346 = lean_box(0); + lean_dec_ref(x_349); + x_354 = lean_box(0); } -if (lean_is_scalar(x_346)) { - x_347 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_354)) { + x_355 = lean_alloc_ctor(1, 2, 0); } else { - x_347 = x_346; + x_355 = x_354; } -lean_ctor_set(x_347, 0, x_344); -lean_ctor_set(x_347, 1, x_345); -return x_347; +lean_ctor_set(x_355, 0, x_352); +lean_ctor_set(x_355, 1, x_353); +return x_355; } } } } else { -lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; uint8_t x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; uint8_t x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; -x_348 = lean_ctor_get(x_28, 1); -lean_inc(x_348); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; +x_356 = lean_ctor_get(x_28, 1); +lean_inc(x_356); lean_dec(x_28); -x_349 = l_Lean_Elab_Command_getRef(x_2, x_3, x_348); -x_350 = lean_ctor_get(x_349, 0); -lean_inc(x_350); -x_351 = lean_ctor_get(x_349, 1); -lean_inc(x_351); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - x_352 = x_349; +x_357 = l_Lean_Elab_Command_getRef(x_2, x_3, x_356); +x_358 = lean_ctor_get(x_357, 0); +lean_inc(x_358); +x_359 = lean_ctor_get(x_357, 1); +lean_inc(x_359); +if (lean_is_exclusive(x_357)) { + lean_ctor_release(x_357, 0); + lean_ctor_release(x_357, 1); + x_360 = x_357; } else { - lean_dec_ref(x_349); - x_352 = lean_box(0); -} -x_353 = 0; -x_354 = l_Lean_SourceInfo_fromRef(x_350, x_353); -lean_dec(x_350); -x_355 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_351); -x_356 = lean_ctor_get(x_355, 0); -lean_inc(x_356); -x_357 = lean_ctor_get(x_355, 1); -lean_inc(x_357); -if (lean_is_exclusive(x_355)) { - lean_ctor_release(x_355, 0); - lean_ctor_release(x_355, 1); - x_358 = x_355; + lean_dec_ref(x_357); + x_360 = lean_box(0); +} +x_361 = 0; +x_362 = l_Lean_SourceInfo_fromRef(x_358, x_361); +lean_dec(x_358); +x_363 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_359); +x_364 = lean_ctor_get(x_363, 0); +lean_inc(x_364); +x_365 = lean_ctor_get(x_363, 1); +lean_inc(x_365); +if (lean_is_exclusive(x_363)) { + lean_ctor_release(x_363, 0); + lean_ctor_release(x_363, 1); + x_366 = x_363; } else { - lean_dec_ref(x_355); - x_358 = lean_box(0); -} -x_359 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_357); -x_360 = lean_ctor_get(x_359, 0); -lean_inc(x_360); -x_361 = lean_ctor_get(x_359, 1); -lean_inc(x_361); -if (lean_is_exclusive(x_359)) { - lean_ctor_release(x_359, 0); - lean_ctor_release(x_359, 1); - x_362 = x_359; + lean_dec_ref(x_363); + x_366 = lean_box(0); +} +x_367 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_365); +x_368 = lean_ctor_get(x_367, 0); +lean_inc(x_368); +x_369 = lean_ctor_get(x_367, 1); +lean_inc(x_369); +if (lean_is_exclusive(x_367)) { + lean_ctor_release(x_367, 0); + lean_ctor_release(x_367, 1); + x_370 = x_367; } else { - lean_dec_ref(x_359); - x_362 = lean_box(0); -} -x_363 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; -x_364 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; -lean_inc(x_354); -x_365 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_365, 0, x_354); -lean_ctor_set(x_365, 1, x_363); -lean_ctor_set(x_365, 2, x_364); -x_366 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; -lean_inc(x_354); -if (lean_is_scalar(x_362)) { - x_367 = lean_alloc_ctor(2, 2, 0); + lean_dec_ref(x_367); + x_370 = lean_box(0); +} +x_371 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_372 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; +lean_inc(x_362); +x_373 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_373, 0, x_362); +lean_ctor_set(x_373, 1, x_371); +lean_ctor_set(x_373, 2, x_372); +x_374 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +lean_inc(x_362); +if (lean_is_scalar(x_370)) { + x_375 = lean_alloc_ctor(2, 2, 0); } else { - x_367 = x_362; - lean_ctor_set_tag(x_367, 2); -} -lean_ctor_set(x_367, 0, x_354); -lean_ctor_set(x_367, 1, x_366); -x_368 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; -x_369 = l_Lean_Name_append(x_368, x_17); -x_370 = 1; -x_371 = l_Lean_mkIdentFrom(x_9, x_369, x_370); + x_375 = x_370; + lean_ctor_set_tag(x_375, 2); +} +lean_ctor_set(x_375, 0, x_362); +lean_ctor_set(x_375, 1, x_374); +x_376 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; +x_377 = l_Lean_Name_append(x_376, x_17); +x_378 = 1; +x_379 = l_Lean_mkIdentFrom(x_9, x_377, x_378); lean_dec(x_9); -x_372 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -if (lean_is_scalar(x_358)) { - x_373 = lean_alloc_ctor(1, 2, 0); +x_380 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +if (lean_is_scalar(x_366)) { + x_381 = lean_alloc_ctor(1, 2, 0); } else { - x_373 = x_358; - lean_ctor_set_tag(x_373, 1); -} -lean_ctor_set(x_373, 0, x_371); -lean_ctor_set(x_373, 1, x_372); -x_374 = lean_array_mk(x_373); -x_375 = lean_box(2); -x_376 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; -x_377 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_377, 0, x_375); -lean_ctor_set(x_377, 1, x_376); -lean_ctor_set(x_377, 2, x_374); -x_378 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; -lean_inc(x_354); -if (lean_is_scalar(x_352)) { - x_379 = lean_alloc_ctor(2, 2, 0); + x_381 = x_366; + lean_ctor_set_tag(x_381, 1); +} +lean_ctor_set(x_381, 0, x_379); +lean_ctor_set(x_381, 1, x_380); +x_382 = lean_array_mk(x_381); +x_383 = lean_box(2); +x_384 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_385 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_385, 0, x_383); +lean_ctor_set(x_385, 1, x_384); +lean_ctor_set(x_385, 2, x_382); +x_386 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +lean_inc(x_362); +if (lean_is_scalar(x_360)) { + x_387 = lean_alloc_ctor(2, 2, 0); } else { - x_379 = x_352; - lean_ctor_set_tag(x_379, 2); -} -lean_ctor_set(x_379, 0, x_354); -lean_ctor_set(x_379, 1, x_378); -x_380 = l_Lean_addMacroScope(x_360, x_16, x_356); -x_381 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; -x_382 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; -lean_inc(x_354); -x_383 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_383, 0, x_354); -lean_ctor_set(x_383, 1, x_381); -lean_ctor_set(x_383, 2, x_380); -lean_ctor_set(x_383, 3, x_382); -x_384 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -lean_inc(x_354); -x_385 = l_Lean_Syntax_node2(x_354, x_384, x_379, x_383); -lean_inc(x_354); -x_386 = l_Lean_Syntax_node1(x_354, x_363, x_385); -x_387 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -lean_inc(x_365); -lean_inc(x_354); -x_388 = l_Lean_Syntax_node2(x_354, x_387, x_365, x_386); -x_389 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -lean_inc(x_354); -x_390 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_390, 0, x_354); -lean_ctor_set(x_390, 1, x_389); -x_391 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; -lean_inc(x_354); + x_387 = x_360; + lean_ctor_set_tag(x_387, 2); +} +lean_ctor_set(x_387, 0, x_362); +lean_ctor_set(x_387, 1, x_386); +x_388 = l_Lean_addMacroScope(x_368, x_16, x_364); +x_389 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; +x_390 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; +lean_inc(x_362); +x_391 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_391, 0, x_362); +lean_ctor_set(x_391, 1, x_389); +lean_ctor_set(x_391, 2, x_388); +lean_ctor_set(x_391, 3, x_390); +x_392 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_362); +x_393 = l_Lean_Syntax_node2(x_362, x_392, x_387, x_391); +lean_inc(x_362); +x_394 = l_Lean_Syntax_node1(x_362, x_371, x_393); +x_395 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_373); +lean_inc(x_362); +x_396 = l_Lean_Syntax_node2(x_362, x_395, x_373, x_394); +x_397 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +lean_inc(x_362); +x_398 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_398, 0, x_362); +lean_ctor_set(x_398, 1, x_397); +x_399 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; +lean_inc(x_362); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_391); -lean_ctor_set(x_18, 0, x_354); -x_392 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; -lean_inc(x_354); -x_393 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_393, 0, x_354); -lean_ctor_set(x_393, 1, x_392); -x_394 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; -lean_inc(x_393); +lean_ctor_set(x_18, 1, x_399); +lean_ctor_set(x_18, 0, x_362); +x_400 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; +lean_inc(x_362); +x_401 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_401, 0, x_362); +lean_ctor_set(x_401, 1, x_400); +x_402 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +lean_inc(x_401); lean_inc(x_18); -lean_inc(x_354); -x_395 = l_Lean_Syntax_node2(x_354, x_394, x_18, x_393); -x_396 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; -lean_inc(x_365); -lean_inc(x_354); -x_397 = l_Lean_Syntax_node1(x_354, x_396, x_365); -x_398 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_365, 3); -lean_inc(x_354); -x_399 = l_Lean_Syntax_node6(x_354, x_398, x_18, x_365, x_365, x_397, x_365, x_393); -x_400 = l_Lean_Elab_Term_toParserDescr_process___closed__2; -lean_inc(x_354); -x_401 = l_Lean_Syntax_node2(x_354, x_400, x_395, x_399); -x_402 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; -lean_inc_n(x_365, 2); -lean_inc(x_354); -x_403 = l_Lean_Syntax_node2(x_354, x_402, x_365, x_365); -x_404 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -lean_inc(x_365); -lean_inc(x_354); -x_405 = l_Lean_Syntax_node4(x_354, x_404, x_390, x_401, x_403, x_365); -x_406 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; -lean_inc(x_365); -lean_inc(x_354); -x_407 = l_Lean_Syntax_node5(x_354, x_406, x_367, x_377, x_388, x_405, x_365); +lean_inc(x_362); +x_403 = l_Lean_Syntax_node2(x_362, x_402, x_18, x_401); +x_404 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; +lean_inc(x_373); +lean_inc(x_362); +x_405 = l_Lean_Syntax_node1(x_362, x_404, x_373); +x_406 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_373); +lean_inc(x_362); +x_407 = l_Lean_Syntax_node1(x_362, x_406, x_373); +x_408 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_373, 2); +lean_inc(x_362); +x_409 = l_Lean_Syntax_node6(x_362, x_408, x_18, x_373, x_405, x_407, x_373, x_401); +x_410 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +lean_inc(x_362); +x_411 = l_Lean_Syntax_node2(x_362, x_410, x_403, x_409); +x_412 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +lean_inc_n(x_373, 2); +lean_inc(x_362); +x_413 = l_Lean_Syntax_node2(x_362, x_412, x_373, x_373); +x_414 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +lean_inc(x_373); +lean_inc(x_362); +x_415 = l_Lean_Syntax_node4(x_362, x_414, x_398, x_411, x_413, x_373); +x_416 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +lean_inc(x_373); +lean_inc(x_362); +x_417 = l_Lean_Syntax_node5(x_362, x_416, x_375, x_385, x_396, x_415, x_373); if (lean_obj_tag(x_19) == 0) { -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; -x_408 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; -lean_inc(x_354); -x_409 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_409, 0, x_354); -lean_ctor_set(x_409, 1, x_363); -lean_ctor_set(x_409, 2, x_408); -x_410 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_365, 4); -lean_inc(x_354); -x_411 = l_Lean_Syntax_node6(x_354, x_410, x_409, x_365, x_365, x_365, x_365, x_365); -x_412 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_413 = l_Lean_Syntax_node2(x_354, x_412, x_411, x_407); +lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_418 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; +lean_inc(x_362); +x_419 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_419, 0, x_362); +lean_ctor_set(x_419, 1, x_371); +lean_ctor_set(x_419, 2, x_418); +x_420 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_373, 4); +lean_inc(x_362); +x_421 = l_Lean_Syntax_node6(x_362, x_420, x_419, x_373, x_373, x_373, x_373, x_373); +x_422 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_423 = l_Lean_Syntax_node2(x_362, x_422, x_421, x_417); lean_inc(x_3); lean_inc(x_2); -x_414 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_361); -if (lean_obj_tag(x_414) == 0) +x_424 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_369); +if (lean_obj_tag(x_424) == 0) { -lean_object* x_415; lean_object* x_416; -x_415 = lean_ctor_get(x_414, 1); -lean_inc(x_415); -lean_dec(x_414); -x_416 = l_Lean_Elab_Command_elabCommand(x_413, x_2, x_3, x_415); -return x_416; +lean_object* x_425; lean_object* x_426; +x_425 = lean_ctor_get(x_424, 1); +lean_inc(x_425); +lean_dec(x_424); +x_426 = l_Lean_Elab_Command_elabCommand(x_423, x_2, x_3, x_425); +return x_426; } else { -lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; -lean_dec(x_413); +lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +lean_dec(x_423); lean_dec(x_3); lean_dec(x_2); -x_417 = lean_ctor_get(x_414, 0); -lean_inc(x_417); -x_418 = lean_ctor_get(x_414, 1); -lean_inc(x_418); -if (lean_is_exclusive(x_414)) { - lean_ctor_release(x_414, 0); - lean_ctor_release(x_414, 1); - x_419 = x_414; +x_427 = lean_ctor_get(x_424, 0); +lean_inc(x_427); +x_428 = lean_ctor_get(x_424, 1); +lean_inc(x_428); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + x_429 = x_424; } else { - lean_dec_ref(x_414); - x_419 = lean_box(0); + lean_dec_ref(x_424); + x_429 = lean_box(0); } -if (lean_is_scalar(x_419)) { - x_420 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_429)) { + x_430 = lean_alloc_ctor(1, 2, 0); } else { - x_420 = x_419; + x_430 = x_429; } -lean_ctor_set(x_420, 0, x_417); -lean_ctor_set(x_420, 1, x_418); -return x_420; +lean_ctor_set(x_430, 0, x_427); +lean_ctor_set(x_430, 1, x_428); +return x_430; } } else { -lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; -x_421 = lean_ctor_get(x_19, 0); -lean_inc(x_421); +lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; +x_431 = lean_ctor_get(x_19, 0); +lean_inc(x_431); lean_dec(x_19); -x_422 = l_Array_mkArray1___rarg(x_421); -x_423 = l_Array_append___rarg(x_364, x_422); -lean_dec(x_422); -lean_inc(x_354); -x_424 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_424, 0, x_354); -lean_ctor_set(x_424, 1, x_363); -lean_ctor_set(x_424, 2, x_423); -x_425 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_365, 4); -lean_inc(x_354); -x_426 = l_Lean_Syntax_node6(x_354, x_425, x_424, x_365, x_365, x_365, x_365, x_365); -x_427 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_428 = l_Lean_Syntax_node2(x_354, x_427, x_426, x_407); +x_432 = l_Array_mkArray1___rarg(x_431); +x_433 = l_Array_append___rarg(x_372, x_432); +lean_dec(x_432); +lean_inc(x_362); +x_434 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_434, 0, x_362); +lean_ctor_set(x_434, 1, x_371); +lean_ctor_set(x_434, 2, x_433); +x_435 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_373, 4); +lean_inc(x_362); +x_436 = l_Lean_Syntax_node6(x_362, x_435, x_434, x_373, x_373, x_373, x_373, x_373); +x_437 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_438 = l_Lean_Syntax_node2(x_362, x_437, x_436, x_417); lean_inc(x_3); lean_inc(x_2); -x_429 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_361); -if (lean_obj_tag(x_429) == 0) +x_439 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_369); +if (lean_obj_tag(x_439) == 0) { -lean_object* x_430; lean_object* x_431; -x_430 = lean_ctor_get(x_429, 1); -lean_inc(x_430); -lean_dec(x_429); -x_431 = l_Lean_Elab_Command_elabCommand(x_428, x_2, x_3, x_430); -return x_431; +lean_object* x_440; lean_object* x_441; +x_440 = lean_ctor_get(x_439, 1); +lean_inc(x_440); +lean_dec(x_439); +x_441 = l_Lean_Elab_Command_elabCommand(x_438, x_2, x_3, x_440); +return x_441; } else { -lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; -lean_dec(x_428); +lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; +lean_dec(x_438); lean_dec(x_3); lean_dec(x_2); -x_432 = lean_ctor_get(x_429, 0); -lean_inc(x_432); -x_433 = lean_ctor_get(x_429, 1); -lean_inc(x_433); -if (lean_is_exclusive(x_429)) { - lean_ctor_release(x_429, 0); - lean_ctor_release(x_429, 1); - x_434 = x_429; +x_442 = lean_ctor_get(x_439, 0); +lean_inc(x_442); +x_443 = lean_ctor_get(x_439, 1); +lean_inc(x_443); +if (lean_is_exclusive(x_439)) { + lean_ctor_release(x_439, 0); + lean_ctor_release(x_439, 1); + x_444 = x_439; } else { - lean_dec_ref(x_429); - x_434 = lean_box(0); + lean_dec_ref(x_439); + x_444 = lean_box(0); } -if (lean_is_scalar(x_434)) { - x_435 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_444)) { + x_445 = lean_alloc_ctor(1, 2, 0); } else { - x_435 = x_434; + x_445 = x_444; } -lean_ctor_set(x_435, 0, x_432); -lean_ctor_set(x_435, 1, x_433); -return x_435; +lean_ctor_set(x_445, 0, x_442); +lean_ctor_set(x_445, 1, x_443); +return x_445; } } } } else { -uint8_t x_436; +uint8_t x_446; lean_dec(x_19); lean_dec(x_17); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_436 = !lean_is_exclusive(x_25); -if (x_436 == 0) +x_446 = !lean_is_exclusive(x_25); +if (x_446 == 0) { -lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; -x_437 = lean_ctor_get(x_25, 0); -x_438 = lean_ctor_get(x_2, 6); -lean_inc(x_438); +lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; +x_447 = lean_ctor_get(x_25, 0); +x_448 = lean_ctor_get(x_2, 6); +lean_inc(x_448); lean_dec(x_2); -x_439 = lean_io_error_to_string(x_437); -x_440 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_440, 0, x_439); -x_441 = l_Lean_MessageData_ofFormat(x_440); -lean_ctor_set(x_18, 1, x_441); -lean_ctor_set(x_18, 0, x_438); +x_449 = lean_io_error_to_string(x_447); +x_450 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_450, 0, x_449); +x_451 = l_Lean_MessageData_ofFormat(x_450); +lean_ctor_set(x_18, 1, x_451); +lean_ctor_set(x_18, 0, x_448); lean_ctor_set(x_25, 0, x_18); return x_25; } else { -lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -x_442 = lean_ctor_get(x_25, 0); -x_443 = lean_ctor_get(x_25, 1); -lean_inc(x_443); -lean_inc(x_442); +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_452 = lean_ctor_get(x_25, 0); +x_453 = lean_ctor_get(x_25, 1); +lean_inc(x_453); +lean_inc(x_452); lean_dec(x_25); -x_444 = lean_ctor_get(x_2, 6); -lean_inc(x_444); +x_454 = lean_ctor_get(x_2, 6); +lean_inc(x_454); lean_dec(x_2); -x_445 = lean_io_error_to_string(x_442); -x_446 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_446, 0, x_445); -x_447 = l_Lean_MessageData_ofFormat(x_446); -lean_ctor_set(x_18, 1, x_447); -lean_ctor_set(x_18, 0, x_444); -x_448 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_448, 0, x_18); -lean_ctor_set(x_448, 1, x_443); -return x_448; +x_455 = lean_io_error_to_string(x_452); +x_456 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_456, 0, x_455); +x_457 = l_Lean_MessageData_ofFormat(x_456); +lean_ctor_set(x_18, 1, x_457); +lean_ctor_set(x_18, 0, x_454); +x_458 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_458, 0, x_18); +lean_ctor_set(x_458, 1, x_453); +return x_458; } } } else { -lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; -x_449 = lean_ctor_get(x_18, 0); -x_450 = lean_ctor_get(x_18, 1); -lean_inc(x_450); -lean_inc(x_449); +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; +x_459 = lean_ctor_get(x_18, 0); +x_460 = lean_ctor_get(x_18, 1); +lean_inc(x_460); +lean_inc(x_459); lean_dec(x_18); -x_451 = lean_ctor_get(x_449, 0); -lean_inc(x_451); -lean_dec(x_449); +x_461 = lean_ctor_get(x_459, 0); +lean_inc(x_461); +lean_dec(x_459); lean_inc(x_17); lean_inc(x_10); -x_452 = l_Lean_Parser_registerParserCategory(x_451, x_15, x_10, x_20, x_17, x_450); -if (lean_obj_tag(x_452) == 0) +x_462 = l_Lean_Parser_registerParserCategory(x_461, x_15, x_10, x_20, x_17, x_460); +if (lean_obj_tag(x_462) == 0) { -lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; uint8_t x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; -x_453 = lean_ctor_get(x_452, 0); -lean_inc(x_453); -x_454 = lean_ctor_get(x_452, 1); -lean_inc(x_454); -lean_dec(x_452); -x_455 = l_Lean_setEnv___at_Lean_Elab_Command_elabDeclareSyntaxCat___spec__1(x_453, x_2, x_3, x_454); -x_456 = lean_ctor_get(x_455, 1); -lean_inc(x_456); -if (lean_is_exclusive(x_455)) { - lean_ctor_release(x_455, 0); - lean_ctor_release(x_455, 1); - x_457 = x_455; -} else { - lean_dec_ref(x_455); - x_457 = lean_box(0); -} -x_458 = l_Lean_Elab_Command_getRef(x_2, x_3, x_456); -x_459 = lean_ctor_get(x_458, 0); -lean_inc(x_459); -x_460 = lean_ctor_get(x_458, 1); -lean_inc(x_460); -if (lean_is_exclusive(x_458)) { - lean_ctor_release(x_458, 0); - lean_ctor_release(x_458, 1); - x_461 = x_458; -} else { - lean_dec_ref(x_458); - x_461 = lean_box(0); -} -x_462 = 0; -x_463 = l_Lean_SourceInfo_fromRef(x_459, x_462); -lean_dec(x_459); -x_464 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_460); -x_465 = lean_ctor_get(x_464, 0); -lean_inc(x_465); -x_466 = lean_ctor_get(x_464, 1); +lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; uint8_t x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; uint8_t x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; +x_463 = lean_ctor_get(x_462, 0); +lean_inc(x_463); +x_464 = lean_ctor_get(x_462, 1); +lean_inc(x_464); +lean_dec(x_462); +x_465 = l_Lean_setEnv___at_Lean_Elab_Command_elabDeclareSyntaxCat___spec__1(x_463, x_2, x_3, x_464); +x_466 = lean_ctor_get(x_465, 1); lean_inc(x_466); -if (lean_is_exclusive(x_464)) { - lean_ctor_release(x_464, 0); - lean_ctor_release(x_464, 1); - x_467 = x_464; +if (lean_is_exclusive(x_465)) { + lean_ctor_release(x_465, 0); + lean_ctor_release(x_465, 1); + x_467 = x_465; } else { - lean_dec_ref(x_464); + lean_dec_ref(x_465); x_467 = lean_box(0); } -x_468 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_466); +x_468 = l_Lean_Elab_Command_getRef(x_2, x_3, x_466); x_469 = lean_ctor_get(x_468, 0); lean_inc(x_469); x_470 = lean_ctor_get(x_468, 1); @@ -17128,274 +17223,307 @@ if (lean_is_exclusive(x_468)) { lean_dec_ref(x_468); x_471 = lean_box(0); } -x_472 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; -x_473 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; -lean_inc(x_463); -x_474 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_474, 0, x_463); -lean_ctor_set(x_474, 1, x_472); -lean_ctor_set(x_474, 2, x_473); -x_475 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; -lean_inc(x_463); -if (lean_is_scalar(x_471)) { - x_476 = lean_alloc_ctor(2, 2, 0); +x_472 = 0; +x_473 = l_Lean_SourceInfo_fromRef(x_469, x_472); +lean_dec(x_469); +x_474 = l_Lean_Elab_Command_getCurrMacroScope(x_2, x_3, x_470); +x_475 = lean_ctor_get(x_474, 0); +lean_inc(x_475); +x_476 = lean_ctor_get(x_474, 1); +lean_inc(x_476); +if (lean_is_exclusive(x_474)) { + lean_ctor_release(x_474, 0); + lean_ctor_release(x_474, 1); + x_477 = x_474; } else { - x_476 = x_471; - lean_ctor_set_tag(x_476, 2); -} -lean_ctor_set(x_476, 0, x_463); -lean_ctor_set(x_476, 1, x_475); -x_477 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; -x_478 = l_Lean_Name_append(x_477, x_17); -x_479 = 1; -x_480 = l_Lean_mkIdentFrom(x_9, x_478, x_479); -lean_dec(x_9); -x_481 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; -if (lean_is_scalar(x_467)) { - x_482 = lean_alloc_ctor(1, 2, 0); + lean_dec_ref(x_474); + x_477 = lean_box(0); +} +x_478 = l_Lean_Elab_Command_getMainModule___rarg(x_3, x_476); +x_479 = lean_ctor_get(x_478, 0); +lean_inc(x_479); +x_480 = lean_ctor_get(x_478, 1); +lean_inc(x_480); +if (lean_is_exclusive(x_478)) { + lean_ctor_release(x_478, 0); + lean_ctor_release(x_478, 1); + x_481 = x_478; +} else { + lean_dec_ref(x_478); + x_481 = lean_box(0); +} +x_482 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__17; +x_483 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__10; +lean_inc(x_473); +x_484 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_484, 0, x_473); +lean_ctor_set(x_484, 1, x_482); +lean_ctor_set(x_484, 2, x_483); +x_485 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__27; +lean_inc(x_473); +if (lean_is_scalar(x_481)) { + x_486 = lean_alloc_ctor(2, 2, 0); } else { - x_482 = x_467; - lean_ctor_set_tag(x_482, 1); -} -lean_ctor_set(x_482, 0, x_480); -lean_ctor_set(x_482, 1, x_481); -x_483 = lean_array_mk(x_482); -x_484 = lean_box(2); -x_485 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; -x_486 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_486, 0, x_484); + x_486 = x_481; + lean_ctor_set_tag(x_486, 2); +} +lean_ctor_set(x_486, 0, x_473); lean_ctor_set(x_486, 1, x_485); -lean_ctor_set(x_486, 2, x_483); -x_487 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; -lean_inc(x_463); -if (lean_is_scalar(x_461)) { - x_488 = lean_alloc_ctor(2, 2, 0); +x_487 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__3; +x_488 = l_Lean_Name_append(x_487, x_17); +x_489 = 1; +x_490 = l_Lean_mkIdentFrom(x_9, x_488, x_489); +lean_dec(x_9); +x_491 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__31; +if (lean_is_scalar(x_477)) { + x_492 = lean_alloc_ctor(1, 2, 0); } else { - x_488 = x_461; - lean_ctor_set_tag(x_488, 2); -} -lean_ctor_set(x_488, 0, x_463); -lean_ctor_set(x_488, 1, x_487); -x_489 = l_Lean_addMacroScope(x_469, x_16, x_465); -x_490 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; -x_491 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; -lean_inc(x_463); -x_492 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_492, 0, x_463); -lean_ctor_set(x_492, 1, x_490); -lean_ctor_set(x_492, 2, x_489); -lean_ctor_set(x_492, 3, x_491); -x_493 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; -lean_inc(x_463); -x_494 = l_Lean_Syntax_node2(x_463, x_493, x_488, x_492); -lean_inc(x_463); -x_495 = l_Lean_Syntax_node1(x_463, x_472, x_494); -x_496 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; -lean_inc(x_474); -lean_inc(x_463); -x_497 = l_Lean_Syntax_node2(x_463, x_496, x_474, x_495); -x_498 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; -lean_inc(x_463); -if (lean_is_scalar(x_457)) { - x_499 = lean_alloc_ctor(2, 2, 0); + x_492 = x_477; + lean_ctor_set_tag(x_492, 1); +} +lean_ctor_set(x_492, 0, x_490); +lean_ctor_set(x_492, 1, x_491); +x_493 = lean_array_mk(x_492); +x_494 = lean_box(2); +x_495 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__29; +x_496 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_496, 0, x_494); +lean_ctor_set(x_496, 1, x_495); +lean_ctor_set(x_496, 2, x_493); +x_497 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__36; +lean_inc(x_473); +if (lean_is_scalar(x_471)) { + x_498 = lean_alloc_ctor(2, 2, 0); } else { - x_499 = x_457; - lean_ctor_set_tag(x_499, 2); + x_498 = x_471; + lean_ctor_set_tag(x_498, 2); +} +lean_ctor_set(x_498, 0, x_473); +lean_ctor_set(x_498, 1, x_497); +x_499 = l_Lean_addMacroScope(x_479, x_16, x_475); +x_500 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__5; +x_501 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; +lean_inc(x_473); +x_502 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_502, 0, x_473); +lean_ctor_set(x_502, 1, x_500); +lean_ctor_set(x_502, 2, x_499); +lean_ctor_set(x_502, 3, x_501); +x_503 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__35; +lean_inc(x_473); +x_504 = l_Lean_Syntax_node2(x_473, x_503, x_498, x_502); +lean_inc(x_473); +x_505 = l_Lean_Syntax_node1(x_473, x_482, x_504); +x_506 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__33; +lean_inc(x_484); +lean_inc(x_473); +x_507 = l_Lean_Syntax_node2(x_473, x_506, x_484, x_505); +x_508 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__46; +lean_inc(x_473); +if (lean_is_scalar(x_467)) { + x_509 = lean_alloc_ctor(2, 2, 0); +} else { + x_509 = x_467; + lean_ctor_set_tag(x_509, 2); } -lean_ctor_set(x_499, 0, x_463); -lean_ctor_set(x_499, 1, x_498); -x_500 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; -lean_inc(x_463); -x_501 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_501, 0, x_463); -lean_ctor_set(x_501, 1, x_500); -x_502 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; -lean_inc(x_463); -x_503 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_503, 0, x_463); -lean_ctor_set(x_503, 1, x_502); -x_504 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; -lean_inc(x_503); -lean_inc(x_501); -lean_inc(x_463); -x_505 = l_Lean_Syntax_node2(x_463, x_504, x_501, x_503); -x_506 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; -lean_inc(x_474); -lean_inc(x_463); -x_507 = l_Lean_Syntax_node1(x_463, x_506, x_474); -x_508 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; -lean_inc_n(x_474, 3); -lean_inc(x_463); -x_509 = l_Lean_Syntax_node6(x_463, x_508, x_501, x_474, x_474, x_507, x_474, x_503); -x_510 = l_Lean_Elab_Term_toParserDescr_process___closed__2; -lean_inc(x_463); -x_511 = l_Lean_Syntax_node2(x_463, x_510, x_505, x_509); -x_512 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; -lean_inc_n(x_474, 2); -lean_inc(x_463); -x_513 = l_Lean_Syntax_node2(x_463, x_512, x_474, x_474); -x_514 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; -lean_inc(x_474); -lean_inc(x_463); -x_515 = l_Lean_Syntax_node4(x_463, x_514, x_499, x_511, x_513, x_474); -x_516 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; -lean_inc(x_474); -lean_inc(x_463); -x_517 = l_Lean_Syntax_node5(x_463, x_516, x_476, x_486, x_497, x_515, x_474); +lean_ctor_set(x_509, 0, x_473); +lean_ctor_set(x_509, 1, x_508); +x_510 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__12; +lean_inc(x_473); +x_511 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_511, 0, x_473); +lean_ctor_set(x_511, 1, x_510); +x_512 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__13; +lean_inc(x_473); +x_513 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_513, 0, x_473); +lean_ctor_set(x_513, 1, x_512); +x_514 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__11; +lean_inc(x_513); +lean_inc(x_511); +lean_inc(x_473); +x_515 = l_Lean_Syntax_node2(x_473, x_514, x_511, x_513); +x_516 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__17; +lean_inc(x_484); +lean_inc(x_473); +x_517 = l_Lean_Syntax_node1(x_473, x_516, x_484); +x_518 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19; +lean_inc(x_484); +lean_inc(x_473); +x_519 = l_Lean_Syntax_node1(x_473, x_518, x_484); +x_520 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__15; +lean_inc_n(x_484, 2); +lean_inc(x_473); +x_521 = l_Lean_Syntax_node6(x_473, x_520, x_511, x_484, x_517, x_519, x_484, x_513); +x_522 = l_Lean_Elab_Term_toParserDescr_process___closed__2; +lean_inc(x_473); +x_523 = l_Lean_Syntax_node2(x_473, x_522, x_515, x_521); +x_524 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__76; +lean_inc_n(x_484, 2); +lean_inc(x_473); +x_525 = l_Lean_Syntax_node2(x_473, x_524, x_484, x_484); +x_526 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__45; +lean_inc(x_484); +lean_inc(x_473); +x_527 = l_Lean_Syntax_node4(x_473, x_526, x_509, x_523, x_525, x_484); +x_528 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__26; +lean_inc(x_484); +lean_inc(x_473); +x_529 = l_Lean_Syntax_node5(x_473, x_528, x_486, x_496, x_507, x_527, x_484); if (lean_obj_tag(x_19) == 0) { -lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; -x_518 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__18; -lean_inc(x_463); -x_519 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_519, 0, x_463); -lean_ctor_set(x_519, 1, x_472); -lean_ctor_set(x_519, 2, x_518); -x_520 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_474, 4); -lean_inc(x_463); -x_521 = l_Lean_Syntax_node6(x_463, x_520, x_519, x_474, x_474, x_474, x_474, x_474); -x_522 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_523 = l_Lean_Syntax_node2(x_463, x_522, x_521, x_517); +lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +x_530 = l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20; +lean_inc(x_473); +x_531 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_531, 0, x_473); +lean_ctor_set(x_531, 1, x_482); +lean_ctor_set(x_531, 2, x_530); +x_532 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_484, 4); +lean_inc(x_473); +x_533 = l_Lean_Syntax_node6(x_473, x_532, x_531, x_484, x_484, x_484, x_484, x_484); +x_534 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_535 = l_Lean_Syntax_node2(x_473, x_534, x_533, x_529); lean_inc(x_3); lean_inc(x_2); -x_524 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_470); -if (lean_obj_tag(x_524) == 0) +x_536 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_480); +if (lean_obj_tag(x_536) == 0) { -lean_object* x_525; lean_object* x_526; -x_525 = lean_ctor_get(x_524, 1); -lean_inc(x_525); -lean_dec(x_524); -x_526 = l_Lean_Elab_Command_elabCommand(x_523, x_2, x_3, x_525); -return x_526; +lean_object* x_537; lean_object* x_538; +x_537 = lean_ctor_get(x_536, 1); +lean_inc(x_537); +lean_dec(x_536); +x_538 = l_Lean_Elab_Command_elabCommand(x_535, x_2, x_3, x_537); +return x_538; } else { -lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; -lean_dec(x_523); +lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; +lean_dec(x_535); lean_dec(x_3); lean_dec(x_2); -x_527 = lean_ctor_get(x_524, 0); -lean_inc(x_527); -x_528 = lean_ctor_get(x_524, 1); -lean_inc(x_528); -if (lean_is_exclusive(x_524)) { - lean_ctor_release(x_524, 0); - lean_ctor_release(x_524, 1); - x_529 = x_524; +x_539 = lean_ctor_get(x_536, 0); +lean_inc(x_539); +x_540 = lean_ctor_get(x_536, 1); +lean_inc(x_540); +if (lean_is_exclusive(x_536)) { + lean_ctor_release(x_536, 0); + lean_ctor_release(x_536, 1); + x_541 = x_536; } else { - lean_dec_ref(x_524); - x_529 = lean_box(0); + lean_dec_ref(x_536); + x_541 = lean_box(0); } -if (lean_is_scalar(x_529)) { - x_530 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_541)) { + x_542 = lean_alloc_ctor(1, 2, 0); } else { - x_530 = x_529; + x_542 = x_541; } -lean_ctor_set(x_530, 0, x_527); -lean_ctor_set(x_530, 1, x_528); -return x_530; +lean_ctor_set(x_542, 0, x_539); +lean_ctor_set(x_542, 1, x_540); +return x_542; } } else { -lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; -x_531 = lean_ctor_get(x_19, 0); -lean_inc(x_531); +lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; +x_543 = lean_ctor_get(x_19, 0); +lean_inc(x_543); lean_dec(x_19); -x_532 = l_Array_mkArray1___rarg(x_531); -x_533 = l_Array_append___rarg(x_473, x_532); -lean_dec(x_532); -lean_inc(x_463); -x_534 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_534, 0, x_463); -lean_ctor_set(x_534, 1, x_472); -lean_ctor_set(x_534, 2, x_533); -x_535 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; -lean_inc_n(x_474, 4); -lean_inc(x_463); -x_536 = l_Lean_Syntax_node6(x_463, x_535, x_534, x_474, x_474, x_474, x_474, x_474); -x_537 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; -x_538 = l_Lean_Syntax_node2(x_463, x_537, x_536, x_517); +x_544 = l_Array_mkArray1___rarg(x_543); +x_545 = l_Array_append___rarg(x_483, x_544); +lean_dec(x_544); +lean_inc(x_473); +x_546 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_546, 0, x_473); +lean_ctor_set(x_546, 1, x_482); +lean_ctor_set(x_546, 2, x_545); +x_547 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__9; +lean_inc_n(x_484, 4); +lean_inc(x_473); +x_548 = l_Lean_Syntax_node6(x_473, x_547, x_546, x_484, x_484, x_484, x_484, x_484); +x_549 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__7; +x_550 = l_Lean_Syntax_node2(x_473, x_549, x_548, x_529); lean_inc(x_3); lean_inc(x_2); -x_539 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_470); -if (lean_obj_tag(x_539) == 0) +x_551 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser(x_10, x_2, x_3, x_480); +if (lean_obj_tag(x_551) == 0) { -lean_object* x_540; lean_object* x_541; -x_540 = lean_ctor_get(x_539, 1); -lean_inc(x_540); -lean_dec(x_539); -x_541 = l_Lean_Elab_Command_elabCommand(x_538, x_2, x_3, x_540); -return x_541; +lean_object* x_552; lean_object* x_553; +x_552 = lean_ctor_get(x_551, 1); +lean_inc(x_552); +lean_dec(x_551); +x_553 = l_Lean_Elab_Command_elabCommand(x_550, x_2, x_3, x_552); +return x_553; } else { -lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; -lean_dec(x_538); +lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; +lean_dec(x_550); lean_dec(x_3); lean_dec(x_2); -x_542 = lean_ctor_get(x_539, 0); -lean_inc(x_542); -x_543 = lean_ctor_get(x_539, 1); -lean_inc(x_543); -if (lean_is_exclusive(x_539)) { - lean_ctor_release(x_539, 0); - lean_ctor_release(x_539, 1); - x_544 = x_539; +x_554 = lean_ctor_get(x_551, 0); +lean_inc(x_554); +x_555 = lean_ctor_get(x_551, 1); +lean_inc(x_555); +if (lean_is_exclusive(x_551)) { + lean_ctor_release(x_551, 0); + lean_ctor_release(x_551, 1); + x_556 = x_551; } else { - lean_dec_ref(x_539); - x_544 = lean_box(0); + lean_dec_ref(x_551); + x_556 = lean_box(0); } -if (lean_is_scalar(x_544)) { - x_545 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_556)) { + x_557 = lean_alloc_ctor(1, 2, 0); } else { - x_545 = x_544; + x_557 = x_556; } -lean_ctor_set(x_545, 0, x_542); -lean_ctor_set(x_545, 1, x_543); -return x_545; +lean_ctor_set(x_557, 0, x_554); +lean_ctor_set(x_557, 1, x_555); +return x_557; } } } else { -lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; +lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_dec(x_19); lean_dec(x_17); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_546 = lean_ctor_get(x_452, 0); -lean_inc(x_546); -x_547 = lean_ctor_get(x_452, 1); -lean_inc(x_547); -if (lean_is_exclusive(x_452)) { - lean_ctor_release(x_452, 0); - lean_ctor_release(x_452, 1); - x_548 = x_452; +x_558 = lean_ctor_get(x_462, 0); +lean_inc(x_558); +x_559 = lean_ctor_get(x_462, 1); +lean_inc(x_559); +if (lean_is_exclusive(x_462)) { + lean_ctor_release(x_462, 0); + lean_ctor_release(x_462, 1); + x_560 = x_462; } else { - lean_dec_ref(x_452); - x_548 = lean_box(0); + lean_dec_ref(x_462); + x_560 = lean_box(0); } -x_549 = lean_ctor_get(x_2, 6); -lean_inc(x_549); +x_561 = lean_ctor_get(x_2, 6); +lean_inc(x_561); lean_dec(x_2); -x_550 = lean_io_error_to_string(x_546); -x_551 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_551, 0, x_550); -x_552 = l_Lean_MessageData_ofFormat(x_551); -x_553 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_553, 0, x_549); -lean_ctor_set(x_553, 1, x_552); -if (lean_is_scalar(x_548)) { - x_554 = lean_alloc_ctor(1, 2, 0); +x_562 = lean_io_error_to_string(x_558); +x_563 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_563, 0, x_562); +x_564 = l_Lean_MessageData_ofFormat(x_563); +x_565 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_565, 0, x_561); +lean_ctor_set(x_565, 1, x_564); +if (lean_is_scalar(x_560)) { + x_566 = lean_alloc_ctor(1, 2, 0); } else { - x_554 = x_548; + x_566 = x_560; } -lean_ctor_set(x_554, 0, x_553); -lean_ctor_set(x_554, 1, x_547); -return x_554; +lean_ctor_set(x_566, 0, x_565); +lean_ctor_set(x_566, 1, x_559); +return x_566; } } } @@ -25869,7 +25997,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1() { _start: { lean_object* x_1; @@ -25877,17 +26005,17 @@ x_1 = lean_mk_string_unchecked("defaultInstance", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -25897,27 +26025,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3; x_2 = l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4; x_2 = l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6() { _start: { lean_object* x_1; @@ -25925,17 +26053,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8() { _start: { lean_object* x_1; @@ -25943,47 +26071,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9; x_2 = l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10; x_2 = l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11; x_2 = l_Lean_Elab_Term_checkLeftRec___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13() { _start: { lean_object* x_1; @@ -25991,33 +26119,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14; -x_2 = lean_unsigned_to_nat(11419u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14; +x_2 = lean_unsigned_to_nat(11445u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -26234,6 +26362,8 @@ l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__25 = _init_l_Lean_El lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__25); l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__26 = _init_l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__26(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__26); +l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_isValidAtom___closed__1); l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1 = _init_l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1); l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2 = _init_l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__2(); @@ -26630,6 +26760,10 @@ l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19 = _init_l_Lean_Elab_Comman lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__19); l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20(); lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__20); +l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21(); +lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__21); +l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22 = _init_l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22(); +lean_mark_persistent(l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__22); l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__1 = _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__1); l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__2 = _init_l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat__1___closed__2(); @@ -26832,37 +26966,37 @@ l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__5 = _init_l_ lean_mark_persistent(l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__5); l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6 = _init_l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6(); lean_mark_persistent(l_Lean_Elab_Command_expandNoKindMacroRulesAux___lambda__1___closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__14); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419____closed__15); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11419_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445____closed__15); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Syntax___hyg_11445_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Basic.c index 016473f0f8fa..c7ab36da71ec 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -21,6 +21,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryCatchRestore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); @@ -31,6 +32,7 @@ static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrA LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__9; static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__1; extern lean_object* l_Lean_profiler; @@ -43,7 +45,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1___closed__3; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_evalTactic_eval___spec__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_admitGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -68,6 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SavedState_restore___boxed(lean_obje extern lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshot; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_evalTactic_eval___spec__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Tactic_closeMainGoal___lambda__1___closed__1; @@ -79,7 +82,6 @@ lean_object* lean_io_promise_new(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,6 +122,7 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_appendGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMacroExpansion___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainGoal_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_eval___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,11 +143,13 @@ static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___ static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__2; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_popMainGoal___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -159,6 +164,7 @@ lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_withTacticInfoContext___spec__1(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6; lean_object* lean_io_promise_result(lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_setMVarUserName(lean_object*, lean_object*, lean_object*); @@ -172,6 +178,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_replaceMainGoal___boxed(lean_object* uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic(lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -180,7 +187,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg(lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__6___boxed__const__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkInitialTacticInfo___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___at_Lean_Elab_Tactic_evalTactic_eval___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,7 +241,6 @@ static lean_object* l_Lean_Elab_Tactic_evalTactic_throwExs___closed__2; static lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549_(lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,6 +254,7 @@ lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getGoals(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -262,9 +268,7 @@ static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___at_Lean_Elab_Tactic_evalTactic_eval___spec__4___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__6___closed__6; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4; static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590_(lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tagUntaggedGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_withMacroExpansion___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -289,14 +293,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_adaptExpander___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_withMacroExpansion___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeUsingOrAdmit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); @@ -304,7 +306,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); @@ -324,6 +325,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_eval(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedTacticM___rarg(lean_object*); lean_object* lean_io_mono_nanos_now(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__6___closed__4; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedTacticM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___closed__4; @@ -342,12 +344,10 @@ LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_tagUntagged LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__9(lean_object*); lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_closeMainGoal___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_eval___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -378,6 +378,7 @@ static lean_object* l_Lean_Elab_Tactic_run___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Language_Snapshot_Diagnostics_empty; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14; static lean_object* l_Lean_Elab_Tactic_ensureHasNoMVars___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_isAbortTacticException(lean_object*); @@ -390,6 +391,7 @@ lean_object* l_Lean_Elab_mkElabAttribute___rarg(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg___boxed(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; @@ -397,15 +399,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__2___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved(lean_object*); lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_focusAndDone___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13; lean_object* lean_name_append_index_after(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___boxed(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeUsingOrAdmit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -423,8 +425,9 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTac LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11; static lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__4; @@ -468,14 +471,13 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTa static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMainContext(lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withRestoreOrSaveFull___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -520,7 +522,6 @@ LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Elab_Tactic_pruneSolvedGoal lean_object* l_Lean_Language_SnapshotTask_get___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__1; lean_object* lean_io_error_to_string(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3; lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Tactic_pruneSolvedGoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -538,7 +539,7 @@ static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_handl LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withTacticInfoContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1; static lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__2; @@ -576,7 +577,6 @@ lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_obj LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveState(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_dbg_trace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -19391,7 +19391,7 @@ x_14 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_1); lean_ctor_set(x_14, 2, x_2); -x_15 = lean_alloc_ctor(3, 1, 0); +x_15 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_15, 0, x_14); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -23417,7 +23417,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23427,7 +23427,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23437,27 +23437,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2; x_2 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3; x_2 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5() { _start: { lean_object* x_1; @@ -23465,17 +23465,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7() { _start: { lean_object* x_1; @@ -23483,47 +23483,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8; x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9; x_2 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10; x_2 = l_Lean_logAt___at_Lean_Elab_Term_reportUnsolvedGoals___spec__2___lambda__2___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12() { _start: { lean_object* x_1; @@ -23531,17 +23531,17 @@ x_1 = lean_mk_string_unchecked("Basic", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14() { _start: { lean_object* x_1; @@ -23549,54 +23549,54 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15; -x_2 = lean_unsigned_to_nat(8549u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15; +x_2 = lean_unsigned_to_nat(8548u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15; -x_2 = lean_unsigned_to_nat(8590u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15; +x_2 = lean_unsigned_to_nat(8589u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_List_forIn_x27_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -23852,44 +23852,44 @@ l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__1 = _init_l_Lean_Elab_Tactic_get lean_mark_persistent(l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__1); l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2 = _init_l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__6); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__7); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__8); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__9); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__10); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__11); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__12); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__13); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__14); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__15); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549____closed__16); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8549_(lean_io_mk_world()); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__6); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__7); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__8); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__9); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__10); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__11); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__12); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__13); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__14); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__15); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548____closed__16); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8548_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590____closed__1); -if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8590_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589____closed__1); +if (builtin) {res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_8589_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index 5f0c487d5925..7a2a068b10fa 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -21220,7 +21220,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Tactic_ela _start: { lean_object* x_11; lean_object* x_12; -x_11 = lean_alloc_ctor(6, 1, 0); +x_11 = lean_alloc_ctor(7, 1, 0); lean_ctor_set(x_11, 0, x_1); x_12 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Tactic_elabSetOption___spec__3(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; @@ -21434,7 +21434,7 @@ x_30 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_30, 0, x_1); lean_ctor_set(x_30, 1, x_25); lean_ctor_set(x_30, 2, x_29); -x_31 = lean_alloc_ctor(4, 1, 0); +x_31 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_31, 0, x_30); x_32 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Tactic_elabSetOption___spec__3(x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); x_33 = !lean_is_exclusive(x_32); @@ -21829,7 +21829,7 @@ x_122 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_122, 0, x_1); lean_ctor_set(x_122, 1, x_117); lean_ctor_set(x_122, 2, x_121); -x_123 = lean_alloc_ctor(4, 1, 0); +x_123 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_123, 0, x_122); x_124 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Tactic_elabSetOption___spec__3(x_123, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_120); x_125 = lean_ctor_get(x_124, 1); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Config.c b/stage0/stdlib/Lean/Elab/Tactic/Config.c index a8aadf5f1dfb..7d5f2a08a3d7 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Config.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Config.c @@ -202,6 +202,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__247; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__14; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__47; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__commandConfigElab__1___closed__3; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__6; @@ -570,6 +571,7 @@ static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkCo static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__136; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__164; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__7; +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_checkConfigElab___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic___aux__Lean__Elab__Tactic__Config______macroRules__Lean__Elab__Tactic__configElab__1___closed__3; static lean_object* l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigElaborator___closed__168; @@ -2409,7 +2411,7 @@ static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("structInstFieldAbbrev", 21, 21); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -2428,6 +2430,26 @@ return x_5; static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__6() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFieldAbbrev", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews___spec__1___closed__1; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews___spec__1___closed__2; +x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__1; +x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; @@ -2437,19 +2459,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__6; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -2457,7 +2479,7 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11() { _start: { lean_object* x_1; @@ -2465,19 +2487,19 @@ x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews___spec__1___closed__1; x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_mkConfigItemViews___spec__1___closed__2; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__1; -x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13() { _start: { lean_object* x_1; @@ -2485,7 +2507,7 @@ x_1 = lean_mk_string_unchecked(":", 1, 1); return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14() { _start: { lean_object* x_1; @@ -2493,7 +2515,7 @@ x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15() { _start: { lean_object* x_1; @@ -2528,7 +2550,7 @@ x_22 = lean_st_ref_get(x_14, x_15); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; x_24 = lean_ctor_get(x_22, 1); x_25 = lean_ctor_get(x_22, 0); lean_dec(x_25); @@ -2546,9 +2568,9 @@ lean_ctor_set(x_29, 1, x_27); lean_ctor_set(x_29, 2, x_28); x_30 = lean_array_size(x_5); x_31 = 0; -x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__1(x_32, x_30, x_31, x_5); -x_34 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; +x_34 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; x_35 = l_Lean_Syntax_SepArray_ofElems(x_34, x_33); lean_dec(x_33); x_36 = l_Array_append___rarg(x_28, x_35); @@ -2558,324 +2580,339 @@ x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_21); lean_ctor_set(x_37, 1, x_27); lean_ctor_set(x_37, 2, x_36); -x_38 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_38 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_21); +x_39 = l_Lean_Syntax_node1(x_21, x_38, x_37); +x_40 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_29); lean_inc(x_21); -x_39 = l_Lean_Syntax_node1(x_21, x_38, x_29); -x_40 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_41 = l_Lean_Syntax_node1(x_21, x_40, x_29); +x_42 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_21); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_21); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_box(0); -x_43 = l_Lean_mkCIdentFrom(x_42, x_3, x_20); +x_43 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_43, 0, x_21); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_box(0); +x_45 = l_Lean_mkCIdentFrom(x_44, x_3, x_20); lean_inc(x_21); -x_44 = l_Lean_Syntax_node2(x_21, x_27, x_41, x_43); -x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_46 = l_Lean_Syntax_node2(x_21, x_27, x_43, x_45); +x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; lean_inc(x_21); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_21); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_48 = l_Lean_Syntax_node6(x_21, x_47, x_22, x_29, x_37, x_39, x_44, x_46); -x_49 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_50 = l_Lean_NameSet_empty; -x_51 = lean_box(0); -x_52 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_50, x_49, x_51, x_9, x_10, x_11, x_12, x_13, x_14, x_24); -return x_52; +x_48 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_48, 0, x_21); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_50 = l_Lean_Syntax_node6(x_21, x_49, x_22, x_29, x_39, x_41, x_46, x_48); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = l_Lean_NameSet_empty; +x_53 = lean_box(0); +x_54 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_52, x_51, x_53, x_9, x_10, x_11, x_12, x_13, x_14, x_24); +return x_54; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; size_t x_59; size_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_53 = lean_ctor_get(x_22, 1); -lean_inc(x_53); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_55 = lean_ctor_get(x_22, 1); +lean_inc(x_55); lean_dec(x_22); -x_54 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +x_56 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; lean_inc(x_21); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_21); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -x_57 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_21); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +x_59 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; lean_inc(x_21); -x_58 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_58, 0, x_21); -lean_ctor_set(x_58, 1, x_56); -lean_ctor_set(x_58, 2, x_57); -x_59 = lean_array_size(x_5); -x_60 = 0; -x_61 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__1(x_61, x_59, x_60, x_5); -x_63 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_64 = l_Lean_Syntax_SepArray_ofElems(x_63, x_62); -lean_dec(x_62); -x_65 = l_Array_append___rarg(x_57, x_64); +x_60 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_60, 0, x_21); +lean_ctor_set(x_60, 1, x_58); +lean_ctor_set(x_60, 2, x_59); +x_61 = lean_array_size(x_5); +x_62 = 0; +x_63 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__1(x_63, x_61, x_62, x_5); +x_65 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_66 = l_Lean_Syntax_SepArray_ofElems(x_65, x_64); lean_dec(x_64); +x_67 = l_Array_append___rarg(x_59, x_66); +lean_dec(x_66); lean_inc(x_21); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_21); -lean_ctor_set(x_66, 1, x_56); -lean_ctor_set(x_66, 2, x_65); -x_67 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_58); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_21); +lean_ctor_set(x_68, 1, x_58); +lean_ctor_set(x_68, 2, x_67); +x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; lean_inc(x_21); -x_68 = l_Lean_Syntax_node1(x_21, x_67, x_58); -x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_70 = l_Lean_Syntax_node1(x_21, x_69, x_68); +x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_60); lean_inc(x_21); -x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_21); -lean_ctor_set(x_70, 1, x_69); -x_71 = lean_box(0); -x_72 = l_Lean_mkCIdentFrom(x_71, x_3, x_20); +x_72 = l_Lean_Syntax_node1(x_21, x_71, x_60); +x_73 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_21); -x_73 = l_Lean_Syntax_node2(x_21, x_56, x_70, x_72); -x_74 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_74 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_74, 0, x_21); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_box(0); +x_76 = l_Lean_mkCIdentFrom(x_75, x_3, x_20); lean_inc(x_21); -x_75 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_75, 0, x_21); -lean_ctor_set(x_75, 1, x_74); -x_76 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_77 = l_Lean_Syntax_node6(x_21, x_76, x_55, x_58, x_66, x_68, x_73, x_75); -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_77); -x_79 = l_Lean_NameSet_empty; -x_80 = lean_box(0); -x_81 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_79, x_78, x_80, x_9, x_10, x_11, x_12, x_13, x_14, x_53); -return x_81; -} -} -else -{ -uint8_t x_82; -x_82 = !lean_is_exclusive(x_7); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_83 = lean_ctor_get(x_7, 0); -x_84 = lean_ctor_get(x_13, 5); -x_85 = 0; -x_86 = l_Lean_SourceInfo_fromRef(x_84, x_85); -x_87 = lean_st_ref_get(x_14, x_15); -x_88 = !lean_is_exclusive(x_87); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; size_t x_97; size_t x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_89 = lean_ctor_get(x_87, 1); -x_90 = lean_ctor_get(x_87, 0); -lean_dec(x_90); -x_91 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_86); -lean_ctor_set_tag(x_87, 2); -lean_ctor_set(x_87, 1, x_91); -lean_ctor_set(x_87, 0, x_86); -x_92 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_86); -x_93 = l_Lean_Syntax_node1(x_86, x_92, x_83); -x_94 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_86); -x_95 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_95, 0, x_86); -lean_ctor_set(x_95, 1, x_94); -lean_inc(x_86); -x_96 = l_Lean_Syntax_node2(x_86, x_92, x_93, x_95); -x_97 = lean_array_size(x_5); -x_98 = 0; -x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_100 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_99, x_97, x_98, x_5); -x_101 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_102 = l_Lean_Syntax_SepArray_ofElems(x_101, x_100); -lean_dec(x_100); -x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_104 = l_Array_append___rarg(x_103, x_102); -lean_dec(x_102); -lean_inc(x_86); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_86); -lean_ctor_set(x_105, 1, x_92); -lean_ctor_set(x_105, 2, x_104); -lean_inc(x_86); -x_106 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_106, 0, x_86); -lean_ctor_set(x_106, 1, x_92); -lean_ctor_set(x_106, 2, x_103); -x_107 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_86); -x_108 = l_Lean_Syntax_node1(x_86, x_107, x_106); -x_109 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_86); -x_110 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_110, 0, x_86); -lean_ctor_set(x_110, 1, x_109); -x_111 = lean_box(0); -x_112 = l_Lean_mkCIdentFrom(x_111, x_3, x_85); -lean_inc(x_86); -x_113 = l_Lean_Syntax_node2(x_86, x_92, x_110, x_112); -x_114 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_86); -x_115 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_115, 0, x_86); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_117 = l_Lean_Syntax_node6(x_86, x_116, x_87, x_96, x_105, x_108, x_113, x_115); -lean_ctor_set(x_7, 0, x_117); -x_118 = l_Lean_NameSet_empty; -x_119 = lean_box(0); -x_120 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_118, x_7, x_119, x_9, x_10, x_11, x_12, x_13, x_14, x_89); -return x_120; +x_77 = l_Lean_Syntax_node2(x_21, x_58, x_74, x_76); +x_78 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_21); +x_79 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_79, 0, x_21); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_81 = l_Lean_Syntax_node6(x_21, x_80, x_57, x_60, x_70, x_72, x_77, x_79); +x_82 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_82, 0, x_81); +x_83 = l_Lean_NameSet_empty; +x_84 = lean_box(0); +x_85 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_83, x_82, x_84, x_9, x_10, x_11, x_12, x_13, x_14, x_55); +return x_85; +} } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; size_t x_129; size_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_121 = lean_ctor_get(x_87, 1); -lean_inc(x_121); -lean_dec(x_87); -x_122 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_86); -x_123 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_123, 0, x_86); -lean_ctor_set(x_123, 1, x_122); -x_124 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_86); -x_125 = l_Lean_Syntax_node1(x_86, x_124, x_83); -x_126 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_86); -x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_86); -lean_ctor_set(x_127, 1, x_126); -lean_inc(x_86); -x_128 = l_Lean_Syntax_node2(x_86, x_124, x_125, x_127); -x_129 = lean_array_size(x_5); -x_130 = 0; -x_131 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_132 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_131, x_129, x_130, x_5); -x_133 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_134 = l_Lean_Syntax_SepArray_ofElems(x_133, x_132); -lean_dec(x_132); -x_135 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_136 = l_Array_append___rarg(x_135, x_134); -lean_dec(x_134); -lean_inc(x_86); -x_137 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_137, 0, x_86); -lean_ctor_set(x_137, 1, x_124); -lean_ctor_set(x_137, 2, x_136); -lean_inc(x_86); -x_138 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_138, 0, x_86); -lean_ctor_set(x_138, 1, x_124); -lean_ctor_set(x_138, 2, x_135); +uint8_t x_86; +x_86 = !lean_is_exclusive(x_7); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_87 = lean_ctor_get(x_7, 0); +x_88 = lean_ctor_get(x_13, 5); +x_89 = 0; +x_90 = l_Lean_SourceInfo_fromRef(x_88, x_89); +x_91 = lean_st_ref_get(x_14, x_15); +x_92 = !lean_is_exclusive(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; size_t x_101; size_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_93 = lean_ctor_get(x_91, 1); +x_94 = lean_ctor_get(x_91, 0); +lean_dec(x_94); +x_95 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_90); +lean_ctor_set_tag(x_91, 2); +lean_ctor_set(x_91, 1, x_95); +lean_ctor_set(x_91, 0, x_90); +x_96 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_90); +x_97 = l_Lean_Syntax_node1(x_90, x_96, x_87); +x_98 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_90); +x_99 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_99, 0, x_90); +lean_ctor_set(x_99, 1, x_98); +lean_inc(x_90); +x_100 = l_Lean_Syntax_node2(x_90, x_96, x_97, x_99); +x_101 = lean_array_size(x_5); +x_102 = 0; +x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_104 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_103, x_101, x_102, x_5); +x_105 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_106 = l_Lean_Syntax_SepArray_ofElems(x_105, x_104); +lean_dec(x_104); +x_107 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_108 = l_Array_append___rarg(x_107, x_106); +lean_dec(x_106); +lean_inc(x_90); +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_90); +lean_ctor_set(x_109, 1, x_96); +lean_ctor_set(x_109, 2, x_108); +x_110 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_90); +x_111 = l_Lean_Syntax_node1(x_90, x_110, x_109); +lean_inc(x_90); +x_112 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_112, 0, x_90); +lean_ctor_set(x_112, 1, x_96); +lean_ctor_set(x_112, 2, x_107); +x_113 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_90); +x_114 = l_Lean_Syntax_node1(x_90, x_113, x_112); +x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_90); +x_116 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_116, 0, x_90); +lean_ctor_set(x_116, 1, x_115); +x_117 = lean_box(0); +x_118 = l_Lean_mkCIdentFrom(x_117, x_3, x_89); +lean_inc(x_90); +x_119 = l_Lean_Syntax_node2(x_90, x_96, x_116, x_118); +x_120 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_90); +x_121 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_121, 0, x_90); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_123 = l_Lean_Syntax_node6(x_90, x_122, x_91, x_100, x_111, x_114, x_119, x_121); +lean_ctor_set(x_7, 0, x_123); +x_124 = l_Lean_NameSet_empty; +x_125 = lean_box(0); +x_126 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_124, x_7, x_125, x_9, x_10, x_11, x_12, x_13, x_14, x_93); +return x_126; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; size_t x_135; size_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_127 = lean_ctor_get(x_91, 1); +lean_inc(x_127); +lean_dec(x_91); +x_128 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_90); +x_129 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_129, 0, x_90); +lean_ctor_set(x_129, 1, x_128); +x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_90); +x_131 = l_Lean_Syntax_node1(x_90, x_130, x_87); +x_132 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_90); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_90); +lean_ctor_set(x_133, 1, x_132); +lean_inc(x_90); +x_134 = l_Lean_Syntax_node2(x_90, x_130, x_131, x_133); +x_135 = lean_array_size(x_5); +x_136 = 0; +x_137 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_138 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_137, x_135, x_136, x_5); x_139 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_86); -x_140 = l_Lean_Syntax_node1(x_86, x_139, x_138); -x_141 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_86); -x_142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_142, 0, x_86); -lean_ctor_set(x_142, 1, x_141); -x_143 = lean_box(0); -x_144 = l_Lean_mkCIdentFrom(x_143, x_3, x_85); -lean_inc(x_86); -x_145 = l_Lean_Syntax_node2(x_86, x_124, x_142, x_144); -x_146 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_86); -x_147 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_147, 0, x_86); -lean_ctor_set(x_147, 1, x_146); -x_148 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_149 = l_Lean_Syntax_node6(x_86, x_148, x_123, x_128, x_137, x_140, x_145, x_147); -lean_ctor_set(x_7, 0, x_149); -x_150 = l_Lean_NameSet_empty; +x_140 = l_Lean_Syntax_SepArray_ofElems(x_139, x_138); +lean_dec(x_138); +x_141 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_142 = l_Array_append___rarg(x_141, x_140); +lean_dec(x_140); +lean_inc(x_90); +x_143 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_143, 0, x_90); +lean_ctor_set(x_143, 1, x_130); +lean_ctor_set(x_143, 2, x_142); +x_144 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_90); +x_145 = l_Lean_Syntax_node1(x_90, x_144, x_143); +lean_inc(x_90); +x_146 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_146, 0, x_90); +lean_ctor_set(x_146, 1, x_130); +lean_ctor_set(x_146, 2, x_141); +x_147 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_90); +x_148 = l_Lean_Syntax_node1(x_90, x_147, x_146); +x_149 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_90); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_90); +lean_ctor_set(x_150, 1, x_149); x_151 = lean_box(0); -x_152 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_150, x_7, x_151, x_9, x_10, x_11, x_12, x_13, x_14, x_121); -return x_152; +x_152 = l_Lean_mkCIdentFrom(x_151, x_3, x_89); +lean_inc(x_90); +x_153 = l_Lean_Syntax_node2(x_90, x_130, x_150, x_152); +x_154 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_90); +x_155 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_155, 0, x_90); +lean_ctor_set(x_155, 1, x_154); +x_156 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_157 = l_Lean_Syntax_node6(x_90, x_156, x_129, x_134, x_145, x_148, x_153, x_155); +lean_ctor_set(x_7, 0, x_157); +x_158 = l_Lean_NameSet_empty; +x_159 = lean_box(0); +x_160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_158, x_7, x_159, x_9, x_10, x_11, x_12, x_13, x_14, x_127); +return x_160; } } else { -lean_object* x_153; lean_object* x_154; uint8_t x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; size_t x_167; size_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_153 = lean_ctor_get(x_7, 0); -lean_inc(x_153); +lean_object* x_161; lean_object* x_162; uint8_t x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; size_t x_175; size_t x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_161 = lean_ctor_get(x_7, 0); +lean_inc(x_161); lean_dec(x_7); -x_154 = lean_ctor_get(x_13, 5); -x_155 = 0; -x_156 = l_Lean_SourceInfo_fromRef(x_154, x_155); -x_157 = lean_st_ref_get(x_14, x_15); -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_159 = x_157; +x_162 = lean_ctor_get(x_13, 5); +x_163 = 0; +x_164 = l_Lean_SourceInfo_fromRef(x_162, x_163); +x_165 = lean_st_ref_get(x_14, x_15); +x_166 = lean_ctor_get(x_165, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_167 = x_165; } else { - lean_dec_ref(x_157); - x_159 = lean_box(0); + lean_dec_ref(x_165); + x_167 = lean_box(0); } -x_160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_156); -if (lean_is_scalar(x_159)) { - x_161 = lean_alloc_ctor(2, 2, 0); +x_168 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_164); +if (lean_is_scalar(x_167)) { + x_169 = lean_alloc_ctor(2, 2, 0); } else { - x_161 = x_159; - lean_ctor_set_tag(x_161, 2); + x_169 = x_167; + lean_ctor_set_tag(x_169, 2); } -lean_ctor_set(x_161, 0, x_156); -lean_ctor_set(x_161, 1, x_160); -x_162 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_156); -x_163 = l_Lean_Syntax_node1(x_156, x_162, x_153); -x_164 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_156); -x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_156); -lean_ctor_set(x_165, 1, x_164); -lean_inc(x_156); -x_166 = l_Lean_Syntax_node2(x_156, x_162, x_163, x_165); -x_167 = lean_array_size(x_5); -x_168 = 0; -x_169 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_170 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_169, x_167, x_168, x_5); -x_171 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_172 = l_Lean_Syntax_SepArray_ofElems(x_171, x_170); -lean_dec(x_170); -x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_174 = l_Array_append___rarg(x_173, x_172); -lean_dec(x_172); -lean_inc(x_156); -x_175 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_175, 0, x_156); -lean_ctor_set(x_175, 1, x_162); -lean_ctor_set(x_175, 2, x_174); -lean_inc(x_156); -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_156); -lean_ctor_set(x_176, 1, x_162); -lean_ctor_set(x_176, 2, x_173); -x_177 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_156); -x_178 = l_Lean_Syntax_node1(x_156, x_177, x_176); -x_179 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_156); -x_180 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_180, 0, x_156); -lean_ctor_set(x_180, 1, x_179); -x_181 = lean_box(0); -x_182 = l_Lean_mkCIdentFrom(x_181, x_3, x_155); -lean_inc(x_156); -x_183 = l_Lean_Syntax_node2(x_156, x_162, x_180, x_182); -x_184 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_156); -x_185 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_185, 0, x_156); -lean_ctor_set(x_185, 1, x_184); -x_186 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_187 = l_Lean_Syntax_node6(x_156, x_186, x_161, x_166, x_175, x_178, x_183, x_185); -x_188 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_188, 0, x_187); -x_189 = l_Lean_NameSet_empty; -x_190 = lean_box(0); -x_191 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_189, x_188, x_190, x_9, x_10, x_11, x_12, x_13, x_14, x_158); -return x_191; +lean_ctor_set(x_169, 0, x_164); +lean_ctor_set(x_169, 1, x_168); +x_170 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_164); +x_171 = l_Lean_Syntax_node1(x_164, x_170, x_161); +x_172 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_164); +x_173 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_173, 0, x_164); +lean_ctor_set(x_173, 1, x_172); +lean_inc(x_164); +x_174 = l_Lean_Syntax_node2(x_164, x_170, x_171, x_173); +x_175 = lean_array_size(x_5); +x_176 = 0; +x_177 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_178 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__2(x_177, x_175, x_176, x_5); +x_179 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_180 = l_Lean_Syntax_SepArray_ofElems(x_179, x_178); +lean_dec(x_178); +x_181 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_182 = l_Array_append___rarg(x_181, x_180); +lean_dec(x_180); +lean_inc(x_164); +x_183 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_183, 0, x_164); +lean_ctor_set(x_183, 1, x_170); +lean_ctor_set(x_183, 2, x_182); +x_184 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_164); +x_185 = l_Lean_Syntax_node1(x_164, x_184, x_183); +lean_inc(x_164); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_164); +lean_ctor_set(x_186, 1, x_170); +lean_ctor_set(x_186, 2, x_181); +x_187 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_164); +x_188 = l_Lean_Syntax_node1(x_164, x_187, x_186); +x_189 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_164); +x_190 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_190, 0, x_164); +lean_ctor_set(x_190, 1, x_189); +x_191 = lean_box(0); +x_192 = l_Lean_mkCIdentFrom(x_191, x_3, x_163); +lean_inc(x_164); +x_193 = l_Lean_Syntax_node2(x_164, x_170, x_190, x_192); +x_194 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_164); +x_195 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_195, 0, x_164); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_197 = l_Lean_Syntax_node6(x_164, x_196, x_169, x_174, x_185, x_188, x_193, x_195); +x_198 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_198, 0, x_197); +x_199 = l_Lean_NameSet_empty; +x_200 = lean_box(0); +x_201 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1(x_1, x_2, x_8, x_4, x_199, x_198, x_200, x_9, x_10, x_11, x_12, x_13, x_14, x_166); +return x_201; } } } @@ -3360,7 +3397,7 @@ lean_inc(x_18); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_18); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_25 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_18); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_18); @@ -3414,7 +3451,7 @@ x_42 = lean_st_ref_get(x_12, x_21); x_43 = !lean_is_exclusive(x_42); if (x_43 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; x_44 = lean_ctor_get(x_42, 0); lean_dec(x_44); x_45 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; @@ -3429,7 +3466,7 @@ lean_ctor_set(x_48, 0, x_18); lean_ctor_set(x_48, 1, x_47); lean_inc(x_18); x_49 = l_Lean_Syntax_node3(x_18, x_29, x_34, x_48, x_41); -x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +x_50 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; lean_inc(x_18); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_18); @@ -3442,311 +3479,323 @@ x_54 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_54, 0, x_18); lean_ctor_set(x_54, 1, x_29); lean_ctor_set(x_54, 2, x_53); -x_55 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_55 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; lean_inc(x_54); lean_inc(x_18); x_56 = l_Lean_Syntax_node1(x_18, x_55, x_54); +x_57 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_18); -x_57 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); -x_58 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_58 = l_Lean_Syntax_node1(x_18, x_57, x_54); lean_inc(x_18); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_18); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_61 = l_Lean_Syntax_node6(x_18, x_60, x_46, x_52, x_54, x_56, x_57, x_59); -lean_ctor_set(x_5, 0, x_61); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_4); -lean_ctor_set(x_62, 1, x_5); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_3); -lean_ctor_set(x_63, 1, x_62); -x_64 = lean_box(0); +x_59 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); +x_60 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_18); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_18); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_63 = l_Lean_Syntax_node6(x_18, x_62, x_46, x_52, x_56, x_58, x_59, x_61); +lean_ctor_set(x_5, 0, x_63); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_4); +lean_ctor_set(x_64, 1, x_5); x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_63); -lean_ctor_set(x_42, 0, x_65); +lean_ctor_set(x_65, 0, x_3); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_box(0); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_65); +lean_ctor_set(x_42, 0, x_67); return x_42; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_66 = lean_ctor_get(x_42, 1); -lean_inc(x_66); +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_68 = lean_ctor_get(x_42, 1); +lean_inc(x_68); lean_dec(x_42); -x_67 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_18); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_18); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; +x_69 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; lean_inc(x_18); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_18); lean_ctor_set(x_70, 1, x_69); +x_71 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; lean_inc(x_18); -x_71 = l_Lean_Syntax_node3(x_18, x_29, x_34, x_70, x_41); -x_72 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_18); +lean_ctor_set(x_72, 1, x_71); lean_inc(x_18); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_18); -lean_ctor_set(x_73, 1, x_72); +x_73 = l_Lean_Syntax_node3(x_18, x_29, x_34, x_72, x_41); +x_74 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; lean_inc(x_18); -x_74 = l_Lean_Syntax_node2(x_18, x_29, x_71, x_73); -x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_75 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_75, 0, x_18); +lean_ctor_set(x_75, 1, x_74); lean_inc(x_18); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_18); -lean_ctor_set(x_76, 1, x_29); -lean_ctor_set(x_76, 2, x_75); -x_77 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_76); +x_76 = l_Lean_Syntax_node2(x_18, x_29, x_73, x_75); +x_77 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; lean_inc(x_18); -x_78 = l_Lean_Syntax_node1(x_18, x_77, x_76); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_18); +lean_ctor_set(x_78, 1, x_29); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_78); lean_inc(x_18); -x_79 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); -x_80 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_80 = l_Lean_Syntax_node1(x_18, x_79, x_78); +x_81 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_18); -x_81 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_81, 0, x_18); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_83 = l_Lean_Syntax_node6(x_18, x_82, x_68, x_74, x_76, x_78, x_79, x_81); -lean_ctor_set(x_5, 0, x_83); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_4); -lean_ctor_set(x_84, 1, x_5); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_3); +x_82 = l_Lean_Syntax_node1(x_18, x_81, x_78); +lean_inc(x_18); +x_83 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); +x_84 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_18); +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_18); lean_ctor_set(x_85, 1, x_84); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_85); +x_86 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_87 = l_Lean_Syntax_node6(x_18, x_86, x_70, x_76, x_80, x_82, x_83, x_85); +lean_ctor_set(x_5, 0, x_87); x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_66); -return x_88; +lean_ctor_set(x_88, 0, x_4); +lean_ctor_set(x_88, 1, x_5); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_3); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_89); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_68); +return x_92; } } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_89 = lean_ctor_get(x_5, 0); -lean_inc(x_89); +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_93 = lean_ctor_get(x_5, 0); +lean_inc(x_93); lean_dec(x_5); -x_90 = lean_st_ref_get(x_12, x_21); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_92 = x_90; +x_94 = lean_st_ref_get(x_12, x_21); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_96 = x_94; } else { - lean_dec_ref(x_90); - x_92 = lean_box(0); + lean_dec_ref(x_94); + x_96 = lean_box(0); } -x_93 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; lean_inc(x_18); -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_18); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_18); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; lean_inc(x_18); -x_96 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_96, 0, x_18); -lean_ctor_set(x_96, 1, x_95); +x_100 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_100, 0, x_18); +lean_ctor_set(x_100, 1, x_99); lean_inc(x_18); -x_97 = l_Lean_Syntax_node3(x_18, x_29, x_34, x_96, x_89); -x_98 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +x_101 = l_Lean_Syntax_node3(x_18, x_29, x_34, x_100, x_93); +x_102 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; lean_inc(x_18); -x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_18); -lean_ctor_set(x_99, 1, x_98); +x_103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_103, 0, x_18); +lean_ctor_set(x_103, 1, x_102); lean_inc(x_18); -x_100 = l_Lean_Syntax_node2(x_18, x_29, x_97, x_99); -x_101 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_104 = l_Lean_Syntax_node2(x_18, x_29, x_101, x_103); +x_105 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; lean_inc(x_18); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_18); -lean_ctor_set(x_102, 1, x_29); -lean_ctor_set(x_102, 2, x_101); -x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_102); +x_106 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_106, 0, x_18); +lean_ctor_set(x_106, 1, x_29); +lean_ctor_set(x_106, 2, x_105); +x_107 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_106); lean_inc(x_18); -x_104 = l_Lean_Syntax_node1(x_18, x_103, x_102); +x_108 = l_Lean_Syntax_node1(x_18, x_107, x_106); +x_109 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_18); -x_105 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); -x_106 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_110 = l_Lean_Syntax_node1(x_18, x_109, x_106); lean_inc(x_18); -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_18); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_109 = l_Lean_Syntax_node6(x_18, x_108, x_94, x_100, x_102, x_104, x_105, x_107); -x_110 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_110, 0, x_109); -x_111 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_111, 0, x_4); -lean_ctor_set(x_111, 1, x_110); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_3); -lean_ctor_set(x_112, 1, x_111); -x_113 = lean_box(0); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_112); -if (lean_is_scalar(x_92)) { - x_115 = lean_alloc_ctor(0, 2, 0); +x_111 = l_Lean_Syntax_node2(x_18, x_29, x_26, x_28); +x_112 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_18); +x_113 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_113, 0, x_18); +lean_ctor_set(x_113, 1, x_112); +x_114 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_115 = l_Lean_Syntax_node6(x_18, x_114, x_98, x_104, x_108, x_110, x_111, x_113); +x_116 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_116, 0, x_115); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_4); +lean_ctor_set(x_117, 1, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_3); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_box(0); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_118); +if (lean_is_scalar(x_96)) { + x_121 = lean_alloc_ctor(0, 2, 0); } else { - x_115 = x_92; + x_121 = x_96; } -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_91); -return x_115; +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_95); +return x_121; } } } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_116 = lean_ctor_get(x_19, 1); -lean_inc(x_116); -lean_dec(x_19); -x_117 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__3; -lean_inc(x_18); -x_118 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_118, 0, x_18); -lean_ctor_set(x_118, 1, x_117); -x_119 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_18); -x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_18); -lean_ctor_set(x_120, 1, x_119); -x_121 = lean_box(0); -x_122 = l_Lean_mkCIdentFrom(x_121, x_2, x_17); -x_123 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_122 = lean_ctor_get(x_19, 1); lean_inc(x_122); +lean_dec(x_19); +x_123 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__3; lean_inc(x_18); -x_124 = l_Lean_Syntax_node1(x_18, x_123, x_122); -x_125 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__25; +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_18); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_18); x_126 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_126, 0, x_18); lean_ctor_set(x_126, 1, x_125); -x_127 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__2; -lean_inc(x_120); +x_127 = lean_box(0); +x_128 = l_Lean_mkCIdentFrom(x_127, x_2, x_17); +x_129 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_128); lean_inc(x_18); -x_128 = l_Lean_Syntax_node5(x_18, x_127, x_118, x_14, x_120, x_124, x_126); +x_130 = l_Lean_Syntax_node1(x_18, x_129, x_128); +x_131 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__25; +lean_inc(x_18); +x_132 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_132, 0, x_18); +lean_ctor_set(x_132, 1, x_131); +x_133 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__2; +lean_inc(x_126); +lean_inc(x_18); +x_134 = l_Lean_Syntax_node5(x_18, x_133, x_124, x_14, x_126, x_130, x_132); if (lean_obj_tag(x_5) == 0) { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -lean_dec(x_122); -lean_dec(x_120); +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_dec(x_128); +lean_dec(x_126); lean_dec(x_18); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_128); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_4); -lean_ctor_set(x_130, 1, x_129); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_3); -lean_ctor_set(x_131, 1, x_130); -x_132 = lean_box(0); -x_133 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_131); -x_134 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_116); -return x_134; -} -else -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_135 = lean_ctor_get(x_5, 0); -lean_inc(x_135); +x_135 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_135, 0, x_134); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_4); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_3); +lean_ctor_set(x_137, 1, x_136); +x_138 = lean_box(0); +x_139 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_137); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_139); +lean_ctor_set(x_140, 1, x_122); +return x_140; +} +else +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_141 = lean_ctor_get(x_5, 0); +lean_inc(x_141); if (lean_is_exclusive(x_5)) { lean_ctor_release(x_5, 0); - x_136 = x_5; + x_142 = x_5; } else { lean_dec_ref(x_5); - x_136 = lean_box(0); -} -x_137 = lean_st_ref_get(x_12, x_116); -x_138 = lean_ctor_get(x_137, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_137)) { - lean_ctor_release(x_137, 0); - lean_ctor_release(x_137, 1); - x_139 = x_137; + x_142 = lean_box(0); +} +x_143 = lean_st_ref_get(x_12, x_122); +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + lean_ctor_release(x_143, 1); + x_145 = x_143; } else { - lean_dec_ref(x_137); - x_139 = lean_box(0); + lean_dec_ref(x_143); + x_145 = lean_box(0); } -x_140 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +x_146 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; lean_inc(x_18); -x_141 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_141, 0, x_18); -lean_ctor_set(x_141, 1, x_140); -x_142 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; +x_147 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_147, 0, x_18); +lean_ctor_set(x_147, 1, x_146); +x_148 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4___closed__4; lean_inc(x_18); -x_143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_143, 0, x_18); -lean_ctor_set(x_143, 1, x_142); +x_149 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_149, 0, x_18); +lean_ctor_set(x_149, 1, x_148); lean_inc(x_18); -x_144 = l_Lean_Syntax_node3(x_18, x_123, x_128, x_143, x_135); -x_145 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +x_150 = l_Lean_Syntax_node3(x_18, x_129, x_134, x_149, x_141); +x_151 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; lean_inc(x_18); -x_146 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_146, 0, x_18); -lean_ctor_set(x_146, 1, x_145); +x_152 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_152, 0, x_18); +lean_ctor_set(x_152, 1, x_151); lean_inc(x_18); -x_147 = l_Lean_Syntax_node2(x_18, x_123, x_144, x_146); -x_148 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_153 = l_Lean_Syntax_node2(x_18, x_129, x_150, x_152); +x_154 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; lean_inc(x_18); -x_149 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_149, 0, x_18); -lean_ctor_set(x_149, 1, x_123); -lean_ctor_set(x_149, 2, x_148); -x_150 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_149); +x_155 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_155, 0, x_18); +lean_ctor_set(x_155, 1, x_129); +lean_ctor_set(x_155, 2, x_154); +x_156 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_155); lean_inc(x_18); -x_151 = l_Lean_Syntax_node1(x_18, x_150, x_149); +x_157 = l_Lean_Syntax_node1(x_18, x_156, x_155); +x_158 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_18); -x_152 = l_Lean_Syntax_node2(x_18, x_123, x_120, x_122); -x_153 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_159 = l_Lean_Syntax_node1(x_18, x_158, x_155); lean_inc(x_18); -x_154 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_154, 0, x_18); -lean_ctor_set(x_154, 1, x_153); -x_155 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_156 = l_Lean_Syntax_node6(x_18, x_155, x_141, x_147, x_149, x_151, x_152, x_154); -if (lean_is_scalar(x_136)) { - x_157 = lean_alloc_ctor(1, 1, 0); +x_160 = l_Lean_Syntax_node2(x_18, x_129, x_126, x_128); +x_161 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_18); +x_162 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_162, 0, x_18); +lean_ctor_set(x_162, 1, x_161); +x_163 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_164 = l_Lean_Syntax_node6(x_18, x_163, x_147, x_153, x_157, x_159, x_160, x_162); +if (lean_is_scalar(x_142)) { + x_165 = lean_alloc_ctor(1, 1, 0); } else { - x_157 = x_136; -} -lean_ctor_set(x_157, 0, x_156); -x_158 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_158, 0, x_4); -lean_ctor_set(x_158, 1, x_157); -x_159 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_159, 0, x_3); -lean_ctor_set(x_159, 1, x_158); -x_160 = lean_box(0); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_159); -if (lean_is_scalar(x_139)) { - x_162 = lean_alloc_ctor(0, 2, 0); + x_165 = x_142; +} +lean_ctor_set(x_165, 0, x_164); +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_4); +lean_ctor_set(x_166, 1, x_165); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_3); +lean_ctor_set(x_167, 1, x_166); +x_168 = lean_box(0); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_167); +if (lean_is_scalar(x_145)) { + x_170 = lean_alloc_ctor(0, 2, 0); } else { - x_162 = x_139; + x_170 = x_145; } -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_138); -return x_162; +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_144); +return x_170; } } } @@ -3859,375 +3908,377 @@ x_74 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Ela x_75 = lean_name_eq(x_73, x_74); if (x_75 == 0) { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_inc(x_73); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_73); lean_inc(x_2); lean_inc(x_4); -lean_inc(x_73); lean_inc(x_71); -x_76 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_76, 0, x_71); -lean_ctor_set(x_76, 1, x_73); -lean_ctor_set(x_76, 2, x_4); -lean_ctor_set(x_76, 3, x_2); -x_77 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_76, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_78 = lean_ctor_get(x_77, 1); -lean_inc(x_78); -lean_dec(x_77); -x_79 = lean_ctor_get(x_15, 0); +x_77 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_77, 0, x_71); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set(x_77, 2, x_4); +lean_ctor_set(x_77, 3, x_2); +x_78 = l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(x_77, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_79 = lean_ctor_get(x_78, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_15, 1); +lean_dec(x_78); +x_80 = lean_ctor_get(x_15, 0); lean_inc(x_80); -x_81 = lean_ctor_get(x_15, 2); +x_81 = lean_ctor_get(x_15, 1); lean_inc(x_81); -x_82 = lean_ctor_get(x_15, 3); +x_82 = lean_ctor_get(x_15, 2); lean_inc(x_82); -x_83 = lean_ctor_get(x_15, 4); +x_83 = lean_ctor_get(x_15, 3); lean_inc(x_83); -x_84 = lean_ctor_get(x_15, 5); +x_84 = lean_ctor_get(x_15, 4); lean_inc(x_84); -x_85 = lean_ctor_get(x_15, 6); +x_85 = lean_ctor_get(x_15, 5); lean_inc(x_85); -x_86 = lean_ctor_get(x_15, 7); +x_86 = lean_ctor_get(x_15, 6); lean_inc(x_86); -x_87 = lean_ctor_get(x_15, 8); +x_87 = lean_ctor_get(x_15, 7); lean_inc(x_87); -x_88 = lean_ctor_get(x_15, 9); +x_88 = lean_ctor_get(x_15, 8); lean_inc(x_88); -x_89 = lean_ctor_get(x_15, 10); +x_89 = lean_ctor_get(x_15, 9); lean_inc(x_89); -x_90 = lean_ctor_get_uint8(x_15, sizeof(void*)*12); -x_91 = lean_ctor_get(x_15, 11); -lean_inc(x_91); -x_92 = lean_ctor_get_uint8(x_15, sizeof(void*)*12 + 1); -x_93 = l_Lean_replaceRef(x_71, x_84); -lean_dec(x_84); -x_94 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_94, 0, x_79); -lean_ctor_set(x_94, 1, x_80); -lean_ctor_set(x_94, 2, x_81); -lean_ctor_set(x_94, 3, x_82); -lean_ctor_set(x_94, 4, x_83); -lean_ctor_set(x_94, 5, x_93); -lean_ctor_set(x_94, 6, x_85); -lean_ctor_set(x_94, 7, x_86); -lean_ctor_set(x_94, 8, x_87); -lean_ctor_set(x_94, 9, x_88); -lean_ctor_set(x_94, 10, x_89); -lean_ctor_set(x_94, 11, x_91); -lean_ctor_set_uint8(x_94, sizeof(void*)*12, x_90); -lean_ctor_set_uint8(x_94, sizeof(void*)*12 + 1, x_92); +x_90 = lean_ctor_get(x_15, 10); +lean_inc(x_90); +x_91 = lean_ctor_get_uint8(x_15, sizeof(void*)*12); +x_92 = lean_ctor_get(x_15, 11); +lean_inc(x_92); +x_93 = lean_ctor_get_uint8(x_15, sizeof(void*)*12 + 1); +x_94 = l_Lean_replaceRef(x_71, x_85); +lean_dec(x_85); +x_95 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_95, 0, x_80); +lean_ctor_set(x_95, 1, x_81); +lean_ctor_set(x_95, 2, x_82); +lean_ctor_set(x_95, 3, x_83); +lean_ctor_set(x_95, 4, x_84); +lean_ctor_set(x_95, 5, x_94); +lean_ctor_set(x_95, 6, x_86); +lean_ctor_set(x_95, 7, x_87); +lean_ctor_set(x_95, 8, x_88); +lean_ctor_set(x_95, 9, x_89); +lean_ctor_set(x_95, 10, x_90); +lean_ctor_set(x_95, 11, x_92); +lean_ctor_set_uint8(x_95, sizeof(void*)*12, x_91); +lean_ctor_set_uint8(x_95, sizeof(void*)*12 + 1, x_93); lean_inc(x_73); lean_inc(x_2); -x_95 = l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_expandField(x_2, x_73, x_13, x_14, x_94, x_16, x_78); -lean_dec(x_94); -if (lean_obj_tag(x_95) == 0) +x_96 = l___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_expandField(x_2, x_73, x_13, x_14, x_95, x_16, x_79); +lean_dec(x_95); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_96; uint8_t x_97; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get_uint8(x_20, sizeof(void*)*3); -if (x_97 == 0) +lean_object* x_97; uint8_t x_98; +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get_uint8(x_20, sizeof(void*)*3); +if (x_98 == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_dec(x_73); lean_dec(x_52); lean_dec(x_49); -x_98 = lean_ctor_get(x_95, 1); -lean_inc(x_98); -lean_dec(x_95); -x_99 = lean_ctor_get(x_96, 0); +x_99 = lean_ctor_get(x_96, 1); lean_inc(x_99); lean_dec(x_96); -x_100 = lean_box(0); +x_100 = lean_ctor_get(x_97, 0); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_box(0); lean_inc(x_15); lean_inc(x_5); lean_inc(x_2); -x_101 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_99, x_2, x_5, x_20, x_48, x_50, x_51, x_100, x_11, x_12, x_13, x_14, x_15, x_16, x_98); +x_102 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_100, x_2, x_5, x_20, x_48, x_50, x_51, x_101, x_11, x_12, x_13, x_14, x_15, x_16, x_99); lean_dec(x_71); -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); +x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); -lean_dec(x_101); -x_21 = x_102; -x_22 = x_103; +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +lean_dec(x_102); +x_21 = x_103; +x_22 = x_104; goto block_46; } else { -lean_object* x_104; uint8_t x_105; -x_104 = lean_ctor_get(x_95, 1); -lean_inc(x_104); -lean_dec(x_95); -x_105 = !lean_is_exclusive(x_96); -if (x_105 == 0) +lean_object* x_105; uint8_t x_106; +x_105 = lean_ctor_get(x_96, 1); +lean_inc(x_105); +lean_dec(x_96); +x_106 = !lean_is_exclusive(x_97); +if (x_106 == 0) { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_96, 0); -x_107 = lean_ctor_get(x_96, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_97, 0); +x_108 = lean_ctor_get(x_97, 1); lean_inc(x_11); -x_108 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_107, x_11, x_12, x_13, x_14, x_15, x_16, x_104); -if (lean_obj_tag(x_108) == 0) -{ -uint8_t x_109; -x_109 = !lean_is_exclusive(x_108); -if (x_109 == 0) -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; -x_110 = lean_ctor_get(x_108, 0); -x_111 = lean_ctor_get(x_108, 1); -x_112 = l_Lean_ConstantInfo_type(x_110); -lean_dec(x_110); -x_113 = l_Lean_Expr_bindingBody_x21(x_112); -lean_dec(x_112); -x_114 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; -x_115 = lean_expr_eqv(x_113, x_114); +x_109 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_108, x_11, x_12, x_13, x_14, x_15, x_16, x_105); +if (lean_obj_tag(x_109) == 0) +{ +uint8_t x_110; +x_110 = !lean_is_exclusive(x_109); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_111 = lean_ctor_get(x_109, 0); +x_112 = lean_ctor_get(x_109, 1); +x_113 = l_Lean_ConstantInfo_type(x_111); +lean_dec(x_111); +x_114 = l_Lean_Expr_bindingBody_x21(x_113); lean_dec(x_113); -if (x_115 == 0) +x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; +x_116 = lean_expr_eqv(x_114, x_115); +lean_dec(x_114); +if (x_116 == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_106); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_107); lean_dec(x_71); -x_116 = lean_ctor_get(x_20, 0); -lean_inc(x_116); +x_117 = lean_ctor_get(x_20, 0); +lean_inc(x_117); lean_dec(x_20); -x_117 = l_Lean_MessageData_ofName(x_73); -x_118 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; -lean_ctor_set_tag(x_108, 7); -lean_ctor_set(x_108, 1, x_117); -lean_ctor_set(x_108, 0, x_118); -x_119 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; -lean_ctor_set_tag(x_96, 7); -lean_ctor_set(x_96, 1, x_119); -lean_ctor_set(x_96, 0, x_108); +x_118 = l_Lean_MessageData_ofName(x_73); +x_119 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; +lean_ctor_set_tag(x_109, 7); +lean_ctor_set(x_109, 1, x_118); +lean_ctor_set(x_109, 0, x_119); +x_120 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; +lean_ctor_set_tag(x_97, 7); +lean_ctor_set(x_97, 1, x_120); +lean_ctor_set(x_97, 0, x_109); lean_inc(x_15); lean_inc(x_11); -x_120 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_116, x_96, x_11, x_12, x_13, x_14, x_15, x_16, x_111); -lean_dec(x_116); -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); +x_121 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_117, x_97, x_11, x_12, x_13, x_14, x_15, x_16, x_112); +lean_dec(x_117); +x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); -lean_dec(x_120); -x_53 = x_121; -x_54 = x_122; +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_53 = x_122; +x_54 = x_123; goto block_70; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_free_object(x_108); -lean_free_object(x_96); +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_free_object(x_109); +lean_free_object(x_97); lean_dec(x_73); lean_dec(x_52); lean_dec(x_49); -x_123 = lean_box(0); +x_124 = lean_box(0); lean_inc(x_15); lean_inc(x_5); lean_inc(x_2); -x_124 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_106, x_2, x_5, x_20, x_48, x_50, x_51, x_123, x_11, x_12, x_13, x_14, x_15, x_16, x_111); +x_125 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_107, x_2, x_5, x_20, x_48, x_50, x_51, x_124, x_11, x_12, x_13, x_14, x_15, x_16, x_112); lean_dec(x_71); -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); +x_126 = lean_ctor_get(x_125, 0); lean_inc(x_126); -lean_dec(x_124); -x_21 = x_125; -x_22 = x_126; +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_21 = x_126; +x_22 = x_127; goto block_46; } } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_127 = lean_ctor_get(x_108, 0); -x_128 = lean_ctor_get(x_108, 1); +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; +x_128 = lean_ctor_get(x_109, 0); +x_129 = lean_ctor_get(x_109, 1); +lean_inc(x_129); lean_inc(x_128); -lean_inc(x_127); -lean_dec(x_108); -x_129 = l_Lean_ConstantInfo_type(x_127); -lean_dec(x_127); -x_130 = l_Lean_Expr_bindingBody_x21(x_129); -lean_dec(x_129); -x_131 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; -x_132 = lean_expr_eqv(x_130, x_131); +lean_dec(x_109); +x_130 = l_Lean_ConstantInfo_type(x_128); +lean_dec(x_128); +x_131 = l_Lean_Expr_bindingBody_x21(x_130); lean_dec(x_130); -if (x_132 == 0) +x_132 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; +x_133 = lean_expr_eqv(x_131, x_132); +lean_dec(x_131); +if (x_133 == 0) { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_106); +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_107); lean_dec(x_71); -x_133 = lean_ctor_get(x_20, 0); -lean_inc(x_133); +x_134 = lean_ctor_get(x_20, 0); +lean_inc(x_134); lean_dec(x_20); -x_134 = l_Lean_MessageData_ofName(x_73); -x_135 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; -x_136 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_134); -x_137 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; -lean_ctor_set_tag(x_96, 7); -lean_ctor_set(x_96, 1, x_137); -lean_ctor_set(x_96, 0, x_136); +x_135 = l_Lean_MessageData_ofName(x_73); +x_136 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; +x_137 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_135); +x_138 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; +lean_ctor_set_tag(x_97, 7); +lean_ctor_set(x_97, 1, x_138); +lean_ctor_set(x_97, 0, x_137); lean_inc(x_15); lean_inc(x_11); -x_138 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_133, x_96, x_11, x_12, x_13, x_14, x_15, x_16, x_128); -lean_dec(x_133); -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); +x_139 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_134, x_97, x_11, x_12, x_13, x_14, x_15, x_16, x_129); +lean_dec(x_134); +x_140 = lean_ctor_get(x_139, 0); lean_inc(x_140); -lean_dec(x_138); -x_53 = x_139; -x_54 = x_140; +x_141 = lean_ctor_get(x_139, 1); +lean_inc(x_141); +lean_dec(x_139); +x_53 = x_140; +x_54 = x_141; goto block_70; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -lean_free_object(x_96); +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_free_object(x_97); lean_dec(x_73); lean_dec(x_52); lean_dec(x_49); -x_141 = lean_box(0); +x_142 = lean_box(0); lean_inc(x_15); lean_inc(x_5); lean_inc(x_2); -x_142 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_106, x_2, x_5, x_20, x_48, x_50, x_51, x_141, x_11, x_12, x_13, x_14, x_15, x_16, x_128); +x_143 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_107, x_2, x_5, x_20, x_48, x_50, x_51, x_142, x_11, x_12, x_13, x_14, x_15, x_16, x_129); lean_dec(x_71); -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); +x_144 = lean_ctor_get(x_143, 0); lean_inc(x_144); -lean_dec(x_142); -x_21 = x_143; -x_22 = x_144; +x_145 = lean_ctor_get(x_143, 1); +lean_inc(x_145); +lean_dec(x_143); +x_21 = x_144; +x_22 = x_145; goto block_46; } } } else { -lean_object* x_145; lean_object* x_146; -lean_free_object(x_96); -lean_dec(x_106); +lean_object* x_146; lean_object* x_147; +lean_free_object(x_97); +lean_dec(x_107); lean_dec(x_73); lean_dec(x_71); lean_dec(x_20); -x_145 = lean_ctor_get(x_108, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_108, 1); +x_146 = lean_ctor_get(x_109, 0); lean_inc(x_146); -lean_dec(x_108); -x_53 = x_145; -x_54 = x_146; +x_147 = lean_ctor_get(x_109, 1); +lean_inc(x_147); +lean_dec(x_109); +x_53 = x_146; +x_54 = x_147; goto block_70; } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_96, 0); -x_148 = lean_ctor_get(x_96, 1); +lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_148 = lean_ctor_get(x_97, 0); +x_149 = lean_ctor_get(x_97, 1); +lean_inc(x_149); lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_96); +lean_dec(x_97); lean_inc(x_11); -x_149 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_148, x_11, x_12, x_13, x_14, x_15, x_16, x_104); -if (lean_obj_tag(x_149) == 0) +x_150 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_149, x_11, x_12, x_13, x_14, x_15, x_16, x_105); +if (lean_obj_tag(x_150) == 0) { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; -x_150 = lean_ctor_get(x_149, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 1); +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_151 = lean_ctor_get(x_150, 0); lean_inc(x_151); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_152 = x_149; +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_153 = x_150; } else { - lean_dec_ref(x_149); - x_152 = lean_box(0); + lean_dec_ref(x_150); + x_153 = lean_box(0); } -x_153 = l_Lean_ConstantInfo_type(x_150); -lean_dec(x_150); -x_154 = l_Lean_Expr_bindingBody_x21(x_153); -lean_dec(x_153); -x_155 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; -x_156 = lean_expr_eqv(x_154, x_155); +x_154 = l_Lean_ConstantInfo_type(x_151); +lean_dec(x_151); +x_155 = l_Lean_Expr_bindingBody_x21(x_154); lean_dec(x_154); -if (x_156 == 0) +x_156 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__2; +x_157 = lean_expr_eqv(x_155, x_156); +lean_dec(x_155); +if (x_157 == 0) { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -lean_dec(x_147); +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_148); lean_dec(x_71); -x_157 = lean_ctor_get(x_20, 0); -lean_inc(x_157); +x_158 = lean_ctor_get(x_20, 0); +lean_inc(x_158); lean_dec(x_20); -x_158 = l_Lean_MessageData_ofName(x_73); -x_159 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; -if (lean_is_scalar(x_152)) { - x_160 = lean_alloc_ctor(7, 2, 0); +x_159 = l_Lean_MessageData_ofName(x_73); +x_160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__4; +if (lean_is_scalar(x_153)) { + x_161 = lean_alloc_ctor(7, 2, 0); } else { - x_160 = x_152; - lean_ctor_set_tag(x_160, 7); -} -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_158); -x_161 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; -x_162 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); + x_161 = x_153; + lean_ctor_set_tag(x_161, 7); +} +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___closed__6; +x_163 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_163, 0, x_161); +lean_ctor_set(x_163, 1, x_162); lean_inc(x_15); lean_inc(x_11); -x_163 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_157, x_162, x_11, x_12, x_13, x_14, x_15, x_16, x_151); -lean_dec(x_157); -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); +x_164 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_158, x_163, x_11, x_12, x_13, x_14, x_15, x_16, x_152); +lean_dec(x_158); +x_165 = lean_ctor_get(x_164, 0); lean_inc(x_165); -lean_dec(x_163); -x_53 = x_164; -x_54 = x_165; +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_53 = x_165; +x_54 = x_166; goto block_70; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -lean_dec(x_152); +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +lean_dec(x_153); lean_dec(x_73); lean_dec(x_52); lean_dec(x_49); -x_166 = lean_box(0); +x_167 = lean_box(0); lean_inc(x_15); lean_inc(x_5); lean_inc(x_2); -x_167 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_147, x_2, x_5, x_20, x_48, x_50, x_51, x_166, x_11, x_12, x_13, x_14, x_15, x_16, x_151); +x_168 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3(x_71, x_148, x_2, x_5, x_20, x_48, x_50, x_51, x_167, x_11, x_12, x_13, x_14, x_15, x_16, x_152); lean_dec(x_71); -x_168 = lean_ctor_get(x_167, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_167, 1); +x_169 = lean_ctor_get(x_168, 0); lean_inc(x_169); -lean_dec(x_167); -x_21 = x_168; -x_22 = x_169; +x_170 = lean_ctor_get(x_168, 1); +lean_inc(x_170); +lean_dec(x_168); +x_21 = x_169; +x_22 = x_170; goto block_46; } } else { -lean_object* x_170; lean_object* x_171; -lean_dec(x_147); +lean_object* x_171; lean_object* x_172; +lean_dec(x_148); lean_dec(x_73); lean_dec(x_71); lean_dec(x_20); -x_170 = lean_ctor_get(x_149, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_149, 1); +x_171 = lean_ctor_get(x_150, 0); lean_inc(x_171); -lean_dec(x_149); -x_53 = x_170; -x_54 = x_171; +x_172 = lean_ctor_get(x_150, 1); +lean_inc(x_172); +lean_dec(x_150); +x_53 = x_171; +x_54 = x_172; goto block_70; } } @@ -4235,461 +4286,476 @@ goto block_70; } else { -lean_object* x_172; lean_object* x_173; +lean_object* x_173; lean_object* x_174; lean_dec(x_73); lean_dec(x_71); lean_dec(x_20); -x_172 = lean_ctor_get(x_95, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_95, 1); +x_173 = lean_ctor_get(x_96, 0); lean_inc(x_173); -lean_dec(x_95); -x_53 = x_172; -x_54 = x_173; +x_174 = lean_ctor_get(x_96, 1); +lean_inc(x_174); +lean_dec(x_96); +x_53 = x_173; +x_54 = x_174; goto block_70; } } else { -uint8_t x_174; +uint8_t x_175; lean_dec(x_73); lean_dec(x_71); lean_dec(x_52); lean_dec(x_49); -x_174 = l_Array_isEmpty___rarg(x_48); -if (x_174 == 0) +x_175 = l_Array_isEmpty___rarg(x_48); +if (x_175 == 0) { lean_dec(x_50); if (lean_obj_tag(x_51) == 0) { -lean_object* x_175; uint8_t x_176; lean_object* x_177; lean_object* x_178; uint8_t x_179; -x_175 = lean_ctor_get(x_15, 5); -lean_inc(x_175); -x_176 = 0; -x_177 = l_Lean_SourceInfo_fromRef(x_175, x_176); -lean_dec(x_175); -x_178 = lean_st_ref_get(x_16, x_17); -x_179 = !lean_is_exclusive(x_178); -if (x_179 == 0) -{ -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; size_t x_186; size_t x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_180 = lean_ctor_get(x_178, 1); -x_181 = lean_ctor_get(x_178, 0); -lean_dec(x_181); -x_182 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_177); -lean_ctor_set_tag(x_178, 2); -lean_ctor_set(x_178, 1, x_182); -lean_ctor_set(x_178, 0, x_177); -x_183 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -x_184 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -lean_inc(x_177); -x_185 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_185, 0, x_177); -lean_ctor_set(x_185, 1, x_183); -lean_ctor_set(x_185, 2, x_184); -x_186 = lean_array_size(x_48); -x_187 = 0; -x_188 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_189 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__3(x_188, x_186, x_187, x_48); -x_190 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_191 = l_Lean_Syntax_SepArray_ofElems(x_190, x_189); -lean_dec(x_189); -x_192 = l_Array_append___rarg(x_184, x_191); -lean_dec(x_191); -lean_inc(x_177); -x_193 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_193, 0, x_177); -lean_ctor_set(x_193, 1, x_183); -lean_ctor_set(x_193, 2, x_192); -x_194 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_185); -lean_inc(x_177); -x_195 = l_Lean_Syntax_node1(x_177, x_194, x_185); -x_196 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_177); -x_197 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_197, 0, x_177); -lean_ctor_set(x_197, 1, x_196); -x_198 = lean_box(0); +lean_object* x_176; uint8_t x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; +x_176 = lean_ctor_get(x_15, 5); +lean_inc(x_176); +x_177 = 0; +x_178 = l_Lean_SourceInfo_fromRef(x_176, x_177); +lean_dec(x_176); +x_179 = lean_st_ref_get(x_16, x_17); +x_180 = !lean_is_exclusive(x_179); +if (x_180 == 0) +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; size_t x_187; size_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_181 = lean_ctor_get(x_179, 1); +x_182 = lean_ctor_get(x_179, 0); +lean_dec(x_182); +x_183 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_178); +lean_ctor_set_tag(x_179, 2); +lean_ctor_set(x_179, 1, x_183); +lean_ctor_set(x_179, 0, x_178); +x_184 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +x_185 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +lean_inc(x_178); +x_186 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_186, 0, x_178); +lean_ctor_set(x_186, 1, x_184); +lean_ctor_set(x_186, 2, x_185); +x_187 = lean_array_size(x_48); +x_188 = 0; +x_189 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_190 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__3(x_189, x_187, x_188, x_48); +x_191 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_192 = l_Lean_Syntax_SepArray_ofElems(x_191, x_190); +lean_dec(x_190); +x_193 = l_Array_append___rarg(x_185, x_192); +lean_dec(x_192); +lean_inc(x_178); +x_194 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_194, 0, x_178); +lean_ctor_set(x_194, 1, x_184); +lean_ctor_set(x_194, 2, x_193); +x_195 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_178); +x_196 = l_Lean_Syntax_node1(x_178, x_195, x_194); +x_197 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_186); +lean_inc(x_178); +x_198 = l_Lean_Syntax_node1(x_178, x_197, x_186); +x_199 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_178); +x_200 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_200, 0, x_178); +lean_ctor_set(x_200, 1, x_199); +x_201 = lean_box(0); lean_inc(x_2); -x_199 = l_Lean_mkCIdentFrom(x_198, x_2, x_176); -lean_inc(x_177); -x_200 = l_Lean_Syntax_node2(x_177, x_183, x_197, x_199); -x_201 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_177); -x_202 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_202, 0, x_177); -lean_ctor_set(x_202, 1, x_201); -x_203 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_204 = l_Lean_Syntax_node6(x_177, x_203, x_178, x_185, x_193, x_195, x_200, x_202); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_204); -x_206 = l_Lean_NameSet_empty; -x_207 = lean_box(0); +x_202 = l_Lean_mkCIdentFrom(x_201, x_2, x_177); +lean_inc(x_178); +x_203 = l_Lean_Syntax_node2(x_178, x_184, x_200, x_202); +x_204 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_178); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_178); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_207 = l_Lean_Syntax_node6(x_178, x_206, x_179, x_186, x_196, x_198, x_203, x_205); +x_208 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_208, 0, x_207); +x_209 = l_Lean_NameSet_empty; +x_210 = lean_box(0); lean_inc(x_5); lean_inc(x_2); -x_208 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_206, x_205, x_207, x_11, x_12, x_13, x_14, x_15, x_16, x_180); -x_209 = lean_ctor_get(x_208, 0); -lean_inc(x_209); -x_210 = lean_ctor_get(x_208, 1); -lean_inc(x_210); -lean_dec(x_208); -x_21 = x_209; -x_22 = x_210; +x_211 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_209, x_208, x_210, x_11, x_12, x_13, x_14, x_15, x_16, x_181); +x_212 = lean_ctor_get(x_211, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_211, 1); +lean_inc(x_213); +lean_dec(x_211); +x_21 = x_212; +x_22 = x_213; goto block_46; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; size_t x_217; size_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_211 = lean_ctor_get(x_178, 1); -lean_inc(x_211); -lean_dec(x_178); -x_212 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_177); -x_213 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_213, 0, x_177); -lean_ctor_set(x_213, 1, x_212); -x_214 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -x_215 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -lean_inc(x_177); -x_216 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_216, 0, x_177); -lean_ctor_set(x_216, 1, x_214); -lean_ctor_set(x_216, 2, x_215); -x_217 = lean_array_size(x_48); -x_218 = 0; -x_219 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_220 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__3(x_219, x_217, x_218, x_48); -x_221 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_222 = l_Lean_Syntax_SepArray_ofElems(x_221, x_220); -lean_dec(x_220); -x_223 = l_Array_append___rarg(x_215, x_222); -lean_dec(x_222); -lean_inc(x_177); -x_224 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_224, 0, x_177); -lean_ctor_set(x_224, 1, x_214); -lean_ctor_set(x_224, 2, x_223); -x_225 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_216); -lean_inc(x_177); -x_226 = l_Lean_Syntax_node1(x_177, x_225, x_216); -x_227 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_177); -x_228 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_228, 0, x_177); -lean_ctor_set(x_228, 1, x_227); -x_229 = lean_box(0); -lean_inc(x_2); -x_230 = l_Lean_mkCIdentFrom(x_229, x_2, x_176); -lean_inc(x_177); -x_231 = l_Lean_Syntax_node2(x_177, x_214, x_228, x_230); -x_232 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_177); +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; size_t x_220; size_t x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_214 = lean_ctor_get(x_179, 1); +lean_inc(x_214); +lean_dec(x_179); +x_215 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_178); +x_216 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_216, 0, x_178); +lean_ctor_set(x_216, 1, x_215); +x_217 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +x_218 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +lean_inc(x_178); +x_219 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_219, 0, x_178); +lean_ctor_set(x_219, 1, x_217); +lean_ctor_set(x_219, 2, x_218); +x_220 = lean_array_size(x_48); +x_221 = 0; +x_222 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_223 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__3(x_222, x_220, x_221, x_48); +x_224 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_225 = l_Lean_Syntax_SepArray_ofElems(x_224, x_223); +lean_dec(x_223); +x_226 = l_Array_append___rarg(x_218, x_225); +lean_dec(x_225); +lean_inc(x_178); +x_227 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_227, 0, x_178); +lean_ctor_set(x_227, 1, x_217); +lean_ctor_set(x_227, 2, x_226); +x_228 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_178); +x_229 = l_Lean_Syntax_node1(x_178, x_228, x_227); +x_230 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_219); +lean_inc(x_178); +x_231 = l_Lean_Syntax_node1(x_178, x_230, x_219); +x_232 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_178); x_233 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_233, 0, x_177); +lean_ctor_set(x_233, 0, x_178); lean_ctor_set(x_233, 1, x_232); -x_234 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_235 = l_Lean_Syntax_node6(x_177, x_234, x_213, x_216, x_224, x_226, x_231, x_233); -x_236 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_236, 0, x_235); -x_237 = l_Lean_NameSet_empty; -x_238 = lean_box(0); +x_234 = lean_box(0); +lean_inc(x_2); +x_235 = l_Lean_mkCIdentFrom(x_234, x_2, x_177); +lean_inc(x_178); +x_236 = l_Lean_Syntax_node2(x_178, x_217, x_233, x_235); +x_237 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_178); +x_238 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_238, 0, x_178); +lean_ctor_set(x_238, 1, x_237); +x_239 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_240 = l_Lean_Syntax_node6(x_178, x_239, x_216, x_219, x_229, x_231, x_236, x_238); +x_241 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_241, 0, x_240); +x_242 = l_Lean_NameSet_empty; +x_243 = lean_box(0); lean_inc(x_5); lean_inc(x_2); -x_239 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_237, x_236, x_238, x_11, x_12, x_13, x_14, x_15, x_16, x_211); -x_240 = lean_ctor_get(x_239, 0); -lean_inc(x_240); -x_241 = lean_ctor_get(x_239, 1); -lean_inc(x_241); -lean_dec(x_239); -x_21 = x_240; -x_22 = x_241; +x_244 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_242, x_241, x_243, x_11, x_12, x_13, x_14, x_15, x_16, x_214); +x_245 = lean_ctor_get(x_244, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_244, 1); +lean_inc(x_246); +lean_dec(x_244); +x_21 = x_245; +x_22 = x_246; goto block_46; } } else { -uint8_t x_242; -x_242 = !lean_is_exclusive(x_51); -if (x_242 == 0) -{ -lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; -x_243 = lean_ctor_get(x_51, 0); -x_244 = lean_ctor_get(x_15, 5); -lean_inc(x_244); -x_245 = 0; -x_246 = l_Lean_SourceInfo_fromRef(x_244, x_245); -lean_dec(x_244); -x_247 = lean_st_ref_get(x_16, x_17); -x_248 = !lean_is_exclusive(x_247); -if (x_248 == 0) -{ -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; size_t x_257; size_t x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; -x_249 = lean_ctor_get(x_247, 1); -x_250 = lean_ctor_get(x_247, 0); -lean_dec(x_250); -x_251 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_246); -lean_ctor_set_tag(x_247, 2); -lean_ctor_set(x_247, 1, x_251); -lean_ctor_set(x_247, 0, x_246); -x_252 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_246); -x_253 = l_Lean_Syntax_node1(x_246, x_252, x_243); -x_254 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_246); -x_255 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_255, 0, x_246); -lean_ctor_set(x_255, 1, x_254); -lean_inc(x_246); -x_256 = l_Lean_Syntax_node2(x_246, x_252, x_253, x_255); -x_257 = lean_array_size(x_48); -x_258 = 0; -x_259 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_260 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_259, x_257, x_258, x_48); -x_261 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_262 = l_Lean_Syntax_SepArray_ofElems(x_261, x_260); -lean_dec(x_260); -x_263 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_264 = l_Array_append___rarg(x_263, x_262); -lean_dec(x_262); -lean_inc(x_246); -x_265 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_265, 0, x_246); -lean_ctor_set(x_265, 1, x_252); -lean_ctor_set(x_265, 2, x_264); -lean_inc(x_246); -x_266 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_266, 0, x_246); -lean_ctor_set(x_266, 1, x_252); -lean_ctor_set(x_266, 2, x_263); -x_267 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_246); -x_268 = l_Lean_Syntax_node1(x_246, x_267, x_266); -x_269 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_246); -x_270 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_270, 0, x_246); -lean_ctor_set(x_270, 1, x_269); -x_271 = lean_box(0); -lean_inc(x_2); -x_272 = l_Lean_mkCIdentFrom(x_271, x_2, x_245); -lean_inc(x_246); -x_273 = l_Lean_Syntax_node2(x_246, x_252, x_270, x_272); +uint8_t x_247; +x_247 = !lean_is_exclusive(x_51); +if (x_247 == 0) +{ +lean_object* x_248; lean_object* x_249; uint8_t x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; +x_248 = lean_ctor_get(x_51, 0); +x_249 = lean_ctor_get(x_15, 5); +lean_inc(x_249); +x_250 = 0; +x_251 = l_Lean_SourceInfo_fromRef(x_249, x_250); +lean_dec(x_249); +x_252 = lean_st_ref_get(x_16, x_17); +x_253 = !lean_is_exclusive(x_252); +if (x_253 == 0) +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; size_t x_262; size_t x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_254 = lean_ctor_get(x_252, 1); +x_255 = lean_ctor_get(x_252, 0); +lean_dec(x_255); +x_256 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_251); +lean_ctor_set_tag(x_252, 2); +lean_ctor_set(x_252, 1, x_256); +lean_ctor_set(x_252, 0, x_251); +x_257 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_251); +x_258 = l_Lean_Syntax_node1(x_251, x_257, x_248); +x_259 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_251); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_251); +lean_ctor_set(x_260, 1, x_259); +lean_inc(x_251); +x_261 = l_Lean_Syntax_node2(x_251, x_257, x_258, x_260); +x_262 = lean_array_size(x_48); +x_263 = 0; +x_264 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_265 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_264, x_262, x_263, x_48); +x_266 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_267 = l_Lean_Syntax_SepArray_ofElems(x_266, x_265); +lean_dec(x_265); +x_268 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_269 = l_Array_append___rarg(x_268, x_267); +lean_dec(x_267); +lean_inc(x_251); +x_270 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_270, 0, x_251); +lean_ctor_set(x_270, 1, x_257); +lean_ctor_set(x_270, 2, x_269); +x_271 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_251); +x_272 = l_Lean_Syntax_node1(x_251, x_271, x_270); +lean_inc(x_251); +x_273 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_273, 0, x_251); +lean_ctor_set(x_273, 1, x_257); +lean_ctor_set(x_273, 2, x_268); x_274 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_246); -x_275 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_275, 0, x_246); -lean_ctor_set(x_275, 1, x_274); -x_276 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_277 = l_Lean_Syntax_node6(x_246, x_276, x_247, x_256, x_265, x_268, x_273, x_275); -lean_ctor_set(x_51, 0, x_277); -x_278 = l_Lean_NameSet_empty; -x_279 = lean_box(0); +lean_inc(x_251); +x_275 = l_Lean_Syntax_node1(x_251, x_274, x_273); +x_276 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_251); +x_277 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_277, 0, x_251); +lean_ctor_set(x_277, 1, x_276); +x_278 = lean_box(0); +lean_inc(x_2); +x_279 = l_Lean_mkCIdentFrom(x_278, x_2, x_250); +lean_inc(x_251); +x_280 = l_Lean_Syntax_node2(x_251, x_257, x_277, x_279); +x_281 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_251); +x_282 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_282, 0, x_251); +lean_ctor_set(x_282, 1, x_281); +x_283 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_284 = l_Lean_Syntax_node6(x_251, x_283, x_252, x_261, x_272, x_275, x_280, x_282); +lean_ctor_set(x_51, 0, x_284); +x_285 = l_Lean_NameSet_empty; +x_286 = lean_box(0); lean_inc(x_5); lean_inc(x_2); -x_280 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_278, x_51, x_279, x_11, x_12, x_13, x_14, x_15, x_16, x_249); -x_281 = lean_ctor_get(x_280, 0); -lean_inc(x_281); -x_282 = lean_ctor_get(x_280, 1); -lean_inc(x_282); -lean_dec(x_280); -x_21 = x_281; -x_22 = x_282; +x_287 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_285, x_51, x_286, x_11, x_12, x_13, x_14, x_15, x_16, x_254); +x_288 = lean_ctor_get(x_287, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_287, 1); +lean_inc(x_289); +lean_dec(x_287); +x_21 = x_288; +x_22 = x_289; goto block_46; } else { -lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; size_t x_291; size_t x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; -x_283 = lean_ctor_get(x_247, 1); -lean_inc(x_283); -lean_dec(x_247); -x_284 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_246); -x_285 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_285, 0, x_246); -lean_ctor_set(x_285, 1, x_284); -x_286 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_246); -x_287 = l_Lean_Syntax_node1(x_246, x_286, x_243); -x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_246); -x_289 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_289, 0, x_246); -lean_ctor_set(x_289, 1, x_288); -lean_inc(x_246); -x_290 = l_Lean_Syntax_node2(x_246, x_286, x_287, x_289); -x_291 = lean_array_size(x_48); -x_292 = 0; -x_293 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_294 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_293, x_291, x_292, x_48); -x_295 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_296 = l_Lean_Syntax_SepArray_ofElems(x_295, x_294); -lean_dec(x_294); -x_297 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_298 = l_Array_append___rarg(x_297, x_296); -lean_dec(x_296); -lean_inc(x_246); -x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_246); -lean_ctor_set(x_299, 1, x_286); -lean_ctor_set(x_299, 2, x_298); -lean_inc(x_246); -x_300 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_300, 0, x_246); -lean_ctor_set(x_300, 1, x_286); -lean_ctor_set(x_300, 2, x_297); -x_301 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_246); -x_302 = l_Lean_Syntax_node1(x_246, x_301, x_300); -x_303 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_246); -x_304 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_304, 0, x_246); -lean_ctor_set(x_304, 1, x_303); -x_305 = lean_box(0); +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; size_t x_298; size_t x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; +x_290 = lean_ctor_get(x_252, 1); +lean_inc(x_290); +lean_dec(x_252); +x_291 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_251); +x_292 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_292, 0, x_251); +lean_ctor_set(x_292, 1, x_291); +x_293 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_251); +x_294 = l_Lean_Syntax_node1(x_251, x_293, x_248); +x_295 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_251); +x_296 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_296, 0, x_251); +lean_ctor_set(x_296, 1, x_295); +lean_inc(x_251); +x_297 = l_Lean_Syntax_node2(x_251, x_293, x_294, x_296); +x_298 = lean_array_size(x_48); +x_299 = 0; +x_300 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_301 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_300, x_298, x_299, x_48); +x_302 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_303 = l_Lean_Syntax_SepArray_ofElems(x_302, x_301); +lean_dec(x_301); +x_304 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_305 = l_Array_append___rarg(x_304, x_303); +lean_dec(x_303); +lean_inc(x_251); +x_306 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_306, 0, x_251); +lean_ctor_set(x_306, 1, x_293); +lean_ctor_set(x_306, 2, x_305); +x_307 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_251); +x_308 = l_Lean_Syntax_node1(x_251, x_307, x_306); +lean_inc(x_251); +x_309 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_309, 0, x_251); +lean_ctor_set(x_309, 1, x_293); +lean_ctor_set(x_309, 2, x_304); +x_310 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_251); +x_311 = l_Lean_Syntax_node1(x_251, x_310, x_309); +x_312 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_251); +x_313 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_313, 0, x_251); +lean_ctor_set(x_313, 1, x_312); +x_314 = lean_box(0); lean_inc(x_2); -x_306 = l_Lean_mkCIdentFrom(x_305, x_2, x_245); -lean_inc(x_246); -x_307 = l_Lean_Syntax_node2(x_246, x_286, x_304, x_306); -x_308 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_246); -x_309 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_309, 0, x_246); -lean_ctor_set(x_309, 1, x_308); -x_310 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_311 = l_Lean_Syntax_node6(x_246, x_310, x_285, x_290, x_299, x_302, x_307, x_309); -lean_ctor_set(x_51, 0, x_311); -x_312 = l_Lean_NameSet_empty; -x_313 = lean_box(0); +x_315 = l_Lean_mkCIdentFrom(x_314, x_2, x_250); +lean_inc(x_251); +x_316 = l_Lean_Syntax_node2(x_251, x_293, x_313, x_315); +x_317 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_251); +x_318 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_318, 0, x_251); +lean_ctor_set(x_318, 1, x_317); +x_319 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_320 = l_Lean_Syntax_node6(x_251, x_319, x_292, x_297, x_308, x_311, x_316, x_318); +lean_ctor_set(x_51, 0, x_320); +x_321 = l_Lean_NameSet_empty; +x_322 = lean_box(0); lean_inc(x_5); lean_inc(x_2); -x_314 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_312, x_51, x_313, x_11, x_12, x_13, x_14, x_15, x_16, x_283); -x_315 = lean_ctor_get(x_314, 0); -lean_inc(x_315); -x_316 = lean_ctor_get(x_314, 1); -lean_inc(x_316); -lean_dec(x_314); -x_21 = x_315; -x_22 = x_316; +x_323 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_321, x_51, x_322, x_11, x_12, x_13, x_14, x_15, x_16, x_290); +x_324 = lean_ctor_get(x_323, 0); +lean_inc(x_324); +x_325 = lean_ctor_get(x_323, 1); +lean_inc(x_325); +lean_dec(x_323); +x_21 = x_324; +x_22 = x_325; goto block_46; } } else { -lean_object* x_317; lean_object* x_318; uint8_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; size_t x_331; size_t x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; -x_317 = lean_ctor_get(x_51, 0); -lean_inc(x_317); +lean_object* x_326; lean_object* x_327; uint8_t x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; size_t x_340; size_t x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +x_326 = lean_ctor_get(x_51, 0); +lean_inc(x_326); lean_dec(x_51); -x_318 = lean_ctor_get(x_15, 5); -lean_inc(x_318); -x_319 = 0; -x_320 = l_Lean_SourceInfo_fromRef(x_318, x_319); -lean_dec(x_318); -x_321 = lean_st_ref_get(x_16, x_17); -x_322 = lean_ctor_get(x_321, 1); -lean_inc(x_322); -if (lean_is_exclusive(x_321)) { - lean_ctor_release(x_321, 0); - lean_ctor_release(x_321, 1); - x_323 = x_321; +x_327 = lean_ctor_get(x_15, 5); +lean_inc(x_327); +x_328 = 0; +x_329 = l_Lean_SourceInfo_fromRef(x_327, x_328); +lean_dec(x_327); +x_330 = lean_st_ref_get(x_16, x_17); +x_331 = lean_ctor_get(x_330, 1); +lean_inc(x_331); +if (lean_is_exclusive(x_330)) { + lean_ctor_release(x_330, 0); + lean_ctor_release(x_330, 1); + x_332 = x_330; } else { - lean_dec_ref(x_321); - x_323 = lean_box(0); + lean_dec_ref(x_330); + x_332 = lean_box(0); } -x_324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_320); -if (lean_is_scalar(x_323)) { - x_325 = lean_alloc_ctor(2, 2, 0); +x_333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_329); +if (lean_is_scalar(x_332)) { + x_334 = lean_alloc_ctor(2, 2, 0); } else { - x_325 = x_323; - lean_ctor_set_tag(x_325, 2); -} -lean_ctor_set(x_325, 0, x_320); -lean_ctor_set(x_325, 1, x_324); -x_326 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_320); -x_327 = l_Lean_Syntax_node1(x_320, x_326, x_317); -x_328 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_320); -x_329 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_329, 0, x_320); -lean_ctor_set(x_329, 1, x_328); -lean_inc(x_320); -x_330 = l_Lean_Syntax_node2(x_320, x_326, x_327, x_329); -x_331 = lean_array_size(x_48); -x_332 = 0; -x_333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; -x_334 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_333, x_331, x_332, x_48); -x_335 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_336 = l_Lean_Syntax_SepArray_ofElems(x_335, x_334); -lean_dec(x_334); -x_337 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; -x_338 = l_Array_append___rarg(x_337, x_336); -lean_dec(x_336); -lean_inc(x_320); -x_339 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_339, 0, x_320); -lean_ctor_set(x_339, 1, x_326); -lean_ctor_set(x_339, 2, x_338); -lean_inc(x_320); -x_340 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_340, 0, x_320); -lean_ctor_set(x_340, 1, x_326); -lean_ctor_set(x_340, 2, x_337); -x_341 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_320); -x_342 = l_Lean_Syntax_node1(x_320, x_341, x_340); -x_343 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_320); -x_344 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_344, 0, x_320); -lean_ctor_set(x_344, 1, x_343); -x_345 = lean_box(0); + x_334 = x_332; + lean_ctor_set_tag(x_334, 2); +} +lean_ctor_set(x_334, 0, x_329); +lean_ctor_set(x_334, 1, x_333); +x_335 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_329); +x_336 = l_Lean_Syntax_node1(x_329, x_335, x_326); +x_337 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_329); +x_338 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_338, 0, x_329); +lean_ctor_set(x_338, 1, x_337); +lean_inc(x_329); +x_339 = l_Lean_Syntax_node2(x_329, x_335, x_336, x_338); +x_340 = lean_array_size(x_48); +x_341 = 0; +x_342 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__9; +x_343 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__4(x_342, x_340, x_341, x_48); +x_344 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_345 = l_Lean_Syntax_SepArray_ofElems(x_344, x_343); +lean_dec(x_343); +x_346 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__8; +x_347 = l_Array_append___rarg(x_346, x_345); +lean_dec(x_345); +lean_inc(x_329); +x_348 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_348, 0, x_329); +lean_ctor_set(x_348, 1, x_335); +lean_ctor_set(x_348, 2, x_347); +x_349 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_329); +x_350 = l_Lean_Syntax_node1(x_329, x_349, x_348); +lean_inc(x_329); +x_351 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_351, 0, x_329); +lean_ctor_set(x_351, 1, x_335); +lean_ctor_set(x_351, 2, x_346); +x_352 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_329); +x_353 = l_Lean_Syntax_node1(x_329, x_352, x_351); +x_354 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_329); +x_355 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_355, 0, x_329); +lean_ctor_set(x_355, 1, x_354); +x_356 = lean_box(0); lean_inc(x_2); -x_346 = l_Lean_mkCIdentFrom(x_345, x_2, x_319); -lean_inc(x_320); -x_347 = l_Lean_Syntax_node2(x_320, x_326, x_344, x_346); -x_348 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_320); -x_349 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_349, 0, x_320); -lean_ctor_set(x_349, 1, x_348); -x_350 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_351 = l_Lean_Syntax_node6(x_320, x_350, x_325, x_330, x_339, x_342, x_347, x_349); -x_352 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_352, 0, x_351); -x_353 = l_Lean_NameSet_empty; -x_354 = lean_box(0); +x_357 = l_Lean_mkCIdentFrom(x_356, x_2, x_328); +lean_inc(x_329); +x_358 = l_Lean_Syntax_node2(x_329, x_335, x_355, x_357); +x_359 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_329); +x_360 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_360, 0, x_329); +lean_ctor_set(x_360, 1, x_359); +x_361 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_362 = l_Lean_Syntax_node6(x_329, x_361, x_334, x_339, x_350, x_353, x_358, x_360); +x_363 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_363, 0, x_362); +x_364 = l_Lean_NameSet_empty; +x_365 = lean_box(0); lean_inc(x_5); lean_inc(x_2); -x_355 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_353, x_352, x_354, x_11, x_12, x_13, x_14, x_15, x_16, x_322); -x_356 = lean_ctor_get(x_355, 0); -lean_inc(x_356); -x_357 = lean_ctor_get(x_355, 1); -lean_inc(x_357); -lean_dec(x_355); -x_21 = x_356; -x_22 = x_357; +x_366 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_5, x_364, x_363, x_365, x_11, x_12, x_13, x_14, x_15, x_16, x_331); +x_367 = lean_ctor_get(x_366, 0); +lean_inc(x_367); +x_368 = lean_ctor_get(x_366, 1); +lean_inc(x_368); +lean_dec(x_366); +x_21 = x_367; +x_22 = x_368; goto block_46; } } } else { -lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; -x_358 = lean_box(0); +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +x_369 = lean_box(0); lean_inc(x_2); -x_359 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_48, x_50, x_51, x_358, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_360 = lean_ctor_get(x_359, 0); -lean_inc(x_360); -x_361 = lean_ctor_get(x_359, 1); -lean_inc(x_361); -lean_dec(x_359); -x_21 = x_360; -x_22 = x_361; +x_370 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__4(x_20, x_2, x_48, x_50, x_51, x_369, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_371 = lean_ctor_get(x_370, 0); +lean_inc(x_371); +x_372 = lean_ctor_get(x_370, 1); +lean_inc(x_372); +lean_dec(x_370); +x_21 = x_371; +x_22 = x_372; goto block_46; } } @@ -7243,7 +7309,7 @@ x_51 = lean_st_ref_get(x_18, x_23); x_52 = !lean_is_exclusive(x_51); if (x_52 == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; size_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; size_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; x_53 = lean_ctor_get(x_51, 1); x_54 = lean_ctor_get(x_51, 0); lean_dec(x_54); @@ -7264,13 +7330,13 @@ lean_inc(x_10); lean_ctor_set_tag(x_22, 1); lean_ctor_set(x_22, 1, x_10); lean_ctor_set(x_22, 0, x_58); -x_59 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +x_59 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 0, x_59); x_60 = lean_array_size(x_25); x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_21, x_60, x_8, x_25); lean_dec(x_21); -x_62 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; +x_62 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; x_63 = l_Lean_Syntax_SepArray_ofElems(x_62, x_61); lean_dec(x_61); x_64 = l_Array_append___rarg(x_12, x_63); @@ -7280,253 +7346,265 @@ x_65 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_65, 0, x_50); lean_ctor_set(x_65, 1, x_56); lean_ctor_set(x_65, 2, x_64); -x_66 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_66 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_50); +x_67 = l_Lean_Syntax_node1(x_50, x_66, x_65); +x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_57); lean_inc(x_50); -x_67 = l_Lean_Syntax_node1(x_50, x_66, x_57); -x_68 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_69 = l_Lean_Syntax_node1(x_50, x_68, x_57); +x_70 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_50); -x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_50); -lean_ctor_set(x_69, 1, x_68); -x_70 = lean_box(0); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_50); +lean_ctor_set(x_71, 1, x_70); +x_72 = lean_box(0); lean_inc(x_2); -x_71 = l_Lean_mkCIdentFrom(x_70, x_2, x_49); +x_73 = l_Lean_mkCIdentFrom(x_72, x_2, x_49); lean_inc(x_50); -x_72 = l_Lean_Syntax_node2(x_50, x_56, x_69, x_71); -x_73 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_74 = l_Lean_Syntax_node2(x_50, x_56, x_71, x_73); +x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; lean_inc(x_50); -x_74 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_74, 0, x_50); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_76 = l_Lean_Syntax_node6(x_50, x_75, x_51, x_57, x_65, x_67, x_72, x_74); -x_30 = x_76; +x_76 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_76, 0, x_50); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_78 = l_Lean_Syntax_node6(x_50, x_77, x_51, x_57, x_67, x_69, x_74, x_76); +x_30 = x_78; x_31 = x_53; goto block_47; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_77 = lean_ctor_get(x_51, 1); -lean_inc(x_77); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; size_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_79 = lean_ctor_get(x_51, 1); +lean_inc(x_79); lean_dec(x_51); -x_78 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +x_80 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; lean_inc(x_50); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_50); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_12); -lean_inc(x_50); -x_81 = lean_alloc_ctor(1, 3, 0); +x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_50); lean_ctor_set(x_81, 1, x_80); -lean_ctor_set(x_81, 2, x_12); -x_82 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +x_82 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_12); +lean_inc(x_50); +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_50); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_12); +x_84 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); lean_ctor_set_tag(x_22, 1); lean_ctor_set(x_22, 1, x_10); -lean_ctor_set(x_22, 0, x_82); -x_83 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_ctor_set(x_22, 0, x_84); +x_85 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 0, x_83); -x_84 = lean_array_size(x_25); -x_85 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_21, x_84, x_8, x_25); +lean_ctor_set(x_21, 0, x_85); +x_86 = lean_array_size(x_25); +x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_21, x_86, x_8, x_25); lean_dec(x_21); -x_86 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_87 = l_Lean_Syntax_SepArray_ofElems(x_86, x_85); -lean_dec(x_85); -x_88 = l_Array_append___rarg(x_12, x_87); +x_88 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_89 = l_Lean_Syntax_SepArray_ofElems(x_88, x_87); lean_dec(x_87); +x_90 = l_Array_append___rarg(x_12, x_89); +lean_dec(x_89); lean_inc(x_50); -x_89 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_89, 0, x_50); -lean_ctor_set(x_89, 1, x_80); -lean_ctor_set(x_89, 2, x_88); -x_90 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_81); +x_91 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_91, 0, x_50); +lean_ctor_set(x_91, 1, x_82); +lean_ctor_set(x_91, 2, x_90); +x_92 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_50); +x_93 = l_Lean_Syntax_node1(x_50, x_92, x_91); +x_94 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_83); lean_inc(x_50); -x_91 = l_Lean_Syntax_node1(x_50, x_90, x_81); -x_92 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_95 = l_Lean_Syntax_node1(x_50, x_94, x_83); +x_96 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_50); -x_93 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_93, 0, x_50); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_box(0); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_50); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_box(0); lean_inc(x_2); -x_95 = l_Lean_mkCIdentFrom(x_94, x_2, x_49); +x_99 = l_Lean_mkCIdentFrom(x_98, x_2, x_49); lean_inc(x_50); -x_96 = l_Lean_Syntax_node2(x_50, x_80, x_93, x_95); -x_97 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_100 = l_Lean_Syntax_node2(x_50, x_82, x_97, x_99); +x_101 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; lean_inc(x_50); -x_98 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_98, 0, x_50); -lean_ctor_set(x_98, 1, x_97); -x_99 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_100 = l_Lean_Syntax_node6(x_50, x_99, x_79, x_81, x_89, x_91, x_96, x_98); -x_30 = x_100; -x_31 = x_77; +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_50); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_104 = l_Lean_Syntax_node6(x_50, x_103, x_81, x_83, x_93, x_95, x_100, x_102); +x_30 = x_104; +x_31 = x_79; goto block_47; } } else { -lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_101 = lean_ctor_get(x_28, 0); -lean_inc(x_101); +lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_105 = lean_ctor_get(x_28, 0); +lean_inc(x_105); lean_dec(x_28); -x_102 = lean_ctor_get(x_17, 5); -lean_inc(x_102); -x_103 = 0; -x_104 = l_Lean_SourceInfo_fromRef(x_102, x_103); -lean_dec(x_102); -x_105 = lean_st_ref_get(x_18, x_23); -x_106 = !lean_is_exclusive(x_105); -if (x_106 == 0) +x_106 = lean_ctor_get(x_17, 5); +lean_inc(x_106); +x_107 = 0; +x_108 = l_Lean_SourceInfo_fromRef(x_106, x_107); +lean_dec(x_106); +x_109 = lean_st_ref_get(x_18, x_23); +x_110 = !lean_is_exclusive(x_109); +if (x_110 == 0) { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; size_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_107 = lean_ctor_get(x_105, 1); -x_108 = lean_ctor_get(x_105, 0); -lean_dec(x_108); -x_109 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_104); -lean_ctor_set_tag(x_105, 2); -lean_ctor_set(x_105, 1, x_109); -lean_ctor_set(x_105, 0, x_104); -x_110 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_104); -x_111 = l_Lean_Syntax_node1(x_104, x_110, x_101); -x_112 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_104); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; size_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_111 = lean_ctor_get(x_109, 1); +x_112 = lean_ctor_get(x_109, 0); +lean_dec(x_112); +x_113 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_108); +lean_ctor_set_tag(x_109, 2); +lean_ctor_set(x_109, 1, x_113); +lean_ctor_set(x_109, 0, x_108); +x_114 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_108); +x_115 = l_Lean_Syntax_node1(x_108, x_114, x_105); +x_116 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_108); lean_ctor_set_tag(x_22, 2); -lean_ctor_set(x_22, 1, x_112); -lean_ctor_set(x_22, 0, x_104); -lean_inc(x_104); -x_113 = l_Lean_Syntax_node2(x_104, x_110, x_111, x_22); -x_114 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +lean_ctor_set(x_22, 1, x_116); +lean_ctor_set(x_22, 0, x_108); +lean_inc(x_108); +x_117 = l_Lean_Syntax_node2(x_108, x_114, x_115, x_22); +x_118 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 1, x_10); -lean_ctor_set(x_21, 0, x_114); -x_115 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_116 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_21); -x_117 = lean_array_size(x_25); -x_118 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_116, x_117, x_8, x_25); -lean_dec(x_116); -x_119 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_120 = l_Lean_Syntax_SepArray_ofElems(x_119, x_118); -lean_dec(x_118); -lean_inc(x_12); -x_121 = l_Array_append___rarg(x_12, x_120); +lean_ctor_set(x_21, 0, x_118); +x_119 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_21); +x_121 = lean_array_size(x_25); +x_122 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_120, x_121, x_8, x_25); lean_dec(x_120); -lean_inc(x_104); -x_122 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_122, 0, x_104); -lean_ctor_set(x_122, 1, x_110); -lean_ctor_set(x_122, 2, x_121); -lean_inc(x_104); -x_123 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_123, 0, x_104); -lean_ctor_set(x_123, 1, x_110); -lean_ctor_set(x_123, 2, x_12); -x_124 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_104); -x_125 = l_Lean_Syntax_node1(x_104, x_124, x_123); -x_126 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_104); -x_127 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_127, 0, x_104); -lean_ctor_set(x_127, 1, x_126); -x_128 = lean_box(0); +x_123 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_124 = l_Lean_Syntax_SepArray_ofElems(x_123, x_122); +lean_dec(x_122); +lean_inc(x_12); +x_125 = l_Array_append___rarg(x_12, x_124); +lean_dec(x_124); +lean_inc(x_108); +x_126 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_126, 0, x_108); +lean_ctor_set(x_126, 1, x_114); +lean_ctor_set(x_126, 2, x_125); +x_127 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_108); +x_128 = l_Lean_Syntax_node1(x_108, x_127, x_126); +lean_inc(x_108); +x_129 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_129, 0, x_108); +lean_ctor_set(x_129, 1, x_114); +lean_ctor_set(x_129, 2, x_12); +x_130 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_108); +x_131 = l_Lean_Syntax_node1(x_108, x_130, x_129); +x_132 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_108); +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_108); +lean_ctor_set(x_133, 1, x_132); +x_134 = lean_box(0); lean_inc(x_2); -x_129 = l_Lean_mkCIdentFrom(x_128, x_2, x_103); -lean_inc(x_104); -x_130 = l_Lean_Syntax_node2(x_104, x_110, x_127, x_129); -x_131 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_104); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_104); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_134 = l_Lean_Syntax_node6(x_104, x_133, x_105, x_113, x_122, x_125, x_130, x_132); -x_30 = x_134; -x_31 = x_107; +x_135 = l_Lean_mkCIdentFrom(x_134, x_2, x_107); +lean_inc(x_108); +x_136 = l_Lean_Syntax_node2(x_108, x_114, x_133, x_135); +x_137 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_108); +x_138 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_138, 0, x_108); +lean_ctor_set(x_138, 1, x_137); +x_139 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_140 = l_Lean_Syntax_node6(x_108, x_139, x_109, x_117, x_128, x_131, x_136, x_138); +x_30 = x_140; +x_31 = x_111; goto block_47; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; size_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_135 = lean_ctor_get(x_105, 1); -lean_inc(x_135); -lean_dec(x_105); -x_136 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_104); -x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_104); -lean_ctor_set(x_137, 1, x_136); -x_138 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_104); -x_139 = l_Lean_Syntax_node1(x_104, x_138, x_101); -x_140 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_104); +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; size_t x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_141 = lean_ctor_get(x_109, 1); +lean_inc(x_141); +lean_dec(x_109); +x_142 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_108); +x_143 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_143, 0, x_108); +lean_ctor_set(x_143, 1, x_142); +x_144 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_108); +x_145 = l_Lean_Syntax_node1(x_108, x_144, x_105); +x_146 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_108); lean_ctor_set_tag(x_22, 2); -lean_ctor_set(x_22, 1, x_140); -lean_ctor_set(x_22, 0, x_104); -lean_inc(x_104); -x_141 = l_Lean_Syntax_node2(x_104, x_138, x_139, x_22); -x_142 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +lean_ctor_set(x_22, 1, x_146); +lean_ctor_set(x_22, 0, x_108); +lean_inc(x_108); +x_147 = l_Lean_Syntax_node2(x_108, x_144, x_145, x_22); +x_148 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 1, x_10); -lean_ctor_set(x_21, 0, x_142); -x_143 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_144 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_21); -x_145 = lean_array_size(x_25); -x_146 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_144, x_145, x_8, x_25); -lean_dec(x_144); -x_147 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_148 = l_Lean_Syntax_SepArray_ofElems(x_147, x_146); -lean_dec(x_146); +lean_ctor_set(x_21, 0, x_148); +x_149 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_21); +x_151 = lean_array_size(x_25); +x_152 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_150, x_151, x_8, x_25); +lean_dec(x_150); +x_153 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_154 = l_Lean_Syntax_SepArray_ofElems(x_153, x_152); +lean_dec(x_152); lean_inc(x_12); -x_149 = l_Array_append___rarg(x_12, x_148); -lean_dec(x_148); -lean_inc(x_104); -x_150 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_150, 0, x_104); -lean_ctor_set(x_150, 1, x_138); -lean_ctor_set(x_150, 2, x_149); -lean_inc(x_104); -x_151 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_151, 0, x_104); -lean_ctor_set(x_151, 1, x_138); -lean_ctor_set(x_151, 2, x_12); -x_152 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_104); -x_153 = l_Lean_Syntax_node1(x_104, x_152, x_151); -x_154 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_104); -x_155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_155, 0, x_104); -lean_ctor_set(x_155, 1, x_154); -x_156 = lean_box(0); +x_155 = l_Array_append___rarg(x_12, x_154); +lean_dec(x_154); +lean_inc(x_108); +x_156 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_156, 0, x_108); +lean_ctor_set(x_156, 1, x_144); +lean_ctor_set(x_156, 2, x_155); +x_157 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_108); +x_158 = l_Lean_Syntax_node1(x_108, x_157, x_156); +lean_inc(x_108); +x_159 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_159, 0, x_108); +lean_ctor_set(x_159, 1, x_144); +lean_ctor_set(x_159, 2, x_12); +x_160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_108); +x_161 = l_Lean_Syntax_node1(x_108, x_160, x_159); +x_162 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_108); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_108); +lean_ctor_set(x_163, 1, x_162); +x_164 = lean_box(0); lean_inc(x_2); -x_157 = l_Lean_mkCIdentFrom(x_156, x_2, x_103); -lean_inc(x_104); -x_158 = l_Lean_Syntax_node2(x_104, x_138, x_155, x_157); -x_159 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_104); -x_160 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_160, 0, x_104); -lean_ctor_set(x_160, 1, x_159); -x_161 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_162 = l_Lean_Syntax_node6(x_104, x_161, x_137, x_141, x_150, x_153, x_158, x_160); -x_30 = x_162; -x_31 = x_135; +x_165 = l_Lean_mkCIdentFrom(x_164, x_2, x_107); +lean_inc(x_108); +x_166 = l_Lean_Syntax_node2(x_108, x_144, x_163, x_165); +x_167 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_108); +x_168 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_168, 0, x_108); +lean_ctor_set(x_168, 1, x_167); +x_169 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_170 = l_Lean_Syntax_node6(x_108, x_169, x_143, x_147, x_158, x_161, x_166, x_168); +x_30 = x_170; +x_31 = x_141; goto block_47; } } @@ -7602,539 +7680,551 @@ return x_46; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = lean_ctor_get(x_22, 1); -lean_inc(x_163); +lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_171 = lean_ctor_get(x_22, 1); +lean_inc(x_171); lean_dec(x_22); -if (lean_obj_tag(x_163) == 0) -{ -lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; size_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_182 = lean_ctor_get(x_17, 5); -lean_inc(x_182); -x_183 = 0; -x_184 = l_Lean_SourceInfo_fromRef(x_182, x_183); -lean_dec(x_182); -x_185 = lean_st_ref_get(x_18, x_23); -x_186 = lean_ctor_get(x_185, 1); -lean_inc(x_186); -if (lean_is_exclusive(x_185)) { - lean_ctor_release(x_185, 0); - lean_ctor_release(x_185, 1); - x_187 = x_185; +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_190; uint8_t x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; size_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_190 = lean_ctor_get(x_17, 5); +lean_inc(x_190); +x_191 = 0; +x_192 = l_Lean_SourceInfo_fromRef(x_190, x_191); +lean_dec(x_190); +x_193 = lean_st_ref_get(x_18, x_23); +x_194 = lean_ctor_get(x_193, 1); +lean_inc(x_194); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_195 = x_193; } else { - lean_dec_ref(x_185); - x_187 = lean_box(0); + lean_dec_ref(x_193); + x_195 = lean_box(0); } -x_188 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_184); -if (lean_is_scalar(x_187)) { - x_189 = lean_alloc_ctor(2, 2, 0); +x_196 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_192); +if (lean_is_scalar(x_195)) { + x_197 = lean_alloc_ctor(2, 2, 0); } else { - x_189 = x_187; - lean_ctor_set_tag(x_189, 2); + x_197 = x_195; + lean_ctor_set_tag(x_197, 2); } -lean_ctor_set(x_189, 0, x_184); -lean_ctor_set(x_189, 1, x_188); -x_190 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_ctor_set(x_197, 0, x_192); +lean_ctor_set(x_197, 1, x_196); +x_198 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; lean_inc(x_12); -lean_inc(x_184); -x_191 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_191, 0, x_184); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_191, 2, x_12); -x_192 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +lean_inc(x_192); +x_199 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_199, 0, x_192); +lean_ctor_set(x_199, 1, x_198); +lean_ctor_set(x_199, 2, x_12); +x_200 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); -x_193 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_193, 0, x_192); -lean_ctor_set(x_193, 1, x_10); -x_194 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +x_201 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_10); +x_202 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; lean_ctor_set_tag(x_21, 1); -lean_ctor_set(x_21, 1, x_193); -lean_ctor_set(x_21, 0, x_194); -x_195 = lean_array_size(x_25); -x_196 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_21, x_195, x_8, x_25); +lean_ctor_set(x_21, 1, x_201); +lean_ctor_set(x_21, 0, x_202); +x_203 = lean_array_size(x_25); +x_204 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_21, x_203, x_8, x_25); lean_dec(x_21); -x_197 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_198 = l_Lean_Syntax_SepArray_ofElems(x_197, x_196); -lean_dec(x_196); -x_199 = l_Array_append___rarg(x_12, x_198); -lean_dec(x_198); -lean_inc(x_184); -x_200 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_200, 0, x_184); -lean_ctor_set(x_200, 1, x_190); -lean_ctor_set(x_200, 2, x_199); -x_201 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_191); -lean_inc(x_184); -x_202 = l_Lean_Syntax_node1(x_184, x_201, x_191); -x_203 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_184); -x_204 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_204, 0, x_184); -lean_ctor_set(x_204, 1, x_203); -x_205 = lean_box(0); +x_205 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_206 = l_Lean_Syntax_SepArray_ofElems(x_205, x_204); +lean_dec(x_204); +x_207 = l_Array_append___rarg(x_12, x_206); +lean_dec(x_206); +lean_inc(x_192); +x_208 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_208, 0, x_192); +lean_ctor_set(x_208, 1, x_198); +lean_ctor_set(x_208, 2, x_207); +x_209 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_192); +x_210 = l_Lean_Syntax_node1(x_192, x_209, x_208); +x_211 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_199); +lean_inc(x_192); +x_212 = l_Lean_Syntax_node1(x_192, x_211, x_199); +x_213 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_192); +x_214 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_214, 0, x_192); +lean_ctor_set(x_214, 1, x_213); +x_215 = lean_box(0); lean_inc(x_2); -x_206 = l_Lean_mkCIdentFrom(x_205, x_2, x_183); -lean_inc(x_184); -x_207 = l_Lean_Syntax_node2(x_184, x_190, x_204, x_206); -x_208 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_184); -x_209 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_209, 0, x_184); -lean_ctor_set(x_209, 1, x_208); -x_210 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_211 = l_Lean_Syntax_node6(x_184, x_210, x_189, x_191, x_200, x_202, x_207, x_209); -x_164 = x_211; -x_165 = x_186; -goto block_181; -} -else -{ -lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; size_t x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_212 = lean_ctor_get(x_163, 0); -lean_inc(x_212); -lean_dec(x_163); -x_213 = lean_ctor_get(x_17, 5); -lean_inc(x_213); -x_214 = 0; -x_215 = l_Lean_SourceInfo_fromRef(x_213, x_214); -lean_dec(x_213); -x_216 = lean_st_ref_get(x_18, x_23); -x_217 = lean_ctor_get(x_216, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_218 = x_216; +x_216 = l_Lean_mkCIdentFrom(x_215, x_2, x_191); +lean_inc(x_192); +x_217 = l_Lean_Syntax_node2(x_192, x_198, x_214, x_216); +x_218 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_192); +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_192); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_221 = l_Lean_Syntax_node6(x_192, x_220, x_197, x_199, x_210, x_212, x_217, x_219); +x_172 = x_221; +x_173 = x_194; +goto block_189; +} +else +{ +lean_object* x_222; lean_object* x_223; uint8_t x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; size_t x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_222 = lean_ctor_get(x_171, 0); +lean_inc(x_222); +lean_dec(x_171); +x_223 = lean_ctor_get(x_17, 5); +lean_inc(x_223); +x_224 = 0; +x_225 = l_Lean_SourceInfo_fromRef(x_223, x_224); +lean_dec(x_223); +x_226 = lean_st_ref_get(x_18, x_23); +x_227 = lean_ctor_get(x_226, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_228 = x_226; } else { - lean_dec_ref(x_216); - x_218 = lean_box(0); + lean_dec_ref(x_226); + x_228 = lean_box(0); } -x_219 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_215); -if (lean_is_scalar(x_218)) { - x_220 = lean_alloc_ctor(2, 2, 0); +x_229 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_225); +if (lean_is_scalar(x_228)) { + x_230 = lean_alloc_ctor(2, 2, 0); } else { - x_220 = x_218; - lean_ctor_set_tag(x_220, 2); -} -lean_ctor_set(x_220, 0, x_215); -lean_ctor_set(x_220, 1, x_219); -x_221 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_215); -x_222 = l_Lean_Syntax_node1(x_215, x_221, x_212); -x_223 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_215); -x_224 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_224, 0, x_215); -lean_ctor_set(x_224, 1, x_223); -lean_inc(x_215); -x_225 = l_Lean_Syntax_node2(x_215, x_221, x_222, x_224); -x_226 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; + x_230 = x_228; + lean_ctor_set_tag(x_230, 2); +} +lean_ctor_set(x_230, 0, x_225); +lean_ctor_set(x_230, 1, x_229); +x_231 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_225); +x_232 = l_Lean_Syntax_node1(x_225, x_231, x_222); +x_233 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_225); +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_225); +lean_ctor_set(x_234, 1, x_233); +lean_inc(x_225); +x_235 = l_Lean_Syntax_node2(x_225, x_231, x_232, x_234); +x_236 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); lean_ctor_set_tag(x_21, 1); lean_ctor_set(x_21, 1, x_10); -lean_ctor_set(x_21, 0, x_226); -x_227 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_228 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_228, 0, x_227); -lean_ctor_set(x_228, 1, x_21); -x_229 = lean_array_size(x_25); -x_230 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_228, x_229, x_8, x_25); -lean_dec(x_228); -x_231 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_232 = l_Lean_Syntax_SepArray_ofElems(x_231, x_230); -lean_dec(x_230); +lean_ctor_set(x_21, 0, x_236); +x_237 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_238 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_238, 0, x_237); +lean_ctor_set(x_238, 1, x_21); +x_239 = lean_array_size(x_25); +x_240 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_238, x_239, x_8, x_25); +lean_dec(x_238); +x_241 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_242 = l_Lean_Syntax_SepArray_ofElems(x_241, x_240); +lean_dec(x_240); lean_inc(x_12); -x_233 = l_Array_append___rarg(x_12, x_232); -lean_dec(x_232); -lean_inc(x_215); -x_234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_234, 0, x_215); -lean_ctor_set(x_234, 1, x_221); -lean_ctor_set(x_234, 2, x_233); -lean_inc(x_215); -x_235 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_235, 0, x_215); -lean_ctor_set(x_235, 1, x_221); -lean_ctor_set(x_235, 2, x_12); -x_236 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_215); -x_237 = l_Lean_Syntax_node1(x_215, x_236, x_235); -x_238 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_215); -x_239 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_239, 0, x_215); -lean_ctor_set(x_239, 1, x_238); -x_240 = lean_box(0); +x_243 = l_Array_append___rarg(x_12, x_242); +lean_dec(x_242); +lean_inc(x_225); +x_244 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_244, 0, x_225); +lean_ctor_set(x_244, 1, x_231); +lean_ctor_set(x_244, 2, x_243); +x_245 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_225); +x_246 = l_Lean_Syntax_node1(x_225, x_245, x_244); +lean_inc(x_225); +x_247 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_247, 0, x_225); +lean_ctor_set(x_247, 1, x_231); +lean_ctor_set(x_247, 2, x_12); +x_248 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_225); +x_249 = l_Lean_Syntax_node1(x_225, x_248, x_247); +x_250 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_225); +x_251 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_251, 0, x_225); +lean_ctor_set(x_251, 1, x_250); +x_252 = lean_box(0); lean_inc(x_2); -x_241 = l_Lean_mkCIdentFrom(x_240, x_2, x_214); -lean_inc(x_215); -x_242 = l_Lean_Syntax_node2(x_215, x_221, x_239, x_241); -x_243 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_215); -x_244 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_244, 0, x_215); -lean_ctor_set(x_244, 1, x_243); -x_245 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_246 = l_Lean_Syntax_node6(x_215, x_245, x_220, x_225, x_234, x_237, x_242, x_244); -x_164 = x_246; -x_165 = x_217; -goto block_181; -} -block_181: -{ -lean_object* x_166; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; -x_166 = l_Lean_Expr_const___override(x_2, x_10); -x_167 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_167, 0, x_166); -x_168 = 1; -x_169 = lean_box(x_168); -x_170 = lean_box(x_168); -x_171 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_171, 0, x_164); -lean_closure_set(x_171, 1, x_167); -lean_closure_set(x_171, 2, x_169); -lean_closure_set(x_171, 3, x_170); -lean_closure_set(x_171, 4, x_11); -x_172 = 1; +x_253 = l_Lean_mkCIdentFrom(x_252, x_2, x_224); +lean_inc(x_225); +x_254 = l_Lean_Syntax_node2(x_225, x_231, x_251, x_253); +x_255 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_225); +x_256 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_256, 0, x_225); +lean_ctor_set(x_256, 1, x_255); +x_257 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_258 = l_Lean_Syntax_node6(x_225, x_257, x_230, x_235, x_246, x_249, x_254, x_256); +x_172 = x_258; +x_173 = x_227; +goto block_189; +} +block_189: +{ +lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_181; +x_174 = l_Lean_Expr_const___override(x_2, x_10); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_174); +x_176 = 1; +x_177 = lean_box(x_176); +x_178 = lean_box(x_176); +x_179 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_179, 0, x_172); +lean_closure_set(x_179, 1, x_175); +lean_closure_set(x_179, 2, x_177); +lean_closure_set(x_179, 3, x_178); +lean_closure_set(x_179, 4, x_11); +x_180 = 1; lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -x_173 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_171, x_172, x_13, x_14, x_15, x_16, x_17, x_18, x_165); -if (lean_obj_tag(x_173) == 0) +x_181 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_179, x_180, x_13, x_14, x_15, x_16, x_17, x_18, x_173); +if (lean_obj_tag(x_181) == 0) { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); -lean_inc(x_175); -lean_dec(x_173); -x_176 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_174, x_13, x_14, x_15, x_16, x_17, x_18, x_175); +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_182, x_13, x_14, x_15, x_16, x_17, x_18, x_183); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -return x_176; +return x_184; } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_177 = lean_ctor_get(x_173, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_173, 1); -lean_inc(x_178); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_179 = x_173; +x_185 = lean_ctor_get(x_181, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_181, 1); +lean_inc(x_186); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_187 = x_181; } else { - lean_dec_ref(x_173); - x_179 = lean_box(0); + lean_dec_ref(x_181); + x_187 = lean_box(0); } -if (lean_is_scalar(x_179)) { - x_180 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_187)) { + x_188 = lean_alloc_ctor(1, 2, 0); } else { - x_180 = x_179; + x_188 = x_187; } -lean_ctor_set(x_180, 0, x_177); -lean_ctor_set(x_180, 1, x_178); -return x_180; +lean_ctor_set(x_188, 0, x_185); +lean_ctor_set(x_188, 1, x_186); +return x_188; } } } } else { -lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_247 = lean_ctor_get(x_21, 0); -lean_inc(x_247); +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_259 = lean_ctor_get(x_21, 0); +lean_inc(x_259); lean_dec(x_21); -x_248 = lean_ctor_get(x_22, 1); -lean_inc(x_248); +x_260 = lean_ctor_get(x_22, 1); +lean_inc(x_260); if (lean_is_exclusive(x_22)) { lean_ctor_release(x_22, 0); lean_ctor_release(x_22, 1); - x_249 = x_22; + x_261 = x_22; } else { lean_dec_ref(x_22); - x_249 = lean_box(0); -} -if (lean_obj_tag(x_248) == 0) -{ -lean_object* x_268; uint8_t x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; size_t x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -x_268 = lean_ctor_get(x_17, 5); -lean_inc(x_268); -x_269 = 0; -x_270 = l_Lean_SourceInfo_fromRef(x_268, x_269); -lean_dec(x_268); -x_271 = lean_st_ref_get(x_18, x_23); -x_272 = lean_ctor_get(x_271, 1); -lean_inc(x_272); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - lean_ctor_release(x_271, 1); - x_273 = x_271; + x_261 = lean_box(0); +} +if (lean_obj_tag(x_260) == 0) +{ +lean_object* x_280; uint8_t x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; size_t x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_280 = lean_ctor_get(x_17, 5); +lean_inc(x_280); +x_281 = 0; +x_282 = l_Lean_SourceInfo_fromRef(x_280, x_281); +lean_dec(x_280); +x_283 = lean_st_ref_get(x_18, x_23); +x_284 = lean_ctor_get(x_283, 1); +lean_inc(x_284); +if (lean_is_exclusive(x_283)) { + lean_ctor_release(x_283, 0); + lean_ctor_release(x_283, 1); + x_285 = x_283; } else { - lean_dec_ref(x_271); - x_273 = lean_box(0); + lean_dec_ref(x_283); + x_285 = lean_box(0); } -x_274 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_270); -if (lean_is_scalar(x_273)) { - x_275 = lean_alloc_ctor(2, 2, 0); +x_286 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_282); +if (lean_is_scalar(x_285)) { + x_287 = lean_alloc_ctor(2, 2, 0); } else { - x_275 = x_273; - lean_ctor_set_tag(x_275, 2); + x_287 = x_285; + lean_ctor_set_tag(x_287, 2); } -lean_ctor_set(x_275, 0, x_270); -lean_ctor_set(x_275, 1, x_274); -x_276 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_ctor_set(x_287, 0, x_282); +lean_ctor_set(x_287, 1, x_286); +x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; lean_inc(x_12); -lean_inc(x_270); -x_277 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_277, 0, x_270); -lean_ctor_set(x_277, 1, x_276); -lean_ctor_set(x_277, 2, x_12); -x_278 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +lean_inc(x_282); +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_282); +lean_ctor_set(x_289, 1, x_288); +lean_ctor_set(x_289, 2, x_12); +x_290 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); -if (lean_is_scalar(x_249)) { - x_279 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_261)) { + x_291 = lean_alloc_ctor(1, 2, 0); } else { - x_279 = x_249; - lean_ctor_set_tag(x_279, 1); -} -lean_ctor_set(x_279, 0, x_278); -lean_ctor_set(x_279, 1, x_10); -x_280 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_281 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_281, 0, x_280); -lean_ctor_set(x_281, 1, x_279); -x_282 = lean_array_size(x_247); -x_283 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_281, x_282, x_8, x_247); -lean_dec(x_281); -x_284 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_285 = l_Lean_Syntax_SepArray_ofElems(x_284, x_283); -lean_dec(x_283); -x_286 = l_Array_append___rarg(x_12, x_285); -lean_dec(x_285); -lean_inc(x_270); -x_287 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_287, 0, x_270); -lean_ctor_set(x_287, 1, x_276); -lean_ctor_set(x_287, 2, x_286); -x_288 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_277); -lean_inc(x_270); -x_289 = l_Lean_Syntax_node1(x_270, x_288, x_277); -x_290 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_270); -x_291 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_291, 0, x_270); -lean_ctor_set(x_291, 1, x_290); -x_292 = lean_box(0); + x_291 = x_261; + lean_ctor_set_tag(x_291, 1); +} +lean_ctor_set(x_291, 0, x_290); +lean_ctor_set(x_291, 1, x_10); +x_292 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_293 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_293, 0, x_292); +lean_ctor_set(x_293, 1, x_291); +x_294 = lean_array_size(x_259); +x_295 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__6(x_293, x_294, x_8, x_259); +lean_dec(x_293); +x_296 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_297 = l_Lean_Syntax_SepArray_ofElems(x_296, x_295); +lean_dec(x_295); +x_298 = l_Array_append___rarg(x_12, x_297); +lean_dec(x_297); +lean_inc(x_282); +x_299 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_299, 0, x_282); +lean_ctor_set(x_299, 1, x_288); +lean_ctor_set(x_299, 2, x_298); +x_300 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_282); +x_301 = l_Lean_Syntax_node1(x_282, x_300, x_299); +x_302 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_289); +lean_inc(x_282); +x_303 = l_Lean_Syntax_node1(x_282, x_302, x_289); +x_304 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_282); +x_305 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_305, 0, x_282); +lean_ctor_set(x_305, 1, x_304); +x_306 = lean_box(0); lean_inc(x_2); -x_293 = l_Lean_mkCIdentFrom(x_292, x_2, x_269); -lean_inc(x_270); -x_294 = l_Lean_Syntax_node2(x_270, x_276, x_291, x_293); -x_295 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_270); -x_296 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_296, 0, x_270); -lean_ctor_set(x_296, 1, x_295); -x_297 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_298 = l_Lean_Syntax_node6(x_270, x_297, x_275, x_277, x_287, x_289, x_294, x_296); -x_250 = x_298; -x_251 = x_272; -goto block_267; -} -else -{ -lean_object* x_299; lean_object* x_300; uint8_t x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; size_t x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_299 = lean_ctor_get(x_248, 0); -lean_inc(x_299); -lean_dec(x_248); -x_300 = lean_ctor_get(x_17, 5); -lean_inc(x_300); -x_301 = 0; -x_302 = l_Lean_SourceInfo_fromRef(x_300, x_301); -lean_dec(x_300); -x_303 = lean_st_ref_get(x_18, x_23); -x_304 = lean_ctor_get(x_303, 1); -lean_inc(x_304); -if (lean_is_exclusive(x_303)) { - lean_ctor_release(x_303, 0); - lean_ctor_release(x_303, 1); - x_305 = x_303; +x_307 = l_Lean_mkCIdentFrom(x_306, x_2, x_281); +lean_inc(x_282); +x_308 = l_Lean_Syntax_node2(x_282, x_288, x_305, x_307); +x_309 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_282); +x_310 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_310, 0, x_282); +lean_ctor_set(x_310, 1, x_309); +x_311 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_312 = l_Lean_Syntax_node6(x_282, x_311, x_287, x_289, x_301, x_303, x_308, x_310); +x_262 = x_312; +x_263 = x_284; +goto block_279; +} +else +{ +lean_object* x_313; lean_object* x_314; uint8_t x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; size_t x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +x_313 = lean_ctor_get(x_260, 0); +lean_inc(x_313); +lean_dec(x_260); +x_314 = lean_ctor_get(x_17, 5); +lean_inc(x_314); +x_315 = 0; +x_316 = l_Lean_SourceInfo_fromRef(x_314, x_315); +lean_dec(x_314); +x_317 = lean_st_ref_get(x_18, x_23); +x_318 = lean_ctor_get(x_317, 1); +lean_inc(x_318); +if (lean_is_exclusive(x_317)) { + lean_ctor_release(x_317, 0); + lean_ctor_release(x_317, 1); + x_319 = x_317; } else { - lean_dec_ref(x_303); - x_305 = lean_box(0); + lean_dec_ref(x_317); + x_319 = lean_box(0); } -x_306 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; -lean_inc(x_302); -if (lean_is_scalar(x_305)) { - x_307 = lean_alloc_ctor(2, 2, 0); +x_320 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__3; +lean_inc(x_316); +if (lean_is_scalar(x_319)) { + x_321 = lean_alloc_ctor(2, 2, 0); } else { - x_307 = x_305; - lean_ctor_set_tag(x_307, 2); -} -lean_ctor_set(x_307, 0, x_302); -lean_ctor_set(x_307, 1, x_306); -x_308 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; -lean_inc(x_302); -x_309 = l_Lean_Syntax_node1(x_302, x_308, x_299); -x_310 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; -lean_inc(x_302); -if (lean_is_scalar(x_249)) { - x_311 = lean_alloc_ctor(2, 2, 0); + x_321 = x_319; + lean_ctor_set_tag(x_321, 2); +} +lean_ctor_set(x_321, 0, x_316); +lean_ctor_set(x_321, 1, x_320); +x_322 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__7; +lean_inc(x_316); +x_323 = l_Lean_Syntax_node1(x_316, x_322, x_313); +x_324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15; +lean_inc(x_316); +if (lean_is_scalar(x_261)) { + x_325 = lean_alloc_ctor(2, 2, 0); } else { - x_311 = x_249; - lean_ctor_set_tag(x_311, 2); + x_325 = x_261; + lean_ctor_set_tag(x_325, 2); } -lean_ctor_set(x_311, 0, x_302); -lean_ctor_set(x_311, 1, x_310); -lean_inc(x_302); -x_312 = l_Lean_Syntax_node2(x_302, x_308, x_309, x_311); -x_313 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; +lean_ctor_set(x_325, 0, x_316); +lean_ctor_set(x_325, 1, x_324); +lean_inc(x_316); +x_326 = l_Lean_Syntax_node2(x_316, x_322, x_323, x_325); +x_327 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__1___closed__3; lean_inc(x_10); -x_314 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_10); -x_315 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; -x_316 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_316, 0, x_315); -lean_ctor_set(x_316, 1, x_314); -x_317 = lean_array_size(x_247); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_316, x_317, x_8, x_247); -lean_dec(x_316); -x_319 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__8; -x_320 = l_Lean_Syntax_SepArray_ofElems(x_319, x_318); -lean_dec(x_318); +x_328 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_10); +x_329 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__7; +x_330 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_328); +x_331 = lean_array_size(x_259); +x_332 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__7(x_330, x_331, x_8, x_259); +lean_dec(x_330); +x_333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_334 = l_Lean_Syntax_SepArray_ofElems(x_333, x_332); +lean_dec(x_332); lean_inc(x_12); -x_321 = l_Array_append___rarg(x_12, x_320); -lean_dec(x_320); -lean_inc(x_302); -x_322 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_322, 0, x_302); -lean_ctor_set(x_322, 1, x_308); -lean_ctor_set(x_322, 2, x_321); -lean_inc(x_302); -x_323 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_323, 0, x_302); -lean_ctor_set(x_323, 1, x_308); -lean_ctor_set(x_323, 2, x_12); -x_324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; -lean_inc(x_302); -x_325 = l_Lean_Syntax_node1(x_302, x_324, x_323); -x_326 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; -lean_inc(x_302); -x_327 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_327, 0, x_302); -lean_ctor_set(x_327, 1, x_326); -x_328 = lean_box(0); +x_335 = l_Array_append___rarg(x_12, x_334); +lean_dec(x_334); +lean_inc(x_316); +x_336 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_336, 0, x_316); +lean_ctor_set(x_336, 1, x_322); +lean_ctor_set(x_336, 2, x_335); +x_337 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__5; +lean_inc(x_316); +x_338 = l_Lean_Syntax_node1(x_316, x_337, x_336); +lean_inc(x_316); +x_339 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_339, 0, x_316); +lean_ctor_set(x_339, 1, x_322); +lean_ctor_set(x_339, 2, x_12); +x_340 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +lean_inc(x_316); +x_341 = l_Lean_Syntax_node1(x_316, x_340, x_339); +x_342 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; +lean_inc(x_316); +x_343 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_343, 0, x_316); +lean_ctor_set(x_343, 1, x_342); +x_344 = lean_box(0); lean_inc(x_2); -x_329 = l_Lean_mkCIdentFrom(x_328, x_2, x_301); -lean_inc(x_302); -x_330 = l_Lean_Syntax_node2(x_302, x_308, x_327, x_329); -x_331 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; -lean_inc(x_302); -x_332 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_332, 0, x_302); -lean_ctor_set(x_332, 1, x_331); -x_333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; -x_334 = l_Lean_Syntax_node6(x_302, x_333, x_307, x_312, x_322, x_325, x_330, x_332); -x_250 = x_334; -x_251 = x_304; -goto block_267; -} -block_267: -{ -lean_object* x_252; lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; uint8_t x_258; lean_object* x_259; -x_252 = l_Lean_Expr_const___override(x_2, x_10); -x_253 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_253, 0, x_252); -x_254 = 1; -x_255 = lean_box(x_254); -x_256 = lean_box(x_254); -x_257 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_257, 0, x_250); -lean_closure_set(x_257, 1, x_253); -lean_closure_set(x_257, 2, x_255); -lean_closure_set(x_257, 3, x_256); -lean_closure_set(x_257, 4, x_11); -x_258 = 1; +x_345 = l_Lean_mkCIdentFrom(x_344, x_2, x_315); +lean_inc(x_316); +x_346 = l_Lean_Syntax_node2(x_316, x_322, x_343, x_345); +x_347 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; +lean_inc(x_316); +x_348 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_348, 0, x_316); +lean_ctor_set(x_348, 1, x_347); +x_349 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__2; +x_350 = l_Lean_Syntax_node6(x_316, x_349, x_321, x_326, x_338, x_341, x_346, x_348); +x_262 = x_350; +x_263 = x_318; +goto block_279; +} +block_279: +{ +lean_object* x_264; lean_object* x_265; uint8_t x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; lean_object* x_271; +x_264 = l_Lean_Expr_const___override(x_2, x_10); +x_265 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_265, 0, x_264); +x_266 = 1; +x_267 = lean_box(x_266); +x_268 = lean_box(x_266); +x_269 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_269, 0, x_262); +lean_closure_set(x_269, 1, x_265); +lean_closure_set(x_269, 2, x_267); +lean_closure_set(x_269, 3, x_268); +lean_closure_set(x_269, 4, x_11); +x_270 = 1; lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); -x_259 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_257, x_258, x_13, x_14, x_15, x_16, x_17, x_18, x_251); -if (lean_obj_tag(x_259) == 0) +x_271 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_269, x_270, x_13, x_14, x_15, x_16, x_17, x_18, x_263); +if (lean_obj_tag(x_271) == 0) { -lean_object* x_260; lean_object* x_261; lean_object* x_262; -x_260 = lean_ctor_get(x_259, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_259, 1); -lean_inc(x_261); -lean_dec(x_259); -x_262 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_260, x_13, x_14, x_15, x_16, x_17, x_18, x_261); +lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_272 = lean_ctor_get(x_271, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_271, 1); +lean_inc(x_273); +lean_dec(x_271); +x_274 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_272, x_13, x_14, x_15, x_16, x_17, x_18, x_273); lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -return x_262; +return x_274; } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); -x_263 = lean_ctor_get(x_259, 0); -lean_inc(x_263); -x_264 = lean_ctor_get(x_259, 1); -lean_inc(x_264); -if (lean_is_exclusive(x_259)) { - lean_ctor_release(x_259, 0); - lean_ctor_release(x_259, 1); - x_265 = x_259; +x_275 = lean_ctor_get(x_271, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_271, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + lean_ctor_release(x_271, 1); + x_277 = x_271; } else { - lean_dec_ref(x_259); - x_265 = lean_box(0); + lean_dec_ref(x_271); + x_277 = lean_box(0); } -if (lean_is_scalar(x_265)) { - x_266 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_266 = x_265; + x_278 = x_277; } -lean_ctor_set(x_266, 0, x_263); -lean_ctor_set(x_266, 1, x_264); -return x_266; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } } } else { -uint8_t x_335; +uint8_t x_351; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -8145,23 +8235,23 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_2); -x_335 = !lean_is_exclusive(x_20); -if (x_335 == 0) +x_351 = !lean_is_exclusive(x_20); +if (x_351 == 0) { return x_20; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; -x_336 = lean_ctor_get(x_20, 0); -x_337 = lean_ctor_get(x_20, 1); -lean_inc(x_337); -lean_inc(x_336); +lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_352 = lean_ctor_get(x_20, 0); +x_353 = lean_ctor_get(x_20, 1); +lean_inc(x_353); +lean_inc(x_352); lean_dec(x_20); -x_338 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_338, 0, x_336); -lean_ctor_set(x_338, 1, x_337); -return x_338; +x_354 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); +return x_354; } } } @@ -11173,11 +11263,11 @@ x_17 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_17, 0, x_12); lean_ctor_set(x_17, 1, x_15); lean_ctor_set(x_17, 2, x_16); -x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__10; +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; lean_inc(x_17); lean_inc(x_12); x_19 = l_Lean_Syntax_node1(x_12, x_18, x_17); -x_20 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__11; +x_20 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13; lean_inc(x_12); x_21 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_21, 0, x_12); @@ -11185,7 +11275,7 @@ lean_ctor_set(x_21, 1, x_20); lean_inc(x_3); lean_inc(x_12); x_22 = l_Lean_Syntax_node2(x_12, x_15, x_21, x_3); -x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12; +x_23 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14; lean_inc(x_12); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_12); @@ -13249,6 +13339,10 @@ l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__12); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__13); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__14); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__2___closed__15); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__1(); lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__1); l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Tactic_Config_0__Lean_Elab_Tactic_elabConfig___spec__5___lambda__3___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Ext.c b/stage0/stdlib/Lean/Elab/Tactic/Ext.c index c1dedf5093fe..7145a91b68da 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Ext.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Ext.c @@ -13,50 +13,51 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_beqKey____x40_Lean_Meta_DiscrTreeTypes___hyg_101_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__1; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1(lean_object*, size_t, size_t, uint64_t); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__12; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__2(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___closed__1; -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20; size_t lean_usize_shift_right(size_t, size_t); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__32; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__2___closed__1; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,19 +65,18 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__35; lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_declRangeExt; lean_object* l_Lean_Elab_Tactic_RCases_rintro(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__59; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__25; @@ -92,32 +92,32 @@ static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___closed__2; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__55; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___closed__5; uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___at_Lean_Elab_Tactic_Ext_withExtN___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIff(lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__22; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3(lean_object*, size_t, lean_object*); lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__44; static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__4; static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -126,196 +126,200 @@ lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_ob lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_instBEqExtTheorem___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__4; size_t lean_usize_mul(size_t, size_t); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_instInhabited(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__21; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Elab_Tactic_Ext_extExt_config___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_getExtTheorems___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_Trie_foldValuesM___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6; lean_object* l_Lean_MVarId_getType_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__3; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_eraseCore(lean_object*, lean_object*); -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__3(size_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getMainGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___at_Lean_Elab_Tactic_Ext_extCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__19; lean_object* lean_environment_find(lean_object*, lean_object*); lean_object* l_ReaderT_read___at_Lean_Macro_instMonadRefMacroM___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_docString__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4; lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Tactic_Ext_mkExtType___spec__1(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__11; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Ext_evalExt___spec__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__36; static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__2___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__1(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_removeUnused___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1; lean_object* l_Lean_Exception_toMessageData(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__10; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__7; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1; lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__56; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___at_Lean_Elab_Tactic_Ext_extCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__21; lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__4___closed__1; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_extCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__4; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____boxed(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___at_Lean_Elab_Tactic_Ext_withExtN___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12; static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__4___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__2(lean_object*, uint8_t, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3; lean_object* l_Lean_Elab_Tactic_RCases_expandRIntroPats(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__5; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21; lean_object* l_Lean_FVarId_getBinderInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__34; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22; static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Ext_evalExt___spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3___closed__1; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_getExtTheorems(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___closed__6; lean_object* l_Array_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__17; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__50; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1(lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__3; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__4___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__1; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorems___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____boxed(lean_object*, lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; -static lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__31; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__1; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__15; lean_object* lean_nat_div(lean_object*, lean_object*); @@ -324,74 +328,79 @@ static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___sp lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27; static lean_object* l_Lean_Elab_Tactic_Ext_instHashableExtTheorem___closed__1; lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__5; +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1___closed__2; LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Tactic_Ext_mkExtType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__18; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__6___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__8; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13; lean_object* l_Lean_MessageData_ofFormat(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__2(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__2; lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__33; -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__15; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__11; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19; lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps(lean_object*); lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__2___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__13; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__63; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1(lean_object*); -static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__60; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__6___closed__4; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_liftCommandElabM___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkAndN(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__3___closed__2; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__5; -LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10; lean_object* lean_array_to_list(lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_getExtTheorems___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); @@ -403,84 +412,80 @@ static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__57; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_DiscrTree_Key_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15; static lean_object* l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4; +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__13; static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18; +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__1(lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272_(lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5; -LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__4; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__54; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__3___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__2; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3; +static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorems; -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__14; lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__5; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__37; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__17; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__5; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Tactic_Ext_evalExt___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_extCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___at_Lean_Elab_Tactic_Ext_extCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__4; @@ -490,21 +495,27 @@ static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___close LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtType(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extExt_config; lean_object* l_Lean_Expr_appFn_x21(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_instHashableExtTheorem; lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__46; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__52; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__6(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___at_Lean_Elab_Tactic_Ext_withExtN___spec__3(lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__53; lean_object* l_Lean_Meta_DiscrTree_mkPath(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); @@ -514,46 +525,45 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___at_Lean_Elab_Tactic_E lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Meta_mkSimpCongrTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__2(lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_withExtHyps___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384_(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__47; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__8; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_insertionSort_swapLoop___at_Lean_Elab_Tactic_Ext_getExtTheorems___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExt1___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___at_Lean_Elab_Tactic_Ext_extCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__1___closed__2; lean_object* l_instMonadControlTOfPure___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__2___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6; lean_object* l_Lean_throwError___at_Lean_Meta_mkSimpCongrTheorem___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__1(lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_getExtTheorems___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21; static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExt1___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__43; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; @@ -561,34 +571,40 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___c static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__2(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1___closed__4; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6; static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_tryIntros___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__3; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__26; LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__2(lean_object*, lean_object*, size_t, size_t, uint8_t); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__4(lean_object*); -static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25; static lean_object* l_Lean_Elab_Tactic_Ext_getExtTheorems___closed__2; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__7; lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_DiscrTree_instInhabitedKey; static lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__6___closed__1; static lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__2; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11; uint8_t l_Lean_Name_isAtomic(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__16; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__4___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__2; lean_object* lean_string_length(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__6; -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__2(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__6; lean_object* l_Lean_TSyntax_getNat(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -599,54 +615,55 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__65; static lean_object* l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__3(lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Elab_Tactic_Ext_evalExt___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__58; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Ext_evalExt___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__1; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7; lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___at_Lean_Elab_Tactic_Ext_extCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange_x3f___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__6; lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_reprKey____x40_Lean_Meta_DiscrTreeTypes___hyg_355_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__39; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem___boxed(lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__2; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1; uint64_t l_Lean_Name_hash___override(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679_(lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__4___closed__2; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem(lean_object*); LEAN_EXPORT lean_object* l_Array_insertionSort_traverse___at_Lean_Elab_Tactic_Ext_getExtTheorems___spec__1(lean_object*, lean_object*, lean_object*); @@ -659,72 +676,66 @@ lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object lean_object* l_Array_insertAt_x21___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLevelErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1; uint8_t l_Lean_isStructure(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8; lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____boxed(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__2___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_DiscrTree_Trie_foldValuesM___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__1(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__6___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__3; lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Tactic_RCases_linter_unusedRCasesPattern; -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__49; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1; extern lean_object* l_Lean_Core_instMonadCoreM; lean_object* l_Lean_MVarId_apply(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15; static lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___closed__1; extern lean_object* l_Lean_protectedExt; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9; extern lean_object* l_Lean_instInhabitedDeclarationRanges; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12(lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__4(lean_object*, lean_object*, size_t, size_t, uint8_t); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___at_Lean_Elab_Tactic_Ext_extCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__42; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_extCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalRepeat_x27___spec__1___rarg(lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11; +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__45; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extExtension; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___rarg___lambda__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_extCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -732,76 +743,74 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalApplyExtTheorem___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExt1___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint64_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266_(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__27; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__19; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Ext_withExtHyps___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___closed__3; +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1(lean_object*, size_t, size_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRangesFromSyntax___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__64; +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__4___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2; +static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__29; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__6; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__5; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__62; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5; +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__4; lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__2; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__7; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_evalExt___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__41; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__40; lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_mkExtType___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_instBEqExtTheorem; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__48; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__3; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30; lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_arrowDomainsN___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExtN(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____lambda__1(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_withExt1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__5(lean_object*, lean_object*, size_t, size_t, uint8_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_withExtHyps___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); @@ -811,29 +820,21 @@ lean_object* l_Lean_MVarId_intros(lean_object*, lean_object*, lean_object*, lean lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_addInstanceEntry___spec__18(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); +static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19; uint8_t l_Lean_Exception_isRuntime(lean_object*); -static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__6; -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7; lean_object* l_Lean_Meta_DiscrTree_getMatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__9; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__30; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem_declRange__1___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__61; -static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15; static lean_object* l_Lean_Elab_Tactic_Ext_extCore___closed__1; uint8_t l_Lean_Expr_isForall(lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_tryIntros___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673_(lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29; -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__3___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__20; @@ -847,22 +848,23 @@ lean_object* l_Lean_Elab_Tactic_replaceMainGoal(lean_object*, lean_object*, lean lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__7; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Tactic_Ext_extCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5; uint64_t l_Lean_Meta_DiscrTree_Key_hash(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__6; -LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__22; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__1; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7; +static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2; lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Match_instInhabitedMatchEqnsExtState___spec__1; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_applyExtTheoremAt___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_mkExtIffType___spec__1___closed__2; static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___lambda__1___closed__28; @@ -870,9 +872,9 @@ lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(l static lean_object* l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt_declRange__1___closed__7; static lean_object* l_Lean_Elab_Tactic_Ext_mkExtIffType___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25; LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__3(lean_object*, lean_object*, size_t, size_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104_(lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_realizeExtIffTheorem___closed__3; lean_object* l_Lean_Elab_Term_withDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4483,7 +4485,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -4503,19 +4505,39 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2__ _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("..", 2, 2); +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; +x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__3; +x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__24; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("..", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28() { _start: { lean_object* x_1; @@ -4523,7 +4545,7 @@ x_1 = lean_mk_string_unchecked(";", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29() { _start: { lean_object* x_1; @@ -4531,19 +4553,19 @@ x_1 = lean_mk_string_unchecked("intros", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31() { _start: { lean_object* x_1; @@ -4551,19 +4573,19 @@ x_1 = lean_mk_string_unchecked("substEqs", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; +x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33() { _start: { lean_object* x_1; @@ -4571,7 +4593,7 @@ x_1 = lean_mk_string_unchecked("subst_eqs", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34() { _start: { lean_object* x_1; @@ -4579,19 +4601,19 @@ x_1 = lean_mk_string_unchecked("tacticRfl", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32; +x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36() { _start: { lean_object* x_1; @@ -4642,7 +4664,7 @@ x_25 = lean_st_ref_get(x_8, x_18); x_26 = !lean_is_exclusive(x_25); if (x_26 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; x_27 = lean_ctor_get(x_25, 1); x_28 = lean_ctor_get(x_25, 0); lean_dec(x_28); @@ -4670,330 +4692,342 @@ x_37 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_37, 0, x_15); lean_ctor_set(x_37, 1, x_36); lean_ctor_set(x_37, 2, x_32); -x_38 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__24; +x_38 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +lean_inc(x_37); lean_inc(x_15); -x_39 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_39, 0, x_15); -lean_ctor_set(x_39, 1, x_38); +x_39 = l_Lean_Syntax_node1(x_15, x_38, x_37); +x_40 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; lean_inc(x_15); -x_40 = l_Lean_Syntax_node1(x_15, x_36, x_39); -x_41 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_15); +lean_ctor_set(x_41, 1, x_40); lean_inc(x_15); -x_42 = l_Lean_Syntax_node1(x_15, x_41, x_40); +x_42 = l_Lean_Syntax_node1(x_15, x_36, x_41); x_43 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; lean_inc(x_15); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_15); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; -lean_inc_n(x_37, 3); +x_44 = l_Lean_Syntax_node1(x_15, x_43, x_42); +x_45 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; lean_inc(x_15); -x_46 = l_Lean_Syntax_node6(x_15, x_45, x_35, x_37, x_37, x_42, x_37, x_44); -lean_inc(x_46); -x_47 = lean_array_push(x_33, x_46); -x_48 = lean_array_push(x_47, x_46); +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_15); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; +lean_inc_n(x_37, 2); lean_inc(x_15); -x_49 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_49, 0, x_15); -lean_ctor_set(x_49, 1, x_36); -lean_ctor_set(x_49, 2, x_48); -x_50 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; +x_48 = l_Lean_Syntax_node6(x_15, x_47, x_35, x_37, x_39, x_44, x_37, x_46); +lean_inc(x_48); +x_49 = lean_array_push(x_33, x_48); +x_50 = lean_array_push(x_49, x_48); lean_inc(x_15); -x_51 = l_Lean_Syntax_node2(x_15, x_50, x_31, x_49); -x_52 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; +x_51 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_51, 0, x_15); +lean_ctor_set(x_51, 1, x_36); +lean_ctor_set(x_51, 2, x_50); +x_52 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; lean_inc(x_15); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_15); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +x_53 = l_Lean_Syntax_node2(x_15, x_52, x_31, x_51); +x_54 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; lean_inc(x_15); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_15); lean_ctor_set(x_55, 1, x_54); -x_56 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; +x_56 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; lean_inc(x_15); -x_57 = l_Lean_Syntax_node2(x_15, x_56, x_55, x_37); -x_58 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31; +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_15); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; lean_inc(x_15); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_15); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; +x_59 = l_Lean_Syntax_node2(x_15, x_58, x_57, x_37); +x_60 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; lean_inc(x_15); -x_61 = l_Lean_Syntax_node1(x_15, x_60, x_59); -x_62 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_15); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32; lean_inc(x_15); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_15); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; +x_63 = l_Lean_Syntax_node1(x_15, x_62, x_61); +x_64 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; lean_inc(x_15); -x_65 = l_Lean_Syntax_node1(x_15, x_64, x_63); -lean_inc_n(x_53, 2); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_15); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; lean_inc(x_15); -x_66 = l_Lean_Syntax_node7(x_15, x_36, x_51, x_53, x_57, x_53, x_61, x_53, x_65); -x_67 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; +x_67 = l_Lean_Syntax_node1(x_15, x_66, x_65); +lean_inc_n(x_55, 2); lean_inc(x_15); -x_68 = l_Lean_Syntax_node1(x_15, x_67, x_66); -x_69 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +x_68 = l_Lean_Syntax_node7(x_15, x_36, x_53, x_55, x_59, x_55, x_63, x_55, x_67); +x_69 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; lean_inc(x_15); x_70 = l_Lean_Syntax_node1(x_15, x_69, x_68); -x_71 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; -x_72 = l_Lean_Syntax_node2(x_15, x_71, x_25, x_70); -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_2); -x_74 = lean_box(0); -x_75 = 1; -x_76 = l_Lean_Elab_Term_elabTermEnsuringType(x_72, x_73, x_75, x_14, x_74, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -return x_76; +x_71 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +lean_inc(x_15); +x_72 = l_Lean_Syntax_node1(x_15, x_71, x_70); +x_73 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; +x_74 = l_Lean_Syntax_node2(x_15, x_73, x_25, x_72); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_2); +x_76 = lean_box(0); +x_77 = 1; +x_78 = l_Lean_Elab_Term_elabTermEnsuringType(x_74, x_75, x_77, x_14, x_76, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +return x_78; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; -x_77 = lean_ctor_get(x_25, 1); -lean_inc(x_77); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; +x_79 = lean_ctor_get(x_25, 1); +lean_inc(x_79); lean_dec(x_25); -x_78 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; -lean_inc(x_15); -x_79 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_79, 0, x_15); -lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__17; +x_80 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; lean_inc(x_15); x_81 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_81, 0, x_15); lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__7; -x_83 = l_Array_append___rarg(x_82, x_24); -lean_dec(x_24); -x_84 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__21; +x_82 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__17; lean_inc(x_15); -x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_15); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__16; +x_83 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_83, 0, x_15); +lean_ctor_set(x_83, 1, x_82); +x_84 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__7; +x_85 = l_Array_append___rarg(x_84, x_24); +lean_dec(x_24); +x_86 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__21; lean_inc(x_15); -x_87 = lean_alloc_ctor(1, 3, 0); +x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_15); lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_82); -x_88 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__24; +x_88 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__16; lean_inc(x_15); -x_89 = lean_alloc_ctor(2, 2, 0); +x_89 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_89, 0, x_15); lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_89, 2, x_84); +x_90 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +lean_inc(x_89); lean_inc(x_15); -x_90 = l_Lean_Syntax_node1(x_15, x_86, x_89); -x_91 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +x_91 = l_Lean_Syntax_node1(x_15, x_90, x_89); +x_92 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; lean_inc(x_15); -x_92 = l_Lean_Syntax_node1(x_15, x_91, x_90); -x_93 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_15); +lean_ctor_set(x_93, 1, x_92); lean_inc(x_15); -x_94 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_94, 0, x_15); -lean_ctor_set(x_94, 1, x_93); -x_95 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; -lean_inc_n(x_87, 3); +x_94 = l_Lean_Syntax_node1(x_15, x_88, x_93); +x_95 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; lean_inc(x_15); -x_96 = l_Lean_Syntax_node6(x_15, x_95, x_85, x_87, x_87, x_92, x_87, x_94); -lean_inc(x_96); -x_97 = lean_array_push(x_83, x_96); -x_98 = lean_array_push(x_97, x_96); +x_96 = l_Lean_Syntax_node1(x_15, x_95, x_94); +x_97 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; lean_inc(x_15); -x_99 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_99, 0, x_15); -lean_ctor_set(x_99, 1, x_86); -lean_ctor_set(x_99, 2, x_98); -x_100 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_15); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; +lean_inc_n(x_89, 2); lean_inc(x_15); -x_101 = l_Lean_Syntax_node2(x_15, x_100, x_81, x_99); -x_102 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; +x_100 = l_Lean_Syntax_node6(x_15, x_99, x_87, x_89, x_91, x_96, x_89, x_98); +lean_inc(x_100); +x_101 = lean_array_push(x_85, x_100); +x_102 = lean_array_push(x_101, x_100); lean_inc(x_15); -x_103 = lean_alloc_ctor(2, 2, 0); +x_103 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_103, 0, x_15); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +lean_ctor_set(x_103, 1, x_88); +lean_ctor_set(x_103, 2, x_102); +x_104 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; lean_inc(x_15); -x_105 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_105, 0, x_15); -lean_ctor_set(x_105, 1, x_104); +x_105 = l_Lean_Syntax_node2(x_15, x_104, x_83, x_103); x_106 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; lean_inc(x_15); -x_107 = l_Lean_Syntax_node2(x_15, x_106, x_105, x_87); -x_108 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31; +x_107 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_107, 0, x_15); +lean_ctor_set(x_107, 1, x_106); +x_108 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; lean_inc(x_15); x_109 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_109, 0, x_15); lean_ctor_set(x_109, 1, x_108); x_110 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; lean_inc(x_15); -x_111 = l_Lean_Syntax_node1(x_15, x_110, x_109); -x_112 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; +x_111 = l_Lean_Syntax_node2(x_15, x_110, x_109, x_89); +x_112 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; lean_inc(x_15); x_113 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_113, 0, x_15); lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; +x_114 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32; lean_inc(x_15); x_115 = l_Lean_Syntax_node1(x_15, x_114, x_113); -lean_inc_n(x_103, 2); +x_116 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; lean_inc(x_15); -x_116 = l_Lean_Syntax_node7(x_15, x_86, x_101, x_103, x_107, x_103, x_111, x_103, x_115); -x_117 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; +x_117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_117, 0, x_15); +lean_ctor_set(x_117, 1, x_116); +x_118 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; lean_inc(x_15); -x_118 = l_Lean_Syntax_node1(x_15, x_117, x_116); -x_119 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +x_119 = l_Lean_Syntax_node1(x_15, x_118, x_117); +lean_inc_n(x_107, 2); lean_inc(x_15); -x_120 = l_Lean_Syntax_node1(x_15, x_119, x_118); -x_121 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; -x_122 = l_Lean_Syntax_node2(x_15, x_121, x_79, x_120); -x_123 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_123, 0, x_2); -x_124 = lean_box(0); -x_125 = 1; -x_126 = l_Lean_Elab_Term_elabTermEnsuringType(x_122, x_123, x_125, x_14, x_124, x_3, x_4, x_5, x_6, x_7, x_8, x_77); -return x_126; +x_120 = l_Lean_Syntax_node7(x_15, x_88, x_105, x_107, x_111, x_107, x_115, x_107, x_119); +x_121 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; +lean_inc(x_15); +x_122 = l_Lean_Syntax_node1(x_15, x_121, x_120); +x_123 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +lean_inc(x_15); +x_124 = l_Lean_Syntax_node1(x_15, x_123, x_122); +x_125 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; +x_126 = l_Lean_Syntax_node2(x_15, x_125, x_81, x_124); +x_127 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_127, 0, x_2); +x_128 = lean_box(0); +x_129 = 1; +x_130 = l_Lean_Elab_Term_elabTermEnsuringType(x_126, x_127, x_129, x_14, x_128, x_3, x_4, x_5, x_6, x_7, x_8, x_79); +return x_130; } } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; lean_object* x_185; -x_127 = lean_ctor_get(x_16, 1); -lean_inc(x_127); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; lean_object* x_191; +x_131 = lean_ctor_get(x_16, 1); +lean_inc(x_131); lean_dec(x_16); -x_128 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__6; +x_132 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__6; lean_inc(x_15); -x_129 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_129, 0, x_15); -lean_ctor_set(x_129, 1, x_128); -x_130 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__5; +x_133 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_133, 0, x_15); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__5; lean_inc(x_15); -x_131 = l_Lean_Syntax_node1(x_15, x_130, x_129); -x_132 = lean_ctor_get(x_11, 1); -lean_inc(x_132); +x_135 = l_Lean_Syntax_node1(x_15, x_134, x_133); +x_136 = lean_ctor_get(x_11, 1); +lean_inc(x_136); lean_dec(x_11); -x_133 = lean_mk_array(x_132, x_131); -x_134 = lean_st_ref_get(x_8, x_127); -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_136 = x_134; +x_137 = lean_mk_array(x_136, x_135); +x_138 = lean_st_ref_get(x_8, x_131); +x_139 = lean_ctor_get(x_138, 1); +lean_inc(x_139); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_140 = x_138; } else { - lean_dec_ref(x_134); - x_136 = lean_box(0); + lean_dec_ref(x_138); + x_140 = lean_box(0); } -x_137 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; +x_141 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__9; lean_inc(x_15); -if (lean_is_scalar(x_136)) { - x_138 = lean_alloc_ctor(2, 2, 0); +if (lean_is_scalar(x_140)) { + x_142 = lean_alloc_ctor(2, 2, 0); } else { - x_138 = x_136; - lean_ctor_set_tag(x_138, 2); + x_142 = x_140; + lean_ctor_set_tag(x_142, 2); } -lean_ctor_set(x_138, 0, x_15); -lean_ctor_set(x_138, 1, x_137); -x_139 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__17; -lean_inc(x_15); -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_15); -lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__7; -x_142 = l_Array_append___rarg(x_141, x_133); -lean_dec(x_133); -x_143 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__21; +lean_ctor_set(x_142, 0, x_15); +lean_ctor_set(x_142, 1, x_141); +x_143 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__17; lean_inc(x_15); x_144 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_144, 0, x_15); lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__16; -lean_inc(x_15); -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_15); -lean_ctor_set(x_146, 1, x_145); -lean_ctor_set(x_146, 2, x_141); -x_147 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__24; +x_145 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__7; +x_146 = l_Array_append___rarg(x_145, x_137); +lean_dec(x_137); +x_147 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__21; lean_inc(x_15); x_148 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_148, 0, x_15); lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__16; lean_inc(x_15); -x_149 = l_Lean_Syntax_node1(x_15, x_145, x_148); -x_150 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +x_150 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_150, 0, x_15); +lean_ctor_set(x_150, 1, x_149); +lean_ctor_set(x_150, 2, x_145); +x_151 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__23; +lean_inc(x_150); lean_inc(x_15); -x_151 = l_Lean_Syntax_node1(x_15, x_150, x_149); -x_152 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; +x_152 = l_Lean_Syntax_node1(x_15, x_151, x_150); +x_153 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; lean_inc(x_15); -x_153 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_153, 0, x_15); -lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; -lean_inc_n(x_146, 3); +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_15); +lean_ctor_set(x_154, 1, x_153); lean_inc(x_15); -x_155 = l_Lean_Syntax_node6(x_15, x_154, x_144, x_146, x_146, x_151, x_146, x_153); -lean_inc(x_155); -x_156 = lean_array_push(x_142, x_155); -x_157 = lean_array_push(x_156, x_155); +x_155 = l_Lean_Syntax_node1(x_15, x_149, x_154); +x_156 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__25; lean_inc(x_15); -x_158 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_158, 0, x_15); -lean_ctor_set(x_158, 1, x_145); -lean_ctor_set(x_158, 2, x_157); -x_159 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; +x_157 = l_Lean_Syntax_node1(x_15, x_156, x_155); +x_158 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; lean_inc(x_15); -x_160 = l_Lean_Syntax_node2(x_15, x_159, x_140, x_158); -x_161 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; +x_159 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_159, 0, x_15); +lean_ctor_set(x_159, 1, x_158); +x_160 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__20; +lean_inc_n(x_150, 2); lean_inc(x_15); -x_162 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_162, 0, x_15); -lean_ctor_set(x_162, 1, x_161); -x_163 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +x_161 = l_Lean_Syntax_node6(x_15, x_160, x_148, x_150, x_152, x_157, x_150, x_159); +lean_inc(x_161); +x_162 = lean_array_push(x_146, x_161); +x_163 = lean_array_push(x_162, x_161); lean_inc(x_15); -x_164 = lean_alloc_ctor(2, 2, 0); +x_164 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_164, 0, x_15); -lean_ctor_set(x_164, 1, x_163); -x_165 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; +lean_ctor_set(x_164, 1, x_149); +lean_ctor_set(x_164, 2, x_163); +x_165 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; lean_inc(x_15); -x_166 = l_Lean_Syntax_node2(x_15, x_165, x_164, x_146); -x_167 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__31; +x_166 = l_Lean_Syntax_node2(x_15, x_165, x_144, x_164); +x_167 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; lean_inc(x_15); x_168 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_168, 0, x_15); lean_ctor_set(x_168, 1, x_167); -x_169 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; +x_169 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; lean_inc(x_15); -x_170 = l_Lean_Syntax_node1(x_15, x_169, x_168); -x_171 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; +x_170 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_170, 0, x_15); +lean_ctor_set(x_170, 1, x_169); +x_171 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; lean_inc(x_15); -x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_15); -lean_ctor_set(x_172, 1, x_171); +x_172 = l_Lean_Syntax_node2(x_15, x_171, x_170, x_150); x_173 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; lean_inc(x_15); -x_174 = l_Lean_Syntax_node1(x_15, x_173, x_172); -lean_inc_n(x_162, 2); +x_174 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_174, 0, x_15); +lean_ctor_set(x_174, 1, x_173); +x_175 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__32; +lean_inc(x_15); +x_176 = l_Lean_Syntax_node1(x_15, x_175, x_174); +x_177 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; +lean_inc(x_15); +x_178 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_178, 0, x_15); +lean_ctor_set(x_178, 1, x_177); +x_179 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; lean_inc(x_15); -x_175 = l_Lean_Syntax_node7(x_15, x_145, x_160, x_162, x_166, x_162, x_170, x_162, x_174); -x_176 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; +x_180 = l_Lean_Syntax_node1(x_15, x_179, x_178); +lean_inc_n(x_168, 2); lean_inc(x_15); -x_177 = l_Lean_Syntax_node1(x_15, x_176, x_175); -x_178 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +x_181 = l_Lean_Syntax_node7(x_15, x_149, x_166, x_168, x_172, x_168, x_176, x_168, x_180); +x_182 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__14; lean_inc(x_15); -x_179 = l_Lean_Syntax_node1(x_15, x_178, x_177); -x_180 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; -x_181 = l_Lean_Syntax_node2(x_15, x_180, x_138, x_179); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_2); -x_183 = lean_box(0); -x_184 = 1; -x_185 = l_Lean_Elab_Term_elabTermEnsuringType(x_181, x_182, x_184, x_14, x_183, x_3, x_4, x_5, x_6, x_7, x_8, x_135); -return x_185; +x_183 = l_Lean_Syntax_node1(x_15, x_182, x_181); +x_184 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__12; +lean_inc(x_15); +x_185 = l_Lean_Syntax_node1(x_15, x_184, x_183); +x_186 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__8; +x_187 = l_Lean_Syntax_node2(x_15, x_186, x_142, x_185); +x_188 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_188, 0, x_2); +x_189 = lean_box(0); +x_190 = 1; +x_191 = l_Lean_Elab_Term_elabTermEnsuringType(x_187, x_188, x_190, x_14, x_189, x_3, x_4, x_5, x_6, x_7, x_8, x_139); +return x_191; } } else { -uint8_t x_186; +uint8_t x_192; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -5001,23 +5035,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_186 = !lean_is_exclusive(x_10); -if (x_186 == 0) +x_192 = !lean_is_exclusive(x_10); +if (x_192 == 0) { return x_10; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_10, 0); -x_188 = lean_ctor_get(x_10, 1); -lean_inc(x_188); -lean_inc(x_187); +lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_193 = lean_ctor_get(x_10, 0); +x_194 = lean_ctor_get(x_10, 1); +lean_inc(x_194); +lean_inc(x_193); lean_dec(x_10); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -return x_189; +x_195 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +return x_195; } } } @@ -6997,7 +7031,7 @@ lean_inc(x_12); lean_ctor_set_tag(x_14, 2); lean_ctor_set(x_14, 1, x_20); lean_ctor_set(x_14, 0, x_12); -x_21 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +x_21 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; lean_inc(x_12); x_22 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_22, 0, x_12); @@ -7009,7 +7043,7 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_12); lean_ctor_set(x_25, 1, x_23); lean_ctor_set(x_25, 2, x_24); -x_26 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; +x_26 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; lean_inc(x_25); lean_inc(x_12); x_27 = l_Lean_Syntax_node2(x_12, x_26, x_22, x_25); @@ -7092,7 +7126,7 @@ x_63 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; lean_inc(x_56); lean_inc(x_12); x_64 = l_Lean_Syntax_node2(x_12, x_63, x_56, x_62); -x_65 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; +x_65 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); @@ -7141,12 +7175,12 @@ lean_inc(x_12); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_12); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; +x_86 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; lean_inc(x_12); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_12); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; +x_88 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; lean_inc(x_12); x_89 = l_Lean_Syntax_node1(x_12, x_88, x_87); lean_inc(x_12); @@ -7367,7 +7401,7 @@ lean_inc(x_12); x_193 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_193, 0, x_12); lean_ctor_set(x_193, 1, x_192); -x_194 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__27; +x_194 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__29; lean_inc(x_12); x_195 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_195, 0, x_12); @@ -7379,7 +7413,7 @@ x_198 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_198, 0, x_12); lean_ctor_set(x_198, 1, x_196); lean_ctor_set(x_198, 2, x_197); -x_199 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; +x_199 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__30; lean_inc(x_198); lean_inc(x_12); x_200 = l_Lean_Syntax_node2(x_12, x_199, x_195, x_198); @@ -7462,7 +7496,7 @@ x_236 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__18; lean_inc(x_229); lean_inc(x_12); x_237 = l_Lean_Syntax_node2(x_12, x_236, x_229, x_235); -x_238 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__26; +x_238 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__28; lean_inc(x_12); x_239 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_239, 0, x_12); @@ -7511,12 +7545,12 @@ lean_inc(x_12); x_258 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_258, 0, x_12); lean_ctor_set(x_258, 1, x_257); -x_259 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34; +x_259 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36; lean_inc(x_12); x_260 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_260, 0, x_12); lean_ctor_set(x_260, 1, x_259); -x_261 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33; +x_261 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35; lean_inc(x_12); x_262 = l_Lean_Syntax_node1(x_12, x_261, x_260); lean_inc(x_12); @@ -8923,7 +8957,7 @@ x_1 = l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -8932,7 +8966,7 @@ x_3 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_reprKey____x40 return x_3; } } -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -8986,7 +9020,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9020,13 +9054,13 @@ lean_inc(x_8); lean_dec(x_1); x_9 = lean_unsigned_to_nat(0u); x_10 = l___private_Lean_Meta_DiscrTreeTypes_0__Lean_Meta_DiscrTree_reprKey____x40_Lean_Meta_DiscrTreeTypes___hyg_355_(x_8, x_9); -x_11 = l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__3(x_2, x_10, x_4); +x_11 = l_List_foldl___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__3(x_2, x_10, x_4); return x_11; } } } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1() { _start: { lean_object* x_1; @@ -9034,29 +9068,29 @@ x_1 = lean_mk_string_unchecked("declName", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2; +x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4() { _start: { lean_object* x_1; @@ -9064,29 +9098,29 @@ x_1 = lean_mk_string_unchecked(" := ", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3; -x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3; +x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -9095,7 +9129,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -9105,7 +9139,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9() { _start: { lean_object* x_1; @@ -9113,17 +9147,17 @@ x_1 = lean_mk_string_unchecked("priority", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11() { _start: { lean_object* x_1; @@ -9131,17 +9165,17 @@ x_1 = lean_mk_string_unchecked("keys", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -9150,7 +9184,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14() { _start: { lean_object* x_1; @@ -9158,35 +9192,35 @@ x_1 = lean_mk_string_unchecked("{ ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18() { _start: { lean_object* x_1; @@ -9194,21 +9228,21 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9216,7 +9250,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21() { _start: { lean_object* x_1; @@ -9224,35 +9258,35 @@ x_1 = lean_mk_string_unchecked("#[", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25() { _start: { lean_object* x_1; @@ -9260,17 +9294,17 @@ x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27() { _start: { lean_object* x_1; @@ -9278,33 +9312,33 @@ x_1 = lean_mk_string_unchecked("#[]", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13; -x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13; +x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30() { +static lean_object* _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29; +x_1 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -9312,7 +9346,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; @@ -9320,7 +9354,7 @@ x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_unsigned_to_nat(0u); x_5 = l_Lean_Name_reprPrec(x_3, x_4); -x_6 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7; +x_6 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -9328,11 +9362,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6; +x_10 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8; +x_12 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -9340,11 +9374,11 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10; +x_16 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5; +x_18 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5; x_19 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -9368,7 +9402,7 @@ lean_ctor_set(x_26, 1, x_12); x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_14); -x_28 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12; +x_28 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -9385,17 +9419,17 @@ if (x_33 == 0) { lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_34 = lean_array_to_list(x_31); -x_35 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20; -x_36 = l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____spec__2(x_34, x_35); -x_37 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24; +x_35 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20; +x_36 = l_Std_Format_joinSep___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____spec__2(x_34, x_35); +x_37 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24; x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_37); lean_ctor_set(x_38, 1, x_36); -x_39 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26; +x_39 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26; x_40 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); -x_41 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23; +x_41 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23; x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_40); @@ -9403,7 +9437,7 @@ x_43 = 1; x_44 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_44, 0, x_42); lean_ctor_set_uint8(x_44, sizeof(void*)*1, x_43); -x_45 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13; +x_45 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13; x_46 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_46, 0, x_45); lean_ctor_set(x_46, 1, x_44); @@ -9413,15 +9447,15 @@ lean_ctor_set_uint8(x_47, sizeof(void*)*1, x_8); x_48 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_48, 0, x_30); lean_ctor_set(x_48, 1, x_47); -x_49 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17; +x_49 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17; x_50 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19; +x_51 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19; x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16; +x_53 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16; x_54 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -9434,19 +9468,19 @@ else { lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_dec(x_31); -x_56 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30; +x_56 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30; x_57 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_57, 0, x_30); lean_ctor_set(x_57, 1, x_56); -x_58 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17; +x_58 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17; x_59 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_59, 0, x_58); lean_ctor_set(x_59, 1, x_57); -x_60 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19; +x_60 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19; x_61 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); -x_62 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16; +x_62 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16; x_63 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_63, 0, x_62); lean_ctor_set(x_63, 1, x_61); @@ -9457,11 +9491,11 @@ return x_64; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098_(x_1, x_2); +x_3 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -9470,7 +9504,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____boxed), 2, 0); return x_1; } } @@ -9482,7 +9516,7 @@ x_1 = l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; uint8_t x_9; @@ -9523,7 +9557,7 @@ return x_17; } } } -LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -9567,18 +9601,18 @@ return x_16; else { uint8_t x_17; -x_17 = l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1(x_5, x_8, lean_box(0), x_5, x_8, x_13, lean_box(0)); +x_17 = l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1(x_5, x_8, lean_box(0), x_5, x_8, x_13, lean_box(0)); return x_17; } } } } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; lean_object* x_9; -x_8 = l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Array_isEqvAux___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -9587,11 +9621,11 @@ x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176_(x_1, x_2); +x_3 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -9602,7 +9636,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_instBEqExtTheorem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182____boxed), 2, 0); return x_1; } } @@ -9614,7 +9648,7 @@ x_1 = l_Lean_Elab_Tactic_Ext_instBEqExtTheorem___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { _start: { uint8_t x_5; @@ -9638,7 +9672,7 @@ return x_4; } } } -LEAN_EXPORT uint64_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -9678,14 +9712,14 @@ size_t x_17; size_t x_18; uint64_t x_19; uint64_t x_20; x_17 = 0; x_18 = lean_usize_of_nat(x_11); lean_dec(x_11); -x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1(x_4, x_17, x_18, x_10); +x_19 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1(x_4, x_17, x_18, x_10); x_20 = lean_uint64_mix_hash(x_9, x_19); return x_20; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; uint64_t x_7; uint64_t x_8; lean_object* x_9; @@ -9695,17 +9729,17 @@ x_6 = lean_unbox_usize(x_3); lean_dec(x_3); x_7 = lean_unbox_uint64(x_4); lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____spec__1(x_1, x_5, x_6, x_7); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____spec__1(x_1, x_5, x_6, x_7); lean_dec(x_1); x_9 = lean_box_uint64(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266_(x_1); +x_2 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -9715,7 +9749,7 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_instHashableExtTheorem___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3266____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_hashExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3272____boxed), 1, 0); return x_1; } } @@ -9770,7 +9804,7 @@ x_1 = l_Lean_Elab_Tactic_Ext_extExt_config___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; uint8_t x_7; @@ -9812,7 +9846,7 @@ return x_15; } } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__1() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__1() { _start: { size_t x_1; size_t x_2; size_t x_3; @@ -9822,17 +9856,17 @@ x_3 = lean_usize_shift_left(x_1, x_2); return x_3; } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2() { _start: { size_t x_1; size_t x_2; size_t x_3; x_1 = 1; -x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__1; +x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__1; x_3 = lean_usize_sub(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -9844,7 +9878,7 @@ if (x_4 == 0) lean_object* x_5; size_t x_6; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_5 = lean_ctor_get(x_1, 0); x_6 = 5; -x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2; +x_7 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2; x_8 = lean_usize_land(x_2, x_7); x_9 = lean_usize_to_nat(x_8); x_10 = lean_box(2); @@ -9905,7 +9939,7 @@ x_20 = lean_ctor_get(x_1, 0); lean_inc(x_20); lean_dec(x_1); x_21 = 5; -x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2; +x_22 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2; x_23 = lean_usize_land(x_2, x_22); x_24 = lean_usize_to_nat(x_23); x_25 = lean_box(2); @@ -9967,24 +10001,24 @@ x_37 = lean_ctor_get(x_1, 1); lean_inc(x_37); lean_dec(x_1); x_38 = lean_unsigned_to_nat(0u); -x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4(x_36, x_37, lean_box(0), x_38, x_3); +x_39 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4(x_36, x_37, lean_box(0), x_38, x_3); lean_dec(x_37); lean_dec(x_36); return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2(lean_object* x_1, lean_object* x_2) { _start: { uint64_t x_3; size_t x_4; lean_object* x_5; x_3 = l_Lean_Meta_DiscrTree_Key_hash(x_2); x_4 = lean_uint64_to_usize(x_3); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3(x_1, x_4, x_2); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3(x_1, x_4, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; uint8_t x_8; @@ -10011,7 +10045,7 @@ x_17 = lean_usize_shift_right(x_12, x_16); x_18 = lean_unsigned_to_nat(1u); x_19 = lean_nat_add(x_5, x_18); lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_6, x_17, x_1, x_9, x_10); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_6, x_17, x_1, x_9, x_10); x_4 = lean_box(0); x_5 = x_19; x_6 = x_20; @@ -10019,7 +10053,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -10111,7 +10145,7 @@ return x_29; } } } -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1() { _start: { lean_object* x_1; @@ -10119,7 +10153,7 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10132,7 +10166,7 @@ lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_7 = lean_ctor_get(x_1, 0); x_8 = 1; x_9 = 5; -x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2; +x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2; x_11 = lean_usize_land(x_2, x_10); x_12 = lean_usize_to_nat(x_11); x_13 = lean_array_get_size(x_7); @@ -10232,7 +10266,7 @@ lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_3 x_35 = lean_ctor_get(x_15, 0); x_36 = lean_usize_shift_right(x_2, x_9); x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_35, x_36, x_37, x_4, x_5); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_35, x_36, x_37, x_4, x_5); lean_ctor_set(x_15, 0, x_38); x_39 = lean_array_fset(x_17, x_12, x_15); lean_dec(x_12); @@ -10247,7 +10281,7 @@ lean_inc(x_40); lean_dec(x_15); x_41 = lean_usize_shift_right(x_2, x_9); x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_40, x_41, x_42, x_4, x_5); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_40, x_41, x_42, x_4, x_5); x_44 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_44, 0, x_43); x_45 = lean_array_fset(x_17, x_12, x_44); @@ -10278,7 +10312,7 @@ lean_inc(x_48); lean_dec(x_1); x_49 = 1; x_50 = 5; -x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2; +x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2; x_52 = lean_usize_land(x_2, x_51); x_53 = lean_usize_to_nat(x_52); x_54 = lean_array_get_size(x_48); @@ -10363,7 +10397,7 @@ if (lean_is_exclusive(x_57)) { } x_73 = lean_usize_shift_right(x_2, x_50); x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_71, x_73, x_74, x_4, x_5); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_71, x_73, x_74, x_4, x_5); if (lean_is_scalar(x_72)) { x_76 = lean_alloc_ctor(1, 1, 0); } else { @@ -10400,7 +10434,7 @@ if (x_82 == 0) { lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__8(x_1, x_83, x_4, x_5); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__8(x_1, x_83, x_4, x_5); x_85 = 7; x_86 = lean_usize_dec_le(x_85, x_3); if (x_86 == 0) @@ -10418,8 +10452,8 @@ lean_inc(x_90); x_91 = lean_ctor_get(x_84, 1); lean_inc(x_91); lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7(x_3, x_90, x_91, lean_box(0), x_83, x_92); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7(x_3, x_90, x_91, lean_box(0), x_83, x_92); lean_dec(x_91); lean_dec(x_90); return x_93; @@ -10446,7 +10480,7 @@ x_96 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_96, 0, x_94); lean_ctor_set(x_96, 1, x_95); x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__8(x_96, x_97, x_4, x_5); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__8(x_96, x_97, x_4, x_5); x_99 = 7; x_100 = lean_usize_dec_le(x_99, x_3); if (x_100 == 0) @@ -10464,8 +10498,8 @@ lean_inc(x_104); x_105 = lean_ctor_get(x_98, 1); lean_inc(x_105); lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7(x_3, x_104, x_105, lean_box(0), x_97, x_106); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7(x_3, x_104, x_105, lean_box(0), x_97, x_106); lean_dec(x_105); lean_dec(x_104); return x_107; @@ -10483,18 +10517,18 @@ return x_98; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint64_t x_4; size_t x_5; size_t x_6; lean_object* x_7; x_4 = l_Lean_Meta_DiscrTree_Key_hash(x_2); x_5 = lean_uint64_to_usize(x_4); x_6 = 1; -x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_1, x_5, x_6, x_2, x_3); +x_7 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_1, x_5, x_6, x_2, x_3); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; @@ -10512,7 +10546,7 @@ else { lean_object* x_7; uint8_t x_8; x_7 = lean_array_fget(x_1, x_3); -x_8 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3176_(x_2, x_7); +x_8 = l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_beqExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3182_(x_2, x_7); lean_dec(x_7); if (x_8 == 0) { @@ -10533,7 +10567,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; @@ -10586,7 +10620,7 @@ x_24 = lean_ctor_get(x_19, 0); lean_dec(x_24); x_25 = lean_unsigned_to_nat(1u); x_26 = lean_nat_add(x_3, x_25); -x_27 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_26, x_23); +x_27 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_26, x_23); lean_dec(x_26); lean_ctor_set(x_19, 1, x_27); lean_ctor_set(x_19, 0, x_4); @@ -10602,7 +10636,7 @@ lean_inc(x_29); lean_dec(x_19); x_30 = lean_unsigned_to_nat(1u); x_31 = lean_nat_add(x_3, x_30); -x_32 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_31, x_29); +x_32 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_31, x_29); lean_dec(x_31); x_33 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_33, 0, x_4); @@ -10653,7 +10687,7 @@ return x_43; } } } -LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; @@ -10704,7 +10738,7 @@ x_21 = lean_ctor_get(x_16, 0); lean_dec(x_21); x_22 = lean_unsigned_to_nat(1u); x_23 = lean_nat_add(x_3, x_22); -x_24 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_23, x_20); +x_24 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_23, x_20); lean_dec(x_23); lean_ctor_set(x_16, 1, x_24); lean_ctor_set(x_16, 0, x_4); @@ -10719,7 +10753,7 @@ lean_inc(x_26); lean_dec(x_16); x_27 = lean_unsigned_to_nat(1u); x_28 = lean_nat_add(x_3, x_27); -x_29 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_28, x_26); +x_29 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_28, x_26); lean_dec(x_28); x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_4); @@ -10774,7 +10808,7 @@ x_44 = lean_ctor_get(x_40, 1); x_45 = lean_ctor_get(x_40, 0); lean_dec(x_45); x_46 = lean_nat_add(x_3, x_37); -x_47 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_46, x_44); +x_47 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_46, x_44); lean_dec(x_46); lean_ctor_set(x_40, 1, x_47); lean_ctor_set(x_40, 0, x_4); @@ -10789,7 +10823,7 @@ x_49 = lean_ctor_get(x_40, 1); lean_inc(x_49); lean_dec(x_40); x_50 = lean_nat_add(x_3, x_37); -x_51 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_50, x_49); +x_51 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_50, x_49); lean_dec(x_50); x_52 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_52, 0, x_4); @@ -10807,7 +10841,7 @@ x_54 = lean_array_get_size(x_5); x_55 = lean_unsigned_to_nat(1u); x_56 = lean_nat_sub(x_54, x_55); lean_dec(x_54); -x_57 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_56); +x_57 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_8, x_56); return x_57; } } @@ -10862,7 +10896,7 @@ return x_72; } } } -static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1() { +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -10873,7 +10907,7 @@ lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -10890,7 +10924,7 @@ if (x_9 == 0) { lean_object* x_10; lean_object* x_11; x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__10(x_6, x_2, x_10); +x_11 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__10(x_6, x_2, x_10); lean_ctor_set(x_4, 0, x_11); return x_4; } @@ -10898,12 +10932,12 @@ else { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_array_fget(x_1, x_3); -x_13 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1; +x_13 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1; lean_inc(x_12); x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11(x_1, x_2, x_3, x_12, x_7, x_14); +x_15 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11(x_1, x_2, x_3, x_12, x_7, x_14); lean_ctor_set(x_4, 1, x_15); return x_4; } @@ -10923,7 +10957,7 @@ if (x_19 == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = lean_unsigned_to_nat(0u); -x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__10(x_16, x_2, x_20); +x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__10(x_16, x_2, x_20); x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_17); @@ -10933,12 +10967,12 @@ else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_23 = lean_array_fget(x_1, x_3); -x_24 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1; +x_24 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1; lean_inc(x_23); x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11(x_1, x_2, x_3, x_23, x_17, x_25); +x_26 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11(x_1, x_2, x_3, x_23, x_17, x_25); x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_16); lean_ctor_set(x_27, 1, x_26); @@ -10947,7 +10981,7 @@ return x_27; } } } -static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1() { +static lean_object* _init_l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1() { _start: { lean_object* x_1; @@ -10955,16 +10989,16 @@ x_1 = l_Lean_Meta_DiscrTree_instInhabited(lean_box(0)); return x_1; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13(lean_object* x_1) { +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1; +x_2 = l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1; x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1() { +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1() { _start: { lean_object* x_1; @@ -10972,7 +11006,7 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.DiscrTree", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2() { +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2() { _start: { lean_object* x_1; @@ -10980,7 +11014,7 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.DiscrTree.insertCore", 30, 30); return x_1; } } -static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3() { +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3() { _start: { lean_object* x_1; @@ -10988,20 +11022,20 @@ x_1 = lean_mk_string_unchecked("invalid key sequence", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4() { +static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1; -x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2; +x_1 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1; +x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2; x_3 = lean_unsigned_to_nat(484u); x_4 = lean_unsigned_to_nat(23u); -x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3; +x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -11019,13 +11053,13 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = l_Lean_Meta_DiscrTree_instInhabitedKey; x_9 = l_outOfBounds___rarg(x_8); lean_inc(x_1); -x_10 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2(x_1, x_9); +x_10 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2(x_1, x_9); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_unsigned_to_nat(1u); x_12 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_2, x_3, x_11); -x_13 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(x_1, x_9, x_12); +x_13 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(x_1, x_9, x_12); return x_13; } else @@ -11035,8 +11069,8 @@ x_14 = lean_ctor_get(x_10, 0); lean_inc(x_14); lean_dec(x_10); x_15 = lean_unsigned_to_nat(1u); -x_16 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_2, x_3, x_15, x_14); -x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(x_1, x_9, x_16); +x_16 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_2, x_3, x_15, x_14); +x_17 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(x_1, x_9, x_16); return x_17; } } @@ -11045,13 +11079,13 @@ else lean_object* x_18; lean_object* x_19; x_18 = lean_array_fget(x_2, x_6); lean_inc(x_1); -x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2(x_1, x_18); +x_19 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2(x_1, x_18); if (lean_obj_tag(x_19) == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = lean_unsigned_to_nat(1u); x_21 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___rarg(x_2, x_3, x_20); -x_22 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(x_1, x_18, x_21); +x_22 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(x_1, x_18, x_21); return x_22; } else @@ -11061,8 +11095,8 @@ x_23 = lean_ctor_get(x_19, 0); lean_inc(x_23); lean_dec(x_19); x_24 = lean_unsigned_to_nat(1u); -x_25 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_2, x_3, x_24, x_23); -x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__5(x_1, x_18, x_25); +x_25 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_2, x_3, x_24, x_23); +x_26 = l_Lean_PersistentHashMap_insert___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__5(x_1, x_18, x_25); return x_26; } } @@ -11072,13 +11106,13 @@ else lean_object* x_27; lean_object* x_28; lean_dec(x_3); lean_dec(x_1); -x_27 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4; -x_28 = l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13(x_27); +x_27 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4; +x_28 = l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13(x_27); return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -11091,7 +11125,7 @@ x_5 = lean_ctor_get(x_1, 1); x_6 = lean_ctor_get(x_2, 2); lean_inc(x_6); lean_inc(x_2); -x_7 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1(x_4, x_6, x_2); +x_7 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1(x_4, x_6, x_2); lean_dec(x_6); x_8 = lean_ctor_get(x_2, 0); lean_inc(x_8); @@ -11113,7 +11147,7 @@ lean_dec(x_1); x_12 = lean_ctor_get(x_2, 2); lean_inc(x_12); lean_inc(x_2); -x_13 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1(x_10, x_12, x_2); +x_13 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1(x_10, x_12, x_2); lean_dec(x_12); x_14 = lean_ctor_get(x_2, 0); lean_inc(x_14); @@ -11127,7 +11161,7 @@ return x_16; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1() { _start: { lean_object* x_1; @@ -11135,7 +11169,7 @@ x_1 = lean_mk_string_unchecked("Elab", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2() { _start: { lean_object* x_1; @@ -11143,7 +11177,7 @@ x_1 = lean_mk_string_unchecked("Ext", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3() { _start: { lean_object* x_1; @@ -11151,28 +11185,28 @@ x_1 = lean_mk_string_unchecked("extExtension", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; -x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; +x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____lambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6() { _start: { lean_object* x_1; @@ -11180,14 +11214,14 @@ x_1 = lean_alloc_closure((void*)(l_id___rarg___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5; x_3 = l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorems___closed__1; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -11196,59 +11230,59 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7; x_3 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__4(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__4(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; lean_object* x_5; x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3(x_1, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3(x_1, x_4, x_3); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__2(x_1, x_2); +x_3 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { size_t x_7; lean_object* x_8; x_7 = lean_unbox_usize(x_1); lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__7(x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__7(x_7, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { size_t x_6; size_t x_7; lean_object* x_8; @@ -11256,45 +11290,45 @@ x_6 = lean_unbox_usize(x_2); lean_dec(x_2); x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6(x_1, x_6, x_7, x_4, x_5); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6(x_1, x_6, x_7, x_4, x_5); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__11(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Array_binInsertM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__11(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9(x_1, x_2, x_3, x_4); +x_5 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1(x_1, x_2, x_3); +x_4 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1(x_1, x_2, x_3); lean_dec(x_2); return x_4; } @@ -12171,7 +12205,7 @@ x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -12259,7 +12293,7 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -12278,7 +12312,7 @@ lean_object* x_11; lean_object* x_12; x_11 = lean_ctor_get(x_3, 6); lean_dec(x_11); lean_ctor_set(x_3, 6, x_9); -x_12 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3(x_2, x_3, x_4, x_8); +x_12 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3(x_2, x_3, x_4, x_8); return x_12; } else @@ -12316,12 +12350,12 @@ lean_ctor_set(x_23, 7, x_19); lean_ctor_set(x_23, 8, x_20); lean_ctor_set(x_23, 9, x_21); lean_ctor_set_uint8(x_23, sizeof(void*)*10, x_22); -x_24 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3(x_2, x_23, x_4, x_8); +x_24 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3(x_2, x_23, x_4, x_8); return x_24; } } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1() { _start: { lean_object* x_1; @@ -12329,7 +12363,7 @@ x_1 = lean_mk_string_unchecked("runtime", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2() { _start: { lean_object* x_1; @@ -12337,17 +12371,17 @@ x_1 = lean_mk_string_unchecked("maxRecDepth", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1; -x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1; +x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -12357,32 +12391,32 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4; x_2 = l_Lean_MessageData_ofFormat(x_1); return x_2; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6() { +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3; -x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3; +x_2 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5; x_3 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6; +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -12392,7 +12426,7 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1() { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1() { _start: { lean_object* x_1; @@ -12400,38 +12434,38 @@ x_1 = l_Lean_Elab_unsupportedSyntaxExceptionId; return x_1; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2() { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg), 1, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; @@ -12600,7 +12634,7 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; lean_object* x_6; lean_object* x_7; @@ -12612,7 +12646,7 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; @@ -12623,7 +12657,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; @@ -12634,7 +12668,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; @@ -12666,23 +12700,23 @@ x_16 = lean_ctor_get(x_14, 3); lean_inc(x_16); lean_dec(x_14); lean_inc(x_8); -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1___boxed), 4, 1); +x_17 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1___boxed), 4, 1); lean_closure_set(x_17, 0, x_8); lean_inc(x_12); x_18 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed), 3, 1); lean_closure_set(x_18, 0, x_12); lean_inc(x_8); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2___boxed), 4, 1); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2___boxed), 4, 1); lean_closure_set(x_19, 0, x_8); lean_inc(x_16); lean_inc(x_12); lean_inc(x_8); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3___boxed), 6, 3); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3___boxed), 6, 3); lean_closure_set(x_20, 0, x_8); lean_closure_set(x_20, 1, x_12); lean_closure_set(x_20, 2, x_16); lean_inc(x_8); -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4___boxed), 6, 3); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4___boxed), 6, 3); lean_closure_set(x_21, 0, x_8); lean_closure_set(x_21, 1, x_12); lean_closure_set(x_21, 2, x_16); @@ -12870,7 +12904,7 @@ lean_object* x_81; lean_object* x_82; lean_object* x_83; x_81 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_81, 0, x_78); x_82 = l_Lean_MessageData_ofFormat(x_81); -x_83 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2(x_77, x_82, x_2, x_3, x_37); +x_83 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2(x_77, x_82, x_2, x_3, x_37); lean_dec(x_77); return x_83; } @@ -12878,7 +12912,7 @@ else { lean_object* x_84; lean_dec(x_78); -x_84 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4(x_77, x_2, x_3, x_37); +x_84 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4(x_77, x_2, x_3, x_37); lean_dec(x_2); return x_84; } @@ -12887,7 +12921,7 @@ else { lean_object* x_85; lean_dec(x_2); -x_85 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg(x_37); +x_85 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg(x_37); return x_85; } } @@ -13024,7 +13058,7 @@ lean_object* x_122; lean_object* x_123; lean_object* x_124; x_122 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_122, 0, x_119); x_123 = l_Lean_MessageData_ofFormat(x_122); -x_124 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2(x_118, x_123, x_2, x_3, x_87); +x_124 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2(x_118, x_123, x_2, x_3, x_87); lean_dec(x_118); return x_124; } @@ -13032,7 +13066,7 @@ else { lean_object* x_125; lean_dec(x_119); -x_125 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4(x_118, x_2, x_3, x_87); +x_125 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4(x_118, x_2, x_3, x_87); lean_dec(x_2); return x_125; } @@ -13041,14 +13075,14 @@ else { lean_object* x_126; lean_dec(x_2); -x_126 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg(x_87); +x_126 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg(x_87); return x_126; } } } } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -13244,7 +13278,7 @@ return x_69; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -13313,7 +13347,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; @@ -13324,7 +13358,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -13384,7 +13418,7 @@ if (x_24 == 0) { lean_object* x_25; lean_object* x_26; x_25 = lean_box(0); -x_26 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1(x_1, x_2, x_25, x_3, x_4, x_5); +x_26 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1(x_1, x_2, x_25, x_3, x_4, x_5); return x_26; } else @@ -13423,7 +13457,7 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -13435,7 +13469,7 @@ lean_ctor_set(x_5, 1, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; @@ -13447,7 +13481,7 @@ lean_ctor_set(x_6, 1, x_3); return x_6; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1() { _start: { lean_object* x_1; @@ -13455,17 +13489,17 @@ x_1 = lean_mk_string_unchecked("prioDefault", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3() { _start: { lean_object* x_1; @@ -13473,16 +13507,16 @@ x_1 = lean_mk_string_unchecked("default", 7, 7); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3; lean_inc(x_1); x_5 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_4); -x_6 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2; +x_6 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2; x_7 = l_Lean_Syntax_node1(x_1, x_6, x_5); x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); @@ -13490,7 +13524,7 @@ lean_ctor_set(x_8, 1, x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -13511,7 +13545,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -13573,7 +13607,7 @@ return x_24; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1() { _start: { lean_object* x_1; @@ -13581,67 +13615,67 @@ x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at_Lean_Macro_instMonadRefMacr return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2; x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4; x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6; x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -13688,13 +13722,13 @@ lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); lean_dec(x_26); -x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__4), 4, 1); +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__4), 4, 1); lean_closure_set(x_29, 0, x_2); -x_30 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7; +x_30 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7; x_31 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg), 4, 2); lean_closure_set(x_31, 0, x_30); lean_closure_set(x_31, 1, x_29); -x_32 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___boxed), 4, 1); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___boxed), 4, 1); lean_closure_set(x_32, 0, x_31); x_33 = 1; x_34 = l_Lean_liftCommandElabM___rarg(x_32, x_33, x_10, x_11, x_28); @@ -13713,7 +13747,7 @@ lean_ctor_set(x_37, 1, x_35); lean_ctor_set(x_37, 2, x_27); x_38 = l_Lean_Elab_Tactic_Ext_getExtTheorems___closed__1; lean_inc(x_10); -x_39 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6(x_38, x_37, x_4, x_8, x_9, x_10, x_11, x_36); +x_39 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6(x_38, x_37, x_4, x_8, x_9, x_10, x_11, x_36); lean_dec(x_9); lean_dec(x_8); if (x_5 == 0) @@ -13752,7 +13786,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_39, 1); lean_inc(x_46); lean_dec(x_39); -x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5___boxed), 5, 2); +x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5___boxed), 5, 2); lean_closure_set(x_47, 0, x_6); lean_closure_set(x_47, 1, x_3); x_48 = l_Lean_liftCommandElabM___rarg(x_47, x_33, x_10, x_11, x_46); @@ -13924,13 +13958,13 @@ lean_inc(x_87); x_88 = lean_ctor_get(x_86, 1); lean_inc(x_88); lean_dec(x_86); -x_89 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__4), 4, 1); +x_89 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__4), 4, 1); lean_closure_set(x_89, 0, x_2); -x_90 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7; +x_90 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7; x_91 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Macro_instMonadRefMacroM___spec__2___rarg), 4, 2); lean_closure_set(x_91, 0, x_90); lean_closure_set(x_91, 1, x_89); -x_92 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___boxed), 4, 1); +x_92 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___boxed), 4, 1); lean_closure_set(x_92, 0, x_91); x_93 = 1; x_94 = l_Lean_liftCommandElabM___rarg(x_92, x_93, x_10, x_11, x_88); @@ -13949,7 +13983,7 @@ lean_ctor_set(x_97, 1, x_95); lean_ctor_set(x_97, 2, x_87); x_98 = l_Lean_Elab_Tactic_Ext_getExtTheorems___closed__1; lean_inc(x_10); -x_99 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6(x_98, x_97, x_4, x_8, x_9, x_10, x_11, x_96); +x_99 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6(x_98, x_97, x_4, x_8, x_9, x_10, x_11, x_96); lean_dec(x_9); lean_dec(x_8); if (x_5 == 0) @@ -13985,7 +14019,7 @@ lean_object* x_104; lean_object* x_105; lean_object* x_106; x_104 = lean_ctor_get(x_99, 1); lean_inc(x_104); lean_dec(x_99); -x_105 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5___boxed), 5, 2); +x_105 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5___boxed), 5, 2); lean_closure_set(x_105, 0, x_6); lean_closure_set(x_105, 1, x_3); x_106 = l_Lean_liftCommandElabM___rarg(x_105, x_93, x_10, x_11, x_104); @@ -14106,7 +14140,7 @@ return x_122; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1() { _start: { lean_object* x_1; @@ -14114,16 +14148,16 @@ x_1 = lean_mk_string_unchecked("@[ext] attribute only applies to structures and return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; @@ -14199,7 +14233,7 @@ lean_dec(x_38); lean_inc(x_37); x_39 = l_Lean_MessageData_ofExpr(x_37); x_40 = l_Lean_indentD(x_39); -x_41 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2; +x_41 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2; lean_ctor_set_tag(x_33, 7); lean_ctor_set(x_33, 1, x_40); lean_ctor_set(x_33, 0, x_41); @@ -14310,7 +14344,7 @@ else lean_object* x_64; lean_object* x_65; lean_dec(x_31); x_64 = lean_box(0); -x_65 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(x_49, x_2, x_6, x_3, x_4, x_5, x_64, x_8, x_9, x_10, x_11, x_35); +x_65 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(x_49, x_2, x_6, x_3, x_4, x_5, x_64, x_8, x_9, x_10, x_11, x_35); return x_65; } } @@ -14325,7 +14359,7 @@ lean_dec(x_33); lean_inc(x_66); x_67 = l_Lean_MessageData_ofExpr(x_66); x_68 = l_Lean_indentD(x_67); -x_69 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2; +x_69 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2; x_70 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_68); @@ -14440,7 +14474,7 @@ else lean_object* x_93; lean_object* x_94; lean_dec(x_31); x_93 = lean_box(0); -x_94 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(x_78, x_2, x_6, x_3, x_4, x_5, x_93, x_8, x_9, x_10, x_11, x_35); +x_94 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(x_78, x_2, x_6, x_3, x_4, x_5, x_93, x_8, x_9, x_10, x_11, x_35); return x_94; } } @@ -14469,7 +14503,7 @@ if (lean_is_exclusive(x_95)) { lean_inc(x_97); x_99 = l_Lean_MessageData_ofExpr(x_97); x_100 = l_Lean_indentD(x_99); -x_101 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2; +x_101 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2; if (lean_is_scalar(x_98)) { x_102 = lean_alloc_ctor(7, 2, 0); } else { @@ -14589,7 +14623,7 @@ else lean_object* x_126; lean_object* x_127; lean_dec(x_104); x_126 = lean_box(0); -x_127 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(x_111, x_2, x_6, x_3, x_4, x_5, x_126, x_8, x_9, x_10, x_11, x_96); +x_127 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(x_111, x_2, x_6, x_3, x_4, x_5, x_126, x_8, x_9, x_10, x_11, x_96); return x_127; } } @@ -14705,7 +14739,7 @@ if (lean_is_exclusive(x_153)) { lean_inc(x_156); x_158 = l_Lean_MessageData_ofExpr(x_156); x_159 = l_Lean_indentD(x_158); -x_160 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2; +x_160 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2; if (lean_is_scalar(x_157)) { x_161 = lean_alloc_ctor(7, 2, 0); } else { @@ -14830,7 +14864,7 @@ else lean_object* x_185; lean_object* x_186; lean_dec(x_163); x_185 = lean_box(0); -x_186 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(x_170, x_2, x_6, x_3, x_4, x_5, x_185, x_8, x_9, x_10, x_11, x_155); +x_186 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(x_170, x_2, x_6, x_3, x_4, x_5, x_185, x_8, x_9, x_10, x_11, x_155); return x_186; } } @@ -14900,7 +14934,7 @@ return x_194; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -14962,7 +14996,7 @@ return x_25; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1() { _start: { lean_object* x_1; @@ -14970,16 +15004,16 @@ x_1 = lean_mk_string_unchecked("unexpected 'flat' configuration on @[ext] theore return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; @@ -15069,7 +15103,7 @@ if (lean_obj_tag(x_7) == 0) { lean_object* x_24; lean_object* x_25; x_24 = lean_box(0); -x_25 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7(x_2, x_17, x_3, x_18, x_1, x_4, x_24, x_8, x_9, x_10, x_11, x_21); +x_25 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7(x_2, x_17, x_3, x_18, x_1, x_4, x_24, x_8, x_9, x_10, x_11, x_21); lean_dec(x_2); return x_25; } @@ -15081,8 +15115,8 @@ lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); x_26 = lean_ctor_get(x_7, 0); -x_27 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2; -x_28 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7(x_26, x_27, x_8, x_9, x_10, x_11, x_21); +x_27 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2; +x_28 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7(x_26, x_27, x_8, x_9, x_10, x_11, x_21); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); @@ -15111,7 +15145,7 @@ else lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; x_33 = lean_box(x_19); lean_inc(x_1); -x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8___boxed), 6, 3); +x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8___boxed), 6, 3); lean_closure_set(x_34, 0, x_1); lean_closure_set(x_34, 1, x_4); lean_closure_set(x_34, 2, x_33); @@ -15126,7 +15160,7 @@ x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); lean_dec(x_36); x_39 = lean_box(0); -x_40 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7(x_2, x_17, x_3, x_18, x_1, x_37, x_39, x_8, x_9, x_10, x_11, x_38); +x_40 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7(x_2, x_17, x_3, x_18, x_1, x_37, x_39, x_8, x_9, x_10, x_11, x_38); lean_dec(x_2); return x_40; } @@ -15165,7 +15199,7 @@ return x_44; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1() { _start: { lean_object* x_1; @@ -15173,16 +15207,16 @@ x_1 = lean_mk_string_unchecked("invalid syntax for 'ext' attribute", 34, 34); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3() { _start: { lean_object* x_1; @@ -15190,7 +15224,7 @@ x_1 = lean_mk_string_unchecked("Attr", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4() { _start: { lean_object* x_1; @@ -15198,19 +15232,19 @@ x_1 = lean_mk_string_unchecked("extFlat", 7, 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4; +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -15230,7 +15264,7 @@ lean_dec(x_13); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_17 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; +x_17 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; x_18 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_17, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); @@ -15244,7 +15278,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = lean_unsigned_to_nat(0u); x_20 = l_Lean_Syntax_getArg(x_13, x_19); lean_dec(x_13); -x_21 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5; +x_21 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5; lean_inc(x_20); x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); if (x_22 == 0) @@ -15254,7 +15288,7 @@ lean_dec(x_20); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_23 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; +x_23 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; x_24 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_23, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_9); @@ -15271,7 +15305,7 @@ lean_dec(x_20); x_27 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_27, 0, x_26); x_28 = lean_box(0); -x_29 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9(x_1, x_2, x_3, x_4, x_6, x_28, x_27, x_7, x_8, x_9, x_10, x_11); +x_29 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9(x_1, x_2, x_3, x_4, x_6, x_28, x_27, x_7, x_8, x_9, x_10, x_11); lean_dec(x_27); return x_29; } @@ -15283,25 +15317,25 @@ lean_object* x_30; lean_object* x_31; lean_dec(x_13); x_30 = lean_box(0); lean_inc(x_2); -x_31 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9(x_1, x_2, x_3, x_4, x_6, x_30, x_2, x_7, x_8, x_9, x_10, x_11); +x_31 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9(x_1, x_2, x_3, x_4, x_6, x_30, x_2, x_7, x_8, x_9, x_10, x_11); lean_dec(x_2); return x_31; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3; +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3; x_4 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2() { _start: { uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; @@ -15326,12 +15360,12 @@ lean_ctor_set_uint8(x_5, 12, x_2); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2; x_3 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__6; x_4 = l_Lean_Elab_Tactic_Ext_withExtHyps___rarg___lambda__2___closed__1; x_5 = lean_unsigned_to_nat(0u); @@ -15348,7 +15382,7 @@ lean_ctor_set_uint8(x_7, sizeof(void*)*6 + 1, x_6); return x_7; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15367,7 +15401,7 @@ lean_ctor_set(x_3, 8, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -15380,15 +15414,15 @@ lean_ctor_set(x_2, 3, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4; x_3 = l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__5; x_4 = l_Lean_Elab_Tactic_Ext_mkExtType___closed__5; -x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5; +x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5; x_6 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_3); @@ -15398,7 +15432,7 @@ lean_ctor_set(x_6, 4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7() { _start: { lean_object* x_1; @@ -15406,27 +15440,27 @@ x_1 = lean_mk_string_unchecked("extIff", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__2; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7; +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; -x_7 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1; +x_7 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1; lean_inc(x_2); x_8 = l_Lean_Syntax_isOfKind(x_2, x_7); x_9 = lean_box(0); -x_10 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6; +x_10 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6; x_11 = lean_st_mk_ref(x_10, x_6); if (x_8 == 0) { @@ -15463,8 +15497,8 @@ if (x_12 == 0) lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_dec(x_2); lean_dec(x_1); -x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; -x_16 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; +x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; +x_16 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; x_17 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_15, x_16, x_13, x_4, x_5, x_14); lean_dec(x_5); lean_dec(x_4); @@ -15505,8 +15539,8 @@ lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_dec(x_23); lean_dec(x_2); lean_dec(x_1); -x_26 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; -x_27 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; +x_26 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; +x_27 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; x_28 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_26, x_27, x_13, x_4, x_5, x_14); lean_dec(x_5); lean_dec(x_4); @@ -15536,7 +15570,7 @@ lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; x_33 = lean_unsigned_to_nat(0u); x_34 = l_Lean_Syntax_getArg(x_23, x_33); lean_dec(x_23); -x_35 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8; +x_35 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8; lean_inc(x_34); x_36 = l_Lean_Syntax_isOfKind(x_34, x_35); if (x_36 == 0) @@ -15545,8 +15579,8 @@ lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_dec(x_34); lean_dec(x_2); lean_dec(x_1); -x_37 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2; -x_38 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; +x_37 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2; +x_38 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; x_39 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_37, x_38, x_13, x_4, x_5, x_14); lean_dec(x_5); lean_dec(x_4); @@ -15579,9 +15613,9 @@ lean_dec(x_34); x_46 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_46, 0, x_45); x_47 = lean_box(0); -x_48 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; +x_48 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; lean_inc(x_13); -x_49 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10(x_2, x_9, x_3, x_1, x_47, x_46, x_48, x_13, x_4, x_5, x_14); +x_49 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10(x_2, x_9, x_3, x_1, x_47, x_46, x_48, x_13, x_4, x_5, x_14); lean_dec(x_46); if (lean_obj_tag(x_49) == 0) { @@ -15645,9 +15679,9 @@ else lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_dec(x_23); x_61 = lean_box(0); -x_62 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3; +x_62 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3; lean_inc(x_13); -x_63 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10(x_2, x_9, x_3, x_1, x_61, x_9, x_62, x_13, x_4, x_5, x_14); +x_63 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10(x_2, x_9, x_3, x_1, x_61, x_9, x_62, x_13, x_4, x_5, x_14); if (lean_obj_tag(x_63) == 0) { lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; @@ -15707,14 +15741,14 @@ return x_74; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12(lean_object* x_1, lean_object* x_2) { _start: { lean_inc(x_1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -15731,7 +15765,7 @@ x_9 = l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorems; x_10 = l_Lean_Elab_Tactic_Ext_getExtTheorems___closed__1; x_11 = l_Lean_ScopedEnvExtension_getState___rarg(x_9, x_10, x_8); lean_dec(x_8); -x_12 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8(x_11, x_1, x_2, x_3, x_7); +x_12 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8(x_11, x_1, x_2, x_3, x_7); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; @@ -15753,7 +15787,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean x_19 = lean_ctor_get(x_16, 0); x_20 = lean_ctor_get(x_16, 4); lean_dec(x_20); -x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12___boxed), 2, 1); +x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12___boxed), 2, 1); lean_closure_set(x_21, 0, x_13); x_22 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_10, x_19, x_21); x_23 = l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__2; @@ -15799,7 +15833,7 @@ lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); lean_dec(x_16); -x_37 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12___boxed), 2, 1); +x_37 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12___boxed), 2, 1); lean_closure_set(x_37, 0, x_13); x_38 = l_Lean_ScopedEnvExtension_modifyState___rarg(x_10, x_31, x_37); x_39 = l_Lean_addDeclarationRanges___at_Lean_Elab_Tactic_Ext_realizeExtTheorem___spec__5___closed__2; @@ -15857,7 +15891,7 @@ return x_49; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15867,37 +15901,37 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5() { _start: { lean_object* x_1; @@ -15905,17 +15939,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7() { _start: { lean_object* x_1; @@ -15923,57 +15957,57 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10; x_2 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13() { _start: { lean_object* x_1; @@ -15981,27 +16015,27 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14; -x_2 = lean_unsigned_to_nat(3673u); +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14; +x_2 = lean_unsigned_to_nat(3679u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -16011,7 +16045,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17() { _start: { lean_object* x_1; @@ -16019,13 +16053,13 @@ x_1 = lean_mk_string_unchecked("Marks a theorem as an extensionality theorem", 4 return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16; +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17; x_4 = 0; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -16035,29 +16069,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21() { +static lean_object* _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20; +x_1 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19; +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -16065,120 +16099,120 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5(x_1, x_2); +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; lean_object* x_10; x_9 = lean_unbox(x_3); lean_dec(x_3); -x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__6(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_ScopedEnvExtension_add___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__6(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); @@ -16186,67 +16220,67 @@ lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__8(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__8(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__1(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__2(x_1, x_2, x_3); +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__2(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3(x_1, x_2, x_3); +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3(x_1, x_2, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__5(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__5(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; uint8_t x_14; lean_object* x_15; @@ -16254,12 +16288,12 @@ x_13 = lean_unbox(x_4); lean_dec(x_4); x_14 = lean_unbox(x_5); lean_dec(x_5); -x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6(x_1, x_2, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6(x_1, x_2, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; uint8_t x_14; lean_object* x_15; @@ -16267,74 +16301,74 @@ x_13 = lean_unbox(x_3); lean_dec(x_3); x_14 = lean_unbox(x_4); lean_dec(x_4); -x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__8(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__8(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_3); lean_dec(x_3); -x_14 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_14 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { uint8_t x_12; lean_object* x_13; x_12 = lean_unbox(x_3); lean_dec(x_3); -x_13 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_13 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); lean_dec(x_5); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__12(x_1, x_2); +x_3 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__12(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__13(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__13(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); return x_5; @@ -17859,9 +17893,9 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheore { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_5 = l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -17880,9 +17914,9 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheore { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_5 = l___regBuiltin_Lean_Elab_Tactic_Ext_evalApplyExtTheorem__1___closed__3; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -20738,9 +20772,9 @@ static lean_object* _init_l_Lean_Elab_Tactic_Ext_evalExt___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_5 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__5___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -20911,9 +20945,9 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__1; -x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1; +x_2 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1; x_3 = l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__10; -x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2; +x_4 = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2; x_5 = l___regBuiltin_Lean_Elab_Tactic_Ext_evalExt__1___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -21240,6 +21274,10 @@ l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33 = _init_l_Lean lean_mark_persistent(l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__33); l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34 = _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__34); +l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35 = _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__35); +l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36 = _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__2___closed__36); l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__3___closed__1); l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4___closed__1 = _init_l_Lean_Elab_Tactic_Ext_realizeExtTheorem___lambda__4___closed__1(); @@ -21414,66 +21452,66 @@ l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem___closed__1 = _init_l_Lean_Elab_T lean_mark_persistent(l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem___closed__1); l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem = _init_l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_instInhabitedExtTheorem); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__1); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__2); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__3); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__4); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__5); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__6); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__7); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__8); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__9); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__10); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__11); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__12); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__13); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__14); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__15); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__16); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__17); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__18); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__19); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__20); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__21); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__22); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__23); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__24); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__25); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__26); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__27); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__28); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__29); -l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30(); -lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3098____closed__30); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__1); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__2); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__3); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__4); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__5); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__6); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__7); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__8); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__9); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__10); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__11); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__12); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__13); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__14); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__15); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__16); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__17); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__18); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__19); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__20); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__21); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__22); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__23); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__24); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__25); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__26); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__27); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__28); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__29); +l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30 = _init_l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30(); +lean_mark_persistent(l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_3104____closed__30); l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1 = _init_l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_instReprExtTheorem___closed__1); l_Lean_Elab_Tactic_Ext_instReprExtTheorem = _init_l_Lean_Elab_Tactic_Ext_instReprExtTheorem(); @@ -21494,37 +21532,37 @@ l_Lean_Elab_Tactic_Ext_extExt_config___closed__1 = _init_l_Lean_Elab_Tactic_Ext_ lean_mark_persistent(l_Lean_Elab_Tactic_Ext_extExt_config___closed__1); l_Lean_Elab_Tactic_Ext_extExt_config = _init_l_Lean_Elab_Tactic_Ext_extExt_config(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_extExt_config); -l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__1(); -l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__3___closed__2(); -l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__6___closed__1); -l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__9___closed__1); -l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1 = _init_l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__13___closed__1); -l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__1); -l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__2); -l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3(); -lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__3); -l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4(); -lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____spec__1___closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__5); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__6); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378____closed__7); -if (builtin) {res = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3378_(lean_io_mk_world()); +l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__1(); +l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__3___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__6___closed__1); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__9___closed__1); +l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1 = _init_l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__13___closed__1); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__1); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__2); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__3); +l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4 = _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4(); +lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____spec__1___closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__5); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__6); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384____closed__7); +if (builtin) {res = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3384_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Tactic_Ext_extExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_extExtension); @@ -21537,119 +21575,119 @@ l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__1 = _init_l_Lean_Elab_ lean_mark_persistent(l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__1); l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__2 = _init_l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_Ext_ExtTheorems_erase___rarg___closed__2); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__1); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__2); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__3); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__4); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__5); -l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6(); -lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__4___closed__6); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____spec__5___rarg___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__3___closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__5); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__6); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__6___closed__7); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__7___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__9___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__10___closed__5); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__5); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__6); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__7); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____lambda__11___closed__8); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__1); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__2); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__3); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__4); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__5); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__6); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__7); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__8); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__9); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__10); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__11); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__12); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__13); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__14); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__15); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__16); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__17); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__18); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__19); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__20); -l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21(); -lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673____closed__21); -if (builtin) {res = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3673_(lean_io_mk_world()); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__1); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__2); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__3); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__4); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__5); +l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6(); +lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__4___closed__6); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__1); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____spec__5___rarg___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__3___closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__5); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__6); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__6___closed__7); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__7___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__9___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__10___closed__5); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__5); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__6); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__7); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____lambda__11___closed__8); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__1); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__2); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__3); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__4); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__5); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__6); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__7); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__8); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__9); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__10); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__11); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__12); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__13); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__14); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__15); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__16); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__17); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__18); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__19); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__20); +l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21 = _init_l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21(); +lean_mark_persistent(l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679____closed__21); +if (builtin) {res = l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_3679_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___lambda__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Tactic_Ext_applyExtTheoremAt___spec__1___lambda__1___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index ca9471b76996..eb5df50e553c 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -24,7 +24,6 @@ lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_hasNoImplicitLambdaAnnotation___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withoutModifyingStateWithInfoAndMessages(lean_object*, lean_object*); lean_object* l_StateRefT_x27_instMonadExceptOf___rarg(lean_object*); @@ -33,7 +32,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_levelMVarToParam(lean_object*, lean_ob static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutSavingRecAppSyntax___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__4; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__4; lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); @@ -41,13 +39,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos(lean_obje LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__4___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElimImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20184_(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3; LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Term_expandDeclId___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Term_expandDeclId___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4; static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Term_expandDeclId___spec__10___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_setLevelNames___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorryImp(lean_object*); @@ -64,10 +62,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwMVarError___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLevelNames___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3; static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1___closed__1; static lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -98,7 +96,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars___boxed(lean_ob LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__26___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -126,6 +123,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_Elab_Term_instAddErrorMessageContextTermElabM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_commitIfNoErrors_x3f(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1; static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Elab_Term_exposeLevelMVars___spec__2___boxed__const__1; @@ -148,19 +146,19 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa LEAN_EXPORT lean_object* l_Lean_Elab_Term_getInfoTreeWithContext_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_transform___at_Lean_Elab_Term_exposeLevelMVars___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTacticMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Term_isLetRecAuxMVar___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandDeclId___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21___closed__2; static lean_object* l_Lean_Elab_Term_mkSaveInfoAnnotation___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__4; lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20279_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedSyntheticMVarDecl; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -175,8 +173,10 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__4; extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1; uint8_t l_Array_contains___at_Lean_Meta_arrowDomainsN___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -206,7 +206,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName_process___lambda__1___boxe LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Elab_Term_exposeLevelMVars___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkAuxName___at_Lean_Elab_Term_mkAuxName___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkAuxName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -223,9 +223,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeEl LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwTypeMismatchError___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticReuse___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158_(lean_object*); static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__2; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_3777____closed__2; static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__3; @@ -257,7 +255,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambd LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_blockImplicitLambda(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,6 +262,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerSyntheticMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -274,15 +272,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___lambda__2___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isNoImplicitLambda___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6; lean_object* l_String_removeLeadingSpaces(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2; lean_object* l_Lean_Core_getMessageLog___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDeclId___closed__1; lean_object* l_Lean_Core_checkSystem(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_resolveId_x3f___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Term_resolveLocalName___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12; LEAN_EXPORT lean_object* l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__36___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -292,16 +290,18 @@ static lean_object* l_Lean_Elab_Term_ensureType___lambda__1___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_traceAtCmdPos___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__32(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind___boxed(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPartialTermInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -315,6 +315,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBound lean_object* l_Lean_Syntax_getArgs(lean_object*); lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getBodyInfo_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName_process___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -329,7 +330,6 @@ static lean_object* l_Lean_Elab_Term_mkAuxName___closed__2; lean_object* l_Lean_Meta_getMVarsAtDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_exceptionToSorry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_10939____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -359,10 +359,9 @@ lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Core_transform_visit___at_Lean_Elab_Term_exposeLevelMVars___spec__2___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_containsPendingMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_MVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__46(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull(lean_object*); @@ -392,6 +391,7 @@ lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, l static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___closed__4; LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__12(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__9; static lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__2___closed__6; @@ -412,6 +412,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplici uint8_t l_Lean_Linter_getLinterValue(lean_object*, lean_object*); lean_object* l_Lean_instAddMessageContextOfMonadLift___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -445,12 +446,15 @@ static lean_object* l_Lean_Elab_Term_resolveName_process___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__9; LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___rarg___lambda__1___boxed(lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__7; lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); @@ -463,7 +467,6 @@ extern lean_object* l_Lean_maxRecDepth; lean_object* l_Lean_CollectLevelMVars_main(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -481,6 +484,7 @@ uint8_t l_Lean_Syntax_isAtom(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__3___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__49(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_levelMVarToParam___lambda__1(lean_object*, lean_object*); @@ -519,7 +523,6 @@ static lean_object* l_Lean_Elab_addBuiltinIncrementalElab___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_3777____closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at_Lean_Elab_Term_resolveLocalName___spec__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasSyntheticSorry(lean_object*); static lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshotData___closed__1; @@ -557,7 +560,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscr LEAN_EXPORT uint8_t l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit___closed__1; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__2; @@ -573,7 +575,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticReuse___rarg(lean_object* extern lean_object* l_Lean_Linter_instInhabitedDeprecationEntry; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_abortCommandExceptionId; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedSyntheticMVarDecl___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -624,7 +625,8 @@ lean_object* l_Array_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__4; static lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshotData___closed__3; static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__3___closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12; uint8_t l_List_elem___at_Lean_addAliasEntry___spec__16(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -635,7 +637,6 @@ lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_obje LEAN_EXPORT lean_object* l_Lean_Elab_Term_withExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -664,6 +665,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkInstMVar(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__14___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1; lean_object* l_Lean_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instToSnapshotTreeTacticParsedSnapshot_go(lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); @@ -674,7 +676,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotComp LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8___boxed(lean_object*); lean_object* l_ReaderT_instMonadFunctor(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -685,7 +686,6 @@ uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__7(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withMacroExpansion(lean_object*); -static lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -697,7 +697,6 @@ lean_object* l_Lean_Meta_isType(lean_object*, lean_object*, lean_object*, lean_o lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName_process(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__28___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__9(lean_object*, lean_object*, size_t, size_t); @@ -714,7 +713,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse(lean_object LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedCache; LEAN_EXPORT lean_object* l_Lean_Core_transform_visit___at_Lean_Elab_Term_exposeLevelMVars___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7; static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -730,10 +728,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_go___boxed(lean_ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__6; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instToSnapshotTreeTacticFinishedSnapshot(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__9(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_exprToSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwTypeMismatchError___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO(lean_object*); @@ -742,8 +740,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_wit lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ppGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10; static lean_object* l_Lean_Elab_Term_getInfoTreeWithContext_x3f___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveParentDeclInfoContext___at_Lean_Elab_Term_withDeclName___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__5; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__5; @@ -752,11 +752,10 @@ static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__23(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__4; static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyAttributesAt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -767,6 +766,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutMacroStackAtErr___rarg(lean_obj lean_object* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_isTacticOrPostponedHole_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); @@ -805,7 +805,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplic static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withReuseContext___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1; static size_t l_Lean_Elab_Tactic_instToSnapshotTreeTacticParsedSnapshot_go___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); @@ -817,7 +816,7 @@ lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__37___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6; +static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -835,8 +834,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar___boxed(lean LEAN_EXPORT lean_object* l_Lean_profileitM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_run_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_addBuiltinIncrementalElab(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg___closed__1; @@ -877,7 +874,6 @@ static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__2; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__2___closed__2; static lean_object* l_Lean_Elab_Tactic_instToSnapshotTreeTacticFinishedSnapshot___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__1(uint8_t, lean_object*); @@ -891,9 +887,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_coerceToSort_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_liftLevelM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004_(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_MVarIdMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Term_resolveLocalName___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -908,7 +902,6 @@ static lean_object* l_Lean_Elab_Term_throwErrorIfErrors___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instTypeNameBodyInfo; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237_(lean_object*); static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__12; lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at_Lean_Elab_Term_resolveLocalName___spec__7___boxed(lean_object*, lean_object*); @@ -925,7 +918,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_loop___lambda__1___bo static lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_levelMVarToParam___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6; lean_object* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicitApp___closed__1; @@ -952,6 +944,7 @@ lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_builtinIncrementalElabs; @@ -959,9 +952,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_instToSnap LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveParentDeclInfoContext___at_Lean_Elab_Term_withDeclName___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot___closed__3; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveId_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -974,7 +965,6 @@ static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__3; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__8; lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__5; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5; lean_object* l_Lean_extractMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); @@ -996,7 +986,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__24(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__2; static lean_object* l_Lean_Elab_Term_instInhabitedTermElabM___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1020,7 +1009,6 @@ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__52___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -1063,6 +1051,7 @@ extern lean_object* l_Lean_diagnostics; uint8_t l_Lean_Syntax_isIdent(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveId_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__46___boxed(lean_object*, lean_object*); @@ -1070,8 +1059,6 @@ lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTacticBlock___boxed(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__4___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___lambda__2___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15; extern lean_object* l_Lean_Elab_abortTermExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___closed__3; @@ -1081,7 +1068,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutMacroStackAtErr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTacticMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_withoutErrToSorry___rarg___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_exposeLevelMVars___spec__7___rarg___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11; uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveLocalName_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1093,6 +1079,7 @@ extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkCoe___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getSyntheticMVarDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_addTermInfo___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1100,6 +1087,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Tactic_beqCacheKey__ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind(lean_object*); extern lean_object* l_Std_Format_defWidth; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens___closed__2; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorInfo___closed__1; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__1; @@ -1107,6 +1095,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedLevelMVarErrorInfo; lean_object* l_Lean_MacroScopesView_review(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_setInfoState___at_Lean_Elab_Term_SavedState_restore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1118,23 +1107,26 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_containsPendingMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_instToSnapshotTreeTacticParsedSnapshot_go___spec__1___closed__1; static lean_object* l_Lean_Elab_Term_resolveName_process___closed__3; static lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Term_exprToSyntax___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_throwErrorIfErrors___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5; LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Term_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_resolveName_process___closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__37(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkAuxName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyAttributesAt___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__6(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); static size_t l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__5; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13; extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_exposeLevelMVars___spec__7(lean_object*); @@ -1163,19 +1155,21 @@ lean_object* l_List_tail_x21___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___lambda__2___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_10939____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureHasTypeWithErrorMsgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwTypeMismatchError___spec__1(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshotData___closed__2; @@ -1198,6 +1192,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe(lean_object* LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__2(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_10939____closed__3; extern lean_object* l_Lean_Linter_deprecatedAttr; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1222,7 +1217,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___lambda__2___boxed(lean_object* LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Term_mkCoe___closed__1; lean_object* l_Lean_instAddErrorMessageContextOfAddMessageContextOfMonad___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51___closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__22(lean_object*, lean_object*); @@ -1230,7 +1224,6 @@ static lean_object* l_Lean_Elab_Term_mkTacticMVar___closed__1; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__4; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___lambda__1___closed__5; static double l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_hasMacroScopes(lean_object*); @@ -1247,27 +1240,34 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLevel___boxed(lean_object*, lean_o static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___rarg___lambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7; lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPartialTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_NameSet_empty; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit___closed__2; static lean_object* l_Lean_Elab_Term_withExpectedType___closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadTermElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585_(lean_object*); static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_withDeclName___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__3___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__11; uint8_t l_Lean_Name_isAtomic(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedTermElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_instToSnapshotTreeTacticParsedSnapshot_go___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_LVal_isFieldName(lean_object*); @@ -1280,6 +1280,7 @@ static lean_object* l_Lean_Elab_Term_getInfoTreeWithContext_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_getSyntheticMVarDecl_x3f___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Term_expandDeclId___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1303,7 +1304,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit_loop___rarg(lean static lean_object* l_Lean_Core_transform___at_Lean_Elab_Term_exposeLevelMVars___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1312,7 +1312,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_LevelMVarErrorInfo_logError___lambda__ static lean_object* l_Lean_Elab_Term_instMonadParentDeclTermElabM___closed__1; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg(lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); @@ -1331,12 +1331,10 @@ static lean_object* l_Lean_Elab_Term_exposeLevelMVars___closed__2; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__6; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_exposeLevelMVars___spec__7___rarg___closed__5; lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveState___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3; static lean_object* l_List_forIn_x27_loop___at_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___spec__4___closed__1; static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedTacticFinishedSnapshot; @@ -1368,7 +1366,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassigne LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___lambda__2___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_isIncrementalElab___rarg___closed__1; @@ -1388,7 +1385,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___rarg___lambda__1(lean_o LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_localDeclDependsOn___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__4___closed__2; +static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___boxed(lean_object*); static double l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_commitIfDidNotPostpone___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1406,7 +1405,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_Lean_Expr_setOption___at_Lean_Expr_setPPExplicit___spec__1(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyAttributes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__20(lean_object*, lean_object*, lean_object*); @@ -1466,10 +1464,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f(lean_object* static lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__5; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__2___boxed(lean_object*, lean_object*); uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_MetavarContext_0__Lean_MetavarContext_MkBinding_visit___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Term_exposeLevelMVars___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withDeclName___spec__2(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__5; @@ -1488,7 +1487,6 @@ static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkTypeMismatchError(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__42(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Linter_linter_deprecated; lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1499,7 +1497,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImpli LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_isIncrementalElab___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_addDotCompletionInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticIncrementality(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__2; @@ -1517,12 +1514,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_incrementalAttr; LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveState___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveName_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__1; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isNoImplicitLambda___closed__1; static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Term_expandDeclId___spec__5___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Term_withMacroExpansion___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAuxDecl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1532,7 +1532,6 @@ lean_object* l___private_Lean_ToExpr_0__Lean_Name_toExprAux(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerLevelMVarErrorInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withNarrowedArgTacticReuse___rarg___lambda__1(lean_object*, lean_object*); lean_object* l___private_Lean_MonadEnv_0__Lean_mkAuxNameAux(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instBEqCacheKey; static lean_object* l_List_foldl___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__2___closed__1; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__3; @@ -1540,6 +1539,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadParentDeclTermElabM; static lean_object* l_Lean_Elab_Term_extraMsgToMsg___closed__1; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkCoe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__47(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_PersistentArray_get_x21___rarg(lean_object*, lean_object*, lean_object*); @@ -1549,7 +1549,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImpli static lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVarError___spec__2___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_observing___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099_(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_resolveName_x27___closed__1; @@ -1573,13 +1573,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift; size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_getMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3; uint8_t l_Lean_Expr_hasFVar(lean_object*); static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__8; static lean_object* l_Lean_Elab_isIncrementalElab___rarg___lambda__3___closed__1; @@ -1588,7 +1587,6 @@ static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__6; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__4; lean_object* l_Lean_Language_SnapshotTask_map___rarg(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1; lean_object* l_Lean_MessageLog_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_withAutoBoundImplicit_loop___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedTacticParsedSnapshotData(lean_object*); @@ -1609,6 +1607,7 @@ size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_transform___at_Lean_Elab_Term_exposeLevelMVars___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_extraMsgToMsg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__1; extern lean_object* l_Lean_trace_profiler; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1626,6 +1625,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go___boxed(lean_objec lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__3; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1; extern lean_object* l_Lean_instInhabitedSyntax; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -1634,6 +1634,7 @@ static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___lambda__1___closed LEAN_EXPORT lean_object* l_Lean_Elab_Term_getSyntheticMVarDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1651,14 +1652,18 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_with LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__21___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_logUnassignedLevelMVarsUsingErrorInfos___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__5; static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__6; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4; lean_object* l_Lean_Meta_mkFreshTypeMVar(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint(lean_object*); static lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1668,8 +1673,8 @@ LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Elab_Term_logUnassignedUsingError lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3; static lean_object* l_List_foldl___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__2___closed__3; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4; static lean_object* l_Lean_Elab_Term_LevelMVarErrorInfo_logError___lambda__1___closed__1; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__3; @@ -1688,7 +1693,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringLVal(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1700,8 +1704,10 @@ extern lean_object* l_Lean_Meta_instAddMessageContextMetaM; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___closed__3; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__16(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Term_exposeLevelMVars___lambda__1___closed__1; lean_object* l_Lean_RBNode_find___at_Lean_NameMap_find_x3f___spec__1___rarg(lean_object*, lean_object*); @@ -1721,6 +1727,7 @@ lean_object* l_Lean_realizeGlobalName(lean_object*, lean_object*, lean_object*, LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3; static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__1; static lean_object* l_Lean_Elab_Term_getInfoTreeWithContext_x3f___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Term_withMacroExpansion___spec__1(lean_object*); @@ -1752,20 +1759,17 @@ uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withDeclName___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_SavedState_restore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2; static lean_object* l_Lean_Elab_Term_expandDeclId___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__1; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getInfoTreeWithContext_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarArgName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwMVarError___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Term_expandDeclId___spec__5___closed__3; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1773,17 +1777,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpo LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Term_expandDeclId___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedState___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVarError___spec__2___rarg(lean_object*); uint8_t l_Lean_Elab_isValidAutoBoundImplicitName(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Elab_Term_exposeLevelMVars___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkBodyInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__10___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__2; lean_object* l_String_toSubstring_x27(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4; LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1791,7 +1793,6 @@ static lean_object* l_Lean_Elab_Term_mkCoe___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_withRestoreOrSaveFull___rarg___closed__15; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_exposeLevelMVars___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_dbg_trace(lean_object*, lean_object*); @@ -1815,6 +1816,7 @@ static double l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__2___lambda_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_removeSaveInfoAnnotation(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLetRecsToLift___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutTacticReuse___rarg___lambda__1(uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Term_expandDeclId___spec__9___closed__2; @@ -1825,7 +1827,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__8; static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__8; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332_(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_add(lean_object*, lean_object*); @@ -21286,7 +21288,7 @@ x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_1); lean_ctor_set(x_12, 2, x_2); -x_13 = lean_alloc_ctor(3, 1, 0); +x_13 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_13, 0, x_12); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_13); @@ -32457,7 +32459,7 @@ lean_dec(x_16); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_6); lean_ctor_set(x_23, 1, x_7); -x_24 = lean_alloc_ctor(2, 1, 0); +x_24 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_24, 0, x_23); x_25 = l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2(x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_22); lean_dec(x_13); @@ -32523,7 +32525,7 @@ lean_dec(x_16); x_36 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_36, 0, x_6); lean_ctor_set(x_36, 1, x_7); -x_37 = lean_alloc_ctor(2, 1, 0); +x_37 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_37, 0, x_36); x_38 = l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2(x_37, x_8, x_9, x_10, x_11, x_12, x_13, x_35); lean_dec(x_13); @@ -32698,7 +32700,7 @@ lean_inc(x_4); lean_inc(x_31); lean_ctor_set(x_36, 1, x_4); lean_ctor_set(x_36, 0, x_31); -lean_ctor_set_tag(x_18, 2); +lean_ctor_set_tag(x_18, 3); lean_ctor_set(x_18, 0, x_36); x_51 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_51, 0, x_18); @@ -32791,7 +32793,7 @@ lean_inc(x_31); x_75 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_75, 0, x_31); lean_ctor_set(x_75, 1, x_4); -lean_ctor_set_tag(x_18, 2); +lean_ctor_set_tag(x_18, 3); lean_ctor_set(x_18, 0, x_75); x_76 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_76, 0, x_18); @@ -32958,7 +32960,7 @@ lean_inc(x_4); lean_inc(x_93); lean_ctor_set(x_104, 1, x_4); lean_ctor_set(x_104, 0, x_93); -lean_ctor_set_tag(x_2, 2); +lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_104); x_119 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_119, 0, x_2); @@ -33051,7 +33053,7 @@ lean_inc(x_93); x_143 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_143, 0, x_93); lean_ctor_set(x_143, 1, x_4); -lean_ctor_set_tag(x_2, 2); +lean_ctor_set_tag(x_2, 3); lean_ctor_set(x_2, 0, x_143); x_144 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_144, 0, x_2); @@ -33215,7 +33217,7 @@ if (lean_is_scalar(x_174)) { } lean_ctor_set(x_182, 0, x_93); lean_ctor_set(x_182, 1, x_4); -x_183 = lean_alloc_ctor(2, 1, 0); +x_183 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_183, 0, x_182); x_184 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_184, 0, x_183); @@ -33370,7 +33372,7 @@ if (lean_is_scalar(x_215)) { } lean_ctor_set(x_223, 0, x_203); lean_ctor_set(x_223, 1, x_4); -x_224 = lean_alloc_ctor(2, 1, 0); +x_224 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_224, 0, x_223); x_225 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); lean_closure_set(x_225, 0, x_224); @@ -33551,10 +33553,10 @@ if (lean_is_scalar(x_260)) { lean_ctor_set(x_268, 0, x_242); lean_ctor_set(x_268, 1, x_4); if (lean_is_scalar(x_245)) { - x_269 = lean_alloc_ctor(2, 1, 0); + x_269 = lean_alloc_ctor(3, 1, 0); } else { x_269 = x_245; - lean_ctor_set_tag(x_269, 2); + lean_ctor_set_tag(x_269, 3); } lean_ctor_set(x_269, 0, x_268); x_270 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); @@ -33766,10 +33768,10 @@ if (lean_is_scalar(x_312)) { lean_ctor_set(x_320, 0, x_300); lean_ctor_set(x_320, 1, x_4); if (lean_is_scalar(x_296)) { - x_321 = lean_alloc_ctor(2, 1, 0); + x_321 = lean_alloc_ctor(3, 1, 0); } else { x_321 = x_296; - lean_ctor_set_tag(x_321, 2); + lean_ctor_set_tag(x_321, 3); } lean_ctor_set(x_321, 0, x_320); x_322 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); @@ -33953,10 +33955,10 @@ if (lean_is_scalar(x_358)) { lean_ctor_set(x_366, 0, x_339); lean_ctor_set(x_366, 1, x_4); if (lean_is_scalar(x_342)) { - x_367 = lean_alloc_ctor(2, 1, 0); + x_367 = lean_alloc_ctor(3, 1, 0); } else { x_367 = x_342; - lean_ctor_set_tag(x_367, 2); + lean_ctor_set_tag(x_367, 3); } lean_ctor_set(x_367, 0, x_366); x_368 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__1___boxed), 9, 1); @@ -46820,731 +46822,780 @@ lean_dec(x_7); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPartialTermInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_st_ref_get(x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 6); +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_7, 1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); -lean_dec(x_12); -if (x_13 == 0) +x_14 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +lean_ctor_set(x_14, 2, x_3); +x_15 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +else { -lean_object* x_14; lean_object* x_15; -lean_dec(x_2); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -return x_15; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_4); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_4, 0); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_3); +lean_ctor_set_tag(x_4, 2); +lean_ctor_set(x_4, 0, x_19); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_4); +lean_ctor_set(x_20, 1, x_11); +return x_20; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +lean_dec(x_4); +x_22 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_3); +x_23 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_11); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkPartialTermInfo___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Elab_Term_mkPartialTermInfo(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); -x_17 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(x_8, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_20 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_20) == 0) +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -lean_inc(x_8); -lean_inc(x_21); -x_23 = lean_apply_8(x_2, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_st_ref_take(x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 6); +lean_inc(x_12); +if (lean_obj_tag(x_2) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_st_ref_take(x_8, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 6); -lean_inc(x_28); -if (lean_obj_tag(x_24) == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_26); -if (x_29 == 0) +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_10, 0); +lean_dec(x_15); +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_26, 1); -x_31 = lean_ctor_get(x_26, 0); -lean_dec(x_31); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_11, 6); +lean_dec(x_17); +x_18 = !lean_is_exclusive(x_12); +if (x_18 == 0) { -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_27, 6); -lean_dec(x_33); -x_34 = !lean_is_exclusive(x_28); -if (x_34 == 0) +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_12, 1); +x_20 = lean_ctor_get(x_2, 0); +lean_inc(x_20); +lean_dec(x_2); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 1, x_19); +lean_ctor_set(x_10, 0, x_20); +x_21 = l_Lean_PersistentArray_push___rarg(x_1, x_10); +lean_ctor_set(x_12, 1, x_21); +x_22 = lean_st_ref_set(x_8, x_11, x_14); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_35 = lean_ctor_get(x_28, 1); -x_36 = lean_ctor_get(x_24, 0); -lean_inc(x_36); +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_35); -lean_ctor_set(x_26, 0, x_36); -x_37 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -lean_ctor_set(x_28, 1, x_37); -x_38 = lean_st_ref_set(x_8, x_27, x_30); -lean_dec(x_8); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); -lean_dec(x_40); -lean_ctor_set(x_38, 0, x_21); -return x_38; +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; } else { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_21); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; } } else { -uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_44 = lean_ctor_get(x_28, 0); -x_45 = lean_ctor_get(x_28, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_28); -x_46 = lean_ctor_get(x_24, 0); -lean_inc(x_46); -lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_45); -lean_ctor_set(x_26, 0, x_46); -x_47 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -x_48 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_48, 0, x_44); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set_uint8(x_48, sizeof(void*)*2, x_43); -lean_ctor_set(x_27, 6, x_48); -x_49 = lean_st_ref_set(x_8, x_27, x_30); -lean_dec(x_8); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_51 = x_49; +uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_30 = lean_ctor_get(x_12, 0); +x_31 = lean_ctor_get(x_12, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_12); +x_32 = lean_ctor_get(x_2, 0); +lean_inc(x_32); +lean_dec(x_2); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 1, x_31); +lean_ctor_set(x_10, 0, x_32); +x_33 = l_Lean_PersistentArray_push___rarg(x_1, x_10); +x_34 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_34, 0, x_30); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set_uint8(x_34, sizeof(void*)*2, x_29); +lean_ctor_set(x_11, 6, x_34); +x_35 = lean_st_ref_set(x_8, x_11, x_14); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_37 = x_35; } else { - lean_dec_ref(x_49); - x_51 = lean_box(0); + lean_dec_ref(x_35); + x_37 = lean_box(0); } -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 2, 0); +x_38 = lean_box(0); +if (lean_is_scalar(x_37)) { + x_39 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_51; + x_39 = x_37; } -lean_ctor_set(x_52, 0, x_21); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_53 = lean_ctor_get(x_27, 0); -x_54 = lean_ctor_get(x_27, 1); -x_55 = lean_ctor_get(x_27, 2); -x_56 = lean_ctor_get(x_27, 3); -x_57 = lean_ctor_get(x_27, 4); -x_58 = lean_ctor_get(x_27, 5); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_27); -x_59 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_60 = lean_ctor_get(x_28, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_28, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_62 = x_28; +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_40 = lean_ctor_get(x_11, 0); +x_41 = lean_ctor_get(x_11, 1); +x_42 = lean_ctor_get(x_11, 2); +x_43 = lean_ctor_get(x_11, 3); +x_44 = lean_ctor_get(x_11, 4); +x_45 = lean_ctor_get(x_11, 5); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_11); +x_46 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_47 = lean_ctor_get(x_12, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_49 = x_12; } else { - lean_dec_ref(x_28); - x_62 = lean_box(0); + lean_dec_ref(x_12); + x_49 = lean_box(0); } -x_63 = lean_ctor_get(x_24, 0); -lean_inc(x_63); -lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_61); -lean_ctor_set(x_26, 0, x_63); -x_64 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -if (lean_is_scalar(x_62)) { - x_65 = lean_alloc_ctor(0, 2, 1); +x_50 = lean_ctor_get(x_2, 0); +lean_inc(x_50); +lean_dec(x_2); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 1, x_48); +lean_ctor_set(x_10, 0, x_50); +x_51 = l_Lean_PersistentArray_push___rarg(x_1, x_10); +if (lean_is_scalar(x_49)) { + x_52 = lean_alloc_ctor(0, 2, 1); } else { - x_65 = x_62; + x_52 = x_49; } -lean_ctor_set(x_65, 0, x_60); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set_uint8(x_65, sizeof(void*)*2, x_59); -x_66 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_66, 0, x_53); -lean_ctor_set(x_66, 1, x_54); -lean_ctor_set(x_66, 2, x_55); -lean_ctor_set(x_66, 3, x_56); -lean_ctor_set(x_66, 4, x_57); -lean_ctor_set(x_66, 5, x_58); -lean_ctor_set(x_66, 6, x_65); -x_67 = lean_st_ref_set(x_8, x_66, x_30); -lean_dec(x_8); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_69 = x_67; +lean_ctor_set(x_52, 0, x_47); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set_uint8(x_52, sizeof(void*)*2, x_46); +x_53 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_53, 0, x_40); +lean_ctor_set(x_53, 1, x_41); +lean_ctor_set(x_53, 2, x_42); +lean_ctor_set(x_53, 3, x_43); +lean_ctor_set(x_53, 4, x_44); +lean_ctor_set(x_53, 5, x_45); +lean_ctor_set(x_53, 6, x_52); +x_54 = lean_st_ref_set(x_8, x_53, x_14); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_56 = x_54; } else { - lean_dec_ref(x_67); - x_69 = lean_box(0); + lean_dec_ref(x_54); + x_56 = lean_box(0); } -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 2, 0); +x_57 = lean_box(0); +if (lean_is_scalar(x_56)) { + x_58 = lean_alloc_ctor(0, 2, 0); } else { - x_70 = x_69; + x_58 = x_56; } -lean_ctor_set(x_70, 0, x_21); -lean_ctor_set(x_70, 1, x_68); -return x_70; +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_55); +return x_58; } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_71 = lean_ctor_get(x_26, 1); -lean_inc(x_71); -lean_dec(x_26); -x_72 = lean_ctor_get(x_27, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_27, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_27, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_27, 3); -lean_inc(x_75); -x_76 = lean_ctor_get(x_27, 4); -lean_inc(x_76); -x_77 = lean_ctor_get(x_27, 5); -lean_inc(x_77); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - lean_ctor_release(x_27, 2); - lean_ctor_release(x_27, 3); - lean_ctor_release(x_27, 4); - lean_ctor_release(x_27, 5); - lean_ctor_release(x_27, 6); - x_78 = x_27; -} else { - lean_dec_ref(x_27); - x_78 = lean_box(0); -} -x_79 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_80 = lean_ctor_get(x_28, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_28, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_82 = x_28; -} else { - lean_dec_ref(x_28); - x_82 = lean_box(0); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_59 = lean_ctor_get(x_10, 1); +lean_inc(x_59); +lean_dec(x_10); +x_60 = lean_ctor_get(x_11, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_11, 1); +lean_inc(x_61); +x_62 = lean_ctor_get(x_11, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_11, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_11, 4); +lean_inc(x_64); +x_65 = lean_ctor_get(x_11, 5); +lean_inc(x_65); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + lean_ctor_release(x_11, 1); + lean_ctor_release(x_11, 2); + lean_ctor_release(x_11, 3); + lean_ctor_release(x_11, 4); + lean_ctor_release(x_11, 5); + lean_ctor_release(x_11, 6); + x_66 = x_11; +} else { + lean_dec_ref(x_11); + x_66 = lean_box(0); } -x_83 = lean_ctor_get(x_24, 0); -lean_inc(x_83); -lean_dec(x_24); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_81); -x_85 = l_Lean_PersistentArray_push___rarg(x_18, x_84); -if (lean_is_scalar(x_82)) { - x_86 = lean_alloc_ctor(0, 2, 1); +x_67 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_68 = lean_ctor_get(x_12, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_12, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_70 = x_12; } else { - x_86 = x_82; + lean_dec_ref(x_12); + x_70 = lean_box(0); } -lean_ctor_set(x_86, 0, x_80); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set_uint8(x_86, sizeof(void*)*2, x_79); -if (lean_is_scalar(x_78)) { - x_87 = lean_alloc_ctor(0, 7, 0); +x_71 = lean_ctor_get(x_2, 0); +lean_inc(x_71); +lean_dec(x_2); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +x_73 = l_Lean_PersistentArray_push___rarg(x_1, x_72); +if (lean_is_scalar(x_70)) { + x_74 = lean_alloc_ctor(0, 2, 1); } else { - x_87 = x_78; + x_74 = x_70; } -lean_ctor_set(x_87, 0, x_72); -lean_ctor_set(x_87, 1, x_73); -lean_ctor_set(x_87, 2, x_74); -lean_ctor_set(x_87, 3, x_75); -lean_ctor_set(x_87, 4, x_76); -lean_ctor_set(x_87, 5, x_77); -lean_ctor_set(x_87, 6, x_86); -x_88 = lean_st_ref_set(x_8, x_87, x_71); -lean_dec(x_8); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_90 = x_88; +lean_ctor_set(x_74, 0, x_68); +lean_ctor_set(x_74, 1, x_73); +lean_ctor_set_uint8(x_74, sizeof(void*)*2, x_67); +if (lean_is_scalar(x_66)) { + x_75 = lean_alloc_ctor(0, 7, 0); +} else { + x_75 = x_66; +} +lean_ctor_set(x_75, 0, x_60); +lean_ctor_set(x_75, 1, x_61); +lean_ctor_set(x_75, 2, x_62); +lean_ctor_set(x_75, 3, x_63); +lean_ctor_set(x_75, 4, x_64); +lean_ctor_set(x_75, 5, x_65); +lean_ctor_set(x_75, 6, x_74); +x_76 = lean_st_ref_set(x_8, x_75, x_59); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_78 = x_76; } else { - lean_dec_ref(x_88); - x_90 = lean_box(0); + lean_dec_ref(x_76); + x_78 = lean_box(0); } -if (lean_is_scalar(x_90)) { - x_91 = lean_alloc_ctor(0, 2, 0); +x_79 = lean_box(0); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 2, 0); } else { - x_91 = x_90; + x_80 = x_78; } -lean_ctor_set(x_91, 0, x_21); -lean_ctor_set(x_91, 1, x_89); -return x_91; +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +return x_80; } } else { -lean_object* x_92; uint8_t x_93; -x_92 = lean_ctor_get(x_26, 1); -lean_inc(x_92); -lean_dec(x_26); -x_93 = !lean_is_exclusive(x_27); -if (x_93 == 0) +lean_object* x_81; uint8_t x_82; +x_81 = lean_ctor_get(x_10, 1); +lean_inc(x_81); +lean_dec(x_10); +x_82 = !lean_is_exclusive(x_11); +if (x_82 == 0) { -lean_object* x_94; uint8_t x_95; -x_94 = lean_ctor_get(x_27, 6); -lean_dec(x_94); -x_95 = !lean_is_exclusive(x_28); -if (x_95 == 0) +lean_object* x_83; uint8_t x_84; +x_83 = lean_ctor_get(x_11, 6); +lean_dec(x_83); +x_84 = !lean_is_exclusive(x_12); +if (x_84 == 0) { -lean_object* x_96; uint8_t x_97; -x_96 = lean_ctor_get(x_28, 1); -lean_dec(x_96); -x_97 = !lean_is_exclusive(x_24); -if (x_97 == 0) +lean_object* x_85; uint8_t x_86; +x_85 = lean_ctor_get(x_12, 1); +lean_dec(x_85); +x_86 = !lean_is_exclusive(x_2); +if (x_86 == 0) { -lean_object* x_98; lean_object* x_99; uint8_t x_100; -lean_ctor_set_tag(x_24, 2); -x_98 = l_Lean_PersistentArray_push___rarg(x_18, x_24); -lean_ctor_set(x_28, 1, x_98); -x_99 = lean_st_ref_set(x_8, x_27, x_92); -lean_dec(x_8); -x_100 = !lean_is_exclusive(x_99); -if (x_100 == 0) +lean_object* x_87; lean_object* x_88; uint8_t x_89; +lean_ctor_set_tag(x_2, 2); +x_87 = l_Lean_PersistentArray_push___rarg(x_1, x_2); +lean_ctor_set(x_12, 1, x_87); +x_88 = lean_st_ref_set(x_8, x_11, x_81); +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 0) { -lean_object* x_101; -x_101 = lean_ctor_get(x_99, 0); -lean_dec(x_101); -lean_ctor_set(x_99, 0, x_21); -return x_99; +lean_object* x_90; lean_object* x_91; +x_90 = lean_ctor_get(x_88, 0); +lean_dec(x_90); +x_91 = lean_box(0); +lean_ctor_set(x_88, 0, x_91); +return x_88; } else { -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_99, 1); -lean_inc(x_102); -lean_dec(x_99); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_21); -lean_ctor_set(x_103, 1, x_102); -return x_103; +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_88, 1); +lean_inc(x_92); +lean_dec(x_88); +x_93 = lean_box(0); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_92); +return x_94; } } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_104 = lean_ctor_get(x_24, 0); -lean_inc(x_104); -lean_dec(x_24); -x_105 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = l_Lean_PersistentArray_push___rarg(x_18, x_105); -lean_ctor_set(x_28, 1, x_106); -x_107 = lean_st_ref_set(x_8, x_27, x_92); -lean_dec(x_8); -x_108 = lean_ctor_get(x_107, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_109 = x_107; +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_95 = lean_ctor_get(x_2, 0); +lean_inc(x_95); +lean_dec(x_2); +x_96 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_96, 0, x_95); +x_97 = l_Lean_PersistentArray_push___rarg(x_1, x_96); +lean_ctor_set(x_12, 1, x_97); +x_98 = lean_st_ref_set(x_8, x_11, x_81); +x_99 = lean_ctor_get(x_98, 1); +lean_inc(x_99); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_100 = x_98; } else { - lean_dec_ref(x_107); - x_109 = lean_box(0); + lean_dec_ref(x_98); + x_100 = lean_box(0); } -if (lean_is_scalar(x_109)) { - x_110 = lean_alloc_ctor(0, 2, 0); +x_101 = lean_box(0); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 2, 0); } else { - x_110 = x_109; + x_102 = x_100; } -lean_ctor_set(x_110, 0, x_21); -lean_ctor_set(x_110, 1, x_108); -return x_110; +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_99); +return x_102; } } else { -uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_111 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_112 = lean_ctor_get(x_28, 0); -lean_inc(x_112); -lean_dec(x_28); -x_113 = lean_ctor_get(x_24, 0); -lean_inc(x_113); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_114 = x_24; +uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_103 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_104 = lean_ctor_get(x_12, 0); +lean_inc(x_104); +lean_dec(x_12); +x_105 = lean_ctor_get(x_2, 0); +lean_inc(x_105); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_106 = x_2; } else { - lean_dec_ref(x_24); - x_114 = lean_box(0); + lean_dec_ref(x_2); + x_106 = lean_box(0); } -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(2, 1, 0); +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(2, 1, 0); } else { - x_115 = x_114; - lean_ctor_set_tag(x_115, 2); -} -lean_ctor_set(x_115, 0, x_113); -x_116 = l_Lean_PersistentArray_push___rarg(x_18, x_115); -x_117 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_117, 0, x_112); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set_uint8(x_117, sizeof(void*)*2, x_111); -lean_ctor_set(x_27, 6, x_117); -x_118 = lean_st_ref_set(x_8, x_27, x_92); -lean_dec(x_8); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_120 = x_118; + x_107 = x_106; + lean_ctor_set_tag(x_107, 2); +} +lean_ctor_set(x_107, 0, x_105); +x_108 = l_Lean_PersistentArray_push___rarg(x_1, x_107); +x_109 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_109, 0, x_104); +lean_ctor_set(x_109, 1, x_108); +lean_ctor_set_uint8(x_109, sizeof(void*)*2, x_103); +lean_ctor_set(x_11, 6, x_109); +x_110 = lean_st_ref_set(x_8, x_11, x_81); +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_112 = x_110; } else { - lean_dec_ref(x_118); - x_120 = lean_box(0); + lean_dec_ref(x_110); + x_112 = lean_box(0); } -if (lean_is_scalar(x_120)) { - x_121 = lean_alloc_ctor(0, 2, 0); +x_113 = lean_box(0); +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(0, 2, 0); } else { - x_121 = x_120; + x_114 = x_112; } -lean_ctor_set(x_121, 0, x_21); -lean_ctor_set(x_121, 1, x_119); -return x_121; +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_111); +return x_114; } } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_122 = lean_ctor_get(x_27, 0); -x_123 = lean_ctor_get(x_27, 1); -x_124 = lean_ctor_get(x_27, 2); -x_125 = lean_ctor_get(x_27, 3); -x_126 = lean_ctor_get(x_27, 4); -x_127 = lean_ctor_get(x_27, 5); -lean_inc(x_127); -lean_inc(x_126); -lean_inc(x_125); -lean_inc(x_124); -lean_inc(x_123); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_115 = lean_ctor_get(x_11, 0); +x_116 = lean_ctor_get(x_11, 1); +x_117 = lean_ctor_get(x_11, 2); +x_118 = lean_ctor_get(x_11, 3); +x_119 = lean_ctor_get(x_11, 4); +x_120 = lean_ctor_get(x_11, 5); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); +lean_inc(x_116); +lean_inc(x_115); +lean_dec(x_11); +x_121 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_122 = lean_ctor_get(x_12, 0); lean_inc(x_122); -lean_dec(x_27); -x_128 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_129 = lean_ctor_get(x_28, 0); -lean_inc(x_129); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_130 = x_28; +if (lean_is_exclusive(x_12)) { + lean_ctor_release(x_12, 0); + lean_ctor_release(x_12, 1); + x_123 = x_12; } else { - lean_dec_ref(x_28); - x_130 = lean_box(0); + lean_dec_ref(x_12); + x_123 = lean_box(0); } -x_131 = lean_ctor_get(x_24, 0); -lean_inc(x_131); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_132 = x_24; +x_124 = lean_ctor_get(x_2, 0); +lean_inc(x_124); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_125 = x_2; } else { - lean_dec_ref(x_24); - x_132 = lean_box(0); + lean_dec_ref(x_2); + x_125 = lean_box(0); } -if (lean_is_scalar(x_132)) { - x_133 = lean_alloc_ctor(2, 1, 0); +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(2, 1, 0); } else { - x_133 = x_132; - lean_ctor_set_tag(x_133, 2); + x_126 = x_125; + lean_ctor_set_tag(x_126, 2); } -lean_ctor_set(x_133, 0, x_131); -x_134 = l_Lean_PersistentArray_push___rarg(x_18, x_133); -if (lean_is_scalar(x_130)) { - x_135 = lean_alloc_ctor(0, 2, 1); -} else { - x_135 = x_130; -} -lean_ctor_set(x_135, 0, x_129); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set_uint8(x_135, sizeof(void*)*2, x_128); -x_136 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_136, 0, x_122); -lean_ctor_set(x_136, 1, x_123); -lean_ctor_set(x_136, 2, x_124); -lean_ctor_set(x_136, 3, x_125); -lean_ctor_set(x_136, 4, x_126); -lean_ctor_set(x_136, 5, x_127); -lean_ctor_set(x_136, 6, x_135); -x_137 = lean_st_ref_set(x_8, x_136, x_92); -lean_dec(x_8); -x_138 = lean_ctor_get(x_137, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_137)) { - lean_ctor_release(x_137, 0); - lean_ctor_release(x_137, 1); - x_139 = x_137; +lean_ctor_set(x_126, 0, x_124); +x_127 = l_Lean_PersistentArray_push___rarg(x_1, x_126); +if (lean_is_scalar(x_123)) { + x_128 = lean_alloc_ctor(0, 2, 1); } else { - lean_dec_ref(x_137); - x_139 = lean_box(0); + x_128 = x_123; } -if (lean_is_scalar(x_139)) { - x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_122); +lean_ctor_set(x_128, 1, x_127); +lean_ctor_set_uint8(x_128, sizeof(void*)*2, x_121); +x_129 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_129, 0, x_115); +lean_ctor_set(x_129, 1, x_116); +lean_ctor_set(x_129, 2, x_117); +lean_ctor_set(x_129, 3, x_118); +lean_ctor_set(x_129, 4, x_119); +lean_ctor_set(x_129, 5, x_120); +lean_ctor_set(x_129, 6, x_128); +x_130 = lean_st_ref_set(x_8, x_129, x_81); +x_131 = lean_ctor_get(x_130, 1); +lean_inc(x_131); +if (lean_is_exclusive(x_130)) { + lean_ctor_release(x_130, 0); + lean_ctor_release(x_130, 1); + x_132 = x_130; } else { - x_140 = x_139; -} -lean_ctor_set(x_140, 0, x_21); -lean_ctor_set(x_140, 1, x_138); -return x_140; -} + lean_dec_ref(x_130); + x_132 = lean_box(0); } +x_133 = lean_box(0); +if (lean_is_scalar(x_132)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_132; } -else -{ -uint8_t x_141; -lean_dec(x_21); -lean_dec(x_18); -lean_dec(x_8); -x_141 = !lean_is_exclusive(x_23); -if (x_141 == 0) -{ -return x_23; +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_131); +return x_134; } -else -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_23, 0); -x_143 = lean_ctor_get(x_23, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_23); -x_144 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -return x_144; } } } -else +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 6); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_dec(x_3); lean_dec(x_2); -x_145 = lean_ctor_get(x_20, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_20, 1); -lean_inc(x_146); -lean_dec(x_20); -x_147 = lean_st_ref_take(x_8, x_146); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_148, 6); -lean_inc(x_149); -x_150 = lean_ctor_get(x_147, 1); -lean_inc(x_150); -lean_dec(x_147); -x_151 = !lean_is_exclusive(x_148); -if (x_151 == 0) +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_dec(x_11); +x_16 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_16; +} +else { -lean_object* x_152; uint8_t x_153; -x_152 = lean_ctor_get(x_148, 6); -lean_dec(x_152); -x_153 = !lean_is_exclusive(x_149); -if (x_153 == 0) +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(x_9, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_21 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_154; lean_object* x_155; uint8_t x_156; -x_154 = lean_ctor_get(x_149, 1); -lean_dec(x_154); -lean_ctor_set(x_149, 1, x_18); -x_155 = lean_st_ref_set(x_8, x_148, x_150); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_3); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_22); +x_24 = lean_apply_8(x_2, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_24) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(x_19, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +lean_dec(x_9); lean_dec(x_8); -x_156 = !lean_is_exclusive(x_155); -if (x_156 == 0) +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -lean_object* x_157; -x_157 = lean_ctor_get(x_155, 0); -lean_dec(x_157); -lean_ctor_set_tag(x_155, 1); -lean_ctor_set(x_155, 0, x_145); -return x_155; +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_22); +return x_27; } else { -lean_object* x_158; lean_object* x_159; -x_158 = lean_ctor_get(x_155, 1); -lean_inc(x_158); -lean_dec(x_155); -x_159 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_159, 0, x_145); -lean_ctor_set(x_159, 1, x_158); -return x_159; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_22); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -uint8_t x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_160 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); -x_161 = lean_ctor_get(x_149, 0); -lean_inc(x_161); -lean_dec(x_149); -x_162 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_18); -lean_ctor_set_uint8(x_162, sizeof(void*)*2, x_160); -lean_ctor_set(x_148, 6, x_162); -x_163 = lean_st_ref_set(x_8, x_148, x_150); +uint8_t x_32; +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_9); lean_dec(x_8); -x_164 = lean_ctor_get(x_163, 1); -lean_inc(x_164); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_165 = x_163; -} else { - lean_dec_ref(x_163); - x_165 = lean_box(0); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) +{ +return x_24; } -if (lean_is_scalar(x_165)) { - x_166 = lean_alloc_ctor(1, 2, 0); -} else { - x_166 = x_165; - lean_ctor_set_tag(x_166, 1); +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 0); +x_34 = lean_ctor_get(x_24, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_24); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } -lean_ctor_set(x_166, 0, x_145); -lean_ctor_set(x_166, 1, x_164); -return x_166; } } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_167 = lean_ctor_get(x_148, 0); -x_168 = lean_ctor_get(x_148, 1); -x_169 = lean_ctor_get(x_148, 2); -x_170 = lean_ctor_get(x_148, 3); -x_171 = lean_ctor_get(x_148, 4); -x_172 = lean_ctor_get(x_148, 5); -lean_inc(x_172); -lean_inc(x_171); -lean_inc(x_170); -lean_inc(x_169); -lean_inc(x_168); -lean_inc(x_167); -lean_dec(x_148); -x_173 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); -x_174 = lean_ctor_get(x_149, 0); -lean_inc(x_174); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_175 = x_149; -} else { - lean_dec_ref(x_149); - x_175 = lean_box(0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_2); +x_36 = lean_ctor_get(x_21, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_21, 1); +lean_inc(x_37); +lean_dec(x_21); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_38 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_39); +x_42 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(x_19, x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +lean_ctor_set_tag(x_42, 1); +lean_ctor_set(x_42, 0, x_36); +return x_42; } -if (lean_is_scalar(x_175)) { - x_176 = lean_alloc_ctor(0, 2, 1); -} else { - x_176 = x_175; -} -lean_ctor_set(x_176, 0, x_174); -lean_ctor_set(x_176, 1, x_18); -lean_ctor_set_uint8(x_176, sizeof(void*)*2, x_173); -x_177 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_177, 0, x_167); -lean_ctor_set(x_177, 1, x_168); -lean_ctor_set(x_177, 2, x_169); -lean_ctor_set(x_177, 3, x_170); -lean_ctor_set(x_177, 4, x_171); -lean_ctor_set(x_177, 5, x_172); -lean_ctor_set(x_177, 6, x_176); -x_178 = lean_st_ref_set(x_8, x_177, x_150); -lean_dec(x_8); -x_179 = lean_ctor_get(x_178, 1); -lean_inc(x_179); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_180 = x_178; -} else { - lean_dec_ref(x_178); - x_180 = lean_box(0); +else +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_45); +return x_46; } -if (lean_is_scalar(x_180)) { - x_181 = lean_alloc_ctor(1, 2, 0); -} else { - x_181 = x_180; - lean_ctor_set_tag(x_181, 1); } -lean_ctor_set(x_181, 0, x_145); -lean_ctor_set(x_181, 1, x_179); -return x_181; +else +{ +uint8_t x_47; +lean_dec(x_36); +lean_dec(x_19); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) +{ +return x_38; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_38, 0); +x_49 = lean_ctor_get(x_38, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_38); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} } } } @@ -47571,20 +47622,20 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_15; uint8_t x_29; -x_29 = lean_ctor_get_uint8(x_8, sizeof(void*)*9 + 7); -if (x_29 == 0) +lean_object* x_15; uint8_t x_30; +x_30 = lean_ctor_get_uint8(x_8, sizeof(void*)*9 + 7); +if (x_30 == 0) { -lean_object* x_30; -x_30 = lean_box(0); -x_15 = x_30; -goto block_28; +lean_object* x_31; +x_31 = lean_box(0); +x_15 = x_31; +goto block_29; } else { if (x_7 == 0) { -lean_object* x_31; lean_object* x_32; +lean_object* x_32; lean_object* x_33; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -47594,26 +47645,30 @@ lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_31 = l_Lean_mkPatternWithRef(x_2, x_1); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_14); -return x_32; +x_32 = l_Lean_mkPatternWithRef(x_2, x_1); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_14); +return x_33; } else { -lean_object* x_33; -x_33 = lean_box(0); -x_15 = x_33; -goto block_28; +lean_object* x_34; +x_34 = lean_box(0); +x_15 = x_34; +goto block_29; } } -block_28: +block_29: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_dec(x_15); x_16 = lean_box(x_6); +lean_inc(x_4); +lean_inc(x_3); lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_5); x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Term_addTermInfo___lambda__1___boxed), 14, 6); lean_closure_set(x_17, 0, x_5); lean_closure_set(x_17, 1, x_1); @@ -47621,57 +47676,76 @@ lean_closure_set(x_17, 2, x_2); lean_closure_set(x_17, 3, x_3); lean_closure_set(x_17, 4, x_4); lean_closure_set(x_17, 5, x_16); -x_18 = l_Lean_Elab_Term_addTermInfo___closed__1; -x_19 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(x_18, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_19) == 0) +x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkPartialTermInfo___boxed), 11, 4); +lean_closure_set(x_18, 0, x_5); +lean_closure_set(x_18, 1, x_1); +lean_closure_set(x_18, 2, x_3); +lean_closure_set(x_18, 3, x_4); +x_19 = l_Lean_Elab_Term_addTermInfo___closed__1; +x_20 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1(x_19, x_17, x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_20) == 0) { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_21; -x_21 = lean_ctor_get(x_19, 0); -lean_dec(x_21); -lean_ctor_set(x_19, 0, x_2); -return x_19; +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_2); +return x_20; } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_2); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_2); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } else { -uint8_t x_24; +uint8_t x_25; lean_dec(x_2); -x_24 = !lean_is_exclusive(x_19); -if (x_24 == 0) +x_25 = !lean_is_exclusive(x_20); +if (x_25 == 0) { -return x_19; +return x_20; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_19, 0); -x_26 = lean_ctor_get(x_19, 1); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 0); +x_27 = lean_ctor_get(x_20, 1); +lean_inc(x_27); lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_19); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_dec(x_20); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} } } } } +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} } LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: @@ -47768,807 +47842,293 @@ x_15 = l_Lean_Elab_Term_addTermInfo_x27(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = lean_st_ref_get(x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 6); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +x_13 = lean_ctor_get(x_12, 6); +lean_inc(x_13); lean_dec(x_12); -if (x_13 == 0) +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; +lean_dec(x_3); lean_dec(x_2); -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -return x_15; +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_dec(x_11); +x_16 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_16; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_dec(x_10); -x_17 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(x_8, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(x_9, x_17); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_20 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_20) == 0) +x_21 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_3); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_9); lean_inc(x_8); -lean_inc(x_21); -x_23 = lean_apply_8(x_2, x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_st_ref_take(x_8, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 6); -lean_inc(x_28); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_22); +x_24 = lean_apply_8(x_2, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_23); if (lean_obj_tag(x_24) == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_26); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_26, 1); -x_31 = lean_ctor_get(x_26, 0); -lean_dec(x_31); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_27, 6); -lean_dec(x_33); -x_34 = !lean_is_exclusive(x_28); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_35 = lean_ctor_get(x_28, 1); -x_36 = lean_ctor_get(x_24, 0); -lean_inc(x_36); +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_35); -lean_ctor_set(x_26, 0, x_36); -x_37 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -lean_ctor_set(x_28, 1, x_37); -x_38 = lean_st_ref_set(x_8, x_27, x_30); +x_27 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(x_19, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +lean_dec(x_9); lean_dec(x_8); -x_39 = !lean_is_exclusive(x_38); -if (x_39 == 0) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); -lean_dec(x_40); -lean_ctor_set(x_38, 0, x_21); -return x_38; -} -else -{ -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_21); -lean_ctor_set(x_42, 1, x_41); -return x_42; -} -} -else +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_43 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_44 = lean_ctor_get(x_28, 0); -x_45 = lean_ctor_get(x_28, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_28); -x_46 = lean_ctor_get(x_24, 0); -lean_inc(x_46); -lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_45); -lean_ctor_set(x_26, 0, x_46); -x_47 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -x_48 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_48, 0, x_44); -lean_ctor_set(x_48, 1, x_47); -lean_ctor_set_uint8(x_48, sizeof(void*)*2, x_43); -lean_ctor_set(x_27, 6, x_48); -x_49 = lean_st_ref_set(x_8, x_27, x_30); -lean_dec(x_8); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_51 = x_49; -} else { - lean_dec_ref(x_49); - x_51 = lean_box(0); -} -if (lean_is_scalar(x_51)) { - x_52 = lean_alloc_ctor(0, 2, 0); -} else { - x_52 = x_51; -} -lean_ctor_set(x_52, 0, x_21); -lean_ctor_set(x_52, 1, x_50); -return x_52; -} +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_22); +return x_27; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_53 = lean_ctor_get(x_27, 0); -x_54 = lean_ctor_get(x_27, 1); -x_55 = lean_ctor_get(x_27, 2); -x_56 = lean_ctor_get(x_27, 3); -x_57 = lean_ctor_get(x_27, 4); -x_58 = lean_ctor_get(x_27, 5); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_inc(x_53); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); lean_dec(x_27); -x_59 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_60 = lean_ctor_get(x_28, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_28, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_62 = x_28; -} else { - lean_dec_ref(x_28); - x_62 = lean_box(0); -} -x_63 = lean_ctor_get(x_24, 0); -lean_inc(x_63); -lean_dec(x_24); -lean_ctor_set_tag(x_26, 1); -lean_ctor_set(x_26, 1, x_61); -lean_ctor_set(x_26, 0, x_63); -x_64 = l_Lean_PersistentArray_push___rarg(x_18, x_26); -if (lean_is_scalar(x_62)) { - x_65 = lean_alloc_ctor(0, 2, 1); -} else { - x_65 = x_62; -} -lean_ctor_set(x_65, 0, x_60); -lean_ctor_set(x_65, 1, x_64); -lean_ctor_set_uint8(x_65, sizeof(void*)*2, x_59); -x_66 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_66, 0, x_53); -lean_ctor_set(x_66, 1, x_54); -lean_ctor_set(x_66, 2, x_55); -lean_ctor_set(x_66, 3, x_56); -lean_ctor_set(x_66, 4, x_57); -lean_ctor_set(x_66, 5, x_58); -lean_ctor_set(x_66, 6, x_65); -x_67 = lean_st_ref_set(x_8, x_66, x_30); -lean_dec(x_8); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_69 = x_67; -} else { - lean_dec_ref(x_67); - x_69 = lean_box(0); -} -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 2, 0); -} else { - x_70 = x_69; -} -lean_ctor_set(x_70, 0, x_21); -lean_ctor_set(x_70, 1, x_68); -return x_70; -} -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_71 = lean_ctor_get(x_26, 1); -lean_inc(x_71); -lean_dec(x_26); -x_72 = lean_ctor_get(x_27, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_27, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_27, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_27, 3); -lean_inc(x_75); -x_76 = lean_ctor_get(x_27, 4); -lean_inc(x_76); -x_77 = lean_ctor_get(x_27, 5); -lean_inc(x_77); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - lean_ctor_release(x_27, 2); - lean_ctor_release(x_27, 3); - lean_ctor_release(x_27, 4); - lean_ctor_release(x_27, 5); - lean_ctor_release(x_27, 6); - x_78 = x_27; -} else { - lean_dec_ref(x_27); - x_78 = lean_box(0); -} -x_79 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_80 = lean_ctor_get(x_28, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_28, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_82 = x_28; -} else { - lean_dec_ref(x_28); - x_82 = lean_box(0); -} -x_83 = lean_ctor_get(x_24, 0); -lean_inc(x_83); -lean_dec(x_24); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_81); -x_85 = l_Lean_PersistentArray_push___rarg(x_18, x_84); -if (lean_is_scalar(x_82)) { - x_86 = lean_alloc_ctor(0, 2, 1); -} else { - x_86 = x_82; -} -lean_ctor_set(x_86, 0, x_80); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set_uint8(x_86, sizeof(void*)*2, x_79); -if (lean_is_scalar(x_78)) { - x_87 = lean_alloc_ctor(0, 7, 0); -} else { - x_87 = x_78; -} -lean_ctor_set(x_87, 0, x_72); -lean_ctor_set(x_87, 1, x_73); -lean_ctor_set(x_87, 2, x_74); -lean_ctor_set(x_87, 3, x_75); -lean_ctor_set(x_87, 4, x_76); -lean_ctor_set(x_87, 5, x_77); -lean_ctor_set(x_87, 6, x_86); -x_88 = lean_st_ref_set(x_8, x_87, x_71); -lean_dec(x_8); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_90 = x_88; -} else { - lean_dec_ref(x_88); - x_90 = lean_box(0); -} -if (lean_is_scalar(x_90)) { - x_91 = lean_alloc_ctor(0, 2, 0); -} else { - x_91 = x_90; -} -lean_ctor_set(x_91, 0, x_21); -lean_ctor_set(x_91, 1, x_89); -return x_91; +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_22); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } else { -lean_object* x_92; uint8_t x_93; -x_92 = lean_ctor_get(x_26, 1); -lean_inc(x_92); -lean_dec(x_26); -x_93 = !lean_is_exclusive(x_27); -if (x_93 == 0) -{ -lean_object* x_94; uint8_t x_95; -x_94 = lean_ctor_get(x_27, 6); -lean_dec(x_94); -x_95 = !lean_is_exclusive(x_28); -if (x_95 == 0) -{ -lean_object* x_96; uint8_t x_97; -x_96 = lean_ctor_get(x_28, 1); -lean_dec(x_96); -x_97 = !lean_is_exclusive(x_24); -if (x_97 == 0) -{ -lean_object* x_98; lean_object* x_99; uint8_t x_100; -lean_ctor_set_tag(x_24, 2); -x_98 = l_Lean_PersistentArray_push___rarg(x_18, x_24); -lean_ctor_set(x_28, 1, x_98); -x_99 = lean_st_ref_set(x_8, x_27, x_92); +uint8_t x_32; +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_9); lean_dec(x_8); -x_100 = !lean_is_exclusive(x_99); -if (x_100 == 0) -{ -lean_object* x_101; -x_101 = lean_ctor_get(x_99, 0); -lean_dec(x_101); -lean_ctor_set(x_99, 0, x_21); -return x_99; -} -else +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) { -lean_object* x_102; lean_object* x_103; -x_102 = lean_ctor_get(x_99, 1); -lean_inc(x_102); -lean_dec(x_99); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_21); -lean_ctor_set(x_103, 1, x_102); -return x_103; -} +return x_24; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_104 = lean_ctor_get(x_24, 0); -lean_inc(x_104); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_24, 0); +x_34 = lean_ctor_get(x_24, 1); +lean_inc(x_34); +lean_inc(x_33); lean_dec(x_24); -x_105 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = l_Lean_PersistentArray_push___rarg(x_18, x_105); -lean_ctor_set(x_28, 1, x_106); -x_107 = lean_st_ref_set(x_8, x_27, x_92); -lean_dec(x_8); -x_108 = lean_ctor_get(x_107, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_109 = x_107; -} else { - lean_dec_ref(x_107); - x_109 = lean_box(0); -} -if (lean_is_scalar(x_109)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_109; -} -lean_ctor_set(x_110, 0, x_21); -lean_ctor_set(x_110, 1, x_108); -return x_110; -} -} -else -{ -uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_111 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_112 = lean_ctor_get(x_28, 0); -lean_inc(x_112); -lean_dec(x_28); -x_113 = lean_ctor_get(x_24, 0); -lean_inc(x_113); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_114 = x_24; -} else { - lean_dec_ref(x_24); - x_114 = lean_box(0); -} -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(2, 1, 0); -} else { - x_115 = x_114; - lean_ctor_set_tag(x_115, 2); -} -lean_ctor_set(x_115, 0, x_113); -x_116 = l_Lean_PersistentArray_push___rarg(x_18, x_115); -x_117 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_117, 0, x_112); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set_uint8(x_117, sizeof(void*)*2, x_111); -lean_ctor_set(x_27, 6, x_117); -x_118 = lean_st_ref_set(x_8, x_27, x_92); -lean_dec(x_8); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_120 = x_118; -} else { - lean_dec_ref(x_118); - x_120 = lean_box(0); -} -if (lean_is_scalar(x_120)) { - x_121 = lean_alloc_ctor(0, 2, 0); -} else { - x_121 = x_120; -} -lean_ctor_set(x_121, 0, x_21); -lean_ctor_set(x_121, 1, x_119); -return x_121; -} -} -else -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_122 = lean_ctor_get(x_27, 0); -x_123 = lean_ctor_get(x_27, 1); -x_124 = lean_ctor_get(x_27, 2); -x_125 = lean_ctor_get(x_27, 3); -x_126 = lean_ctor_get(x_27, 4); -x_127 = lean_ctor_get(x_27, 5); -lean_inc(x_127); -lean_inc(x_126); -lean_inc(x_125); -lean_inc(x_124); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_27); -x_128 = lean_ctor_get_uint8(x_28, sizeof(void*)*2); -x_129 = lean_ctor_get(x_28, 0); -lean_inc(x_129); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_130 = x_28; -} else { - lean_dec_ref(x_28); - x_130 = lean_box(0); -} -x_131 = lean_ctor_get(x_24, 0); -lean_inc(x_131); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_132 = x_24; -} else { - lean_dec_ref(x_24); - x_132 = lean_box(0); -} -if (lean_is_scalar(x_132)) { - x_133 = lean_alloc_ctor(2, 1, 0); -} else { - x_133 = x_132; - lean_ctor_set_tag(x_133, 2); -} -lean_ctor_set(x_133, 0, x_131); -x_134 = l_Lean_PersistentArray_push___rarg(x_18, x_133); -if (lean_is_scalar(x_130)) { - x_135 = lean_alloc_ctor(0, 2, 1); -} else { - x_135 = x_130; -} -lean_ctor_set(x_135, 0, x_129); -lean_ctor_set(x_135, 1, x_134); -lean_ctor_set_uint8(x_135, sizeof(void*)*2, x_128); -x_136 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_136, 0, x_122); -lean_ctor_set(x_136, 1, x_123); -lean_ctor_set(x_136, 2, x_124); -lean_ctor_set(x_136, 3, x_125); -lean_ctor_set(x_136, 4, x_126); -lean_ctor_set(x_136, 5, x_127); -lean_ctor_set(x_136, 6, x_135); -x_137 = lean_st_ref_set(x_8, x_136, x_92); -lean_dec(x_8); -x_138 = lean_ctor_get(x_137, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_137)) { - lean_ctor_release(x_137, 0); - lean_ctor_release(x_137, 1); - x_139 = x_137; -} else { - lean_dec_ref(x_137); - x_139 = lean_box(0); -} -if (lean_is_scalar(x_139)) { - x_140 = lean_alloc_ctor(0, 2, 0); -} else { - x_140 = x_139; -} -lean_ctor_set(x_140, 0, x_21); -lean_ctor_set(x_140, 1, x_138); -return x_140; +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } } else { -uint8_t x_141; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_2); +x_36 = lean_ctor_get(x_21, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_21, 1); +lean_inc(x_37); lean_dec(x_21); -lean_dec(x_18); -lean_dec(x_8); -x_141 = !lean_is_exclusive(x_23); -if (x_141 == 0) -{ -return x_23; -} -else -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_23, 0); -x_143 = lean_ctor_get(x_23, 1); -lean_inc(x_143); -lean_inc(x_142); -lean_dec(x_23); -x_144 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_144, 0, x_142); -lean_ctor_set(x_144, 1, x_143); -return x_144; -} -} -} -else +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_38 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_39); +x_42 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_addTermInfo___spec__1___lambda__1(x_19, x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_145 = lean_ctor_get(x_20, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_20, 1); -lean_inc(x_146); -lean_dec(x_20); -x_147 = lean_st_ref_take(x_8, x_146); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_148, 6); -lean_inc(x_149); -x_150 = lean_ctor_get(x_147, 1); -lean_inc(x_150); -lean_dec(x_147); -x_151 = !lean_is_exclusive(x_148); -if (x_151 == 0) -{ -lean_object* x_152; uint8_t x_153; -x_152 = lean_ctor_get(x_148, 6); -lean_dec(x_152); -x_153 = !lean_is_exclusive(x_149); -if (x_153 == 0) -{ -lean_object* x_154; lean_object* x_155; uint8_t x_156; -x_154 = lean_ctor_get(x_149, 1); -lean_dec(x_154); -lean_ctor_set(x_149, 1, x_18); -x_155 = lean_st_ref_set(x_8, x_148, x_150); -lean_dec(x_8); -x_156 = !lean_is_exclusive(x_155); -if (x_156 == 0) +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_157; -x_157 = lean_ctor_get(x_155, 0); -lean_dec(x_157); -lean_ctor_set_tag(x_155, 1); -lean_ctor_set(x_155, 0, x_145); -return x_155; +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +lean_ctor_set_tag(x_42, 1); +lean_ctor_set(x_42, 0, x_36); +return x_42; } else { -lean_object* x_158; lean_object* x_159; -x_158 = lean_ctor_get(x_155, 1); -lean_inc(x_158); -lean_dec(x_155); -x_159 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_159, 0, x_145); -lean_ctor_set(x_159, 1, x_158); -return x_159; +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } else { -uint8_t x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_160 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); -x_161 = lean_ctor_get(x_149, 0); -lean_inc(x_161); -lean_dec(x_149); -x_162 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_18); -lean_ctor_set_uint8(x_162, sizeof(void*)*2, x_160); -lean_ctor_set(x_148, 6, x_162); -x_163 = lean_st_ref_set(x_8, x_148, x_150); +uint8_t x_47; +lean_dec(x_36); +lean_dec(x_19); +lean_dec(x_9); lean_dec(x_8); -x_164 = lean_ctor_get(x_163, 1); -lean_inc(x_164); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_165 = x_163; -} else { - lean_dec_ref(x_163); - x_165 = lean_box(0); -} -if (lean_is_scalar(x_165)) { - x_166 = lean_alloc_ctor(1, 2, 0); -} else { - x_166 = x_165; - lean_ctor_set_tag(x_166, 1); -} -lean_ctor_set(x_166, 0, x_145); -lean_ctor_set(x_166, 1, x_164); -return x_166; -} +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) +{ +return x_38; } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; uint8_t x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_167 = lean_ctor_get(x_148, 0); -x_168 = lean_ctor_get(x_148, 1); -x_169 = lean_ctor_get(x_148, 2); -x_170 = lean_ctor_get(x_148, 3); -x_171 = lean_ctor_get(x_148, 4); -x_172 = lean_ctor_get(x_148, 5); -lean_inc(x_172); -lean_inc(x_171); -lean_inc(x_170); -lean_inc(x_169); -lean_inc(x_168); -lean_inc(x_167); -lean_dec(x_148); -x_173 = lean_ctor_get_uint8(x_149, sizeof(void*)*2); -x_174 = lean_ctor_get(x_149, 0); -lean_inc(x_174); -if (lean_is_exclusive(x_149)) { - lean_ctor_release(x_149, 0); - lean_ctor_release(x_149, 1); - x_175 = x_149; -} else { - lean_dec_ref(x_149); - x_175 = lean_box(0); -} -if (lean_is_scalar(x_175)) { - x_176 = lean_alloc_ctor(0, 2, 1); -} else { - x_176 = x_175; -} -lean_ctor_set(x_176, 0, x_174); -lean_ctor_set(x_176, 1, x_18); -lean_ctor_set_uint8(x_176, sizeof(void*)*2, x_173); -x_177 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_177, 0, x_167); -lean_ctor_set(x_177, 1, x_168); -lean_ctor_set(x_177, 2, x_169); -lean_ctor_set(x_177, 3, x_170); -lean_ctor_set(x_177, 4, x_171); -lean_ctor_set(x_177, 5, x_172); -lean_ctor_set(x_177, 6, x_176); -x_178 = lean_st_ref_set(x_8, x_177, x_150); -lean_dec(x_8); -x_179 = lean_ctor_get(x_178, 1); -lean_inc(x_179); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_180 = x_178; -} else { - lean_dec_ref(x_178); - x_180 = lean_box(0); -} -if (lean_is_scalar(x_180)) { - x_181 = lean_alloc_ctor(1, 2, 0); -} else { - x_181 = x_180; - lean_ctor_set_tag(x_181, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_38, 0); +x_49 = lean_ctor_get(x_38, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_38); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } -lean_ctor_set(x_181, 0, x_145); -lean_ctor_set(x_181, 1, x_179); -return x_181; } } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_11; -x_11 = lean_ctor_get_uint8(x_4, sizeof(void*)*9 + 7); -if (x_11 == 0) +uint8_t x_12; +x_12 = lean_ctor_get_uint8(x_5, sizeof(void*)*9 + 7); +if (x_12 == 0) { -lean_object* x_12; +lean_object* x_13; lean_dec(x_1); -x_12 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; +x_13 = l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } else { -lean_object* x_13; +lean_object* x_14; +lean_dec(x_4); lean_dec(x_3); -x_13 = lean_apply_7(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) +x_14 = lean_apply_7(x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -x_16 = l_Lean_mkPatternWithRef(x_15, x_1); -lean_ctor_set(x_13, 0, x_16); -return x_13; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_mkPatternWithRef(x_16, x_1); +lean_ctor_set(x_14, 0, x_17); +return x_14; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_14, 0); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_13); -x_19 = l_Lean_mkPatternWithRef(x_17, x_1); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_dec(x_14); +x_20 = l_Lean_mkPatternWithRef(x_18, x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } else { -uint8_t x_21; +uint8_t x_22; lean_dec(x_1); -x_21 = !lean_is_exclusive(x_13); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_14); +if (x_22 == 0) { -return x_13; +return x_14; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_13, 0); -x_23 = lean_ctor_get(x_13, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +x_24 = lean_ctor_get(x_14, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_13); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_14); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } } } -static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1() { _start: { lean_object* x_1; @@ -48576,23 +48136,23 @@ x_1 = lean_mk_string_unchecked("BodyInfo", 8, 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; x_3 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; -x_4 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1; +x_4 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_() { +static lean_object* _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2; +x_1 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2; return x_1; } } @@ -48600,7 +48160,7 @@ static lean_object* _init_l_Lean_Elab_Term_instTypeNameBodyInfo() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_; +x_1 = l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_; return x_1; } } @@ -48615,7 +48175,7 @@ lean_ctor_set(x_4, 1, x_2); x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_4); -x_6 = lean_alloc_ctor(8, 1, 0); +x_6 = lean_alloc_ctor(9, 1, 0); lean_ctor_set(x_6, 0, x_5); return x_6; } @@ -48623,7 +48183,7 @@ return x_6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_getBodyInfo_x3f(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 8) +if (lean_obj_tag(x_1) == 9) { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = lean_ctor_get(x_1, 0); @@ -48649,46 +48209,79 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_11 = lean_box(0); -x_12 = lean_box(0); -x_13 = 0; -x_14 = l_Lean_Elab_Term_mkTermInfo(x_12, x_1, x_3, x_2, x_11, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_object* x_14; +x_14 = l_Lean_Elab_Term_mkTermInfo(x_1, x_2, x_6, x_3, x_4, x_5, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_box(x_6); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_2); lean_inc(x_1); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___boxed), 9, 2); -lean_closure_set(x_10, 0, x_1); -lean_closure_set(x_10, 1, x_2); -lean_inc(x_1); -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Term_postponeElabTerm___lambda__1___boxed), 10, 2); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_2); -x_12 = l_Lean_Elab_Term_withInfoContext_x27(x_1, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1___boxed), 13, 5); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +lean_closure_set(x_15, 2, x_4); +lean_closure_set(x_15, 3, x_5); +lean_closure_set(x_15, 4, x_14); +lean_inc(x_2); +x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkPartialTermInfo___boxed), 11, 4); +lean_closure_set(x_16, 0, x_1); +lean_closure_set(x_16, 1, x_2); +lean_closure_set(x_16, 2, x_4); +lean_closure_set(x_16, 3, x_5); +x_17 = l_Lean_Elab_Term_withInfoContext_x27(x_2, x_3, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_postponeElabTerm___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_5); +lean_dec(x_5); +x_15 = l_Lean_Elab_Term_withTermInfoContext_x27___lambda__1(x_1, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withTermInfoContext_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_6); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; +x_15 = l_Lean_Elab_Term_withTermInfoContext_x27(x_1, x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +lean_inc(x_2); +lean_inc(x_1); +x_10 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___boxed), 9, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_2); +x_11 = lean_box(0); +x_12 = lean_box(0); +x_13 = 0; +x_14 = l_Lean_Elab_Term_withTermInfoContext_x27(x_12, x_1, x_10, x_2, x_11, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_14; } } static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1() { @@ -49249,22 +48842,6 @@ return x_79; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_box(0); -x_15 = 0; -x_16 = l_Lean_Elab_Term_mkTermInfo(x_13, x_2, x_4, x_3, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_16; -} -} static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___closed__1() { _start: { @@ -49310,39 +48887,42 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; x_20 = lean_ctor_get(x_5, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_5, 1); lean_inc(x_21); lean_dec(x_5); -x_22 = lean_ctor_get(x_20, 1); +x_22 = lean_ctor_get(x_20, 0); lean_inc(x_22); -x_23 = lean_box(x_4); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_box(x_4); lean_inc(x_1); lean_inc(x_3); lean_inc(x_2); -x_24 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___boxed), 12, 5); -lean_closure_set(x_24, 0, x_22); -lean_closure_set(x_24, 1, x_2); -lean_closure_set(x_24, 2, x_3); -lean_closure_set(x_24, 3, x_23); -lean_closure_set(x_24, 4, x_1); -lean_inc(x_3); -lean_inc(x_2); -x_25 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2___boxed), 11, 3); -lean_closure_set(x_25, 0, x_20); -lean_closure_set(x_25, 1, x_2); -lean_closure_set(x_25, 2, x_3); +x_26 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___boxed), 12, 5); +lean_closure_set(x_26, 0, x_24); +lean_closure_set(x_26, 1, x_2); +lean_closure_set(x_26, 2, x_3); +lean_closure_set(x_26, 3, x_25); +lean_closure_set(x_26, 4, x_1); +x_27 = lean_box(0); +x_28 = 0; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); +lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_Term_withInfoContext_x27(x_2, x_24, x_25, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_26) == 0) +x_29 = l_Lean_Elab_Term_withTermInfoContext_x27(x_23, x_2, x_26, x_3, x_27, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_29) == 0) { lean_dec(x_21); lean_dec(x_11); @@ -49354,25 +48934,25 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_26; +return x_29; } else { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -x_30 = l_Lean_Exception_isInterrupt(x_28); +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) { -uint8_t x_31; -x_31 = l_Lean_Exception_isRuntime(x_28); -if (x_31 == 0) +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +x_33 = l_Lean_Exception_isInterrupt(x_31); +if (x_33 == 0) { -if (lean_obj_tag(x_28) == 0) +uint8_t x_34; +x_34 = l_Lean_Exception_isRuntime(x_31); +if (x_34 == 0) +{ +if (lean_obj_tag(x_31) == 0) { lean_dec(x_21); lean_dec(x_11); @@ -49384,17 +48964,17 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_26; +return x_29; } else { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_ctor_get(x_28, 0); -lean_inc(x_32); -x_33 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1; -x_34 = lean_nat_dec_eq(x_32, x_33); -lean_dec(x_32); -if (x_34 == 0) +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = lean_ctor_get(x_31, 0); +lean_inc(x_35); +x_36 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1; +x_37 = lean_nat_dec_eq(x_35, x_36); +lean_dec(x_35); +if (x_37 == 0) { lean_dec(x_21); lean_dec(x_11); @@ -49406,21 +48986,20 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_26; +return x_29; } else { -uint8_t x_35; lean_object* x_36; lean_object* x_37; -lean_free_object(x_26); -lean_dec(x_28); -x_35 = 0; +lean_object* x_38; lean_object* x_39; +lean_free_object(x_29); +lean_dec(x_31); lean_inc(x_1); -x_36 = l_Lean_Elab_Term_SavedState_restore(x_1, x_35, x_6, x_7, x_8, x_9, x_10, x_11, x_29); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -lean_dec(x_36); +x_38 = l_Lean_Elab_Term_SavedState_restore(x_1, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_32); +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +lean_dec(x_38); x_5 = x_21; -x_12 = x_37; +x_12 = x_39; goto _start; } } @@ -49437,7 +49016,7 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_26; +return x_29; } } else @@ -49452,27 +49031,27 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_26; +return x_29; } } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_26, 0); -x_40 = lean_ctor_get(x_26, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_26); -x_41 = l_Lean_Exception_isInterrupt(x_39); -if (x_41 == 0) +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_ctor_get(x_29, 0); +x_42 = lean_ctor_get(x_29, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_29); +x_43 = l_Lean_Exception_isInterrupt(x_41); +if (x_43 == 0) { -uint8_t x_42; -x_42 = l_Lean_Exception_isRuntime(x_39); -if (x_42 == 0) +uint8_t x_44; +x_44 = l_Lean_Exception_isRuntime(x_41); +if (x_44 == 0) { -if (lean_obj_tag(x_39) == 0) +if (lean_obj_tag(x_41) == 0) { -lean_object* x_43; +lean_object* x_45; lean_dec(x_21); lean_dec(x_11); lean_dec(x_10); @@ -49483,22 +49062,22 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_39); -lean_ctor_set(x_43, 1, x_40); -return x_43; +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_42); +return x_45; } else { -lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_44 = lean_ctor_get(x_39, 0); -lean_inc(x_44); -x_45 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1; -x_46 = lean_nat_dec_eq(x_44, x_45); -lean_dec(x_44); -if (x_46 == 0) +lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_46 = lean_ctor_get(x_41, 0); +lean_inc(x_46); +x_47 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1; +x_48 = lean_nat_dec_eq(x_46, x_47); +lean_dec(x_46); +if (x_48 == 0) { -lean_object* x_47; +lean_object* x_49; lean_dec(x_21); lean_dec(x_11); lean_dec(x_10); @@ -49509,30 +49088,29 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_39); -lean_ctor_set(x_47, 1, x_40); -return x_47; +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_41); +lean_ctor_set(x_49, 1, x_42); +return x_49; } else { -uint8_t x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_39); -x_48 = 0; +lean_object* x_50; lean_object* x_51; +lean_dec(x_41); lean_inc(x_1); -x_49 = l_Lean_Elab_Term_SavedState_restore(x_1, x_48, x_6, x_7, x_8, x_9, x_10, x_11, x_40); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -lean_dec(x_49); +x_50 = l_Lean_Elab_Term_SavedState_restore(x_1, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); x_5 = x_21; -x_12 = x_50; +x_12 = x_51; goto _start; } } } else { -lean_object* x_52; +lean_object* x_53; lean_dec(x_21); lean_dec(x_11); lean_dec(x_10); @@ -49543,15 +49121,15 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_39); -lean_ctor_set(x_52, 1, x_40); -return x_52; +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_41); +lean_ctor_set(x_53, 1, x_42); +return x_53; } } else { -lean_object* x_53; +lean_object* x_54; lean_dec(x_21); lean_dec(x_11); lean_dec(x_10); @@ -49562,10 +49140,10 @@ lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_39); -lean_ctor_set(x_53, 1, x_40); -return x_53; +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_41); +lean_ctor_set(x_54, 1, x_42); +return x_54; } } } @@ -49582,20 +49160,6 @@ x_14 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_12; -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -57913,17 +57477,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_box(0); -x_13 = 0; -x_14 = l_Lean_Elab_Term_mkTermInfo(x_1, x_2, x_4, x_3, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; -} -} -static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -57931,7 +57485,7 @@ x_1 = lean_mk_string_unchecked("elaborator", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2() { _start: { lean_object* x_1; @@ -57939,11 +57493,11 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_el return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; -x_12 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1; +x_12 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1; x_13 = l_Lean_Core_checkSystem(x_12, x_9, x_10, x_11); if (lean_obj_tag(x_13) == 0) { @@ -57975,7 +57529,7 @@ lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); lean_dec(x_20); -x_23 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2; +x_23 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2; if (lean_obj_tag(x_21) == 0) { if (x_4 == 0) @@ -58069,7 +57623,7 @@ lean_inc(x_5); x_41 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(x_40, x_5, x_6, x_7, x_8, x_9, x_10, x_22); if (lean_obj_tag(x_41) == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); x_43 = lean_ctor_get(x_41, 1); @@ -58089,61 +57643,58 @@ x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withMacroExpansion___rarg), 1 lean_closure_set(x_47, 0, x_1); lean_closure_set(x_47, 1, x_42); lean_closure_set(x_47, 2, x_46); -lean_inc(x_1); -x_48 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___boxed), 11, 3); -lean_closure_set(x_48, 0, x_38); -lean_closure_set(x_48, 1, x_1); -lean_closure_set(x_48, 2, x_2); +x_48 = lean_box(0); +x_49 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_49 = l_Lean_Elab_Term_withInfoContext_x27(x_1, x_47, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_43); -if (lean_obj_tag(x_49) == 0) +x_50 = l_Lean_Elab_Term_withTermInfoContext_x27(x_38, x_1, x_47, x_2, x_48, x_49, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_apply_8(x_23, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_51); -return x_52; +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_apply_8(x_23, x_51, x_5, x_6, x_7, x_8, x_9, x_10, x_52); +return x_53; } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_53 = !lean_is_exclusive(x_49); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_50); +if (x_54 == 0) { -return x_49; +return x_50; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_49, 0); -x_55 = lean_ctor_get(x_49, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_50, 0); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_49); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_dec(x_50); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } else { -uint8_t x_57; +uint8_t x_58; lean_dec(x_38); lean_dec(x_10); lean_dec(x_9); @@ -58153,30 +57704,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_57 = !lean_is_exclusive(x_41); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_41); +if (x_58 == 0) { return x_41; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_41, 0); -x_59 = lean_ctor_get(x_41, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_41, 0); +x_60 = lean_ctor_get(x_41, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); lean_dec(x_41); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; } } } } else { -uint8_t x_61; +uint8_t x_62; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -58185,29 +57736,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_61 = !lean_is_exclusive(x_20); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_20); +if (x_62 == 0) { return x_20; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_20, 0); -x_63 = lean_ctor_get(x_20, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_20, 0); +x_64 = lean_ctor_get(x_20, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); lean_dec(x_20); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -uint8_t x_65; +uint8_t x_66; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -58216,28 +57767,28 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_65 = !lean_is_exclusive(x_13); -if (x_65 == 0) +x_66 = !lean_is_exclusive(x_13); +if (x_66 == 0) { return x_13; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_13, 0); -x_67 = lean_ctor_get(x_13, 1); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_13, 0); +x_68 = lean_ctor_get(x_13, 1); +lean_inc(x_68); lean_inc(x_67); -lean_inc(x_66); lean_dec(x_13); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; } } } } -LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(lean_object* x_1) { +LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(lean_object* x_1) { _start: { uint8_t x_2; @@ -58245,7 +57796,7 @@ x_2 = 0; return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; @@ -58385,7 +57936,7 @@ static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermA _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed), 1, 0); return x_1; } } @@ -58409,7 +57960,7 @@ lean_closure_set(x_13, 1, x_1); x_14 = lean_box(x_2); x_15 = lean_box(x_3); lean_inc(x_4); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed), 11, 4); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___boxed), 11, 4); lean_closure_set(x_16, 0, x_4); lean_closure_set(x_16, 1, x_1); lean_closure_set(x_16, 2, x_14); @@ -58419,7 +57970,7 @@ x_18 = 1; x_19 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__2; x_20 = l_Lean_Name_toString(x_17, x_18, x_19); x_21 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1; -x_22 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__9), 11, 8); +x_22 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8), 11, 8); lean_closure_set(x_22, 0, x_21); lean_closure_set(x_22, 1, x_13); lean_closure_set(x_22, 2, x_16); @@ -58725,34 +58276,20 @@ return x_14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_12; -x_12 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ uint8_t x_12; uint8_t x_13; lean_object* x_14; x_12 = lean_unbox(x_3); lean_dec(x_3); x_13 = lean_unbox(x_4); lean_dec(x_4); -x_14 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(x_1); +x_2 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; @@ -59025,7 +58562,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDo _start: { lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_ctor(6, 1, 0); +x_9 = lean_alloc_ctor(7, 1, 0); lean_ctor_set(x_9, 0, x_1); x_10 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; @@ -68778,7 +68315,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1() { _start: { lean_object* x_1; @@ -68786,17 +68323,17 @@ x_1 = lean_mk_string_unchecked("letrec", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -68806,27 +68343,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6() { _start: { lean_object* x_1; @@ -68834,17 +68371,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8() { _start: { lean_object* x_1; @@ -68852,47 +68389,47 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13() { _start: { lean_object* x_1; @@ -68900,33 +68437,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14; -x_2 = lean_unsigned_to_nat(17700u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14; +x_2 = lean_unsigned_to_nat(17795u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -69333,7 +68870,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar(lean_object* x_1, lean _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2; +x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2; x_10 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -79157,77 +78694,77 @@ x_3 = lean_alloc_closure((void*)(l_Lean_Elab_withoutModifyingStateWithInfoAndMes return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6; -x_2 = lean_unsigned_to_nat(20004u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6; +x_2 = lean_unsigned_to_nat(20099u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -79237,7 +78774,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -79247,13 +78784,13 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -79269,7 +78806,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8; +x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8; x_11 = l_Lean_registerTraceClass(x_10, x_3, x_4, x_9); if (lean_obj_tag(x_11) == 0) { @@ -79277,7 +78814,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); -x_13 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9; +x_13 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9; x_14 = l_Lean_registerTraceClass(x_13, x_3, x_4, x_12); return x_14; } @@ -79351,7 +78888,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; @@ -79362,7 +78899,7 @@ lean_ctor_set(x_6, 1, x_4); return x_6; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1() { _start: { lean_object* x_1; @@ -79370,17 +78907,17 @@ x_1 = lean_mk_string_unchecked("incremental", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3() { _start: { lean_object* x_1; @@ -79388,18 +78925,18 @@ x_1 = lean_mk_string_unchecked("incrementalAttr", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5() { _start: { lean_object* x_1; @@ -79407,39 +78944,39 @@ x_1 = lean_mk_string_unchecked("Marks an elaborator (tactic or command, currentl return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1___boxed), 4, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6; -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6; +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4; x_6 = 0; x_7 = l_Lean_registerTagAttribute(x_2, x_3, x_4, x_5, x_6, x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20184_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20279_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -79507,7 +79044,7 @@ return x_13; } } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -79515,40 +79052,40 @@ x_1 = lean_mk_string_unchecked("addBuiltinIncrementalElab", 25, 25); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_2 = l_Lean_Elab_Term_withoutTacticIncrementality___rarg___lambda__3___closed__2; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_inc(x_1); x_6 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_1); -x_7 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3; +x_7 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3; x_8 = l_Lean_Expr_app___override(x_7, x_6); x_9 = l_Lean_declareBuiltin(x_1, x_8, x_3, x_4, x_5); return x_9; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -79556,16 +79093,16 @@ x_1 = lean_mk_string_unchecked("invalid attribute 'builtin_incremental', must be return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -79583,7 +79120,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_dec(x_1); -x_11 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2; +x_11 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2; x_12 = l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(x_11, x_4, x_5, x_8); lean_dec(x_4); x_13 = !lean_is_exclusive(x_12); @@ -79609,7 +79146,7 @@ else { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1(x_1, x_17, x_4, x_5, x_8); +x_18 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1(x_1, x_17, x_4, x_5, x_8); lean_dec(x_4); return x_18; } @@ -79640,7 +79177,7 @@ return x_22; } } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1() { _start: { lean_object* x_1; @@ -79648,35 +79185,35 @@ x_1 = lean_mk_string_unchecked("attribute cannot be erased", 26, 26); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2; +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2; x_6 = l_Lean_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(x_5, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6; -x_2 = lean_unsigned_to_nat(20237u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6; +x_2 = lean_unsigned_to_nat(20332u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2() { _start: { lean_object* x_1; @@ -79684,17 +79221,17 @@ x_1 = lean_mk_string_unchecked("builtin_incremental", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4() { _start: { lean_object* x_1; @@ -79702,7 +79239,7 @@ x_1 = lean_mk_string_unchecked("(builtin) ", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -79715,29 +79252,29 @@ lean_dec(x_2); x_4 = lean_ctor_get(x_3, 2); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4; +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4; x_6 = lean_string_append(x_5, x_4); lean_dec(x_4); return x_6; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5; x_2 = l_Lean_Elab_Term_extraMsgToMsg___closed__3; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -79747,29 +79284,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -79777,42 +79314,42 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { uint8_t x_7; lean_object* x_8; x_7 = lean_unbox(x_3); lean_dec(x_3); -x_8 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -79989,23 +79526,23 @@ x_7 = l_Lean_Elab_isIncrementalElab___rarg___lambda__4(x_1, x_2, x_3, x_4, x_6); return x_7; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6; -x_2 = lean_unsigned_to_nat(20490u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6; +x_2 = lean_unsigned_to_nat(20585u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -80616,12 +80153,12 @@ l_Lean_Elab_Term_mkSaveInfoAnnotation___closed__2 = _init_l_Lean_Elab_Term_mkSav lean_mark_persistent(l_Lean_Elab_Term_mkSaveInfoAnnotation___closed__2); l_Lean_Elab_Term_addTermInfo___closed__1 = _init_l_Lean_Elab_Term_addTermInfo___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_addTermInfo___closed__1); -l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1 = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__1); -l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2 = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660____closed__2); -l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_ = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_(); -lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11660_); +l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1 = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__1); +l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2 = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721____closed__2); +l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_ = _init_l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_(); +lean_mark_persistent(l_Lean_Elab_Term_instImpl____x40_Lean_Elab_Term___hyg_11721_); l_Lean_Elab_Term_instTypeNameBodyInfo = _init_l_Lean_Elab_Term_instTypeNameBodyInfo(); lean_mark_persistent(l_Lean_Elab_Term_instTypeNameBodyInfo); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1(); @@ -80738,10 +80275,10 @@ l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__2 lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__2); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__3 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__3(); lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__3); -l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1); -l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2); +l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__1); +l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__6___closed__2); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__2 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__2(); @@ -80771,37 +80308,37 @@ l_Lean_Elab_Term_mkAuxName___closed__1 = _init_l_Lean_Elab_Term_mkAuxName___clos lean_mark_persistent(l_Lean_Elab_Term_mkAuxName___closed__1); l_Lean_Elab_Term_mkAuxName___closed__2 = _init_l_Lean_Elab_Term_mkAuxName___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_mkAuxName___closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17700_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_17795_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__1 = _init_l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__1(); @@ -80960,95 +80497,95 @@ l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3 = _init_l_Lean_Elab_Term_e lean_mark_persistent(l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3); l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4 = _init_l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4(); lean_mark_persistent(l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004____closed__9); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20004_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099____closed__9); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20099_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158____closed__6); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20158_(lean_io_mk_world()); +}l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253____closed__6); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20253_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_incrementalAttr = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_incrementalAttr); lean_dec_ref(res); -}if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20184_(lean_io_mk_world()); +}if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20279_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_builtinIncrementalElabs = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_builtinIncrementalElabs); lean_dec_ref(res); }l_Lean_Elab_addBuiltinIncrementalElab___closed__1 = _init_l_Lean_Elab_addBuiltinIncrementalElab___closed__1(); lean_mark_persistent(l_Lean_Elab_addBuiltinIncrementalElab___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__1___closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__2___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____lambda__3___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__9); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237____closed__10); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20237_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__1___closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__2___closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____lambda__3___closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__9); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332____closed__10); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20332_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Elab_isIncrementalElab___rarg___lambda__3___closed__1 = _init_l_Lean_Elab_isIncrementalElab___rarg___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_isIncrementalElab___rarg___lambda__3___closed__1); l_Lean_Elab_isIncrementalElab___rarg___closed__1 = _init_l_Lean_Elab_isIncrementalElab___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_isIncrementalElab___rarg___closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490____closed__1); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20490_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585____closed__1); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_20585_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Language/Lean.c b/stage0/stdlib/Lean/Language/Lean.c index 3f64a28dcaba..400cfa10876a 100644 --- a/stage0/stdlib/Lean/Language/Lean.c +++ b/stage0/stdlib/Lean/Language/Lean.c @@ -8248,7 +8248,7 @@ x_7 = l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1_ x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_5); -x_9 = lean_alloc_ctor(2, 1, 0); +x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_8); x_10 = l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4; x_11 = lean_alloc_ctor(1, 2, 0); @@ -8276,7 +8276,7 @@ x_15 = l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1 x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_13); -x_17 = lean_alloc_ctor(2, 1, 0); +x_17 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_17, 0, x_16); x_18 = l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__4; x_19 = lean_alloc_ctor(1, 2, 0); @@ -8529,7 +8529,7 @@ lean_inc(x_5); x_34 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_34, 0, x_33); lean_ctor_set(x_34, 1, x_5); -x_35 = lean_alloc_ctor(2, 1, 0); +x_35 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_35, 0, x_34); x_36 = lean_unsigned_to_nat(1u); x_37 = l_Lean_Syntax_getArg(x_5, x_36); @@ -8847,7 +8847,7 @@ lean_inc(x_5); x_155 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_155, 0, x_154); lean_ctor_set(x_155, 1, x_5); -x_156 = lean_alloc_ctor(2, 1, 0); +x_156 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_156, 0, x_155); x_157 = lean_unsigned_to_nat(1u); x_158 = l_Lean_Syntax_getArg(x_5, x_157); diff --git a/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c b/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c index fd6b701223a8..ac94d4f34bc2 100644 --- a/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c +++ b/stage0/stdlib/Lean/Linter/ConstructorAsVariable.c @@ -160,10 +160,10 @@ static lean_object* l_Lean_Linter_constructorNameAsVariable___elambda__1___close static lean_object* l_Lean_Linter_constructorNameAsVariable___elambda__1___closed__4; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1342_(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__9___closed__1; lean_object* l_Lean_Elab_Info_stx(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1340_(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_constructorNameAsVariable___elambda__1___spec__15___closed__3; @@ -4108,7 +4108,7 @@ lean_dec(x_1); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1340_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1342_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -4225,7 +4225,7 @@ l_Lean_Linter_constructorNameAsVariable___closed__3 = _init_l_Lean_Linter_constr lean_mark_persistent(l_Lean_Linter_constructorNameAsVariable___closed__3); l_Lean_Linter_constructorNameAsVariable = _init_l_Lean_Linter_constructorNameAsVariable(); lean_mark_persistent(l_Lean_Linter_constructorNameAsVariable); -if (builtin) {res = l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1340_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Linter_initFn____x40_Lean_Linter_ConstructorAsVariable___hyg_1342_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index d0f3edafa8a1..1ed6df82f1c3 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -5479,7 +5479,7 @@ return x_4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__12___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 10) +if (lean_obj_tag(x_2) == 11) { lean_object* x_4; uint8_t x_5; lean_object* x_6; x_4 = lean_ctor_get(x_2, 0); diff --git a/stage0/stdlib/Lean/Linter/UnusedVariables.c b/stage0/stdlib/Lean/Linter/UnusedVariables.c index a7907655514b..191e69cc5741 100644 --- a/stage0/stdlib/Lean/Linter/UnusedVariables.c +++ b/stage0/stdlib/Lean/Linter/UnusedVariables.c @@ -339,6 +339,7 @@ static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_L extern lean_object* l_Lean_Elab_Term_instTypeNameBodyInfo; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2594____lambda__1___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_426____closed__9; +LEAN_EXPORT lean_object* l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6401_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_987____lambda__1___closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1021____lambda__1___closed__1; uint64_t l_Lean_Expr_hash(lean_object*); @@ -641,7 +642,6 @@ LEAN_EXPORT lean_object* l_Lean_Linter_mkIgnoreFnImpl(lean_object*, lean_object* LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__13(lean_object*, size_t, size_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1316____lambda__4___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1021____lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6359_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_640____closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_visitAssignments_visitNode___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__10(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); @@ -13519,7 +13519,7 @@ lean_ctor_set(x_20, 1, x_7); return x_20; } } -case 8: +case 9: { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_21 = lean_ctor_get(x_3, 0); @@ -13554,467 +13554,485 @@ return x_31; } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_28); -if (x_32 == 0) +lean_object* x_32; +x_32 = lean_ctor_get(x_28, 0); +lean_inc(x_32); +lean_dec(x_28); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_33 = lean_ctor_get(x_28, 0); -x_34 = lean_ctor_get(x_22, 2); -lean_inc(x_34); +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_dec(x_22); -x_35 = l_Lean_instantiateMVarsCore(x_34, x_33); -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -lean_dec(x_35); -x_37 = lean_st_ref_take(x_6, x_7); -x_38 = lean_ctor_get(x_37, 0); +lean_dec(x_2); +lean_dec(x_1); +x_33 = lean_box(0); +x_34 = lean_box(x_5); +x_35 = lean_apply_4(x_8, x_33, x_34, x_6, x_7); +return x_35; +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_32); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_37 = lean_ctor_get(x_32, 0); +x_38 = lean_ctor_get(x_22, 2); lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = !lean_is_exclusive(x_38); -if (x_40 == 0) +lean_dec(x_22); +x_39 = l_Lean_instantiateMVarsCore(x_38, x_37); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_st_ref_take(x_6, x_7); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = !lean_is_exclusive(x_42); +if (x_44 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; -x_41 = lean_ctor_get(x_38, 4); -x_42 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; -x_43 = lean_box(0); -x_44 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_42, x_43, x_36); -x_45 = lean_array_push(x_41, x_44); -lean_ctor_set(x_38, 4, x_45); -x_46 = lean_st_ref_set(x_6, x_38, x_39); -x_47 = lean_ctor_get(x_46, 1); -lean_inc(x_47); -lean_dec(x_46); -x_48 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); -x_49 = l_Lean_PersistentArray_toArray___rarg(x_4); -lean_ctor_set(x_28, 0, x_2); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_45 = lean_ctor_get(x_42, 4); +x_46 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; +x_47 = lean_box(0); +x_48 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_46, x_47, x_40); +x_49 = lean_array_push(x_45, x_48); +lean_ctor_set(x_42, 4, x_49); +x_50 = lean_st_ref_set(x_6, x_42, x_43); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); +x_53 = l_Lean_PersistentArray_toArray___rarg(x_4); +lean_ctor_set(x_32, 0, x_2); if (x_5 == 0) { -lean_object* x_50; -x_50 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_49, x_28, x_48, x_6, x_47); -lean_dec(x_49); -if (lean_obj_tag(x_50) == 0) +lean_object* x_54; +x_54 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_53, x_32, x_52, x_6, x_51); +lean_dec(x_53); +if (lean_obj_tag(x_54) == 0) { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +uint8_t x_55; +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_52; uint8_t x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_50, 0); -lean_dec(x_52); -x_53 = 0; -x_54 = lean_box(x_53); -lean_ctor_set(x_50, 0, x_54); -return x_50; +lean_object* x_56; uint8_t x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_54, 0); +lean_dec(x_56); +x_57 = 0; +x_58 = lean_box(x_57); +lean_ctor_set(x_54, 0, x_58); +return x_54; } else { -lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; -x_55 = lean_ctor_get(x_50, 1); -lean_inc(x_55); -lean_dec(x_50); -x_56 = 0; -x_57 = lean_box(x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_55); -return x_58; +lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; +x_59 = lean_ctor_get(x_54, 1); +lean_inc(x_59); +lean_dec(x_54); +x_60 = 0; +x_61 = lean_box(x_60); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +return x_62; } } else { -uint8_t x_59; -x_59 = !lean_is_exclusive(x_50); -if (x_59 == 0) +uint8_t x_63; +x_63 = !lean_is_exclusive(x_54); +if (x_63 == 0) { -return x_50; +return x_54; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_50, 0); -x_61 = lean_ctor_get(x_50, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_50); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -return x_62; +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_54, 0); +x_65 = lean_ctor_get(x_54, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_54); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } } else { -uint8_t x_63; lean_object* x_64; -x_63 = 1; -x_64 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_49, x_28, x_63, x_6, x_47); -lean_dec(x_49); -if (lean_obj_tag(x_64) == 0) +uint8_t x_67; lean_object* x_68; +x_67 = 1; +x_68 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_53, x_32, x_67, x_6, x_51); +lean_dec(x_53); +if (lean_obj_tag(x_68) == 0) { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_64); -if (x_65 == 0) +uint8_t x_69; +x_69 = !lean_is_exclusive(x_68); +if (x_69 == 0) { -lean_object* x_66; uint8_t x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_64, 0); -lean_dec(x_66); -x_67 = 0; -x_68 = lean_box(x_67); -lean_ctor_set(x_64, 0, x_68); -return x_64; +lean_object* x_70; uint8_t x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_68, 0); +lean_dec(x_70); +x_71 = 0; +x_72 = lean_box(x_71); +lean_ctor_set(x_68, 0, x_72); +return x_68; } else { -lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; -x_69 = lean_ctor_get(x_64, 1); -lean_inc(x_69); -lean_dec(x_64); -x_70 = 0; -x_71 = lean_box(x_70); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -return x_72; +lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_68, 1); +lean_inc(x_73); +lean_dec(x_68); +x_74 = 0; +x_75 = lean_box(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; } } else { -uint8_t x_73; -x_73 = !lean_is_exclusive(x_64); -if (x_73 == 0) +uint8_t x_77; +x_77 = !lean_is_exclusive(x_68); +if (x_77 == 0) { -return x_64; +return x_68; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_64, 0); -x_75 = lean_ctor_get(x_64, 1); -lean_inc(x_75); -lean_inc(x_74); -lean_dec(x_64); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -return x_76; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_68, 0); +x_79 = lean_ctor_get(x_68, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_68); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } } } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; -x_77 = lean_ctor_get(x_38, 0); -x_78 = lean_ctor_get(x_38, 1); -x_79 = lean_ctor_get(x_38, 2); -x_80 = lean_ctor_get(x_38, 3); -x_81 = lean_ctor_get(x_38, 4); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; +x_81 = lean_ctor_get(x_42, 0); +x_82 = lean_ctor_get(x_42, 1); +x_83 = lean_ctor_get(x_42, 2); +x_84 = lean_ctor_get(x_42, 3); +x_85 = lean_ctor_get(x_42, 4); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_38); -x_82 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; -x_83 = lean_box(0); -x_84 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_82, x_83, x_36); -x_85 = lean_array_push(x_81, x_84); -x_86 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_86, 0, x_77); -lean_ctor_set(x_86, 1, x_78); -lean_ctor_set(x_86, 2, x_79); -lean_ctor_set(x_86, 3, x_80); -lean_ctor_set(x_86, 4, x_85); -x_87 = lean_st_ref_set(x_6, x_86, x_39); -x_88 = lean_ctor_get(x_87, 1); -lean_inc(x_88); -lean_dec(x_87); -x_89 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); -x_90 = l_Lean_PersistentArray_toArray___rarg(x_4); -lean_ctor_set(x_28, 0, x_2); +lean_dec(x_42); +x_86 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; +x_87 = lean_box(0); +x_88 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_86, x_87, x_40); +x_89 = lean_array_push(x_85, x_88); +x_90 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_90, 0, x_81); +lean_ctor_set(x_90, 1, x_82); +lean_ctor_set(x_90, 2, x_83); +lean_ctor_set(x_90, 3, x_84); +lean_ctor_set(x_90, 4, x_89); +x_91 = lean_st_ref_set(x_6, x_90, x_43); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +lean_dec(x_91); +x_93 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); +x_94 = l_Lean_PersistentArray_toArray___rarg(x_4); +lean_ctor_set(x_32, 0, x_2); if (x_5 == 0) { -lean_object* x_91; -x_91 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_90, x_28, x_89, x_6, x_88); -lean_dec(x_90); -if (lean_obj_tag(x_91) == 0) +lean_object* x_95; +x_95 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_94, x_32, x_93, x_6, x_92); +lean_dec(x_94); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_92; lean_object* x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; -x_92 = lean_ctor_get(x_91, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_93 = x_91; +lean_object* x_96; lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; } else { - lean_dec_ref(x_91); - x_93 = lean_box(0); + lean_dec_ref(x_95); + x_97 = lean_box(0); } -x_94 = 0; -x_95 = lean_box(x_94); -if (lean_is_scalar(x_93)) { - x_96 = lean_alloc_ctor(0, 2, 0); +x_98 = 0; +x_99 = lean_box(x_98); +if (lean_is_scalar(x_97)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_96 = x_93; + x_100 = x_97; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_92); -return x_96; +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_96); +return x_100; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_97 = lean_ctor_get(x_91, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_91, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_99 = x_91; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_101 = lean_ctor_get(x_95, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_95, 1); +lean_inc(x_102); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_103 = x_95; } else { - lean_dec_ref(x_91); - x_99 = lean_box(0); + lean_dec_ref(x_95); + x_103 = lean_box(0); } -if (lean_is_scalar(x_99)) { - x_100 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_103)) { + x_104 = lean_alloc_ctor(1, 2, 0); } else { - x_100 = x_99; + x_104 = x_103; } -lean_ctor_set(x_100, 0, x_97); -lean_ctor_set(x_100, 1, x_98); -return x_100; +lean_ctor_set(x_104, 0, x_101); +lean_ctor_set(x_104, 1, x_102); +return x_104; } } else { -uint8_t x_101; lean_object* x_102; -x_101 = 1; -x_102 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_90, x_28, x_101, x_6, x_88); -lean_dec(x_90); -if (lean_obj_tag(x_102) == 0) +uint8_t x_105; lean_object* x_106; +x_105 = 1; +x_106 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_94, x_32, x_105, x_6, x_92); +lean_dec(x_94); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; -x_103 = lean_ctor_get(x_102, 1); -lean_inc(x_103); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_104 = x_102; +lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_108 = x_106; } else { - lean_dec_ref(x_102); - x_104 = lean_box(0); + lean_dec_ref(x_106); + x_108 = lean_box(0); } -x_105 = 0; -x_106 = lean_box(x_105); -if (lean_is_scalar(x_104)) { - x_107 = lean_alloc_ctor(0, 2, 0); +x_109 = 0; +x_110 = lean_box(x_109); +if (lean_is_scalar(x_108)) { + x_111 = lean_alloc_ctor(0, 2, 0); } else { - x_107 = x_104; + x_111 = x_108; } -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_103); -return x_107; +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_107); +return x_111; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_108 = lean_ctor_get(x_102, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_102, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_110 = x_102; +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_106, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_106, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_114 = x_106; } else { - lean_dec_ref(x_102); - x_110 = lean_box(0); + lean_dec_ref(x_106); + x_114 = lean_box(0); } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); } else { - x_111 = x_110; + x_115 = x_114; } -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; } } } } else { -lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; -x_112 = lean_ctor_get(x_28, 0); -lean_inc(x_112); -lean_dec(x_28); -x_113 = lean_ctor_get(x_22, 2); -lean_inc(x_113); -lean_dec(x_22); -x_114 = l_Lean_instantiateMVarsCore(x_113, x_112); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -lean_dec(x_114); -x_116 = lean_st_ref_take(x_6, x_7); -x_117 = lean_ctor_get(x_116, 0); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; +x_116 = lean_ctor_get(x_32, 0); +lean_inc(x_116); +lean_dec(x_32); +x_117 = lean_ctor_get(x_22, 2); lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_ctor_get(x_117, 0); +lean_dec(x_22); +x_118 = l_Lean_instantiateMVarsCore(x_117, x_116); +x_119 = lean_ctor_get(x_118, 0); lean_inc(x_119); -x_120 = lean_ctor_get(x_117, 1); -lean_inc(x_120); -x_121 = lean_ctor_get(x_117, 2); +lean_dec(x_118); +x_120 = lean_st_ref_take(x_6, x_7); +x_121 = lean_ctor_get(x_120, 0); lean_inc(x_121); -x_122 = lean_ctor_get(x_117, 3); +x_122 = lean_ctor_get(x_120, 1); lean_inc(x_122); -x_123 = lean_ctor_get(x_117, 4); +lean_dec(x_120); +x_123 = lean_ctor_get(x_121, 0); lean_inc(x_123); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - lean_ctor_release(x_117, 2); - lean_ctor_release(x_117, 3); - lean_ctor_release(x_117, 4); - x_124 = x_117; +x_124 = lean_ctor_get(x_121, 1); +lean_inc(x_124); +x_125 = lean_ctor_get(x_121, 2); +lean_inc(x_125); +x_126 = lean_ctor_get(x_121, 3); +lean_inc(x_126); +x_127 = lean_ctor_get(x_121, 4); +lean_inc(x_127); +if (lean_is_exclusive(x_121)) { + lean_ctor_release(x_121, 0); + lean_ctor_release(x_121, 1); + lean_ctor_release(x_121, 2); + lean_ctor_release(x_121, 3); + lean_ctor_release(x_121, 4); + x_128 = x_121; } else { - lean_dec_ref(x_117); - x_124 = lean_box(0); + lean_dec_ref(x_121); + x_128 = lean_box(0); } -x_125 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; -x_126 = lean_box(0); -x_127 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_125, x_126, x_115); -x_128 = lean_array_push(x_123, x_127); -if (lean_is_scalar(x_124)) { - x_129 = lean_alloc_ctor(0, 5, 0); +x_129 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__9___closed__4; +x_130 = lean_box(0); +x_131 = l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(x_129, x_130, x_119); +x_132 = lean_array_push(x_127, x_131); +if (lean_is_scalar(x_128)) { + x_133 = lean_alloc_ctor(0, 5, 0); } else { - x_129 = x_124; -} -lean_ctor_set(x_129, 0, x_119); -lean_ctor_set(x_129, 1, x_120); -lean_ctor_set(x_129, 2, x_121); -lean_ctor_set(x_129, 3, x_122); -lean_ctor_set(x_129, 4, x_128); -x_130 = lean_st_ref_set(x_6, x_129, x_118); -x_131 = lean_ctor_get(x_130, 1); -lean_inc(x_131); -lean_dec(x_130); -x_132 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); -x_133 = l_Lean_PersistentArray_toArray___rarg(x_4); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_2); + x_133 = x_128; +} +lean_ctor_set(x_133, 0, x_123); +lean_ctor_set(x_133, 1, x_124); +lean_ctor_set(x_133, 2, x_125); +lean_ctor_set(x_133, 3, x_126); +lean_ctor_set(x_133, 4, x_132); +x_134 = lean_st_ref_set(x_6, x_133, x_122); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +lean_dec(x_134); +x_136 = l_Lean_PersistentArray_anyM___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__11(x_4); +x_137 = l_Lean_PersistentArray_toArray___rarg(x_4); +x_138 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_138, 0, x_2); if (x_5 == 0) { -lean_object* x_135; -x_135 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_133, x_134, x_132, x_6, x_131); -lean_dec(x_133); -if (lean_obj_tag(x_135) == 0) +lean_object* x_139; +x_139 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_137, x_138, x_136, x_6, x_135); +lean_dec(x_137); +if (lean_obj_tag(x_139) == 0) { -lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_139; lean_object* x_140; -x_136 = lean_ctor_get(x_135, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_137 = x_135; +lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_141 = x_139; } else { - lean_dec_ref(x_135); - x_137 = lean_box(0); + lean_dec_ref(x_139); + x_141 = lean_box(0); } -x_138 = 0; -x_139 = lean_box(x_138); -if (lean_is_scalar(x_137)) { - x_140 = lean_alloc_ctor(0, 2, 0); +x_142 = 0; +x_143 = lean_box(x_142); +if (lean_is_scalar(x_141)) { + x_144 = lean_alloc_ctor(0, 2, 0); } else { - x_140 = x_137; + x_144 = x_141; } -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_136); -return x_140; +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_140); +return x_144; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_141 = lean_ctor_get(x_135, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_135, 1); -lean_inc(x_142); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_143 = x_135; +lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_145 = lean_ctor_get(x_139, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_139, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_147 = x_139; } else { - lean_dec_ref(x_135); - x_143 = lean_box(0); + lean_dec_ref(x_139); + x_147 = lean_box(0); } -if (lean_is_scalar(x_143)) { - x_144 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_147)) { + x_148 = lean_alloc_ctor(1, 2, 0); } else { - x_144 = x_143; + x_148 = x_147; } -lean_ctor_set(x_144, 0, x_141); -lean_ctor_set(x_144, 1, x_142); -return x_144; +lean_ctor_set(x_148, 0, x_145); +lean_ctor_set(x_148, 1, x_146); +return x_148; } } else { -uint8_t x_145; lean_object* x_146; -x_145 = 1; -x_146 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_133, x_134, x_145, x_6, x_131); -lean_dec(x_133); -if (lean_obj_tag(x_146) == 0) +uint8_t x_149; lean_object* x_150; +x_149 = 1; +x_150 = l_Lean_Linter_UnusedVariables_collectReferences_go(x_1, x_137, x_138, x_149, x_6, x_135); +lean_dec(x_137); +if (lean_obj_tag(x_150) == 0) { -lean_object* x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; -x_147 = lean_ctor_get(x_146, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_148 = x_146; +lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; lean_object* x_155; +x_151 = lean_ctor_get(x_150, 1); +lean_inc(x_151); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_152 = x_150; } else { - lean_dec_ref(x_146); - x_148 = lean_box(0); + lean_dec_ref(x_150); + x_152 = lean_box(0); } -x_149 = 0; -x_150 = lean_box(x_149); -if (lean_is_scalar(x_148)) { - x_151 = lean_alloc_ctor(0, 2, 0); +x_153 = 0; +x_154 = lean_box(x_153); +if (lean_is_scalar(x_152)) { + x_155 = lean_alloc_ctor(0, 2, 0); } else { - x_151 = x_148; + x_155 = x_152; } -lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_147); -return x_151; +lean_ctor_set(x_155, 0, x_154); +lean_ctor_set(x_155, 1, x_151); +return x_155; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_152 = lean_ctor_get(x_146, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_146, 1); -lean_inc(x_153); -if (lean_is_exclusive(x_146)) { - lean_ctor_release(x_146, 0); - lean_ctor_release(x_146, 1); - x_154 = x_146; +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_156 = lean_ctor_get(x_150, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_150, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_158 = x_150; } else { - lean_dec_ref(x_146); - x_154 = lean_box(0); + lean_dec_ref(x_150); + x_158 = lean_box(0); } -if (lean_is_scalar(x_154)) { - x_155 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(1, 2, 0); } else { - x_155 = x_154; + x_159 = x_158; +} +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_157); +return x_159; } -lean_ctor_set(x_155, 0, x_152); -lean_ctor_set(x_155, 1, x_153); -return x_155; } } } @@ -14022,54 +14040,54 @@ return x_155; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_dec(x_22); lean_dec(x_21); lean_dec(x_2); lean_dec(x_1); -x_156 = lean_box(0); -x_157 = lean_box(x_5); -x_158 = lean_apply_4(x_8, x_156, x_157, x_6, x_7); -return x_158; +x_160 = lean_box(0); +x_161 = lean_box(x_5); +x_162 = lean_apply_4(x_8, x_160, x_161, x_6, x_7); +return x_162; } } -case 9: +case 10: { lean_dec(x_2); lean_dec(x_1); if (x_5 == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_3, 0); -lean_inc(x_159); +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_3, 0); +lean_inc(x_163); lean_dec(x_3); -x_160 = lean_box(0); -x_161 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__8(x_159, x_8, x_160, x_5, x_6, x_7); -return x_161; +x_164 = lean_box(0); +x_165 = l_Array_forIn_x27Unsafe_loop___at_Lean_Linter_UnusedVariables_collectReferences_go___spec__24___lambda__8(x_163, x_8, x_164, x_5, x_6, x_7); +return x_165; } else { -uint8_t x_162; lean_object* x_163; lean_object* x_164; +uint8_t x_166; lean_object* x_167; lean_object* x_168; lean_dec(x_6); lean_dec(x_3); -x_162 = 1; -x_163 = lean_box(x_162); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_7); -return x_164; +x_166 = 1; +x_167 = lean_box(x_166); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_7); +return x_168; } } default: { -lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_165 = lean_box(0); -x_166 = lean_box(x_5); -x_167 = lean_apply_4(x_8, x_165, x_166, x_6, x_7); -return x_167; +x_169 = lean_box(0); +x_170 = lean_box(x_5); +x_171 = lean_apply_4(x_8, x_169, x_170, x_6, x_7); +return x_171; } } } @@ -15955,7 +15973,7 @@ return x_22; } else { -if (lean_obj_tag(x_3) == 3) +if (lean_obj_tag(x_3) == 4) { uint8_t x_23; x_23 = !lean_is_exclusive(x_3); @@ -16039,7 +16057,7 @@ lean_ctor_set(x_39, 1, x_8); return x_39; } } -case 3: +case 4: { uint8_t x_40; x_40 = !lean_is_exclusive(x_12); @@ -21434,7 +21452,7 @@ lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6359_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6401_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -22098,7 +22116,7 @@ l_Lean_Linter_UnusedVariables_unusedVariables___closed__3 = _init_l_Lean_Linter_ lean_mark_persistent(l_Lean_Linter_UnusedVariables_unusedVariables___closed__3); l_Lean_Linter_UnusedVariables_unusedVariables = _init_l_Lean_Linter_UnusedVariables_unusedVariables(); lean_mark_persistent(l_Lean_Linter_UnusedVariables_unusedVariables); -if (builtin) {res = l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6359_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Linter_UnusedVariables_initFn____x40_Lean_Linter_UnusedVariables___hyg_6401_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_MessageData_isUnusedVariableWarning___closed__1 = _init_l_Lean_MessageData_isUnusedVariableWarning___closed__1(); diff --git a/stage0/stdlib/Lean/Linter/Util.c b/stage0/stdlib/Lean/Linter/Util.c index 38871709e0b5..7bb1492a5980 100644 --- a/stage0/stdlib/Lean/Linter/Util.c +++ b/stage0/stdlib/Lean/Linter/Util.c @@ -525,7 +525,7 @@ return x_24; } else { -if (lean_obj_tag(x_4) == 3) +if (lean_obj_tag(x_4) == 4) { uint8_t x_25; x_25 = !lean_is_exclusive(x_4); @@ -629,7 +629,7 @@ x_51 = lean_apply_2(x_49, lean_box(0), x_50); return x_51; } } -case 3: +case 4: { uint8_t x_52; x_52 = !lean_is_exclusive(x_10); diff --git a/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c index 64d9cdfa1965..16e5ea875efb 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c @@ -20,27 +20,25 @@ lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_obj static lean_object* l_Lean_Meta_AC_abstractAtoms_go___closed__1; static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool___lambda__3___boxed(lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8; lean_object* l_Lean_mkNatLit(lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9; LEAN_EXPORT lean_object* l_Array_insertionSort_traverse___at_Lean_Meta_AC_toACExpr___spec__7(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof___spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__4; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_AC_toACExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Meta_AC_toACExpr___spec__10___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_AC_acRflTactic___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___lambda__1(lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14; static lean_object* l_Lean_Meta_AC_abstractAtoms_go___lambda__2___closed__3; static lean_object* l_Lean_Meta_AC_getInstance___closed__6; static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_AC_abstractAtoms(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__13; LEAN_EXPORT lean_object* l_Std_HashMap_get_x21___at_Lean_Meta_AC_toACExpr___spec__9___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13; lean_object* l_Lean_mkApp7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_HashMap_get_x21___at_Lean_Meta_AC_toACExpr___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_acNfTargetTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -49,7 +47,6 @@ uint8_t l_Lean_Exception_isInterrupt(lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13; static lean_object* l_Lean_Meta_AC_getInstance___closed__1; size_t lean_uint64_to_usize(uint64_t); lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); @@ -111,7 +108,7 @@ static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange__1___close lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_acNfHypTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9; lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_CollectFVars_visit___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Data_AC_norm___at_Lean_Meta_AC_buildNormProof___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); @@ -130,12 +127,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_instContextInformationProdPreContextArra size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Meta_AC_abstractAtoms_go___lambda__2___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___closed__2; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7; static lean_object* l_Lean_Meta_AC_buildNormProof___lambda__1___closed__1; lean_object* l_Lean_MVarId_getNondepPropHyps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1; lean_object* l_panic___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_abstractAtoms_go___lambda__2___closed__1; @@ -144,6 +139,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__4___boxed( static lean_object* l_Lean_Meta_AC_evalNf0___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_AC_acNfTargetTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7; static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__6; extern lean_object* l_Lean_Meta_Simp_neutralConfig; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange__1(lean_object*); @@ -151,10 +147,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__2___boxed( static lean_object* l_Lean_Meta_AC_getInstance___closed__2; static lean_object* l_Lean_Meta_AC_getInstance___closed__7; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_AC_toACExpr___spec__11(lean_object*, size_t, size_t, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___closed__1; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___closed__4; static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__5; LEAN_EXPORT uint8_t l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool___lambda__3(lean_object*); @@ -227,17 +225,14 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_AC_toACExpr___ static lean_object* l___regBuiltin_Lean_Meta_AC_evalNf0__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___closed__3; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4; static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange__1___closed__1; lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalizedRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Meta_AC_toACExpr___spec__5(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_buildNormProof___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3; LEAN_EXPORT lean_object* l_Lean_Data_AC_evalList___at_Lean_Meta_AC_buildNormProof___spec__6___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_toACExpr___closed__4; static lean_object* l_Lean_Meta_AC_preContext___closed__7; @@ -248,7 +243,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr; lean_object* l_Lean_Data_AC_mergeIdem(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_acRflTactic___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___closed__6; @@ -259,6 +254,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof_convertTarget___boxed(lea static lean_object* l_Lean_Meta_AC_toACExpr___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_AC_preContext___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_applySimpResultToLocalDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -271,16 +267,16 @@ static lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___closed_ LEAN_EXPORT lean_object* l_Lean_Meta_AC_toACExpr_toPreExpr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool___closed__2; lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_AC_abstractAtoms_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Meta_AC_buildNormProof___spec__7___closed__1; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5; static lean_object* l_Lean_Meta_AC_buildNormProof___lambda__1___closed__2; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalExact___spec__1___rarg(lean_object*); static lean_object* l_Lean_Meta_AC_preContext___closed__3; static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__1; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15; static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__3; uint64_t lean_uint64_xor(uint64_t, uint64_t); uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -295,7 +291,6 @@ static lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___lambda_ LEAN_EXPORT lean_object* l_Lean_Meta_AC_getInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_rewriteUnnormalized___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10; static lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___closed__2; LEAN_EXPORT lean_object* l_Array_insertionSort_swapLoop___at_Lean_Meta_AC_toACExpr___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_toACExpr_toPreExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -308,7 +303,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_acRflTactic___rarg___lambda__1(lean_obje LEAN_EXPORT lean_object* l_Lean_Meta_AC_toACExpr_toACExpr(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Meta_AC_toACExpr___spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___closed__11; -LEAN_EXPORT lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580_(lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool; static lean_object* l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_AC_instContextInformationProdPreContextArrayBool___lambda__2(lean_object*); @@ -316,6 +311,8 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Meta_AC_toACExpr___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_instInhabitedPreContext___closed__2; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10; +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11; static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instInhabitedPreContext; lean_object* lean_array_uget(lean_object*, size_t); @@ -325,6 +322,7 @@ static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__5; static lean_object* l_Lean_Meta_AC_buildNormProof___closed__4; static lean_object* l_Lean_Meta_AC_rewriteUnnormalized___closed__2; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___lambda__1___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic__1___closed__6; static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange__1___closed__5; @@ -347,6 +345,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_AC_acNfTargetTactic___lambda__1(lean_object static lean_object* l_Lean_Meta_AC_buildNormProof___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_buildNormProof___closed__6; +LEAN_EXPORT lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572_(lean_object*); lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Meta_AC_toACExpr___spec__10___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_AC_toACExpr_toPreExpr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -358,6 +357,7 @@ lean_object* l_Lean_Data_AC_sort(lean_object*); static lean_object* l_Std_DHashMap_Internal_AssocList_get_x21___at_Lean_Meta_AC_toACExpr___spec__10___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_AC_instEvalInformationPreContextACExpr___lambda__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3; static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__7; static lean_object* l_Lean_Meta_AC_getInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_AC_getInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -365,11 +365,11 @@ static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__2; static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic__1___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_Elab_Tactic_replaceMainGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5; static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__3; static lean_object* l_Lean_Meta_AC_instInhabitedPreContext___closed__1; static lean_object* l_Lean_Meta_AC_buildNormProof_convert___closed__7; static lean_object* l_Lean_Meta_AC_preContext___closed__4; -static lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_AC_post___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AC_getInstance___closed__5; @@ -10202,7 +10202,7 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -10212,27 +10212,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1; x_2 = l_Lean_Meta_AC_getInstance___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2; x_2 = l_Lean_Meta_AC_getInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4() { _start: { lean_object* x_1; @@ -10240,17 +10240,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3; -x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3; +x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6() { _start: { lean_object* x_1; @@ -10258,57 +10258,57 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5; -x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5; +x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7; x_2 = l_Array_mapMUnsafe_map___at_Lean_Meta_AC_buildNormProof_mkContext___spec__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8; x_2 = l_Lean_Meta_AC_getInstance___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9; x_2 = l___regBuiltin_Lean_Meta_AC_acRflTactic__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10; x_2 = l_Lean_Meta_AC_getInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12() { _start: { lean_object* x_1; @@ -10316,17 +10316,17 @@ x_1 = lean_mk_string_unchecked("Main", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11; -x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11; +x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14() { _start: { lean_object* x_1; @@ -10334,33 +10334,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13; -x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14; +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13; +x_2 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16() { +static lean_object* _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15; -x_2 = lean_unsigned_to_nat(6580u); +x_1 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15; +x_2 = lean_unsigned_to_nat(6572u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_AC_getInstance___closed__3; x_3 = 0; -x_4 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16; +x_4 = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -10655,39 +10655,39 @@ lean_mark_persistent(l___regBuiltin_Lean_Meta_AC_evalNf0__1___closed__3); if (builtin) {res = l___regBuiltin_Lean_Meta_AC_evalNf0__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__1); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__2); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__3); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__4); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__5); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__6); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__7); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__8); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__9); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__10); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__11); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__12); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__13); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__14); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__15); -l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16(); -lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580____closed__16); -if (builtin) {res = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6580_(lean_io_mk_world()); +}l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__1); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__2); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__3); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__4); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__5); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__6); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__7); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__8); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__9); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__10); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__11); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__12); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__13); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__14); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__15); +l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16 = _init_l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16(); +lean_mark_persistent(l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572____closed__16); +if (builtin) {res = l_Lean_Meta_AC_initFn____x40_Lean_Meta_Tactic_AC_Main___hyg_6572_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c index 60cebdf2c684..beed735b8c3d 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/FunInd.c +++ b/stage0/stdlib/Lean/Meta/Tactic/FunInd.c @@ -17,6 +17,7 @@ static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__9__ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Tactic_FunInd_deriveUnaryInduction___spec__3___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionCase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_abstractIndependentMVars___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at_Lean_Tactic_FunInd_buildInductionBody___spec__43___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__13___closed__5; @@ -67,9 +68,9 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildIn static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__13___closed__4; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_buildInductionCase___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkPProdFst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_foldAndCollect___spec__30(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); @@ -85,6 +86,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__18___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_foldAndCollect___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Tactic_FunInd_unpackMutualInduction___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -114,6 +116,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductio LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanupAfter_allHeqToEq___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deduplicateIHs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_foldAndCollect___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12; uint8_t l_Lean_Exception_isInterrupt(lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_tell(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +131,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_de LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__26___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10; lean_object* l_Lean_Elab_Structural_RecArgInfo_pickIndicesMajor(lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__10___closed__3; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); @@ -157,18 +161,17 @@ lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_buildInductionBody___spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_projectMutualInduct___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_foldAndCollect___spec__1___lambda__3___closed__1; -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at_Lean_Tactic_FunInd_buildInductionBody___spec__38(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__12___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Tactic_FunInd_buildInductionBody___spec__35___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_buildInductionCase___spec__7___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__4(lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Tactic_FunInd_buildInductionCase___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -176,7 +179,6 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda static lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_localMapM(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Tactic_FunInd_cleanupAfter_allHeqToEq___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_logAt___at_Lean_Tactic_FunInd_buildInductionBody___spec__42___closed__1; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -319,6 +321,7 @@ lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_ static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__6___closed__2; lean_object* l_Lean_InductiveVal_numTypeFormers(lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__9___closed__5; +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3; static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__14___closed__1; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__12___closed__1; @@ -326,6 +329,7 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__1___boxe LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_stripPProdProjs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__19___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Structural_instInhabitedRecArgInfo; @@ -391,6 +395,7 @@ LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_Tactic_FunInd_foldAndCollect___ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Tactic_FunInd_buildInductionCase___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Structural_eqnInfoExt; +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__29___closed__2; @@ -432,6 +437,7 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__2(lean_ob lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_withLetDecls_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_buildInductionBody___spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Tactic_FunInd_foldAndCollect___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Tactic_FunInd_foldAndCollect___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -475,9 +481,7 @@ static lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__3___closed static lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__8___closed__1; static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__13___closed__2; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static double l_Lean_withTraceNode___at_Lean_Tactic_FunInd_foldAndCollect___spec__30___lambda__2___closed__1; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1; lean_object* l_Lean_Meta_kabstract(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -493,15 +497,16 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda static lean_object* l_List_forIn_x27_loop___at_Lean_Tactic_FunInd_foldAndCollect___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Tactic_FunInd_cleanupAfter_allHeqToEq___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__26(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_branch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_tell___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__28(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Tactic_FunInd_abstractIndependentMVars___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___boxed__const__1; @@ -542,7 +547,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Tactic_FunInd_cl LEAN_EXPORT uint8_t l_Lean_logAt___at_Lean_Tactic_FunInd_buildInductionBody___spec__42___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__22___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lean_Meta_MatcherApp_forallAltTelescope_x27___at_Lean_Tactic_FunInd_buildInductionBody___spec__43___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -554,6 +558,7 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__27(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_exec___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__14___closed__2; +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withErasedFVars___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__23___closed__2; static lean_object* l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__2; @@ -576,11 +581,13 @@ static lean_object* l_Lean_logAt___at_Lean_Tactic_FunInd_buildInductionBody___sp lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_withLetDecls___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_getConstInfoDefn___at_Lean_Tactic_FunInd_deriveUnaryInduction___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Tactic_FunInd_foldAndCollect___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__6___closed__11; static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__10___closed__2; +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1; lean_object* l_Lean_Expr_appFnCleanup(lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__11___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_deriveUnaryInduction___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -589,7 +596,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lea LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__25___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Tactic_FunInd_foldAndCollect___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -629,9 +635,9 @@ extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_StateT_pure___at_Lean_Tactic_FunInd_buildInductionBody___spec__47(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6; static lean_object* l_Lean_Tactic_FunInd_removeLamda___rarg___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_withLetDecls_go(lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_exec(lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deduplicateIHs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -643,12 +649,12 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda___rarg___lambda__1(lea LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_foldAndCollect___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1; lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_Tactic_FunInd_buildInductionCase___spec__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__22(lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__3___closed__3; @@ -658,6 +664,7 @@ static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__9__ static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__16___closed__3; lean_object* l_Array_indexOfAux___at_Lean_Meta_getElimExprInfo___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__12___closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); @@ -679,7 +686,6 @@ static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__6___clo lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Tactic_FunInd_withLetDecls___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda___at_Lean_Tactic_FunInd_buildInductionBody___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763_(lean_object*); extern lean_object* l_Lean_warningAsError; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_buildInductionBody___spec__4___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__16___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__17___lambda__2___closed__1; @@ -693,10 +699,8 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_branch___rarg___lambda__1(lean_o LEAN_EXPORT lean_object* l_StateT_lift___at_Lean_Tactic_FunInd_buildInductionBody___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler_threshold; -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__9___closed__1; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__19___lambda__3___closed__4; static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__11___closed__1; @@ -713,11 +717,12 @@ static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___closed__1; uint8_t l_Lean_LocalDecl_isLet(lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__4___closed__3; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Tactic_FunInd_unpackMutualInduction___spec__1___lambda__2___closed__3; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__25___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Tactic_FunInd_buildInductionCase___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1; lean_object* l_Array_range(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_buildInductionBody___spec__4___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__19___lambda__1___boxed(lean_object*, lean_object*); @@ -726,6 +731,7 @@ LEAN_EXPORT lean_object* l_Nat_foldRevM_loop___at_Lean_Tactic_FunInd_foldAndColl static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Tactic_FunInd_foldAndCollect___spec__32___closed__2; static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__23___closed__1; lean_object* l_Lean_Expr_appFn_x21(lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_withUserNames___at_Lean_Tactic_FunInd_buildInductionBody___spec__38___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__16___lambda__4___closed__3; static lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__5___closed__1; @@ -751,7 +757,6 @@ static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at_Lean_Tactic_F LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_buildInductionBody___spec__31(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assert(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__18___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_localMapM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -766,7 +771,6 @@ static lean_object* l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__1 LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_isFunInductName___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3; lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Tactic_FunInd_mkLambdaFVarsMasked___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -789,15 +793,18 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Tactic_FunInd_bui lean_object* l_Lean_throwError___at_Lean_LMVarId_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__28___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Structural_IndGroupInfo_ofInductiveVal(lean_object*); static lean_object* l_panic___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__13___closed__1; +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__27___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Tactic_FunInd_buildInductionBody___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(size_t, size_t, lean_object*); @@ -821,6 +828,7 @@ lean_object* l_Lean_Meta_instantiateLambda___boxed(lean_object*, lean_object*, l static lean_object* l_panic___at_Lean_Tactic_FunInd_foldAndCollect___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9; lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__6___closed__4; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Tactic_FunInd_foldAndCollect___spec__14(lean_object*, lean_object*, size_t, size_t); @@ -832,6 +840,7 @@ lean_object* lean_name_append_index_after(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__39___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionCase___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Tactic_FunInd_buildInductionBody___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -845,7 +854,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_de static lean_object* l_panic___at_Lean_Tactic_FunInd_foldAndCollect___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__39___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__15___closed__3; @@ -905,17 +913,14 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductio LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_buildInductionBody___spec__20___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_mkLambdaFVarsMasked(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mapErrorImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12; static lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___closed__1; static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__13___closed__2; static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__13___closed__2; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_branch___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__6___closed__6; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Tactic_FunInd_buildInductionBody___spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909_(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_abstractIndependentMVars___lambda__3___closed__1; @@ -925,7 +930,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatcherApp_Transform_0__Lea static lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__11___closed__2; extern lean_object* l_Lean_brecOnSuffix; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Tactic_FunInd_foldAndCollect___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Tactic_FunInd_foldAndCollect___spec__3___closed__2; @@ -938,7 +942,6 @@ static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__12___closed__1 LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__1___closed__2; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Tactic_FunInd_foldAndCollect___spec__34(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__27___closed__1; @@ -973,7 +976,6 @@ static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__7___clos LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_abstractIndependentMVars___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__1___closed__6; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2; lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__8___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_maskArray___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -1019,7 +1021,6 @@ LEAN_EXPORT lean_object* l_StateT_pure___at_Lean_Tactic_FunInd_buildInductionBod LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__36(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_foldAndCollect___spec__1___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_buildInductionCase___spec__7___lambda__3___boxed(lean_object**); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_M_localM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1071,7 +1072,6 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_bu LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_withLetDecls_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe___at_Lean_Tactic_FunInd_M_localMapM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__9___closed__1; -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6; static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__7___boxed(lean_object**); @@ -1149,7 +1149,6 @@ lean_object* l_Lean_hasConst___at_Lean_Elab_Structural_getRecArgInfo___spec__2(l static lean_object* l_Lean_Expr_withAppAux___at_Lean_Tactic_FunInd_unpackMutualInduction___spec__1___closed__2; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveInductionStructural___lambda__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanupAfter_cleanupAfter_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__44___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1165,7 +1164,6 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_removeLamda(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Tactic_FunInd_foldAndCollect___spec__22(lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Tactic_FunInd_deriveUnaryInduction___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Tactic_FunInd_foldAndCollect___spec__30___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__8___closed__1; @@ -1186,6 +1184,8 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanPackedArgs___lambda__2___boxe LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Tactic_FunInd_foldAndCollect___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_Meta_addImplicitTargets_collect___spec__1(lean_object*, lean_object*); lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Structural_Positions_groupAndSort___spec__6(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2; +LEAN_EXPORT lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905_(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_assertIHs___spec__1___closed__1; static lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__9___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1250,11 +1250,9 @@ LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__2(l LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_foldAndCollect___lambda__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_beta(lean_object*, lean_object*); lean_object* l_Lean_mkApp5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_cleanupAfter_go(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Structural_Positions_mapMwith___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__12___closed__6; -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_maskArray___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_buildInductionBody___spec__33(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mdataExpr_x21(lean_object*); @@ -1303,14 +1301,16 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_de LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_deriveUnaryInduction___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_106_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__1___closed__1; -static lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2; +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5; static lean_object* l_Lean_Elab_Structural_Positions_groupAndSort___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__3___closed__6; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_isFunInductName___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MatcherApp_transform___at_Lean_Tactic_FunInd_buildInductionBody___spec__30___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionBody___lambda__4___boxed(lean_object**); +static lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1; static lean_object* l_Lean_Tactic_FunInd_unpackMutualInduction___lambda__6___closed__9; LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_buildInductionCase___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__16___at_Lean_Tactic_FunInd_deriveInductionStructural___spec__17___boxed(lean_object**); @@ -76813,7 +76813,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; lean_object* x_6; lean_object* x_7; @@ -76825,15 +76825,15 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2() { _start: { uint8_t x_1; uint8_t x_2; uint8_t x_3; uint8_t x_4; lean_object* x_5; @@ -76858,7 +76858,7 @@ lean_ctor_set_uint8(x_5, 12, x_2); return x_5; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3() { _start: { lean_object* x_1; @@ -76866,21 +76866,21 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Tactic_FunInd_foldAndCollect___spec__32___closed__3; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -76888,13 +76888,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; x_1 = lean_box(0); -x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2; -x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5; +x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2; +x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5; x_4 = l_Lean_Tactic_FunInd_M_run___rarg___closed__1; x_5 = lean_unsigned_to_nat(0u); x_6 = 0; @@ -76910,12 +76910,12 @@ lean_ctor_set_uint8(x_7, sizeof(void*)*6 + 1, x_6); return x_7; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; +x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; x_3 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_1); @@ -76929,22 +76929,22 @@ lean_ctor_set(x_3, 8, x_2); return x_3; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; x_2 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -76953,13 +76953,13 @@ lean_ctor_set(x_2, 3, x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8; -x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; -x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8; +x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; +x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9; x_4 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -76971,11 +76971,11 @@ lean_ctor_set(x_4, 6, x_3); return x_4; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4; +x_1 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4; x_2 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -76984,15 +76984,15 @@ lean_ctor_set(x_2, 3, x_1); return x_2; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = lean_box(0); -x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7; -x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10; +x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7; +x_3 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10; x_4 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Tactic_FunInd_foldAndCollect___spec__32___closed__3; -x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11; +x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11; x_6 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_3); @@ -77002,7 +77002,7 @@ lean_ctor_set(x_6, 4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; uint8_t x_6; @@ -77016,7 +77016,7 @@ x_8 = lean_ctor_get(x_5, 1); x_9 = lean_ctor_get(x_7, 0); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1; +x_10 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1; lean_inc(x_1); x_11 = l_Lean_Tactic_FunInd_isFunInductName(x_9, x_1); lean_dec(x_9); @@ -77040,14 +77040,14 @@ lean_free_object(x_5); x_15 = lean_ctor_get(x_1, 0); lean_inc(x_15); lean_dec(x_1); -x_16 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12; +x_16 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12; x_17 = lean_st_mk_ref(x_16, x_8); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6; +x_20 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6; lean_inc(x_18); x_21 = l_Lean_Tactic_FunInd_deriveInduction(x_15, x_20, x_18, x_2, x_3, x_19); if (lean_obj_tag(x_21) == 0) @@ -77131,7 +77131,7 @@ lean_dec(x_5); x_40 = lean_ctor_get(x_38, 0); lean_inc(x_40); lean_dec(x_38); -x_41 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1; +x_41 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1; lean_inc(x_1); x_42 = l_Lean_Tactic_FunInd_isFunInductName(x_40, x_1); lean_dec(x_40); @@ -77153,14 +77153,14 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean x_46 = lean_ctor_get(x_1, 0); lean_inc(x_46); lean_dec(x_1); -x_47 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12; +x_47 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12; x_48 = lean_st_mk_ref(x_47, x_39); x_49 = lean_ctor_get(x_48, 0); lean_inc(x_49); x_50 = lean_ctor_get(x_48, 1); lean_inc(x_50); lean_dec(x_48); -x_51 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6; +x_51 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6; lean_inc(x_49); x_52 = l_Lean_Tactic_FunInd_deriveInduction(x_46, x_51, x_49, x_2, x_3, x_50); if (lean_obj_tag(x_52) == 0) @@ -77235,7 +77235,7 @@ return x_66; } } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1() { _start: { lean_object* x_1; @@ -77243,19 +77243,19 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_isFunInductName___boxed), return x_1; } } -static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2() { +static lean_object* _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2), 4, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1; +x_2 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1; x_3 = l_Lean_registerReservedNamePredicate(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -77263,7 +77263,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2; +x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2; x_6 = l_Lean_registerReservedNameAction(x_5, x_4); return x_6; } @@ -77291,18 +77291,18 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_5; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1() { _start: { lean_object* x_1; @@ -77310,17 +77310,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1; +x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3() { _start: { lean_object* x_1; @@ -77328,17 +77328,17 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2; -x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2; +x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5() { _start: { lean_object* x_1; @@ -77346,47 +77346,47 @@ x_1 = lean_mk_string_unchecked("Lean", 4, 4); return x_1; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4; -x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4; +x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6; x_2 = l_Lean_Tactic_FunInd_foldAndCollect___lambda__27___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7; x_2 = l_Lean_logAt___at_Lean_Tactic_FunInd_buildInductionBody___spec__42___lambda__2___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8; x_2 = l_Lean_Tactic_FunInd_foldAndCollect___lambda__27___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10() { _start: { lean_object* x_1; @@ -77394,33 +77394,33 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9; -x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10; +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9; +x_2 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12() { +static lean_object* _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11; -x_2 = lean_unsigned_to_nat(16909u); +x_1 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11; +x_2 = lean_unsigned_to_nat(16905u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Tactic_FunInd_foldAndCollect___lambda__27___closed__3; x_3 = 0; -x_4 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12; +x_4 = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -78069,62 +78069,62 @@ l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__1 = _init_l_Lean_Tact lean_mark_persistent(l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__1); l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__2 = _init_l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__2(); lean_mark_persistent(l_Lean_Tactic_FunInd_isFunInductName___lambda__2___closed__2); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__1); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__2); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__3); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__4); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__5); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__6); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__7); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__8); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__9); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__10); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__11); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____lambda__2___closed__12); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__1); -l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2(); -lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763____closed__2); -if (builtin) {res = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16763_(lean_io_mk_world()); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__1); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__2); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__3); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__4); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__5); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__6); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__7); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__8); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__9); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__10); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__11); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____lambda__2___closed__12); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__1); +l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2 = _init_l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2(); +lean_mark_persistent(l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759____closed__2); +if (builtin) {res = l_Lean_Tactic_FunInd_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16759_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__1); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__2); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__3); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__4); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__5); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__6); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__7); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__8); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__9); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__10); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__11); -l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12(); -lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909____closed__12); -if (builtin) {res = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16909_(lean_io_mk_world()); +}l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__1); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__2); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__3); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__4); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__5); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__6); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__7); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__8); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__9); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__10); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__11); +l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12 = _init_l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12(); +lean_mark_persistent(l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905____closed__12); +if (builtin) {res = l_initFn____x40_Lean_Meta_Tactic_FunInd___hyg_16905_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c index f78f83f47fa8..a85639ab9b65 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c +++ b/stage0/stdlib/Lean/Meta/Tactic/LinearArith/Solver.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.LinearArith.Solver -// Imports: Init.Data.Ord Init.Data.Array.DecidableEq Lean.Data.Rat +// Imports: Init.Data.Ord Init.Data.Array.DecidableEq Std.Internal.Rat #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,314 +13,314 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115_(lean_object*, lean_object*); +lean_object* l_Std_Internal_Rat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_get_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Meta_Linear_Cnstr_getBound___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstrKind___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprVar___closed__1; -uint8_t l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__3(lean_object*, lean_object*); lean_object* l_Array_take___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstr; static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__2; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_resolve(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9; extern lean_object* l_Int_instInhabited; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprAssumptionId; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3025_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Linear_getBestBound_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqJustification(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instInhabitedJustification___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_currVar___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstrKind; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedState; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_currVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_Cnstr_isStrict(lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqPoly___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18; uint8_t lean_usize_dec_eq(size_t, size_t); extern uint8_t l_instInhabitedBool; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10; lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623____boxed(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instInhabitedCnstrKind; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023_(uint8_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22; static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_CnstrKind_ofNat(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6; +uint8_t l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(lean_object*, lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Rat_instDecidableLe(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instReprPoly___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqJustification; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Linear_Cnstr_isUnsat___spec__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstrKind(uint8_t, uint8_t); static lean_object* l_Lean_Meta_Linear_instOrdVar___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instBEqCnstr; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_getBestBound_x3f(lean_object*, lean_object*, uint8_t, uint8_t); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedJustification; -lean_object* l_Lean_Rat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481____boxed(lean_object*, lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Meta_Linear_instInhabitedState___closed__1; static lean_object* l_Lean_Meta_Linear_instBEqCnstrKind___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars___boxed(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff(lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220____boxed(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(lean_object*, lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5; static lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3025____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_77____boxed(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Linear_getBestBound_x3f___spec__1(lean_object*, uint8_t, uint8_t, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_take(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004____boxed(lean_object*, lean_object*); -lean_object* l_Lean_Rat_inv(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableLtVar(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instOrdVar; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7; static lean_object* l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12; lean_object* l_outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f(lean_object*, uint8_t, lean_object*, uint8_t); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; static lean_object* l_Lean_Meta_Linear_instReprCnstrKind___closed__1; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587_(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2; static lean_object* l_Lean_Meta_Linear_instReprAssumptionId___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqCnstr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestUpperBound_x3f___boxed(lean_object*); static lean_object* l_Lean_Meta_Linear_instReprJustification___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_scale(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_solve(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at_Lean_Meta_Linear_Cnstr_getBound___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqVar(lean_object*, lean_object*); +lean_object* l_Std_Internal_Rat_mul(lean_object*, lean_object*); lean_object* l_Int_repr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprJustification; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____boxed(lean_object*, lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssumptionId; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqJustification___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedVar; static lean_object* l_Lean_Meta_Linear_instBEqJustification___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstr(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__4; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqPoly(lean_object*, lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11; +uint8_t l_Std_Internal_Rat_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestUpperBound_x3f(lean_object*); -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__1(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedAssignment; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_size(lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_take___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; +lean_object* l_Std_Internal_Rat_ceil(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprPoly; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3026____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25____boxed(lean_object*, lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3; +uint8_t l_Std_Internal_Rat_instDecidableLe(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVarCoeff___boxed(lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getNumVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprCnstrKind; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_combine_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_scale___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVar(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3; lean_object* lean_int_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Linear_Poly_scale___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; lean_object* lean_string_length(lean_object*); lean_object* l_Int_instDecidableEq___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqAssumptionId(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instBEqCnstr___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestLowerBound_x3f(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_getBound(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_instInhabitedAssignment___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__3(lean_object*, lean_object*); lean_object* l_instDecidableEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instLTVar; +lean_object* l_Std_Internal_Rat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_getBestLowerBound_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Cnstr_isStrict___boxed(lean_object*); -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; +uint8_t l_Std_Internal_Rat_isInt(lean_object*); static lean_object* l_Lean_Meta_Linear_instReprCnstr___closed__1; -uint8_t l_Lean_Rat_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13; static lean_object* l_Lean_Meta_Linear_pickAssignment_x3f___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedCnstr; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2; lean_object* l_Std_Format_joinSep___at_Prod_repr___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2; -lean_object* l_Lean_Rat_mul(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_noConfusion___rarg___lambda__1(lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11; +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7; lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_getBestBound_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_int_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6; uint8_t l_Array_instDecidableEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqAssumptionId___boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Linear_Poly_scale___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Cnstr_isUnsat___closed__2; -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3; lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_resolve___closed__1; -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18; -lean_object* l_Lean_Rat_floor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_size(lean_object*); +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005____boxed(lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -lean_object* l_Lean_Rat_ceil(lean_object*); +lean_object* l_Std_Internal_Rat_inv(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_resolve___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instInhabitedPoly; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2; size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____boxed(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14; static lean_object* l_Lean_Meta_Linear_Poly_add___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Assignment_get_x3f___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_77_(lean_object*, lean_object*); lean_object* lean_int_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10; uint8_t lean_int_dec_eq(lean_object*, lean_object*); -uint8_t l_Lean_Rat_isInt(lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3; -lean_object* l_Lean_Rat_sub(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableLtVar___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022_(uint8_t, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_ofNat___boxed(lean_object*); lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Meta_Linear_Poly_eval_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_add_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11; -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480_(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23; -static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instReprVar; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_State_assignCurr(lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_instDecidableEqVar___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Linear_Poly_eval_x3f___closed__1; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Linear_Poly_getMaxVar___boxed(lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +lean_object* l_Std_Internal_Rat_floor(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3026_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Linear_CnstrKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17; +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3; +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__1(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481_(lean_object*, lean_object*); static lean_object* _init_l_Lean_Meta_Linear_instInhabitedVar() { _start: { @@ -329,7 +329,7 @@ x_1 = lean_unsigned_to_nat(0u); return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -359,11 +359,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -374,7 +374,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instOrdVar___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_25____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_ordVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_26____boxed), 2, 0); return x_1; } } @@ -386,7 +386,7 @@ x_1 = l_Lean_Meta_Linear_instOrdVar___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_77_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -394,11 +394,11 @@ x_3 = lean_nat_dec_eq(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_77____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_76_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_77_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -424,7 +424,7 @@ x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1() { _start: { lean_object* x_1; @@ -432,29 +432,29 @@ x_1 = lean_mk_string_unchecked("id", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4() { _start: { lean_object* x_1; @@ -462,29 +462,29 @@ x_1 = lean_mk_string_unchecked(" := ", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -493,7 +493,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8() { _start: { lean_object* x_1; @@ -501,35 +501,35 @@ x_1 = lean_mk_string_unchecked("{ ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12() { _start: { lean_object* x_1; @@ -537,24 +537,24 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_3 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); x_4 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7; x_6 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -562,19 +562,19 @@ x_7 = 0; x_8 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_8, 0, x_6); lean_ctor_set_uint8(x_8, sizeof(void*)*1, x_7); -x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6; +x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6; x_10 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_14 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; +x_15 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; x_16 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -584,11 +584,11 @@ lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_7); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -597,7 +597,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprVar___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____boxed), 2, 0); return x_1; } } @@ -736,7 +736,7 @@ x_1 = l_Lean_Meta_Linear_instInhabitedAssignment___closed__1; return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -745,7 +745,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2() { _start: { lean_object* x_1; @@ -753,21 +753,21 @@ x_1 = lean_mk_string_unchecked(",", 1, 1); return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -775,7 +775,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5() { _start: { lean_object* x_1; @@ -783,35 +783,35 @@ x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9() { _start: { lean_object* x_1; @@ -819,17 +819,17 @@ x_1 = lean_mk_string_unchecked(")", 1, 1); return x_1; } } -static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10() { +static lean_object* _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -839,11 +839,11 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_4 = lean_ctor_get(x_1, 0); x_5 = lean_ctor_get(x_1, 1); -x_6 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_6 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_7 = lean_int_dec_lt(x_4, x_6); x_8 = lean_box(0); x_9 = lean_unsigned_to_nat(0u); -x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(x_5, x_9); +x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167_(x_5, x_9); if (x_7 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; @@ -858,17 +858,17 @@ x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_10); lean_ctor_set(x_13, 1, x_1); x_14 = l_List_reverse___rarg(x_13); -x_15 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; +x_15 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; x_16 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_14, x_15); -x_17 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8; +x_17 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8; x_18 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10; +x_19 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10; x_20 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7; +x_21 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7; x_22 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -893,17 +893,17 @@ x_28 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_28, 0, x_10); lean_ctor_set(x_28, 1, x_1); x_29 = l_List_reverse___rarg(x_28); -x_30 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; +x_30 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; x_31 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_29, x_30); -x_32 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8; +x_32 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8; x_33 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_33, 0, x_32); lean_ctor_set(x_33, 1, x_31); -x_34 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10; +x_34 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10; x_35 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7; +x_36 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7; x_37 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_37, 0, x_36); lean_ctor_set(x_37, 1, x_35); @@ -922,11 +922,11 @@ x_41 = lean_ctor_get(x_1, 1); lean_inc(x_41); lean_inc(x_40); lean_dec(x_1); -x_42 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_42 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_43 = lean_int_dec_lt(x_40, x_42); x_44 = lean_box(0); x_45 = lean_unsigned_to_nat(0u); -x_46 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166_(x_41, x_45); +x_46 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167_(x_41, x_45); if (x_43 == 0) { lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; @@ -941,17 +941,17 @@ x_50 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_50, 0, x_46); lean_ctor_set(x_50, 1, x_49); x_51 = l_List_reverse___rarg(x_50); -x_52 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; +x_52 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; x_53 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_51, x_52); -x_54 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8; +x_54 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8; x_55 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); -x_56 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10; +x_56 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10; x_57 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); -x_58 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7; +x_58 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7; x_59 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_59, 0, x_58); lean_ctor_set(x_59, 1, x_57); @@ -976,17 +976,17 @@ x_66 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_66, 0, x_46); lean_ctor_set(x_66, 1, x_65); x_67 = l_List_reverse___rarg(x_66); -x_68 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; +x_68 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; x_69 = l_Std_Format_joinSep___at_Prod_repr___spec__1(x_67, x_68); -x_70 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8; +x_70 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8; x_71 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_71, 0, x_70); lean_ctor_set(x_71, 1, x_69); -x_72 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10; +x_72 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10; x_73 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7; +x_74 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7; x_75 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_75, 0, x_74); lean_ctor_set(x_75, 1, x_73); @@ -999,16 +999,16 @@ return x_77; } } } -LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_1, x_2); +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -1030,7 +1030,7 @@ lean_ctor_set_tag(x_3, 5); lean_ctor_set(x_3, 1, x_1); lean_ctor_set(x_3, 0, x_2); x_7 = lean_unsigned_to_nat(0u); -x_8 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_5, x_7); +x_8 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_5, x_7); x_9 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_9, 0, x_3); lean_ctor_set(x_9, 1, x_8); @@ -1051,7 +1051,7 @@ x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_2); lean_ctor_set(x_13, 1, x_1); x_14 = lean_unsigned_to_nat(0u); -x_15 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_11, x_14); +x_15 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_11, x_14); x_16 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_16, 0, x_13); lean_ctor_set(x_16, 1, x_15); @@ -1062,7 +1062,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -1085,7 +1085,7 @@ x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); lean_dec(x_1); x_6 = lean_unsigned_to_nat(0u); -x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_5, x_6); +x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_5, x_6); return x_7; } else @@ -1095,14 +1095,14 @@ x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); lean_dec(x_1); x_9 = lean_unsigned_to_nat(0u); -x_10 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_8, x_9); -x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__4(x_2, x_10, x_4); +x_10 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_8, x_9); +x_11 = l_List_foldl___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__4(x_2, x_10, x_4); return x_11; } } } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1() { _start: { lean_object* x_1; @@ -1110,41 +1110,41 @@ x_1 = lean_mk_string_unchecked("val", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -1153,7 +1153,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6() { _start: { lean_object* x_1; @@ -1161,35 +1161,35 @@ x_1 = lean_mk_string_unchecked("#[", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10() { _start: { lean_object* x_1; @@ -1197,17 +1197,17 @@ x_1 = lean_mk_string_unchecked("]", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12() { _start: { lean_object* x_1; @@ -1215,33 +1215,33 @@ x_1 = lean_mk_string_unchecked("#[]", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1249,59 +1249,59 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1309,7 +1309,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -1321,17 +1321,17 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; x_6 = lean_array_to_list(x_1); -x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4; -x_8 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__3(x_6, x_7); -x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9; +x_7 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4; +x_8 = l_Std_Format_joinSep___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__3(x_6, x_7); +x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9; x_10 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11; x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8; x_14 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -1339,7 +1339,7 @@ x_15 = 1; x_16 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_16, 0, x_14); lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_15); -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5; x_18 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -1347,19 +1347,19 @@ x_19 = 0; x_20 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_20, 0, x_18); lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_19); -x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4; +x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4; x_22 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); -x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; +x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; x_24 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); -x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; +x_27 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; x_28 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); @@ -1372,25 +1372,25 @@ else { lean_object* x_30; lean_dec(x_1); -x_30 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20; +x_30 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20; return x_30; } } } -LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2(x_1, x_2); +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -1399,7 +1399,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprPoly___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____boxed), 2, 0); return x_1; } } @@ -1411,7 +1411,7 @@ x_1 = l_Lean_Meta_Linear_instReprPoly___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1() { _start: { lean_object* x_1; @@ -1419,7 +1419,7 @@ x_1 = lean_alloc_closure((void*)(l_Int_instDecidableEq___boxed), 2, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2() { _start: { lean_object* x_1; @@ -1427,32 +1427,32 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Linear_instDecidableEqVar___boxed), return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2; x_3 = lean_alloc_closure((void*)(l_instDecidableEqProd___rarg), 4, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3; +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3; x_4 = l_Array_instDecidableEq___rarg(x_3, x_1, x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1463,7 +1463,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqPoly(lean_object* x_1, lea _start: { uint8_t x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(x_1, x_2); return x_3; } } @@ -1757,7 +1757,7 @@ lean_dec(x_28); x_39 = lean_int_add(x_23, x_27); lean_dec(x_27); lean_dec(x_23); -x_40 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_40 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_41 = lean_int_dec_eq(x_39, x_40); if (x_41 == 0) { @@ -1845,7 +1845,7 @@ lean_dec(x_52); x_65 = lean_int_add(x_23, x_51); lean_dec(x_51); lean_dec(x_23); -x_66 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_66 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_67 = lean_int_dec_eq(x_65, x_66); if (x_67 == 0) { @@ -2073,7 +2073,7 @@ lean_dec(x_41); x_57 = lean_int_add(x_55, x_56); lean_dec(x_56); lean_dec(x_55); -x_58 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_58 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_59 = lean_int_dec_eq(x_57, x_58); if (x_59 == 0) { @@ -2169,7 +2169,7 @@ lean_dec(x_69); x_87 = lean_int_add(x_85, x_86); lean_dec(x_86); lean_dec(x_85); -x_88 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_88 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_89 = lean_int_dec_eq(x_87, x_88); if (x_89 == 0) { @@ -2310,7 +2310,7 @@ lean_dec(x_111); x_130 = lean_int_add(x_128, x_129); lean_dec(x_129); lean_dec(x_128); -x_131 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_131 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_132 = lean_int_dec_eq(x_130, x_131); if (x_132 == 0) { @@ -2446,10 +2446,10 @@ x_20 = lean_array_fget(x_1, x_14); lean_dec(x_14); x_21 = lean_unsigned_to_nat(1u); lean_ctor_set(x_10, 1, x_21); -x_22 = l_Lean_Rat_mul(x_10, x_20); +x_22 = l_Std_Internal_Rat_mul(x_10, x_20); lean_dec(x_20); lean_dec(x_10); -x_23 = l_Lean_Rat_add(x_15, x_22); +x_23 = l_Std_Internal_Rat_add(x_15, x_22); lean_inc(x_4); lean_ctor_set(x_8, 1, x_23); lean_ctor_set(x_8, 0, x_4); @@ -2490,10 +2490,10 @@ x_34 = lean_array_fget(x_1, x_28); lean_dec(x_28); x_35 = lean_unsigned_to_nat(1u); lean_ctor_set(x_10, 1, x_35); -x_36 = l_Lean_Rat_mul(x_10, x_34); +x_36 = l_Std_Internal_Rat_mul(x_10, x_34); lean_dec(x_34); lean_dec(x_10); -x_37 = l_Lean_Rat_add(x_29, x_36); +x_37 = l_Std_Internal_Rat_add(x_29, x_36); lean_inc(x_4); x_38 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_38, 0, x_4); @@ -2552,10 +2552,10 @@ x_51 = lean_unsigned_to_nat(1u); x_52 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_52, 0, x_42); lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_Rat_mul(x_52, x_50); +x_53 = l_Std_Internal_Rat_mul(x_52, x_50); lean_dec(x_50); lean_dec(x_52); -x_54 = l_Lean_Rat_add(x_44, x_53); +x_54 = l_Std_Internal_Rat_add(x_44, x_53); lean_inc(x_4); if (lean_is_scalar(x_45)) { x_55 = lean_alloc_ctor(0, 2, 0); @@ -2587,7 +2587,7 @@ static lean_object* _init_l_Lean_Meta_Linear_Poly_eval_x3f___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_1 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2683,7 +2683,7 @@ x_1 = lean_unsigned_to_nat(0u); return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3025_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3026_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -2691,11 +2691,11 @@ x_3 = lean_nat_dec_eq(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3025____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3026____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3025_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3026_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2721,14 +2721,14 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; x_3 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); x_4 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7; x_6 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -2736,19 +2736,19 @@ x_7 = 0; x_8 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_8, 0, x_6); lean_ctor_set_uint8(x_8, sizeof(void*)*1, x_7); -x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6; +x_9 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6; x_10 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; x_12 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_14 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; +x_15 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; x_16 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -2758,11 +2758,11 @@ lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_7); return x_17; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2771,7 +2771,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprAssumptionId___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116____boxed), 2, 0); return x_1; } } @@ -2801,7 +2801,7 @@ x_1 = l_Lean_Meta_Linear_instInhabitedJustification___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2827,7 +2827,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(x_4, x_8); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_4, x_8); if (x_13 == 0) { uint8_t x_14; @@ -2879,11 +2879,11 @@ return x_22; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2894,7 +2894,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqJustification(lean_object* _start: { uint8_t x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_1, x_2); return x_3; } } @@ -2909,7 +2909,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2935,7 +2935,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622_(x_4, x_8); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623_(x_4, x_8); if (x_13 == 0) { uint8_t x_14; @@ -2987,11 +2987,11 @@ return x_22; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3002,7 +3002,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqJustification___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623____boxed), 2, 0); return x_1; } } @@ -3014,7 +3014,7 @@ x_1 = l_Lean_Meta_Linear_instBEqJustification___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1() { _start: { lean_object* x_1; @@ -3022,21 +3022,21 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.Justification.combine", 38, 38) return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3044,7 +3044,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -3053,7 +3053,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -3062,7 +3062,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6() { _start: { lean_object* x_1; @@ -3070,21 +3070,21 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.Justification.assumption", 41, return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3092,7 +3092,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3109,22 +3109,22 @@ lean_inc(x_6); lean_dec(x_1); x_7 = lean_unsigned_to_nat(1024u); x_8 = lean_nat_dec_le(x_7, x_2); -x_9 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_9 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_10 = lean_int_dec_lt(x_3, x_9); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(x_4, x_7); +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(x_4, x_7); x_12 = lean_int_dec_lt(x_5, x_9); -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(x_6, x_7); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(x_6, x_7); if (x_8 == 0) { lean_object* x_48; -x_48 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; +x_48 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; x_14 = x_48; goto block_47; } else { lean_object* x_49; -x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; +x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; x_14 = x_49; goto block_47; } @@ -3155,7 +3155,7 @@ goto block_41; block_41: { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3; +x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -3233,15 +3233,15 @@ lean_inc(x_50); lean_dec(x_1); x_51 = lean_unsigned_to_nat(1024u); x_52 = lean_nat_dec_le(x_51, x_2); -x_53 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3115_(x_50, x_51); -x_54 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8; +x_53 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprAssumptionId____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3116_(x_50, x_51); +x_54 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8; x_55 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_53); if (x_52 == 0) { lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; -x_56 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; +x_56 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; x_57 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_57, 0, x_56); lean_ctor_set(x_57, 1, x_55); @@ -3255,7 +3255,7 @@ return x_60; else { lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; -x_61 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; +x_61 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; x_62 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_62, 1, x_55); @@ -3269,11 +3269,11 @@ return x_65; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -3282,7 +3282,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprJustification___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____boxed), 2, 0); return x_1; } } @@ -3473,7 +3473,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -3485,7 +3485,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -3493,7 +3493,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004_(x_3, x_4); +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -3502,7 +3502,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005____boxed), 2, 0); return x_1; } } @@ -3514,7 +3514,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstrKind___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1() { _start: { lean_object* x_1; @@ -3522,33 +3522,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.eq", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3556,23 +3556,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3580,7 +3580,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7() { _start: { lean_object* x_1; @@ -3588,33 +3588,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.div", 30, 30); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3622,23 +3622,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3646,7 +3646,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13() { _start: { lean_object* x_1; @@ -3654,33 +3654,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.lt", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3688,23 +3688,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3712,7 +3712,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19() { _start: { lean_object* x_1; @@ -3720,33 +3720,33 @@ x_1 = lean_mk_string_unchecked("Lean.Meta.Linear.CnstrKind.le", 29, 29); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3754,23 +3754,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3778,7 +3778,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -3790,14 +3790,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6; +x_7 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -3810,14 +3810,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10; +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12; +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -3830,14 +3830,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16; +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18; +x_19 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -3850,14 +3850,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22; +x_23 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24; +x_25 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -3865,13 +3865,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022_(x_3, x_2); +x_4 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -3880,7 +3880,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____boxed), 2, 0); return x_1; } } @@ -3898,7 +3898,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instInhabitedCnstr___closed__1() { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 0; x_2 = l_Lean_Meta_Linear_instInhabitedAssignment___closed__1; -x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_3 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_4 = l_Lean_Meta_Linear_instInhabitedJustification___closed__1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_2); @@ -3916,7 +3916,7 @@ x_1 = l_Lean_Meta_Linear_instInhabitedCnstr___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -3938,7 +3938,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(x_4, x_8); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(x_4, x_8); if (x_13 == 0) { uint8_t x_14; @@ -3958,18 +3958,18 @@ return x_16; else { uint8_t x_17; -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3207_(x_6, x_10); +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3208_(x_6, x_10); return x_17; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -3980,7 +3980,7 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Linear_instDecidableEqCnstr(lean_object* x_1, le _start: { uint8_t x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4219_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4220_(x_1, x_2); return x_3; } } @@ -3995,7 +3995,7 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -4007,7 +4007,7 @@ x_7 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); x_10 = lean_ctor_get(x_2, 2); -x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4004_(x_3, x_7); +x_11 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4005_(x_3, x_7); if (x_11 == 0) { uint8_t x_12; @@ -4017,7 +4017,7 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430_(x_4, x_8); +x_13 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431_(x_4, x_8); if (x_13 == 0) { uint8_t x_14; @@ -4037,18 +4037,18 @@ return x_16; else { uint8_t x_17; -x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3622_(x_6, x_10); +x_17 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3623_(x_6, x_10); return x_17; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -4059,7 +4059,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instBEqCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4480____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_beqCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4481____boxed), 2, 0); return x_1; } } @@ -4071,7 +4071,7 @@ x_1 = l_Lean_Meta_Linear_instBEqCnstr___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1() { _start: { lean_object* x_1; @@ -4079,41 +4079,41 @@ x_1 = lean_mk_string_unchecked("kind", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3; -x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3; +x_2 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -4122,7 +4122,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6() { _start: { lean_object* x_1; @@ -4130,17 +4130,17 @@ x_1 = lean_mk_string_unchecked("lhs", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8() { _start: { lean_object* x_1; @@ -4148,17 +4148,17 @@ x_1 = lean_mk_string_unchecked("rhs", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10() { _start: { lean_object* x_1; @@ -4166,24 +4166,24 @@ x_1 = lean_mk_string_unchecked("jst", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11() { +static lean_object* _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022_(x_3, x_4); -x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5; +x_5 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023_(x_3, x_4); +x_6 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -4191,11 +4191,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4; +x_10 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3; +x_12 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -4203,18 +4203,18 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7; +x_16 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5; +x_18 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5; x_19 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); x_20 = lean_ctor_get(x_1, 0); lean_inc(x_20); -x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392_(x_20, x_4); -x_22 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5; +x_21 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393_(x_20, x_4); +x_22 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5; x_23 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -4230,7 +4230,7 @@ lean_ctor_set(x_26, 1, x_12); x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_14); -x_28 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9; +x_28 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -4239,12 +4239,12 @@ lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_18); x_31 = lean_ctor_get(x_1, 1); lean_inc(x_31); -x_32 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_32 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_33 = lean_int_dec_lt(x_31, x_32); x_34 = lean_ctor_get(x_1, 2); lean_inc(x_34); lean_dec(x_1); -x_35 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751_(x_34, x_4); +x_35 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752_(x_34, x_4); x_36 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_36, 0, x_22); lean_ctor_set(x_36, 1, x_35); @@ -4273,7 +4273,7 @@ lean_ctor_set(x_43, 1, x_12); x_44 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_44, 0, x_43); lean_ctor_set(x_44, 1, x_14); -x_45 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11; +x_45 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11; x_46 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); @@ -4283,15 +4283,15 @@ lean_ctor_set(x_47, 1, x_18); x_48 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_48, 0, x_47); lean_ctor_set(x_48, 1, x_37); -x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; +x_49 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; x_50 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_51 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; +x_53 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; x_54 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -4323,7 +4323,7 @@ lean_ctor_set(x_62, 1, x_12); x_63 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_63, 0, x_62); lean_ctor_set(x_63, 1, x_14); -x_64 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11; +x_64 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11; x_65 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_65, 0, x_63); lean_ctor_set(x_65, 1, x_64); @@ -4333,15 +4333,15 @@ lean_ctor_set(x_66, 1, x_18); x_67 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_67, 0, x_66); lean_ctor_set(x_67, 1, x_37); -x_68 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11; +x_68 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11; x_69 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_69, 0, x_68); lean_ctor_set(x_69, 1, x_67); -x_70 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13; +x_70 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13; x_71 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_71, 0, x_69); lean_ctor_set(x_71, 1, x_70); -x_72 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10; +x_72 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10; x_73 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_73, 0, x_72); lean_ctor_set(x_73, 1, x_71); @@ -4352,11 +4352,11 @@ return x_74; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586_(x_1, x_2); +x_3 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4365,7 +4365,7 @@ static lean_object* _init_l_Lean_Meta_Linear_instReprCnstr___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____boxed), 2, 0); return x_1; } } @@ -4449,10 +4449,10 @@ x_17 = lean_array_fget(x_1, x_11); lean_dec(x_11); x_18 = lean_unsigned_to_nat(1u); lean_ctor_set(x_8, 1, x_18); -x_19 = l_Lean_Rat_mul(x_8, x_17); +x_19 = l_Std_Internal_Rat_mul(x_8, x_17); lean_dec(x_17); lean_dec(x_8); -x_20 = l_Lean_Rat_sub(x_5, x_19); +x_20 = l_Std_Internal_Rat_sub(x_5, x_19); x_21 = 1; x_22 = lean_usize_add(x_4, x_21); x_4 = x_22; @@ -4490,10 +4490,10 @@ x_32 = lean_unsigned_to_nat(1u); x_33 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_33, 0, x_24); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Rat_mul(x_33, x_31); +x_34 = l_Std_Internal_Rat_mul(x_33, x_31); lean_dec(x_31); lean_dec(x_33); -x_35 = l_Lean_Rat_sub(x_5, x_34); +x_35 = l_Std_Internal_Rat_sub(x_5, x_34); x_36 = 1; x_37 = lean_usize_add(x_4, x_36); x_4 = x_37; @@ -4543,8 +4543,8 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; x_19 = lean_ctor_get(x_17, 1); lean_dec(x_19); lean_ctor_set(x_17, 1, x_4); -x_20 = l_Lean_Rat_inv(x_17); -x_21 = l_Lean_Rat_mul(x_15, x_20); +x_20 = l_Std_Internal_Rat_inv(x_17); +x_21 = l_Std_Internal_Rat_mul(x_15, x_20); lean_dec(x_20); lean_dec(x_15); return x_21; @@ -4558,8 +4558,8 @@ lean_dec(x_17); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_4); -x_24 = l_Lean_Rat_inv(x_23); -x_25 = l_Lean_Rat_mul(x_15, x_24); +x_24 = l_Std_Internal_Rat_inv(x_23); +x_25 = l_Std_Internal_Rat_mul(x_15, x_24); lean_dec(x_24); lean_dec(x_15); return x_25; @@ -4629,7 +4629,7 @@ static lean_object* _init_l_Lean_Meta_Linear_Cnstr_isUnsat___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Meta_Linear_Cnstr_isUnsat___closed__1; x_2 = l_Lean_Meta_Linear_Cnstr_isUnsat___closed__2; -x_3 = lean_unsigned_to_nat(164u); +x_3 = lean_unsigned_to_nat(165u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Meta_Linear_Cnstr_isUnsat___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4666,7 +4666,7 @@ lean_inc(x_9); x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); -x_12 = l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(x_8, x_11); +x_12 = l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(x_8, x_11); lean_dec(x_11); lean_dec(x_8); if (x_12 == 0) @@ -4704,7 +4704,7 @@ lean_inc(x_20); x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Rat_lt(x_19, x_22); +x_23 = l_Std_Internal_Rat_lt(x_19, x_22); if (x_23 == 0) { uint8_t x_24; lean_object* x_25; @@ -4732,7 +4732,7 @@ lean_inc(x_29); x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_Rat_instDecidableLe(x_28, x_31); +x_32 = l_Std_Internal_Rat_instDecidableLe(x_28, x_31); if (x_32 == 0) { uint8_t x_33; lean_object* x_34; @@ -4783,7 +4783,7 @@ lean_object* x_13; uint8_t x_14; x_13 = lean_ctor_get(x_7, 0); lean_inc(x_13); lean_inc(x_10); -x_14 = l_Lean_Rat_lt(x_10, x_13); +x_14 = l_Std_Internal_Rat_lt(x_10, x_13); if (x_14 == 0) { lean_dec(x_10); @@ -4809,7 +4809,7 @@ lean_object* x_18; uint8_t x_19; x_18 = lean_ctor_get(x_7, 0); lean_inc(x_18); lean_inc(x_10); -x_19 = l_Lean_Rat_lt(x_18, x_10); +x_19 = l_Std_Internal_Rat_lt(x_18, x_10); if (x_19 == 0) { lean_dec(x_10); @@ -4835,7 +4835,7 @@ else if (x_2 == 0) { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_23 = l_Lean_Rat_floor(x_10); +x_23 = l_Std_Internal_Rat_floor(x_10); x_24 = lean_unsigned_to_nat(1u); x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_23); @@ -4843,7 +4843,7 @@ lean_ctor_set(x_25, 1, x_24); x_26 = lean_ctor_get(x_7, 0); lean_inc(x_26); lean_inc(x_25); -x_27 = l_Lean_Rat_lt(x_25, x_26); +x_27 = l_Std_Internal_Rat_lt(x_25, x_26); if (x_27 == 0) { lean_dec(x_25); @@ -4866,7 +4866,7 @@ goto _start; else { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_31 = l_Lean_Rat_ceil(x_10); +x_31 = l_Std_Internal_Rat_ceil(x_10); x_32 = lean_unsigned_to_nat(1u); x_33 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_33, 0, x_31); @@ -4874,7 +4874,7 @@ lean_ctor_set(x_33, 1, x_32); x_34 = lean_ctor_get(x_7, 0); lean_inc(x_34); lean_inc(x_33); -x_35 = l_Lean_Rat_lt(x_34, x_33); +x_35 = l_Std_Internal_Rat_lt(x_34, x_33); if (x_35 == 0) { lean_dec(x_33); @@ -4944,7 +4944,7 @@ else if (x_3 == 0) { lean_object* x_29; lean_object* x_30; -x_29 = l_Lean_Rat_floor(x_10); +x_29 = l_Std_Internal_Rat_floor(x_10); x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_12); @@ -4954,7 +4954,7 @@ goto block_28; else { lean_object* x_31; lean_object* x_32; -x_31 = l_Lean_Rat_ceil(x_10); +x_31 = l_Std_Internal_Rat_ceil(x_10); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_12); @@ -5307,7 +5307,7 @@ static lean_object* _init_l_Lean_Meta_Linear_pickAssignment_x3f___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5320,7 +5320,7 @@ static lean_object* _init_l_Lean_Meta_Linear_pickAssignment_x3f___closed__2() { { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__1; -x_2 = l_Lean_Rat_inv(x_1); +x_2 = l_Std_Internal_Rat_inv(x_1); return x_2; } } @@ -5328,7 +5328,7 @@ static lean_object* _init_l_Lean_Meta_Linear_pickAssignment_x3f___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5; +x_1 = l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5340,13 +5340,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Linear_pickAssignment_x3f(lean_object* x_1, _start: { uint8_t x_5; -x_5 = l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(x_1, x_3); +x_5 = l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(x_1, x_3); if (x_5 == 0) { uint8_t x_6; lean_inc(x_3); lean_inc(x_1); -x_6 = l_Lean_Rat_lt(x_1, x_3); +x_6 = l_Std_Internal_Rat_lt(x_1, x_3); if (x_6 == 0) { lean_object* x_7; @@ -5368,26 +5368,26 @@ return x_8; else { uint8_t x_9; -x_9 = l_Lean_Rat_isInt(x_1); +x_9 = l_Std_Internal_Rat_isInt(x_1); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_inc(x_1); -x_10 = l_Lean_Rat_ceil(x_1); +x_10 = l_Std_Internal_Rat_ceil(x_1); x_11 = lean_unsigned_to_nat(1u); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); lean_inc(x_3); lean_inc(x_12); -x_13 = l_Lean_Rat_lt(x_12, x_3); +x_13 = l_Std_Internal_Rat_lt(x_12, x_3); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_12); -x_14 = l_Lean_Rat_add(x_1, x_3); +x_14 = l_Std_Internal_Rat_add(x_1, x_3); x_15 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__2; -x_16 = l_Lean_Rat_mul(x_14, x_15); +x_16 = l_Std_Internal_Rat_mul(x_14, x_15); lean_dec(x_14); x_17 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -5408,17 +5408,17 @@ else lean_object* x_19; lean_object* x_20; uint8_t x_21; x_19 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; lean_inc(x_1); -x_20 = l_Lean_Rat_add(x_1, x_19); +x_20 = l_Std_Internal_Rat_add(x_1, x_19); lean_inc(x_3); lean_inc(x_20); -x_21 = l_Lean_Rat_lt(x_20, x_3); +x_21 = l_Std_Internal_Rat_lt(x_20, x_3); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_20); -x_22 = l_Lean_Rat_add(x_1, x_3); +x_22 = l_Std_Internal_Rat_add(x_1, x_3); x_23 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__2; -x_24 = l_Lean_Rat_mul(x_22, x_23); +x_24 = l_Std_Internal_Rat_mul(x_22, x_23); lean_dec(x_22); x_25 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_25, 0, x_24); @@ -5623,7 +5623,7 @@ x_39 = l_Array_take___rarg(x_38, x_34); x_40 = lean_ctor_get(x_33, 0); lean_inc(x_40); lean_dec(x_33); -x_41 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_41 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_42 = lean_int_dec_lt(x_40, x_41); lean_dec(x_40); if (x_42 == 0) @@ -5707,7 +5707,7 @@ x_65 = l_Array_take___rarg(x_64, x_34); x_66 = lean_ctor_get(x_33, 0); lean_inc(x_66); lean_dec(x_33); -x_67 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1; +x_67 = l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1; x_68 = lean_int_dec_lt(x_66, x_67); lean_dec(x_66); if (x_68 == 0) @@ -6067,11 +6067,11 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean x_30 = lean_ctor_get(x_25, 0); x_31 = lean_ctor_get(x_25, 1); lean_dec(x_31); -x_32 = l_Lean_Rat_floor(x_30); +x_32 = l_Std_Internal_Rat_floor(x_30); lean_ctor_set(x_25, 1, x_5); lean_ctor_set(x_25, 0, x_32); x_33 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; -x_34 = l_Lean_Rat_sub(x_25, x_33); +x_34 = l_Std_Internal_Rat_sub(x_25, x_33); x_35 = lean_array_push(x_10, x_34); x_36 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_36, 0, x_7); @@ -6088,12 +6088,12 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean x_38 = lean_ctor_get(x_25, 0); lean_inc(x_38); lean_dec(x_25); -x_39 = l_Lean_Rat_floor(x_38); +x_39 = l_Std_Internal_Rat_floor(x_38); x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_39); lean_ctor_set(x_40, 1, x_5); x_41 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; -x_42 = l_Lean_Rat_sub(x_40, x_41); +x_42 = l_Std_Internal_Rat_sub(x_40, x_41); x_43 = lean_array_push(x_10, x_42); x_44 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_44, 0, x_7); @@ -6116,7 +6116,7 @@ lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean x_47 = lean_ctor_get(x_25, 0); x_48 = lean_ctor_get(x_25, 1); lean_dec(x_48); -x_49 = l_Lean_Rat_floor(x_47); +x_49 = l_Std_Internal_Rat_floor(x_47); lean_ctor_set(x_25, 1, x_5); lean_ctor_set(x_25, 0, x_49); x_50 = lean_array_push(x_10, x_25); @@ -6135,7 +6135,7 @@ lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean x_53 = lean_ctor_get(x_25, 0); lean_inc(x_53); lean_dec(x_25); -x_54 = l_Lean_Rat_floor(x_53); +x_54 = l_Std_Internal_Rat_floor(x_53); x_55 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_5); @@ -6184,11 +6184,11 @@ lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean x_71 = lean_ctor_get(x_61, 0); x_72 = lean_ctor_get(x_61, 1); lean_dec(x_72); -x_73 = l_Lean_Rat_ceil(x_71); +x_73 = l_Std_Internal_Rat_ceil(x_71); lean_ctor_set(x_61, 1, x_5); lean_ctor_set(x_61, 0, x_73); x_74 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; -x_75 = l_Lean_Rat_add(x_61, x_74); +x_75 = l_Std_Internal_Rat_add(x_61, x_74); x_76 = lean_array_push(x_10, x_75); lean_ctor_set(x_2, 3, x_76); x_1 = x_6; @@ -6200,12 +6200,12 @@ lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean x_78 = lean_ctor_get(x_61, 0); lean_inc(x_78); lean_dec(x_61); -x_79 = l_Lean_Rat_ceil(x_78); +x_79 = l_Std_Internal_Rat_ceil(x_78); x_80 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_5); x_81 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; -x_82 = l_Lean_Rat_add(x_80, x_81); +x_82 = l_Std_Internal_Rat_add(x_80, x_81); x_83 = lean_array_push(x_10, x_82); lean_ctor_set(x_2, 3, x_83); x_1 = x_6; @@ -6223,7 +6223,7 @@ lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; x_86 = lean_ctor_get(x_61, 0); x_87 = lean_ctor_get(x_61, 1); lean_dec(x_87); -x_88 = l_Lean_Rat_ceil(x_86); +x_88 = l_Std_Internal_Rat_ceil(x_86); lean_ctor_set(x_61, 1, x_5); lean_ctor_set(x_61, 0, x_88); x_89 = lean_array_push(x_10, x_61); @@ -6237,7 +6237,7 @@ lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; x_91 = lean_ctor_get(x_61, 0); lean_inc(x_91); lean_dec(x_61); -x_92 = l_Lean_Rat_ceil(x_91); +x_92 = l_Std_Internal_Rat_ceil(x_91); x_93 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_93, 0, x_92); lean_ctor_set(x_93, 1, x_5); @@ -6270,7 +6270,7 @@ if (lean_is_exclusive(x_61)) { lean_dec_ref(x_61); x_100 = lean_box(0); } -x_101 = l_Lean_Rat_ceil(x_99); +x_101 = l_Std_Internal_Rat_ceil(x_99); if (lean_is_scalar(x_100)) { x_102 = lean_alloc_ctor(0, 2, 0); } else { @@ -6279,7 +6279,7 @@ if (lean_is_scalar(x_100)) { lean_ctor_set(x_102, 0, x_101); lean_ctor_set(x_102, 1, x_5); x_103 = l_Lean_Meta_Linear_pickAssignment_x3f___closed__3; -x_104 = l_Lean_Rat_add(x_102, x_103); +x_104 = l_Std_Internal_Rat_add(x_102, x_103); x_105 = lean_array_push(x_10, x_104); x_106 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_106, 0, x_7); @@ -6304,7 +6304,7 @@ if (lean_is_exclusive(x_61)) { lean_dec_ref(x_61); x_109 = lean_box(0); } -x_110 = l_Lean_Rat_ceil(x_108); +x_110 = l_Std_Internal_Rat_ceil(x_108); if (lean_is_scalar(x_109)) { x_111 = lean_alloc_ctor(0, 2, 0); } else { @@ -6486,7 +6486,7 @@ return x_202; } lean_object* initialize_Init_Data_Ord(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Array_DecidableEq(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_LinearArith_Solver(uint8_t builtin, lean_object* w) { lean_object * res; @@ -6498,7 +6498,7 @@ lean_dec_ref(res); res = initialize_Init_Data_Array_DecidableEq(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Data_Rat(builtin, lean_io_mk_world()); +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Meta_Linear_instInhabitedVar = _init_l_Lean_Meta_Linear_instInhabitedVar(); @@ -6507,32 +6507,32 @@ l_Lean_Meta_Linear_instOrdVar___closed__1 = _init_l_Lean_Meta_Linear_instOrdVar_ lean_mark_persistent(l_Lean_Meta_Linear_instOrdVar___closed__1); l_Lean_Meta_Linear_instOrdVar = _init_l_Lean_Meta_Linear_instOrdVar(); lean_mark_persistent(l_Lean_Meta_Linear_instOrdVar); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__11); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__12); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_166____closed__13); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__12); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprVar____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_167____closed__13); l_Lean_Meta_Linear_instReprVar___closed__1 = _init_l_Lean_Meta_Linear_instReprVar___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprVar___closed__1); l_Lean_Meta_Linear_instReprVar = _init_l_Lean_Meta_Linear_instReprVar(); @@ -6545,76 +6545,76 @@ l_Lean_Meta_Linear_instInhabitedAssignment = _init_l_Lean_Meta_Linear_instInhabi lean_mark_persistent(l_Lean_Meta_Linear_instInhabitedAssignment); l_Lean_Meta_Linear_instInhabitedPoly = _init_l_Lean_Meta_Linear_instInhabitedPoly(); lean_mark_persistent(l_Lean_Meta_Linear_instInhabitedPoly); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__1); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__2); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__3); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__4); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__5); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__6); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__7); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__8); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__9); -l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10(); -lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____spec__2___closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__11); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__12); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__13); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__14); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__15); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__16); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__17); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__18); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__19); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_392____closed__20); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__1); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__2); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__3); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__4); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__5); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__6); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__7); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__8); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__9); +l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10 = _init_l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10(); +lean_mark_persistent(l_Prod_repr___at___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____spec__2___closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__12); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__13); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__14); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__15); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__16); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__17); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__18); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__19); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_393____closed__20); l_Lean_Meta_Linear_instReprPoly___closed__1 = _init_l_Lean_Meta_Linear_instReprPoly___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprPoly___closed__1); l_Lean_Meta_Linear_instReprPoly = _init_l_Lean_Meta_Linear_instReprPoly(); lean_mark_persistent(l_Lean_Meta_Linear_instReprPoly); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_430____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_decEqPoly____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_431____closed__3); l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1 = _init_l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_Poly_getMaxVarCoeff___closed__1); l_Lean_Meta_Linear_Poly_add___closed__1 = _init_l_Lean_Meta_Linear_Poly_add___closed__1(); @@ -6639,22 +6639,22 @@ l_Lean_Meta_Linear_instBEqJustification___closed__1 = _init_l_Lean_Meta_Linear_i lean_mark_persistent(l_Lean_Meta_Linear_instBEqJustification___closed__1); l_Lean_Meta_Linear_instBEqJustification = _init_l_Lean_Meta_Linear_instBEqJustification(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqJustification); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3751____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprJustification____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_3752____closed__8); l_Lean_Meta_Linear_instReprJustification___closed__1 = _init_l_Lean_Meta_Linear_instReprJustification___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprJustification___closed__1); l_Lean_Meta_Linear_instReprJustification = _init_l_Lean_Meta_Linear_instReprJustification(); @@ -6666,54 +6666,54 @@ l_Lean_Meta_Linear_instBEqCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instB lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind___closed__1); l_Lean_Meta_Linear_instBEqCnstrKind = _init_l_Lean_Meta_Linear_instBEqCnstrKind(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstrKind); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__11); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__12); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__13); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__14); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__15); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__16); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__17); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__18); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__19); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__20); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__21); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__22); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__23); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4022____closed__24); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__12); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__13); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__14); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__15); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__16); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__17); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__18); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__19); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__20); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__21); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__22); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__23); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstrKind____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4023____closed__24); l_Lean_Meta_Linear_instReprCnstrKind___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstrKind___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstrKind___closed__1); l_Lean_Meta_Linear_instReprCnstrKind = _init_l_Lean_Meta_Linear_instReprCnstrKind(); @@ -6726,28 +6726,28 @@ l_Lean_Meta_Linear_instBEqCnstr___closed__1 = _init_l_Lean_Meta_Linear_instBEqCn lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr___closed__1); l_Lean_Meta_Linear_instBEqCnstr = _init_l_Lean_Meta_Linear_instBEqCnstr(); lean_mark_persistent(l_Lean_Meta_Linear_instBEqCnstr); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__1); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__2); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__3); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__4); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__5); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__6); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__7); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__8); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__9); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__10); -l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11(); -lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4586____closed__11); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__1); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__2); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__3); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__4); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__5); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__6); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__7); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__8); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__9); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__10); +l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11 = _init_l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11(); +lean_mark_persistent(l___private_Lean_Meta_Tactic_LinearArith_Solver_0__Lean_Meta_Linear_reprCnstr____x40_Lean_Meta_Tactic_LinearArith_Solver___hyg_4587____closed__11); l_Lean_Meta_Linear_instReprCnstr___closed__1 = _init_l_Lean_Meta_Linear_instReprCnstr___closed__1(); lean_mark_persistent(l_Lean_Meta_Linear_instReprCnstr___closed__1); l_Lean_Meta_Linear_instReprCnstr = _init_l_Lean_Meta_Linear_instReprCnstr(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Replace.c b/stage0/stdlib/Lean/Meta/Tactic/Replace.c index 50862aac3d0d..482c29582a68 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Replace.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Replace.c @@ -3125,7 +3125,7 @@ x_35 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_24); lean_ctor_set(x_35, 2, x_31); -lean_ctor_set_tag(x_14, 9); +lean_ctor_set_tag(x_14, 10); lean_ctor_set(x_14, 0, x_35); x_36 = l_Lean_Elab_pushInfoLeaf___at_Lean_MVarId_withReverted___spec__1(x_14, x_7, x_8, x_9, x_10, x_34); x_37 = lean_ctor_get(x_36, 1); @@ -3186,7 +3186,7 @@ x_49 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_24); lean_ctor_set(x_49, 2, x_45); -x_50 = lean_alloc_ctor(9, 1, 0); +x_50 = lean_alloc_ctor(10, 1, 0); lean_ctor_set(x_50, 0, x_49); x_51 = l_Lean_Elab_pushInfoLeaf___at_Lean_MVarId_withReverted___spec__1(x_50, x_7, x_8, x_9, x_10, x_48); x_52 = lean_ctor_get(x_51, 1); @@ -3279,10 +3279,10 @@ lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_60); lean_ctor_set(x_72, 2, x_67); if (lean_is_scalar(x_68)) { - x_73 = lean_alloc_ctor(9, 1, 0); + x_73 = lean_alloc_ctor(10, 1, 0); } else { x_73 = x_68; - lean_ctor_set_tag(x_73, 9); + lean_ctor_set_tag(x_73, 10); } lean_ctor_set(x_73, 0, x_72); x_74 = l_Lean_Elab_pushInfoLeaf___at_Lean_MVarId_withReverted___spec__1(x_73, x_7, x_8, x_9, x_10, x_71); diff --git a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c index ffc64a02a06c..89edfa8a27d3 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/TryThis.c +++ b/stage0/stdlib/Lean/Meta/Tactic/TryThis.c @@ -18,7 +18,6 @@ static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_success___closed_ static uint8_t l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__18; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5; lean_object* l_Lean_FileMap_utf8RangeToLspRange(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__10; static lean_object* l_Lean_Meta_Tactic_TryThis_instInhabitedSuggestion___closed__1; @@ -49,7 +48,6 @@ lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__10; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__7; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -61,6 +59,7 @@ lean_object* l_Lean_PrettyPrinter_delab(lean_object*, lean_object*, lean_object* static lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisWidget; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2; static double l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addTermSuggestion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addExactSuggestionCore___lambda__1___closed__1; @@ -167,7 +166,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addTermSuggestions___boxed(l static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__20; uint8_t l_instDecidableNot___rarg(uint8_t); -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_getIndentAndColumn___lambda__1___boxed(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_String_findLineStart(lean_object*, lean_object*); @@ -211,12 +209,14 @@ static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__1 static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__71; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__54; extern lean_object* l_Lean_instToJsonJson; +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5; static double l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__43; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instToMessageDataSuggestionText(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionText_prettyExtra___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__9; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__23; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__9; @@ -245,7 +245,6 @@ double pow(double, double); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__18; static lean_object* l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__1; -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Tactic_TryThis_addSuggestions___spec__2___closed__1; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); @@ -310,6 +309,7 @@ static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__32; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__11; static lean_object* l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__3; static lean_object* l_List_mapTR_loop___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__2___closed__1; +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_asHypothesis___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_addSuggestions___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -323,9 +323,9 @@ lean_object* l_Lean_Widget_savePanelWidgetInfo(uint64_t, lean_object*, lean_obje static lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestion___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__19; static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586_(lean_object*); static lean_object* l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addSuggestions___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instToMessageDataSuggestion(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_success___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -408,6 +408,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_instCoeHeadTSyntaxConsSyntax lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); double round(double); +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1; +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__17; static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__5; @@ -445,7 +447,6 @@ lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_D lean_object* l_Lean_MessageData_bracket(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed__53; uint8_t l_Lean_Expr_isConst(lean_object*); -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_addExactSuggestions___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__3; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__24; @@ -455,6 +456,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Tactic_TryThis_ad static lean_object* l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__8; lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addExactSuggestionCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore___closed__4; static lean_object* l_Lean_Meta_Tactic_TryThis_Suggestion_toJsonAndInfoM___lambda__3___closed__1; @@ -463,9 +465,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Meta_Tactic_TryThi static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_warning___closed__1; static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__15; static lean_object* l_Lean_Meta_Tactic_TryThis_addExactSuggestions___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584_(lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__3; -static lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_value___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Tactic_TryThis_SuggestionStyle_error___closed__17; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_TryThis_0__Lean_Meta_Tactic_TryThis_addSuggestionCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1168,7 +1168,7 @@ return x_16; LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_tryThisProvider___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_4) == 8) +if (lean_obj_tag(x_4) == 9) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_6 = lean_ctor_get(x_4, 0); @@ -1726,7 +1726,7 @@ lean_dec(x_7); return x_12; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1() { _start: { lean_object* x_1; @@ -1734,7 +1734,7 @@ x_1 = lean_mk_string_unchecked("format", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2() { _start: { lean_object* x_1; @@ -1742,17 +1742,17 @@ x_1 = lean_mk_string_unchecked("inputWidth", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1; -x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2; +x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1; +x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4() { _start: { lean_object* x_1; @@ -1760,7 +1760,7 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5() { _start: { lean_object* x_1; @@ -1768,13 +1768,13 @@ x_1 = lean_mk_string_unchecked("ideal input width", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_unsigned_to_nat(100u); -x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; -x_3 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5; +x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; +x_3 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1782,7 +1782,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7() { +static lean_object* _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; @@ -1790,19 +1790,19 @@ x_1 = l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__1; x_2 = l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__2; x_3 = l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__3; x_4 = l___regBuiltin_Lean_Meta_Tactic_TryThis_tryThisWidget__1___closed__4; -x_5 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1; -x_6 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2; +x_5 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1; +x_6 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2; x_7 = l_Lean_Name_mkStr6(x_1, x_2, x_3, x_4, x_5, x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3; -x_3 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6; -x_4 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7; +x_2 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3; +x_3 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6; +x_4 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7; x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_5____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -3854,10 +3854,10 @@ x_38 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_38, 0, x_37); lean_ctor_set(x_38, 1, x_19); if (lean_is_scalar(x_17)) { - x_39 = lean_alloc_ctor(8, 1, 0); + x_39 = lean_alloc_ctor(9, 1, 0); } else { x_39 = x_17; - lean_ctor_set_tag(x_39, 8); + lean_ctor_set_tag(x_39, 9); } lean_ctor_set(x_39, 0, x_38); x_40 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(x_39, x_8, x_9, x_27); @@ -4137,10 +4137,10 @@ x_140 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_140, 0, x_139); lean_ctor_set(x_140, 1, x_136); if (lean_is_scalar(x_17)) { - x_141 = lean_alloc_ctor(8, 1, 0); + x_141 = lean_alloc_ctor(9, 1, 0); } else { x_141 = x_17; - lean_ctor_set_tag(x_141, 8); + lean_ctor_set_tag(x_141, 9); } lean_ctor_set(x_141, 0, x_140); x_142 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(x_141, x_8, x_9, x_128); @@ -4333,7 +4333,7 @@ static lean_object* _init_l_Lean_Meta_Tactic_TryThis_addSuggestion___closed__1() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -6395,7 +6395,7 @@ static lean_object* _init_l_Lean_Meta_Tactic_TryThis_addHaveSuggestion___closed_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_2 = l_String_toSubstring_x27(x_1); return x_2; } @@ -8200,7 +8200,7 @@ static lean_object* _init_l_List_mapTR_loop___at_Lean_Meta_Tactic_TryThis_addRew _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_1 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -8685,7 +8685,7 @@ lean_ctor_set(x_24, 1, x_23); if (lean_obj_tag(x_5) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_25 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_26 = lean_box(0); x_27 = l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__1(x_6, x_2, x_3, x_25, x_24, x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_10); @@ -8729,7 +8729,7 @@ x_39 = lean_format_pretty(x_35, x_37, x_38, x_38); x_40 = l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__8; x_41 = lean_string_append(x_40, x_39); lean_dec(x_39); -x_42 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_42 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_43 = lean_string_append(x_41, x_42); x_44 = lean_string_append(x_42, x_43); lean_dec(x_43); @@ -8794,7 +8794,7 @@ lean_ctor_set(x_59, 1, x_58); if (lean_obj_tag(x_5) == 0) { lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_60 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_61 = lean_box(0); x_62 = l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__1(x_6, x_2, x_3, x_60, x_59, x_61, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_10); @@ -8838,7 +8838,7 @@ x_74 = lean_format_pretty(x_70, x_72, x_73, x_73); x_75 = l_Lean_Meta_Tactic_TryThis_addRewriteSuggestion___lambda__2___closed__8; x_76 = lean_string_append(x_75, x_74); lean_dec(x_74); -x_77 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4; +x_77 = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4; x_78 = lean_string_append(x_76, x_77); x_79 = lean_string_append(x_77, x_78); lean_dec(x_78); @@ -9497,21 +9497,21 @@ l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__4 = _init_l_Lean_Met lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__4); l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5 = _init_l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5(); lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_delabToRefinableSyntax___closed__5); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__1); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__2); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__3); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__4); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__5); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__6); -l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7(); -lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584____closed__7); -res = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_584_(lean_io_mk_world()); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__1); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__2); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__3); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__4); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__5); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__6); +l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7 = _init_l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7(); +lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586____closed__7); +res = l_Lean_Meta_Tactic_TryThis_initFn____x40_Lean_Meta_Tactic_TryThis___hyg_586_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_Tactic_TryThis_format_inputWidth = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_Tactic_TryThis_format_inputWidth); diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index fd8cfef081fe..ec27cb3d7e88 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -25,7 +25,6 @@ lean_object* l_Lean_Parser_Term_binderTactic_parenthesizer(lean_object*, lean_ob static lean_object* l_Lean_Parser_Command_computedFields___closed__13; static lean_object* l_Lean_Parser_Command_declValSimple___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6; lean_object* l_Lean_Parser_many1_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_register__tactic__tag_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -36,11 +35,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__10; lean_object* l_Lean_PrettyPrinter_Parenthesizer_recover_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33; static lean_object* l_Lean_Parser_Tactic_open___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersT_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_nonrec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__8; @@ -59,7 +58,6 @@ static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__5; static lean_object* l_Lean_Parser_Command_structure___closed__11; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange__1___closed__6; static lean_object* l_Lean_Parser_Command_openSimple___closed__6; @@ -140,6 +138,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_formatter__1___ static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_omit_docString__1(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24; lean_object* l_Lean_Parser_takeUntilFn___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_synth___closed__1; static lean_object* l_Lean_Parser_Command_set__option___closed__10; @@ -149,7 +148,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_eoi; static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__1; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48; static lean_object* l_Lean_Parser_Command_import___closed__1; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange__1___closed__2; @@ -190,7 +188,6 @@ static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_structureTk_parenthesizer__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_in; static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52; static lean_object* l_Lean_Parser_Command_theorem_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_where_parenthesizer__1___closed__1; @@ -206,8 +203,10 @@ static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot__1(lean_object*); static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_example_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48; static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_import_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3; lean_object* l_Lean_Parser_Term_attrInstance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter__1___closed__1; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); @@ -263,7 +262,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_formatter_ static lean_object* l_Lean_Parser_Command_namedPrio___closed__6; static lean_object* l_Lean_Parser_Command_printTacTags_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_tactic__extension___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__11; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_orelse_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -275,7 +273,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_register__tactic__tag_dec static lean_object* l___regBuiltin_Lean_Parser_Command_version_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_addDocString___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_docString__1___closed__1; static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_tactic__extension_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -313,6 +310,7 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__6; static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open__1___closed__1; static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__9; static lean_object* l_Lean_Parser_Term_open_formatter___closed__3; @@ -370,7 +368,6 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___c static lean_object* l_Lean_Parser_Command_declSig___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_omit_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36; static lean_object* l_Lean_Parser_Command_deriving___closed__11; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__6; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__3; @@ -390,7 +387,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_eraseAttr_formatter(lean_object*, static lean_object* l_Lean_Parser_Command_whereStructInst___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Command_opaque_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3092_(lean_object*); static lean_object* l_Lean_Parser_Command_synth_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__12; static lean_object* l_Lean_Parser_Command_theorem___closed__9; @@ -431,6 +427,9 @@ static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange__1___closed__7; +static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11; static lean_object* l_Lean_Parser_Command_definition_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange__1(lean_object*); @@ -586,7 +585,6 @@ static lean_object* l_Lean_Parser_Command_moduleDoc___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__10; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__7; static lean_object* l_Lean_Parser_Command_open___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26; lean_object* l_instDecidableEqChar___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__5; @@ -598,7 +596,6 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___close static lean_object* l_Lean_Parser_Command_structInstBinder___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange__1___closed__3; static lean_object* l_Lean_Parser_Command_openRenaming___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_docString__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange__1___closed__2; @@ -616,7 +613,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange__1_ static lean_object* l_Lean_Parser_Command_end___closed__3; static lean_object* l_Lean_Parser_Command_register__tactic__tag_formatter___closed__2; static lean_object* l_Lean_Parser_Command_computedField___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44; static lean_object* l_Lean_Parser_Command_nonrec___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange__1___closed__2; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__2; @@ -628,6 +624,7 @@ static lean_object* l_Lean_Parser_Command_extends_formatter___closed__5; lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Command_tactic__extension___closed__3; static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_print; @@ -647,6 +644,7 @@ static lean_object* l_Lean_Parser_Command_section___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange__1___closed__5; lean_object* l_Lean_Parser_ppDedent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__19; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_section___closed__2; @@ -658,9 +656,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_register__tactic__tag_par static lean_object* l___regBuiltin_Lean_Parser_Command_where_parenthesizer__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structFields_formatter__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__8; @@ -674,7 +670,6 @@ lean_object* l_Lean_Parser_andthenFn(lean_object*, lean_object*, lean_object*, l static lean_object* l_Lean_Parser_Tactic_set__option_formatter___closed__2; static lean_object* l_Lean_Parser_Command_classInductive___closed__11; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__6; static lean_object* l_Lean_Parser_Command_instance___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__3; @@ -686,7 +681,6 @@ static lean_object* l_Lean_Parser_Command_namedPrio___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers___closed__4; static lean_object* l_Lean_Parser_Command_register__tactic__tag_formatter___closed__7; lean_object* l_Lean_Parser_withOpenFn(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51; lean_object* l_Lean_Parser_withOpen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_abbrev___closed__8; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__2; @@ -789,7 +783,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange__1___ static lean_object* l_Lean_Parser_Command_declId___closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5; static lean_object* l_Lean_Parser_Command_optNamedPrio___closed__1; static lean_object* l_Lean_Parser_Command_printAxioms___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter__1(lean_object*); @@ -802,7 +795,6 @@ static lean_object* l_Lean_Parser_Command_nonrec_formatter___closed__1; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__4; static lean_object* l_Lean_Parser_Command_optDeriving___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__2; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__11; @@ -816,6 +808,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags__1(lean static lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__7; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49; static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange__1___closed__7; static lean_object* l_Lean_Parser_Command_example_formatter___closed__1; @@ -849,7 +842,6 @@ static lean_object* l_Lean_Parser_Tactic_set__option___closed__7; lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_check__failure_parenthesizer___closed__2; lean_object* l_Lean_Parser_setExpected(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_include___closed__8; @@ -863,7 +855,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange__1_ static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_formatter__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_eoi_formatter___closed__2; @@ -958,12 +949,14 @@ static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__7; static lean_object* l_Lean_Parser_Command_quot___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__3; lean_object* l_Lean_Parser_priorityParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_deriving___closed__8; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__19; static lean_object* l_Lean_Parser_Command_ctor___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Command_where_formatter__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20; static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_declaration___closed__8; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__17; @@ -1046,6 +1039,7 @@ static lean_object* l_Lean_Parser_Command_classTk_formatter___closed__2; static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__4; +lean_object* l_Lean_Parser_Term_structInstFields(lean_object*); static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_set__option_docString__1(lean_object*); @@ -1120,6 +1114,7 @@ static lean_object* l_Lean_Parser_Command_openRenaming___closed__7; static lean_object* l_Lean_Parser_Command_classTk_formatter___closed__1; static lean_object* l_Lean_Parser_Term_set__option___closed__1; static lean_object* l_Lean_Parser_Command_universe___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_computedField_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructInst___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Command_set__option_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1185,9 +1180,11 @@ static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25; static lean_object* l_Lean_Parser_Term_quot___closed__13; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eval_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_end_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange__1___closed__5; @@ -1199,6 +1196,7 @@ lean_object* l_Lean_Parser_registerBuiltinNodeKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange__1___closed__7; +lean_object* l_Lean_Parser_Term_structInstFields_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_section_parenthesizer__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_quot; static lean_object* l_Lean_Parser_Command_abbrev___closed__10; @@ -1243,7 +1241,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_formatte LEAN_EXPORT lean_object* l_Lean_Parser_Command_structFields; static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_check__failure___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange__1___closed__4; @@ -1252,7 +1249,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange__1___clo static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_set__option___closed__8; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17; static lean_object* l_Lean_Parser_Command_structure___closed__15; static lean_object* l_Lean_Parser_Command_computedField___closed__7; static lean_object* l_Lean_Parser_Command_optDeriving___closed__2; @@ -1284,11 +1280,11 @@ static lean_object* l_Lean_Parser_Command_optDeriving___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange__1___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7; static lean_object* l_Lean_Parser_Command_import_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__5; static lean_object* l_Lean_Parser_Term_open___closed__1; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28; static lean_object* l_Lean_Parser_Command_structure___closed__20; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_formatter__1___closed__1; @@ -1362,6 +1358,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_declRange__1 static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_unsafe_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange__1___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__5; static lean_object* l_Lean_Parser_Command_quot___closed__14; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__18; @@ -1375,6 +1372,7 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__8; static lean_object* l_Lean_Parser_Command_omit_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange__1___closed__7; static lean_object* l_Lean_Parser_Command_structureTk___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41; static lean_object* l_Lean_Parser_Command_declSig___closed__11; static lean_object* l_Lean_Parser_Command_ctor___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_print_declRange__1___closed__4; @@ -1445,7 +1443,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection static lean_object* l_Lean_Parser_Command_export___closed__3; static lean_object* l_Lean_Parser_Command_opaque___closed__6; static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_formatter__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1___closed__1; @@ -1466,7 +1463,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange__1___closed LEAN_EXPORT lean_object* l_Lean_Parser_Term_precheckedQuot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_partial_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21; static lean_object* l_Lean_Parser_Command_register__tactic__tag_parenthesizer___closed__4; lean_object* l_Lean_Parser_Term_binderIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_declVal_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1484,6 +1480,7 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__18; static lean_object* l_Lean_Parser_Command_printAxioms_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26; static lean_object* l_Lean_Parser_Command_namespace___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange__1___closed__1; static lean_object* l_Lean_Parser_Command_private_parenthesizer___closed__1; @@ -1504,7 +1501,6 @@ static lean_object* l_Lean_Parser_Tactic_set__option___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_printTacTags_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_unsafe_formatter__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot___closed__7; static lean_object* l_Lean_Parser_Command_whereStructField___closed__1; @@ -1537,6 +1533,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter__1___closed__ lean_object* l_Lean_Parser_identWithPartialTrailingDot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structInstBinder___closed__7; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43; static lean_object* l_Lean_Parser_Command_export_formatter___closed__1; static lean_object* l_Lean_Parser_Command_printTacTags___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange__1(lean_object*); @@ -1633,7 +1630,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving__1(lean_obj LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_tactic__extension_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_register__tactic__tag_formatter___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47; static lean_object* l_Lean_Parser_Command_structureTk_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__4; @@ -1675,7 +1671,6 @@ extern lean_object* l_Lean_Parser_Term_attrKind; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_Command_import_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42; lean_object* l_Lean_ppLine_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_where___closed__2; @@ -1705,6 +1700,7 @@ static lean_object* l_Lean_Parser_Command_openScoped___closed__8; lean_object* l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_printTacTags___closed__8; static lean_object* l_Lean_Parser_Command_mutual___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52; static lean_object* l_Lean_Parser_Command_eraseAttr_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange__1___closed__6; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__5; @@ -1728,7 +1724,6 @@ LEAN_EXPORT uint8_t l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipU static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_formatter__1___closed__1; lean_object* l_Lean_Parser_notFollowedBy(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_theorem_formatter__1(lean_object*); @@ -1768,7 +1763,6 @@ static lean_object* l_Lean_Parser_Command_version___closed__5; static lean_object* l_Lean_Parser_Command_definition_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structFields___closed__13; static lean_object* l_Lean_Parser_Command_declVal___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32; LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__1; @@ -1812,12 +1806,14 @@ lean_object* l_Lean_Parser_symbol_formatter(lean_object*, lean_object*, lean_obj static lean_object* l_Lean_Parser_Command_whereStructInst___closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19; static lean_object* l_Lean_Parser_Command_section_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_computedField___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30; LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_set__option___closed__5; static lean_object* l_Lean_Parser_Command_omit___closed__6; @@ -1862,15 +1858,16 @@ static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange__1___closed__2; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49; static lean_object* l_Lean_Parser_Command_inductive___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_evalBang___closed__2; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__18; static lean_object* l_Lean_Parser_Command_tactic__extension___closed__10; static lean_object* l_Lean_Parser_Command_declSig_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedField_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openScoped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_variable___closed__2; @@ -1897,12 +1894,10 @@ static lean_object* l_Lean_Parser_Command_evalBang___closed__5; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12; static lean_object* l_Lean_Parser_Command_openRenaming___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor; static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_parenthesizer__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29; extern lean_object* l_Lean_Parser_Command_docComment; static lean_object* l_Lean_Parser_Command_openRenaming___closed__12; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__2; @@ -1936,7 +1931,6 @@ static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange__1___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_openDecl_docString__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declBody; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_quot_formatter___closed__4; @@ -1975,6 +1969,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_inductive_parenthesizer(lean_obje static lean_object* l___regBuiltin_Lean_Parser_Command_evalBang_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_structure_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18; static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntilWsOrDelim___closed__1; static lean_object* l_Lean_Parser_Command_synth___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__9; @@ -1982,7 +1977,6 @@ static lean_object* l_Lean_Parser_Command_structFields___closed__11; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersF; static lean_object* l_Lean_Parser_Command_deriving___closed__15; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9; static lean_object* l_Lean_Parser_Tactic_set__option_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__2; static lean_object* l_Lean_Parser_Command_tactic__extension___closed__4; @@ -1999,7 +1993,6 @@ static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__15; static lean_object* l_Lean_Parser_Command_attribute___closed__4; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_classInductive___closed__15; @@ -2051,13 +2044,13 @@ static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_initialize___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_optDeriving___closed__4; static lean_object* l_Lean_Parser_Command_version___closed__6; static lean_object* l_Lean_Parser_Command_namespace_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__2; @@ -2072,7 +2065,6 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ lean_object* l_Lean_Parser_many_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__7; lean_object* l_Lean_Parser_Term_attributes_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__5; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__5; @@ -2094,7 +2086,6 @@ static lean_object* l_Lean_Parser_Command_structureTk___closed__7; static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1___closed__7; static lean_object* l_Lean_Parser_Command_ctor___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_axiom___closed__9; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__1; @@ -2154,7 +2145,6 @@ static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_quot___closed__4; static lean_object* l_Lean_Parser_Command_printTacTags_formatter___closed__5; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__1; static lean_object* l_Lean_Parser_Command_opaque___closed__8; static lean_object* l_Lean_Parser_Command_inductive___closed__12; @@ -2221,6 +2211,7 @@ static lean_object* l_Lean_Parser_Term_set__option___closed__9; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_section_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2262,6 +2253,7 @@ static lean_object* l_Lean_Parser_Command_moduleDoc___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__12; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__16; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export__1(lean_object*); static lean_object* l_Lean_Parser_Command_opaque___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable; @@ -2269,6 +2261,7 @@ static lean_object* l_Lean_Parser_Term_open_formatter___closed__2; static lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntil___closed__1; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_skipUntil___elambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe; static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_tactic__extension_formatter___closed__1; @@ -2284,6 +2277,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer(lean_ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_formatter__1___closed__1; @@ -2291,6 +2285,7 @@ static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__13; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__14; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22; static lean_object* l_Lean_Parser_Command_omit___closed__10; lean_object* l_Lean_Parser_checkColGe(lean_object*); static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__9; @@ -2393,7 +2388,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_formatter__ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declId_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_omit_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_declVal_docString__1___closed__3; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__1; @@ -2434,6 +2428,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange__1_ static lean_object* l_Lean_Parser_Command_printEqns_formatter___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_computedField_formatter__1___closed__1; lean_object* l_Lean_ppDedent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3095_(lean_object*); static lean_object* l_Lean_Parser_Command_universe___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_declId___closed__8; @@ -2460,6 +2455,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_end___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53; static lean_object* l_Lean_Parser_Command_include_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_parenthesizer__1(lean_object*); @@ -2483,6 +2479,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec; static lean_object* l_Lean_Parser_Command_declValEqns___closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119_(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_set__option_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange__1___closed__6; @@ -2499,7 +2496,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_parenthesizer_ static lean_object* l_Lean_Parser_Command_exit_formatter___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__7; lean_object* l_Lean_ppSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize; static lean_object* l_Lean_Parser_Command_openDecl___closed__4; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__4; @@ -2525,6 +2521,7 @@ static lean_object* l_Lean_Parser_Command_structCtor___closed__9; static lean_object* l_Lean_Parser_Command_optDeclSig___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_theorem_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structureTk___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange__1___closed__1; @@ -2597,13 +2594,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_formatter__1___clo static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange__1___closed__5; static lean_object* l_Lean_Parser_Command_printTacTags___closed__12; lean_object* l_Lean_Parser_sepBy1_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_section_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDeriving___closed__6; static lean_object* l_Lean_Parser_Command_computedFields___closed__3; static lean_object* l_Lean_Parser_Command_register__tactic__tag___closed__5; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__14; lean_object* l_Lean_Parser_manyIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_formatter__1___closed__1; @@ -2667,7 +2664,6 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__5; static lean_object* l_Lean_Parser_Command_mutual___closed__14; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_theorem_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eval; static lean_object* l_Lean_Parser_Command_whereStructField___closed__5; @@ -2717,6 +2713,7 @@ static lean_object* l_Lean_Parser_Command_omit_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_declValEqns_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_namedPrio___closed__20; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39; static lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_openDecl_formatter___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option_formatter___closed__1; @@ -2757,11 +2754,13 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l_Lean_Parser_Command_addDocString___closed__8; static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__4; lean_object* l_Lean_Parser_ppGroup_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37; static lean_object* l_Lean_Parser_Command_initialize___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_section_docString__1___closed__1; static lean_object* l_Lean_Parser_Command_computedField___closed__3; lean_object* l_Lean_Parser_incQuotDepth(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_formatter__1___closed__2; lean_object* l_Lean_Parser_Term_matchAlts(lean_object*); @@ -2769,7 +2768,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Command_0__Lean_Parser_Command_ extern lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__7; lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35; static lean_object* l_Lean_Parser_Command_evalBang_formatter___closed__4; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_formatter__1___closed__2; @@ -2810,6 +2808,7 @@ lean_object* l_Lean_Parser_Command_commentBody_parenthesizer___boxed(lean_object LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_parenthesizer__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__8; +lean_object* l_Lean_Parser_Term_structInstFields_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_unicodeSymbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_in___closed__2; @@ -2843,6 +2842,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declSig_parenthesizer(lean_object static lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDefDeriving_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_computedFields_formatter___closed__7; @@ -2859,6 +2859,7 @@ static lean_object* l_Lean_Parser_Command_variable_formatter___closed__6; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_ctor___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_include_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21; static lean_object* l_Lean_Parser_Command_namedPrio___closed__13; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange__1___closed__1; @@ -2871,6 +2872,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer__1 static lean_object* l_Lean_Parser_Command_example___closed__8; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50; static lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter__1___closed__1; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; @@ -2896,7 +2898,6 @@ static lean_object* l_Lean_Parser_Command_export_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_private_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_register__tactic__tag__1(lean_object*); static lean_object* l_Lean_Parser_Command_structFields___closed__16; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_in_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_tactic__extension_declRange__1___closed__7; @@ -2985,6 +2986,7 @@ static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___c static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange__1___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer__1(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40; static lean_object* l_Lean_Parser_Command_inductive___closed__17; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_inductive_docString__1(lean_object*); static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; @@ -3017,7 +3019,6 @@ static lean_object* l_Lean_Parser_Command_quot___closed__5; static lean_object* l_Lean_Parser_Command_check___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_protected; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_print_formatter___closed__1; static lean_object* l_Lean_Parser_Command_private___closed__2; @@ -3037,12 +3038,12 @@ static lean_object* l_Lean_Parser_Command_export___closed__9; static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__2; lean_object* l_Lean_Parser_Command_commentBody_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31; static lean_object* l_Lean_Parser_Command_definition___closed__9; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__3; lean_object* l_Lean_Parser_Tactic_tacticSeq_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14; static lean_object* l_Lean_Parser_Command_print___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_computedFields_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__8; @@ -3081,6 +3082,7 @@ static lean_object* l_Lean_Parser_Command_eval_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_formatter__1___closed__1; static lean_object* l_Lean_Parser_Command_initialize___closed__12; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange__1___closed__4; static lean_object* l_Lean_Parser_Command_open_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputable_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3103,10 +3105,11 @@ static lean_object* l_Lean_Parser_Command_namedPrio_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_classTk___closed__1; static lean_object* l_Lean_Parser_Command_set__option___closed__7; static lean_object* l_Lean_Parser_Command_protected_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30; LEAN_EXPORT lean_object* l_Lean_Parser_Command_quot; static lean_object* l_Lean_Parser_Command_partial___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44; static lean_object* l_Lean_Parser_Command_declId___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38; lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_printEqns___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_quot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3167,7 +3170,7 @@ static lean_object* l_Lean_Parser_Command_extends_formatter___closed__4; static lean_object* l_Lean_Parser_Command_printEqns___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval_docString__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_omit_formatter___closed__5; static lean_object* l_Lean_Parser_Command_check__failure___closed__3; @@ -3324,7 +3327,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_docString__1(lea LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34; static lean_object* l_Lean_Parser_Command_definition_parenthesizer___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_omit_formatter__1___closed__1; @@ -3337,6 +3339,7 @@ static lean_object* l_Lean_Parser_Command_definition_formatter___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__8; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__2; static lean_object* l_Lean_Parser_Command_unsafe_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_whereStructInst___closed__30; static lean_object* l_Lean_Parser_Command_end_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_structureTk___closed__1; static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__2; @@ -3348,7 +3351,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange__1___ static lean_object* l_Lean_Parser_Command_printEqns_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declBody___closed__7; static lean_object* l_Lean_Parser_Command_structureTk_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_example_formatter___closed__3; static lean_object* l_Lean_Parser_Command_optionValue_parenthesizer___closed__2; @@ -3376,7 +3378,6 @@ static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_declModifiers_docString__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23; static lean_object* l_Lean_Parser_Command_ctor___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers___closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__8; @@ -3392,7 +3393,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange__1( static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_example_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structureTk_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18; static lean_object* l_Lean_Parser_Command_addDocString___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Tactic_set__option___closed__3; @@ -3427,6 +3427,7 @@ static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__4; static lean_object* l_Lean_Parser_Command_printTacTags_formatter___closed__3; static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_evalBang___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14; static lean_object* l_Lean_Parser_Command_declBody_formatter___closed__3; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_formatter__1(lean_object*); @@ -3493,6 +3494,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange__1___c static lean_object* l_Lean_Parser_Command_structureTk_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_printAxioms___closed__10; static lean_object* l_Lean_Parser_Command_print___closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32; static lean_object* l_Lean_Parser_Command_opaque___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersF_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declValEqns___closed__6; @@ -3662,6 +3664,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structureTk; static lean_object* l_Lean_Parser_Command_mutual___closed__12; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4; static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_set__option___closed__4; static lean_object* l_Lean_Parser_Command_variable___closed__11; @@ -3683,7 +3686,6 @@ static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_definition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__9; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_print_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_attribute_declRange__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_partial_parenthesizer__1(lean_object*); @@ -3713,6 +3715,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_formatter static lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_definition_formatter__1___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29; static lean_object* l___regBuiltin_Lean_Parser_Command_printEqns_declRange__1___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Command_where_parenthesizer___closed__3; @@ -3734,15 +3737,18 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_partial_formatter__1 static lean_object* l_Lean_Parser_Command_namespace___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_evalBang_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility; static lean_object* l_Lean_Parser_Command_eval_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10; static lean_object* l_Lean_Parser_Command_noncomputable___closed__6; static lean_object* l_Lean_Parser_Command_check_parenthesizer___closed__2; lean_object* l_Lean_Parser_notSymbol_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_register__tactic__tag_declRange__1___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eoi_formatter__1(lean_object*); @@ -3763,6 +3769,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_formatter(lean_object*, lean static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_formatter__1___closed__2; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_docString__1___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_include_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_printTacTags___closed__11; @@ -3770,7 +3777,6 @@ static lean_object* l_Lean_Parser_Command_where_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_formatter__1___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange__1___closed__2; static lean_object* l_Lean_Parser_Command_in___closed__1; @@ -8326,60 +8332,69 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__24() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_whereStructInst___closed__23; +x_2 = l_Lean_Parser_Term_structInstFields(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_whereStructInst___closed__24; x_2 = l_Lean_Parser_Command_declValSimple___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__25() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst___closed__5; -x_2 = l_Lean_Parser_Command_whereStructInst___closed__24; +x_2 = l_Lean_Parser_Command_whereStructInst___closed__25; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__26() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Command_whereStructInst___closed__25; +x_2 = l_Lean_Parser_Command_whereStructInst___closed__26; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__27() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_whereStructInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_whereStructInst___closed__26; +x_3 = l_Lean_Parser_Command_whereStructInst___closed__27; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__28() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst___closed__3; -x_2 = l_Lean_Parser_Command_whereStructInst___closed__27; +x_2 = l_Lean_Parser_Command_whereStructInst___closed__28; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__29() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst___closed__2; -x_2 = l_Lean_Parser_Command_whereStructInst___closed__28; +x_2 = l_Lean_Parser_Command_whereStructInst___closed__29; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -8388,7 +8403,7 @@ static lean_object* _init_l_Lean_Parser_Command_whereStructInst() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_whereStructInst___closed__29; +x_1 = l_Lean_Parser_Command_whereStructInst___closed__30; return x_1; } } @@ -13650,8 +13665,18 @@ return x_4; static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_whereStructInst_formatter___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_whereStructInst_formatter___closed__7; x_2 = l_Lean_Parser_Command_declValSimple_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13659,37 +13684,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst_formatter___closed__3; -x_2 = l_Lean_Parser_Command_whereStructInst_formatter___closed__7; +x_2 = l_Lean_Parser_Command_whereStructInst_formatter___closed__8; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst_formatter___closed__2; -x_2 = l_Lean_Parser_Command_whereStructInst_formatter___closed__8; +x_2 = l_Lean_Parser_Command_whereStructInst_formatter___closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_whereStructInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_whereStructInst_formatter___closed__9; +x_3 = l_Lean_Parser_Command_whereStructInst_formatter___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -13702,7 +13727,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst_formatter(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_whereStructInst_formatter___closed__1; -x_7 = l_Lean_Parser_Command_whereStructInst_formatter___closed__10; +x_7 = l_Lean_Parser_Command_whereStructInst_formatter___closed__11; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -19328,8 +19353,18 @@ return x_6; static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__7; x_2 = l_Lean_Parser_Command_declValSimple_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -19337,37 +19372,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_whereStructInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9; +x_3 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -19380,7 +19415,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer(lea { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10; +x_7 = l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11; x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -42138,7 +42173,7 @@ x_1 = l_Lean_Parser_Command_eoi___closed__5; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3092_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3095_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -42199,7 +42234,7 @@ x_1 = l_Lean_Parser_Command_ctor___closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42209,7 +42244,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2() { _start: { lean_object* x_1; @@ -42217,19 +42252,19 @@ x_1 = lean_mk_string_unchecked("declModifiersF", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -42239,7 +42274,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -42249,7 +42284,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -42259,12 +42294,12 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -42273,7 +42308,7 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8() { _start: { lean_object* x_1; @@ -42281,17 +42316,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10() { _start: { lean_object* x_1; @@ -42299,7 +42334,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11() { _start: { lean_object* x_1; @@ -42307,17 +42342,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersF_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13() { _start: { lean_object* x_1; @@ -42325,7 +42360,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14() { _start: { lean_object* x_1; @@ -42333,17 +42368,17 @@ x_1 = lean_mk_string_unchecked("nestedDeclModifiers", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16() { _start: { lean_object* x_1; @@ -42351,19 +42386,19 @@ x_1 = lean_mk_string_unchecked("declModifiersT", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -42373,7 +42408,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19() { _start: { lean_object* x_1; @@ -42381,17 +42416,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_formatter) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21() { _start: { lean_object* x_1; @@ -42399,17 +42434,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Command_declModifiersT_parenthesi return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42419,7 +42454,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -42429,7 +42464,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -42439,7 +42474,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -42449,7 +42484,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -42459,7 +42494,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42469,7 +42504,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -42479,7 +42514,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -42489,7 +42524,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -42499,7 +42534,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -42509,7 +42544,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42519,7 +42554,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -42529,7 +42564,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -42539,7 +42574,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -42549,7 +42584,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -42559,7 +42594,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42569,7 +42604,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -42579,7 +42614,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -42589,7 +42624,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -42599,7 +42634,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -42609,7 +42644,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -42619,7 +42654,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -42629,7 +42664,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -42639,7 +42674,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -42649,7 +42684,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -42659,7 +42694,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48() { _start: { lean_object* x_1; @@ -42667,29 +42702,29 @@ x_1 = lean_mk_string_unchecked("docComment", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Term_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -42699,17 +42734,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -42719,7 +42754,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -42729,15 +42764,15 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4; -x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5; -x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5; +x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) { @@ -42745,8 +42780,8 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); -x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10; -x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10; +x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -42754,8 +42789,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); -x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13; -x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13; +x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -42763,9 +42798,9 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15; -x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17; -x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17; +x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18; x_20 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_5, x_6, x_16); if (lean_obj_tag(x_20) == 0) { @@ -42773,7 +42808,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20; +x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20; x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); if (lean_obj_tag(x_23) == 0) { @@ -42781,7 +42816,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); -x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22; +x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22; x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); if (lean_obj_tag(x_26) == 0) { @@ -42789,10 +42824,10 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); -x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23; x_29 = l_Lean_Parser_Command_declId___closed__2; -x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24; -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25; +x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24; +x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25; x_32 = l_Lean_Parser_registerAlias(x_28, x_29, x_30, x_31, x_6, x_27); if (lean_obj_tag(x_32) == 0) { @@ -42800,7 +42835,7 @@ lean_object* x_33; lean_object* x_34; lean_object* x_35; x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26; +x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26; x_35 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_34, x_33); if (lean_obj_tag(x_35) == 0) { @@ -42808,7 +42843,7 @@ lean_object* x_36; lean_object* x_37; lean_object* x_38; x_36 = lean_ctor_get(x_35, 1); lean_inc(x_36); lean_dec(x_35); -x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27; +x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27; x_38 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_37, x_36); if (lean_obj_tag(x_38) == 0) { @@ -42816,10 +42851,10 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28; x_41 = l_Lean_Parser_Command_declSig___closed__2; -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29; -x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30; +x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29; +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30; x_44 = l_Lean_Parser_registerAlias(x_40, x_41, x_42, x_43, x_6, x_39); if (lean_obj_tag(x_44) == 0) { @@ -42827,7 +42862,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_44, 1); lean_inc(x_45); lean_dec(x_44); -x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31; +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31; x_47 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_40, x_46, x_45); if (lean_obj_tag(x_47) == 0) { @@ -42835,7 +42870,7 @@ lean_object* x_48; lean_object* x_49; lean_object* x_50; x_48 = lean_ctor_get(x_47, 1); lean_inc(x_48); lean_dec(x_47); -x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32; +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32; x_50 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_40, x_49, x_48); if (lean_obj_tag(x_50) == 0) { @@ -42843,10 +42878,10 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33; +x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33; x_53 = l___regBuiltin_Lean_Parser_Command_declVal_docString__1___closed__2; -x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34; -x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35; +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34; +x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35; x_56 = l_Lean_Parser_registerAlias(x_52, x_53, x_54, x_55, x_6, x_51); if (lean_obj_tag(x_56) == 0) { @@ -42854,7 +42889,7 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_56, 1); lean_inc(x_57); lean_dec(x_56); -x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36; +x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_52, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -42862,7 +42897,7 @@ lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_59, 1); lean_inc(x_60); lean_dec(x_59); -x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_52, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -42870,10 +42905,10 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean x_63 = lean_ctor_get(x_62, 1); lean_inc(x_63); lean_dec(x_62); -x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38; +x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38; x_65 = l_Lean_Parser_Command_optDeclSig___closed__2; -x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39; -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40; +x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39; +x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40; x_68 = l_Lean_Parser_registerAlias(x_64, x_65, x_66, x_67, x_6, x_63); if (lean_obj_tag(x_68) == 0) { @@ -42881,7 +42916,7 @@ lean_object* x_69; lean_object* x_70; lean_object* x_71; x_69 = lean_ctor_get(x_68, 1); lean_inc(x_69); lean_dec(x_68); -x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41; +x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41; x_71 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_70, x_69); if (lean_obj_tag(x_71) == 0) { @@ -42889,7 +42924,7 @@ lean_object* x_72; lean_object* x_73; lean_object* x_74; x_72 = lean_ctor_get(x_71, 1); lean_inc(x_72); lean_dec(x_71); -x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42; x_74 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_73, x_72); if (lean_obj_tag(x_74) == 0) { @@ -42897,10 +42932,10 @@ lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean x_75 = lean_ctor_get(x_74, 1); lean_inc(x_75); lean_dec(x_74); -x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43; +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43; x_77 = l___regBuiltin_Lean_Parser_Command_openDecl_docString__1___closed__2; -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44; -x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45; +x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44; +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45; x_80 = l_Lean_Parser_registerAlias(x_76, x_77, x_78, x_79, x_6, x_75); if (lean_obj_tag(x_80) == 0) { @@ -42908,7 +42943,7 @@ lean_object* x_81; lean_object* x_82; lean_object* x_83; x_81 = lean_ctor_get(x_80, 1); lean_inc(x_81); lean_dec(x_80); -x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46; +x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46; x_83 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_76, x_82, x_81); if (lean_obj_tag(x_83) == 0) { @@ -42916,7 +42951,7 @@ lean_object* x_84; lean_object* x_85; lean_object* x_86; x_84 = lean_ctor_get(x_83, 1); lean_inc(x_84); lean_dec(x_83); -x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47; +x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47; x_86 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_76, x_85, x_84); if (lean_obj_tag(x_86) == 0) { @@ -42924,10 +42959,10 @@ lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean x_87 = lean_ctor_get(x_86, 1); lean_inc(x_87); lean_dec(x_86); -x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49; -x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50; -x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51; -x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52; +x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49; +x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50; +x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51; +x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52; x_92 = l_Lean_Parser_registerAlias(x_88, x_89, x_90, x_91, x_6, x_87); if (lean_obj_tag(x_92) == 0) { @@ -42935,7 +42970,7 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; x_93 = lean_ctor_get(x_92, 1); lean_inc(x_93); lean_dec(x_92); -x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53; +x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53; x_95 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_88, x_94, x_93); if (lean_obj_tag(x_95) == 0) { @@ -42943,7 +42978,7 @@ lean_object* x_96; lean_object* x_97; lean_object* x_98; x_96 = lean_ctor_get(x_95, 1); lean_inc(x_96); lean_dec(x_95); -x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54; +x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54; x_98 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_88, x_97, x_96); return x_98; } @@ -46305,6 +46340,8 @@ l_Lean_Parser_Command_whereStructInst___closed__28 = _init_l_Lean_Parser_Command lean_mark_persistent(l_Lean_Parser_Command_whereStructInst___closed__28); l_Lean_Parser_Command_whereStructInst___closed__29 = _init_l_Lean_Parser_Command_whereStructInst___closed__29(); lean_mark_persistent(l_Lean_Parser_Command_whereStructInst___closed__29); +l_Lean_Parser_Command_whereStructInst___closed__30 = _init_l_Lean_Parser_Command_whereStructInst___closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_whereStructInst___closed__30); l_Lean_Parser_Command_whereStructInst = _init_l_Lean_Parser_Command_whereStructInst(); lean_mark_persistent(l_Lean_Parser_Command_whereStructInst); l___regBuiltin_Lean_Parser_Command_declVal_docString__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_declVal_docString__1___closed__1(); @@ -47316,6 +47353,8 @@ l_Lean_Parser_Command_whereStructInst_formatter___closed__9 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_formatter___closed__9); l_Lean_Parser_Command_whereStructInst_formatter___closed__10 = _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__10(); lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_formatter___closed__10); +l_Lean_Parser_Command_whereStructInst_formatter___closed__11 = _init_l_Lean_Parser_Command_whereStructInst_formatter___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_formatter___closed__11); l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter__1___closed__1); l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter__1___closed__2(); @@ -48232,6 +48271,8 @@ l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9 = _init_l_Lean_P lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9); l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10 = _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10(); lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__10); +l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11 = _init_l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__11); l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__1); l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer__1___closed__2(); @@ -52158,122 +52199,122 @@ l_Lean_Parser_Command_eoi___closed__5 = _init_l_Lean_Parser_Command_eoi___closed lean_mark_persistent(l_Lean_Parser_Command_eoi___closed__5); l_Lean_Parser_Command_eoi = _init_l_Lean_Parser_Command_eoi(); lean_mark_persistent(l_Lean_Parser_Command_eoi); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3092_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3095_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Parser_Command_declModifiersF = _init_l_Lean_Parser_Command_declModifiersF(); lean_mark_persistent(l_Lean_Parser_Command_declModifiersF); l_Lean_Parser_Command_declModifiersT = _init_l_Lean_Parser_Command_declModifiersT(); lean_mark_persistent(l_Lean_Parser_Command_declModifiersT); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__50); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__51); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__52); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__53); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116____closed__54); -if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3116_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__50); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__51); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__52); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__53); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119____closed__54); +if (builtin) {res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3119_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Parser_Term_open___closed__1 = _init_l_Lean_Parser_Term_open___closed__1(); diff --git a/stage0/stdlib/Lean/Parser/Term.c b/stage0/stdlib/Lean/Parser/Term.c index 843c51bd6fce..8f943da18637 100644 --- a/stage0/stdlib/Lean/Parser/Term.c +++ b/stage0/stdlib/Lean/Parser/Term.c @@ -47,7 +47,6 @@ static lean_object* l_Lean_Parser_Tactic_quotSeq___closed__4; static lean_object* l_Lean_Parser_Term_matchExprPat_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_let__delayed_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_parenthesizer__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_formatter__1___closed__1; static lean_object* l_Lean_Parser_Term_noErrorIfUnused_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__4; @@ -57,7 +56,6 @@ static lean_object* l_Lean_Parser_Term_let__delayed___closed__5; static lean_object* l_Lean_Parser_Term_funBinder___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_rightact_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_dotIdent___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52; static lean_object* l_Lean_Parser_Term_argument_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_local_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__5; @@ -80,7 +78,6 @@ static lean_object* l_Lean_Parser_Term_binrel___closed__11; static lean_object* l_Lean_Parser_Term_letrec___closed__11; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_docString__1___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_tuple_formatter___closed__11; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__5; @@ -149,12 +146,14 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_formatter__1___closed static lean_object* l_Lean_Parser_Term_borrowed___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_omission; static lean_object* l___regBuiltin_Lean_Parser_Term_omission_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noImplicitLambda; static lean_object* l_Lean_Parser_Term_suffices___closed__1; static lean_object* l_Lean_Parser_Term_inaccessible___closed__9; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__13; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38; static lean_object* l_Lean_Parser_Term_haveI___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_leftact___closed__7; @@ -162,7 +161,6 @@ static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_parenthesizer___close static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__2; lean_object* l_Lean_Parser_charLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45; static lean_object* l_Lean_Parser_Tactic_quot_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noindex_declRange__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroArg_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -172,6 +170,7 @@ static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_depArrow___closed__8; static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__5; static lean_object* l_Lean_Parser_Term_namedArgument___closed__1; static lean_object* l_Lean_Parser_Command_docComment_parenthesizer___closed__1; @@ -242,6 +241,7 @@ static lean_object* l_Lean_Parser_Term_trailing__parser___closed__3; static lean_object* l_Lean_Parser_Term_byTactic_x27_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_letExpr_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_letIdDecl_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43; static lean_object* l_Lean_Parser_Term_structInst___closed__25; static lean_object* l_Lean_Parser_Term_basicFun___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,10 +337,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange__1___cl static lean_object* l_Lean_Parser_Term_omission___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Termination_suffix_parenthesizer__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17; static lean_object* l_Lean_Parser_Term_let__delayed___closed__7; static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_proj_formatter__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20; static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange__1___closed__5; static lean_object* l_Lean_Parser_Term_clear_parenthesizer___closed__7; @@ -434,6 +436,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attrKind_docString__1(l static lean_object* l_Lean_Parser_Term_scoped___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_hole___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32; static lean_object* l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange__1___closed__4; @@ -479,6 +482,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_structInstField_formatter__1 static lean_object* l_Lean_Parser_Term_letRecDecl___closed__5; static lean_object* l_Lean_Parser_Term_typeOf_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_borrowed___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__1; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__13; @@ -701,7 +705,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitBinder___boxed(lean_ob static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_app_formatter__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange__1___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_trailing__parser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter__1___closed__2; @@ -710,7 +713,7 @@ static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__6; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__3; static lean_object* l_Lean_Parser_Term_match_formatter___closed__14; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_assert___closed__1; @@ -731,6 +734,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_borrowed_parenthesizer(lean_object*, static lean_object* l_Lean_Parser_Term_matchExpr___closed__8; static lean_object* l_Lean_Parser_Term_suffices_formatter___closed__6; lean_object* l_Lean_Parser_rawFn(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_trailing__parser___closed__4; @@ -762,6 +766,7 @@ static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sorry_docString__1(lean_object*); static lean_object* l_Lean_Parser_Term_bracketedBinder_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Term_structInstFields___closed__2; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__1; static lean_object* l_Lean_Parser_Term_unreachable___closed__2; static lean_object* l_Lean_Parser_Term_match___closed__10; @@ -853,7 +858,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange__1___closed__ static lean_object* l_Lean_Parser_Tactic_quotSeq_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawCh_parenthesizer___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_scientific_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_docString__1(lean_object*); static lean_object* l_Lean_Parser_Term_rightact___closed__3; static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__5; @@ -888,12 +892,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroDollarArg_parenthesizer(lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_Term_str; static lean_object* l_Lean_Parser_Term_binrel_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__7; +static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__20; static lean_object* l___regBuiltin_Lean_Parser_Term_attributes_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letI_formatter__1___closed__2; static lean_object* l_Lean_Parser_Term_paren___closed__7; static lean_object* l_Lean_Parser_Term_fromTerm___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_declRange__1___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_declRange__1___closed__5; static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_matchExprElseAlt_parenthesizer___closed__1; @@ -937,6 +941,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange__1(lea static lean_object* l_Lean_Parser_Term_pipeProj_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_let__tmp___closed__5; static lean_object* l_Lean_Parser_Term_noErrorIfUnused___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48; lean_object* l_Lean_Parser_interpolatedStr(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange__1___closed__7; static lean_object* l_Lean_Parser_Term_ellipsis_formatter___closed__6; @@ -948,6 +953,7 @@ static lean_object* l_Lean_Parser_Term_syntheticHole_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letExpr___closed__7; static lean_object* l_Lean_Parser_Term_structInst___closed__28; static lean_object* l_Lean_Parser_Term_stateRefT___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33; static lean_object* l_Lean_Parser_Term_fromTerm_formatter___closed__4; static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__5; static lean_object* l_Lean_Parser_Term_unop_parenthesizer___closed__1; @@ -1027,6 +1033,7 @@ static lean_object* l_Lean_Parser_Term_let__delayed_formatter___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binderTactic_parenthesizer__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_leftact_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_185____closed__11; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30; static lean_object* l_Lean_Parser_Term_strictImplicitBinder_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_tuple___closed__14; @@ -1062,7 +1069,6 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__8; static lean_object* l_Lean_Parser_Term_proj_formatter___closed__5; static lean_object* l_Lean_Parser_Term_binop_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40; static lean_object* l_Lean_Parser_Term_completion___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_formatter__1___closed__1; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__2; @@ -1134,6 +1140,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchAltsWhereDecls; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_declRange__1___closed__6; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57; static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_attrInstance___closed__2; @@ -1217,7 +1224,6 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_optIdent___closed__3; static lean_object* l_Lean_Parser_Term_typeAscription___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24; static lean_object* l_Lean_Parser_Term_cdot_formatter___closed__3; static lean_object* l_Lean_Parser_Term_instBinder_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__1; @@ -1276,7 +1282,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_parent static lean_object* l_Lean_Parser_Term_whereDecls_formatter___closed__4; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_noindex_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_forInMacro___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange__1___closed__2; @@ -1305,7 +1310,6 @@ static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__6; static lean_object* l_Lean_Parser_Term_structInst___closed__27; static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__7; static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47; static lean_object* l_Lean_Parser_Term_binop__lazy_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_show_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_typeAscription___closed__16; @@ -1328,6 +1332,7 @@ static lean_object* l_Lean_Parser_Term_declName___closed__1; static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_declRange__1___closed__5; static lean_object* l_Lean_Parser_Term_noImplicitLambda_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42; static lean_object* l_Lean_Parser_Term_leftact_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_forall___closed__12; static lean_object* l_Lean_Parser_Term_letIdDecl_parenthesizer___closed__6; @@ -1338,7 +1343,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_panic_docString__1___closed_ static lean_object* l_Lean_Parser_Term_attrKind___closed__2; static lean_object* l_Lean_Parser_Term_ensureExpectedType___closed__8; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__9; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_formatter__1___closed__1; @@ -1380,7 +1384,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_docString__1(l LEAN_EXPORT lean_object* l_Lean_Parser_Term_sorry; static lean_object* l_Lean_Parser_Term_unsafe_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8; static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__4; static lean_object* l_Lean_Parser_Termination_suffix_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_formatter__1(lean_object*); @@ -1409,6 +1412,7 @@ static lean_object* l_Lean_Parser_Term_fromTerm_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_attrInstance___closed__5; static lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange__1___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40; LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprAlt(lean_object*); static lean_object* l_Lean_Parser_Term_structInstField_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveEqnsDecl; @@ -1466,9 +1470,7 @@ static lean_object* l_Lean_Parser_Command_docComment___closed__10; static lean_object* l_Lean_Parser_Term_declName___closed__8; static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__4; static lean_object* l_Lean_Parser_Term_hole___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15; static lean_object* l_Lean_Parser_Term_optEllipsis_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_funBinder___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter___closed__5; @@ -1547,11 +1549,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_unop_parenthesizer__1___clos lean_object* l_Lean_Parser_ppLine_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Termination_terminationBy___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchExprAlt_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36; lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_app_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_subst___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Termination_terminationBy_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroDollarArg; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_argument_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letIdDecl_formatter__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange__1(lean_object*); @@ -1565,7 +1569,6 @@ static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_haveIdDecl_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Tactic_quot___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_declRange__1___closed__2; static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__5; static lean_object* l_Lean_Parser_Term_local_formatter___closed__1; @@ -1583,6 +1586,7 @@ static lean_object* l_Lean_Parser_Term_attrInstance_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_binderTactic___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields(lean_object*); static lean_object* l_Lean_Parser_Command_docComment___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Term_namedPattern; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange__1___closed__7; @@ -1591,7 +1595,6 @@ static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__6; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_tuple_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_letDecl___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription__1(lean_object*); static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__2; static lean_object* l_Lean_Parser_Termination_terminationBy___closed__4; @@ -1645,7 +1648,6 @@ static lean_object* l_Lean_Parser_Term_namedPattern___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_declRange__1___closed__1; static lean_object* l_Lean_Parser_Term_have___closed__4; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49; static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt___closed__6; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_formatter___closed__5; static lean_object* l_Lean_Parser_Term_attrKind_parenthesizer___closed__2; @@ -1687,6 +1689,7 @@ static lean_object* l_Lean_Parser_Term_panic___closed__8; static lean_object* l_Lean_Parser_Term_have___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeAscription; static lean_object* l_Lean_Parser_Term_binrel___closed__13; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63; LEAN_EXPORT lean_object* l_Lean_Parser_Term_proj; static lean_object* l_Lean_Parser_Term_declName___closed__3; static lean_object* l_Lean_Parser_Term_trueVal___closed__1; @@ -1697,7 +1700,6 @@ static lean_object* l_Lean_Parser_Term_dbgTrace___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_falseVal; static lean_object* l___regBuiltin_Lean_Parser_Term_subst_docString__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__6; @@ -1719,7 +1721,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_subst_docString__1(lean static lean_object* l_Lean_Parser_Termination_terminationBy_x3f___closed__5; static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__3; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3; static lean_object* l_Lean_Parser_Term_prop___closed__6; static lean_object* l_Lean_Parser_Term_haveId_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange__1___closed__4; @@ -1849,6 +1850,7 @@ static lean_object* l_Lean_Parser_Term_structInstField_parenthesizer___closed__5 static lean_object* l_Lean_Parser_Term_motive_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_prop_declRange__1___closed__2; static lean_object* l_Lean_Parser_Term_optIdent___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_unsafe_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_let_docString__1___closed__1; @@ -1904,6 +1906,7 @@ static lean_object* l_Lean_Parser_Term_trailing__parser___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sepBy1IndentSemicolon_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_nofun_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_parenthesizer__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_formatter__1___closed__1; static lean_object* l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__6; @@ -1991,12 +1994,10 @@ static lean_object* l_Lean_Parser_semicolonOrLinebreak_parenthesizer___closed__2 static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__15; static lean_object* l_Lean_Parser_Term_letrec_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13; static lean_object* l_Lean_Parser_Termination_suffix___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_formatter__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_formatter__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_tuple_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange__1___closed__6; static lean_object* l_Lean_Parser_Term_macroDollarArg_formatter___closed__3; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer___closed__1; @@ -2063,7 +2064,6 @@ static lean_object* l_Lean_Parser_Term_binderDefault___closed__6; static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54; static lean_object* l_Lean_Parser_Term_leftact___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optType_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_completion_formatter__1(lean_object*); @@ -2081,6 +2081,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_match___closed__17; @@ -2180,7 +2181,6 @@ static lean_object* l_Lean_Parser_Term_structInstLVal___closed__6; static lean_object* l_Lean_Parser_Term_matchDiscr___closed__5; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__5; static lean_object* l_Lean_Parser_Term_letIdLhs___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__1; static lean_object* l_Lean_Parser_Term_argument___closed__6; static lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter___closed__1; @@ -2204,6 +2204,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_docString__1___ static lean_object* l___regBuiltin_Lean_Parser_Term_basicFun_formatter__1___closed__2; lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_convParser(lean_object*); +static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__20; static lean_object* l_Lean_Parser_Term_letrec_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_withDeclName; static lean_object* l___regBuiltin_Lean_Parser_Term_matchAltsWhereDecls_formatter__1___closed__1; @@ -2331,11 +2332,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_letrec_formatter__1___closed static lean_object* l_Lean_Parser_Term_letExpr_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_forall_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37; static lean_object* l_Lean_Parser_Term_cdot___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange__1___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_noindex___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_parenthesizer__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attrInstance; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___closed__7; static lean_object* l_Lean_Parser_Term_cdot_parenthesizer___closed__4; @@ -2407,6 +2410,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_leftact; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_parenthesizer__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13; static lean_object* l_Lean_Parser_Term_let__delayed___closed__3; static lean_object* l_Lean_Parser_Term_letExpr_formatter___closed__3; static lean_object* l_Lean_Parser_Term_assert___closed__4; @@ -2440,6 +2444,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_macroDollarArg_formatte LEAN_EXPORT lean_object* l_Lean_Parser_Term_generalizingParam; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_omission_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__10; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Term_haveId___closed__5; @@ -2527,6 +2532,7 @@ lean_object* l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(lean_objec LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Termination_terminationBy_docString__1(lean_object*); static lean_object* l_Lean_Parser_Term_matchExpr___closed__11; static lean_object* l_Lean_Parser_Term_typeAscription___closed__15; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60; static lean_object* l_Lean_Parser_Term_leftact___closed__1; static lean_object* l_Lean_Parser_Term_subst___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange__1___closed__4; @@ -2555,6 +2561,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_cdot_docString__1(lean_ static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__14; static lean_object* l_Lean_Parser_Command_docComment_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange__1___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_formatter__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_docString__1___closed__1; @@ -2581,7 +2588,9 @@ static lean_object* l_Lean_Parser_Term_binderTactic_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__3; static lean_object* l_Lean_Parser_Term_byTactic_formatter___closed__5; lean_object* l_Lean_PrettyPrinter_Parenthesizer_parserOfStack_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26; static lean_object* l___regBuiltin_Lean_Parser_Term_haveDecl_docString__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_namedArgument___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_formatter__1___closed__2; @@ -2672,6 +2681,7 @@ static lean_object* l_Lean_Parser_Term_tuple_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__4; @@ -2684,6 +2694,7 @@ static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__7; lean_object* l_Lean_Parser_symbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attrInstance_parenthesizer__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31; static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__1; static lean_object* l_Lean_Parser_Term_matchDiscr_formatter___closed__4; static lean_object* l_Lean_Parser_Term_let_formatter___closed__4; @@ -2713,7 +2724,6 @@ static lean_object* l_Lean_Parser_Term_instBinder_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_letExpr_formatter___closed__7; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__6; static lean_object* l_Lean_Parser_Term_clear_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23; static lean_object* l_Lean_Parser_Term_forInMacro___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Termination_suffix; LEAN_EXPORT lean_object* l_Lean_Parser_Term_scoped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2737,6 +2747,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_trailing__parser; static lean_object* l___regBuiltin_Lean_Parser_Term_fun_declRange__1___closed__4; static lean_object* l_Lean_Parser_Term_basicFun___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29; static lean_object* l_Lean_Parser_Term_letrec___closed__9; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar___closed__4; lean_object* l_Lean_Parser_nameLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2757,7 +2768,7 @@ static lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer___closed__ static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_formatter__1___closed__2; static lean_object* l_Lean_Parser_Term_binderDefault___closed__5; static lean_object* l_Lean_Parser_Term_sorry_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34; static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__4; static lean_object* l_Lean_Parser_Term_show___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_generalizingParam_parenthesizer__1___closed__2; @@ -2816,10 +2827,8 @@ static lean_object* l_Lean_Parser_Term_hole_formatter___closed__1; static lean_object* l_Lean_Parser_Term_unsafe___closed__4; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_pipeProj___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369_(lean_object*); static lean_object* l_Lean_Parser_Term_have___closed__3; static lean_object* l_Lean_Parser_Term_type___closed__17; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_explicitUniv___closed__5; static lean_object* l_Lean_Parser_Term_byTactic___closed__8; @@ -2848,7 +2857,6 @@ static lean_object* l_Lean_Parser_Term_sorry___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sort_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl_declRange__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_declRange__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq__1(lean_object*); static lean_object* l_Lean_Parser_Term_forall___closed__3; @@ -2880,6 +2888,7 @@ static lean_object* l_Lean_Parser_Term_letIdBinder___closed__3; static lean_object* l_Lean_Parser_Term_let_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_letDecl___closed__2; static lean_object* l_Lean_Parser_Term_basicFun_formatter___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54; static lean_object* l_Lean_Parser_Term_letRecDecls_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_sorry_formatter___closed__2; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_formatter___closed__5; @@ -2888,7 +2897,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_darrow; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optSemicolon(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_formatter__1___closed__2; static lean_object* l_Lean_Parser_Term_namedPattern___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__3; static lean_object* l_Lean_Parser_Term_show___closed__4; static lean_object* l_Lean_Parser_Term_ensureTypeOf_parenthesizer___closed__5; @@ -3007,6 +3015,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_docStri LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_generalizingParam___closed__12; static lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12; static lean_object* l_Lean_Parser_Term_letrec_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Term_type_parenthesizer___closed__7; @@ -3145,6 +3154,7 @@ static lean_object* l_Lean_Parser_Term_clear_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_implicitBinder_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_quotSeq_formatter___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_matchExpr__1(lean_object*); static lean_object* l_Lean_Parser_Term_optEllipsis___closed__7; static lean_object* l_Lean_Parser_Term_letDecl___closed__4; @@ -3347,6 +3357,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_haveI_declRange__1___closed_ static lean_object* l_Lean_Parser_Term_optIdent_formatter___closed__1; static lean_object* l_Lean_Parser_Termination_terminationBy___closed__3; static lean_object* l_Lean_Parser_Term_haveI_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_quot_formatter___closed__5; @@ -3369,6 +3380,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq1Indented_formatt static lean_object* l_Lean_Parser_Term_letRecDecl___closed__3; static lean_object* l_Lean_Parser_Term_unop___closed__5; static lean_object* l_Lean_Parser_Term_structInstLVal_parenthesizer___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noindex_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_showTermElabImpl___closed__8; @@ -3397,8 +3409,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_char__1(lean_object*); static lean_object* l_Lean_Parser_Term_noErrorIfUnused_formatter___closed__4; static lean_object* l_Lean_Parser_Term_letPatDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Term_syntheticHole___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Term_motive_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Term_matchExpr_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange__1___closed__7; @@ -3426,7 +3436,6 @@ static lean_object* l_Lean_Parser_Term_letDecl_parenthesizer___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_showTermElabImpl__1(lean_object*); static lean_object* l_Lean_Parser_Term_panic___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_depArrow_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__11; static lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange__1___closed__2; @@ -3515,12 +3524,10 @@ static lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_letDecl_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__3; static lean_object* l_Lean_Parser_Term_byTactic___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__8; static lean_object* l_Lean_Parser_Term_type_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__fun__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; static lean_object* l_Lean_Parser_Term_basicFun___closed__13; static lean_object* l_Lean_Parser_Term_letPatDecl___closed__10; @@ -3611,7 +3618,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_namedArgument_formatter__1__ static lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_structInstLVal_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_declName_declRange__1___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_instBinder___closed__5; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__1; @@ -3670,7 +3676,6 @@ static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__6; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8; static lean_object* l_Lean_Parser_Term_type_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_cdot_docString__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31; static lean_object* l_Lean_Parser_Term_whereDecls___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_rightact_formatter__1___closed__1; static lean_object* l_Lean_Parser_Term_structInstField___closed__9; @@ -3680,13 +3685,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_omission_formatter__1___clos static lean_object* l_Lean_Parser_Term_scoped___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_formatter___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_app_formatter__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder___closed__7; static lean_object* l_Lean_Parser_Term_rightact___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_unsafe___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_fun__1(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__21; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda__1(lean_object*); static lean_object* l_Lean_Parser_Term_attrKind___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange__1___closed__6; @@ -3694,7 +3699,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_proj__1(lean_object*); static lean_object* l_Lean_Parser_Termination_decreasingBy_formatter___closed__6; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_assert_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_declRange__1___closed__4; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__7; static lean_object* l_Lean_Parser_Tactic_quotSeq_parenthesizer___closed__6; @@ -3755,6 +3759,7 @@ static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__2; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange__1(lean_object*); static lean_object* l_Lean_Parser_Term_explicitBinder___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange__1___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_formatter__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicitBinder___boxed(lean_object*); @@ -3849,7 +3854,6 @@ static lean_object* l_Lean_Parser_Term_fromTerm_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_stateRefT___closed__10; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46; static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_subst___closed__4; @@ -3902,7 +3906,6 @@ static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_let__delayed_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_matchExpr_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_unsafe_parenthesizer__1___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_panic_docString__1(lean_object*); static lean_object* l_Lean_Parser_Term_letIdDecl___closed__2; static lean_object* l_Lean_Parser_Term_binop__lazy_formatter___closed__3; @@ -3973,6 +3976,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_declName_formatter(lean_object*, lea static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_optEllipsis_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_formatter__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24; static lean_object* l_Lean_Parser_Term_binderDefault_formatter___closed__1; static lean_object* l_Lean_Parser_Term_typeSpec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_explicit_parenthesizer___closed__1; @@ -4024,6 +4029,7 @@ static lean_object* l_Lean_Parser_Term_explicit_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unreachable_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_parenthesizer__1___closed__2; static lean_object* l_Lean_Parser_Term_rightact___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__7; static lean_object* l_Lean_Parser_Term_ensureTypeOf___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange__1(lean_object*); @@ -4054,7 +4060,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_formatter static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__7; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__2; static lean_object* l_Lean_Parser_Term_forInMacro_x27_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36; static lean_object* l_Lean_Parser_Term_namedPattern___closed__1; static lean_object* l_Lean_Parser_Term_suffices_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__4; @@ -4063,6 +4068,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_parenth static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_letDecl_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange__1___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46; LEAN_EXPORT lean_object* l_Lean_Parser_Term_rightact_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_type___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_parenthesizer__1(lean_object*); @@ -4132,7 +4138,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange__1___cl static lean_object* l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange__1___closed__7; static lean_object* l_Lean_Parser_Term_motive___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_type_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_omission___closed__8; @@ -4219,7 +4224,6 @@ static lean_object* l_Lean_Parser_Term_binderDefault___closed__1; static lean_object* l_Lean_Parser_Termination_suffix_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Term_ident_formatter___closed__2; @@ -4253,9 +4257,11 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty__1( static lean_object* l___regBuiltin_Lean_Parser_Term_matchExpr_parenthesizer__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binrel_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_unicodeSymbol_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_docString__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_anonymousCtor_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47; static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_parenthesizer__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange__1___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let_parenthesizer__1(lean_object*); @@ -4321,6 +4327,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_dotIdent_formatter(lean_object*, lea LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letrec_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_suffices___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange__1___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41; static lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__7; static lean_object* l_Lean_Parser_Term_sort_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_matchAltsWhereDecls_parenthesizer__1___closed__1; @@ -4334,7 +4341,6 @@ static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed_ static lean_object* l_Lean_Parser_Term_haveId_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_haveDecl_formatter___closed__3; static lean_object* l_Lean_Parser_Term_sort_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33; static lean_object* l_Lean_Parser_Term_letPatDecl___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letRecDecl; static lean_object* l___regBuiltin_Lean_Parser_Term_fun_declRange__1___closed__6; @@ -4466,6 +4472,7 @@ static lean_object* l_Lean_Parser_Term_falseVal_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_nomatch_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Command_docComment___closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_dotIdent; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_formatter__1___closed__1; @@ -4541,6 +4548,7 @@ static lean_object* l_Lean_Parser_Term_clear_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_prop_parenthesizer__1(lean_object*); static lean_object* l_Lean_Parser_Term_match___closed__12; static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__6; +static lean_object* l_Lean_Parser_Term_structInstFields___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_have_formatter___closed__6; static lean_object* l_Lean_Parser_Term_forall___closed__2; @@ -4609,7 +4617,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange__1(lean_o static lean_object* l_Lean_Parser_Term_strictImplicitBinder_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_show_declRange__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_suffices__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28; static lean_object* l_Lean_Parser_Termination_suffix_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_formatter__1___closed__1; static lean_object* l_Lean_Parser_Term_letDecl___closed__1; @@ -4684,6 +4691,7 @@ static lean_object* l_Lean_Parser_Term_rightact_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_noindex_declRange__1___closed__7; static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange__1___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_proj_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_inaccessible_formatter___closed__1; @@ -4824,6 +4832,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_cdot_parenthesizer__1(l static lean_object* l_Lean_Parser_Term_type_formatter___closed__3; static lean_object* l_Lean_Parser_Term_syntheticHole___closed__2; static lean_object* l_Lean_Parser_Term_proj___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45; lean_object* l_Lean_Parser_sepBy1_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_ensureTypeOf_formatter___closed__2; lean_object* l_Lean_Parser_lookahead(lean_object*); @@ -4838,7 +4847,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_letExpr_parenthesizer__1___c static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder___closed__4; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_parenthesizer__1___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55; static lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_declRange__1___closed__7; static lean_object* l_Lean_Parser_Term_sort___closed__7; static lean_object* l_Lean_Parser_Term_paren___closed__6; @@ -4996,6 +5004,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_dbgTrace; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__7; static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__5; static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28; static lean_object* l_Lean_Parser_Term_borrowed_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_falseVal_formatter___closed__1; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__7; @@ -5017,6 +5026,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sepBy1IndentSemicolon_formatter(le static lean_object* l_Lean_Parser_Term_leading__parser___closed__9; static lean_object* l_Lean_Parser_Term_dynamicQuot_formatter___closed__7; static lean_object* l_Lean_Parser_Term_explicit___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_docString__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_haveIdDecl_formatter__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_fun_formatter__1___closed__1; @@ -5025,6 +5035,8 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_parenthesiz LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_have_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_declRange__1___closed__2; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_show_formatter___closed__1; static lean_object* l_Lean_Parser_Term_panic___closed__9; @@ -5071,6 +5083,7 @@ static lean_object* l_Lean_Parser_Term_unsafe___closed__6; static lean_object* l_Lean_Parser_Term_basicFun___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticSeq_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_formatter__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52; static lean_object* l_Lean_Parser_Term_ensureExpectedType_formatter___closed__4; extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__2; @@ -5094,6 +5107,7 @@ static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__12; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_assert_docString__1___closed__1; static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11; static lean_object* l_Lean_Parser_Term_tuple_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_basicFun; static lean_object* l_Lean_Parser_Term_attrKind_parenthesizer___closed__1; @@ -5122,12 +5136,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_byTactic; static lean_object* l_Lean_Parser_Term_letIdLhs___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_docComment_formatter__1___closed__2; static lean_object* l_Lean_Parser_Term_matchDiscr_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38; static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_formatter__1___closed__2; static lean_object* l_Lean_Parser_Term_nomatch___closed__5; static lean_object* l_Lean_Parser_Term_show_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstLVal_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_haveI_declRange__1(lean_object*); @@ -5139,6 +5151,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_letExpr_declRange__1___close static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_parenthesizer__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__3; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_parenthesizer__1___closed__1; @@ -5194,6 +5207,7 @@ static lean_object* l_Lean_Parser_Term_let__fun___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_parenthesizer__1___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16; static lean_object* l_Lean_Parser_Term_matchAlt___closed__8; static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__7; static lean_object* l_Lean_Parser_Term_basicFun_formatter___closed__9; @@ -5259,7 +5273,6 @@ static lean_object* l_Lean_Parser_Term_nomatch___closed__8; static lean_object* l_Lean_Parser_Term_matchExprPat_formatter___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_declRange__1(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange__1___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noErrorIfUnused_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_haveId___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon(lean_object*); @@ -5319,6 +5332,7 @@ static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__2; static lean_object* l_Lean_Parser_Term_noErrorIfUnused_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_forInMacro; lean_object* l_Lean_PrettyPrinter_Formatter_categoryParser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61; static lean_object* l_Lean_Parser_Term_haveDecl___closed__3; static lean_object* l_Lean_Parser_Term_funImplicitBinder___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_docString__1(lean_object*); @@ -5327,6 +5341,7 @@ static lean_object* l_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_let__tmp_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__14; static lean_object* l_Lean_Parser_Term_matchAlt___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_declRange__1___closed__2; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer__1___closed__1; @@ -5382,12 +5397,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_fromTerm_parenthesizer__1___ LEAN_EXPORT lean_object* l_Lean_Parser_Term_forInMacro_x27_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_have_formatter___closed__1; lean_object* l_Lean_Parser_ppDedentIfGrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25; static lean_object* l_Lean_Parser_semicolonOrLinebreak_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_let__tmp_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_explicitUniv___closed__14; static lean_object* l_Lean_Parser_Term_cdot_formatter___closed__2; static lean_object* l_Lean_Parser_Term_fun___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letRecDecls_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_stateRefT; static lean_object* l_Lean_Parser_Term_letRecDecl_parenthesizer___closed__4; @@ -5444,6 +5459,7 @@ static lean_object* l_Lean_Parser_Term_binrel___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_parenthesizer__1___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_formatter__1___closed__1; static lean_object* l_Lean_Parser_Termination_terminationBy_x3f_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_parenthesizer__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicitBinder(uint8_t); static lean_object* l_Lean_Parser_Term_leftact_formatter___closed__2; @@ -5510,7 +5526,6 @@ static lean_object* l_Lean_Parser_Term_namedArgument___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attrKind_formatter__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35; static lean_object* l_Lean_Parser_Term_letExpr___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_ellipsis_formatter__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_dynamicQuot; @@ -5558,7 +5573,6 @@ static lean_object* l_Lean_Parser_Term_forInMacro_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_forInMacro_x27_formatter___closed__4; static lean_object* l_Lean_Parser_Term_inaccessible___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_generalizingParam_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32; static lean_object* l_Lean_Parser_Term_match_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_formatter__1___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__2; @@ -5619,7 +5633,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange__1(lea static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_matchExpr_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_attrInstance___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1; static lean_object* l_Lean_Parser_Term_typeSpec___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_declRange__1___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_parenthesizer__1___closed__2; @@ -5645,6 +5658,7 @@ static lean_object* l_Lean_Parser_Term_motive_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_declRange__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_inaccessible__1(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58; static lean_object* l_Lean_Parser_Term_explicit_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange__1___closed__6; static lean_object* l_Lean_Parser_Term_letrec___closed__13; @@ -5672,6 +5686,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_argument___closed__5; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56; static lean_object* l_Lean_Parser_Term_depArrow_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_declRange__1___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_declRange__1___closed__5; @@ -5794,7 +5809,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange__1__ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_docComment_formatter__1(lean_object*); static lean_object* l_Lean_Parser_Term_paren___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_sepByIndentSemicolon_docString__1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_formatter__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_assert_docString__1(lean_object*); static lean_object* l_Lean_Parser_Term_trailing__parser_parenthesizer___closed__4; @@ -5830,6 +5844,7 @@ static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__9; static lean_object* l_Lean_Parser_Term_clear___closed__4; static lean_object* l_Lean_Parser_Term_binderType___closed__4; lean_object* l_Lean_Parser_termParser(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_pushNone_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_clear_formatter___closed__5; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2; @@ -5841,6 +5856,7 @@ static lean_object* l_Lean_Parser_Termination_terminationBy___closed__11; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_matchAlts_formatter___closed__1; static lean_object* l_Lean_Parser_Term_matchExprAlts___closed__6; +static lean_object* l_Lean_Parser_Term_structInst___closed__31; static lean_object* l_Lean_Parser_Term_binderDefault___closed__4; static lean_object* _init_l_Lean_Parser_Command_commentBody___elambda__1___closed__1() { _start: @@ -18899,6 +18915,35 @@ x_1 = l_Lean_Parser_Term_optEllipsis___closed__9; return x_1; } } +static lean_object* _init_l_Lean_Parser_Term_structInstFields___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_structInstFields___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Command_docComment___closed__1; +x_2 = l_Lean_Parser_Command_docComment___closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__14; +x_4 = l_Lean_Parser_Term_structInstFields___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Parser_Term_structInstFields___closed__2; +x_3 = l_Lean_Parser_node(x_2, x_1); +return x_3; +} +} static lean_object* _init_l_Lean_Parser_Term_structInst___closed__1() { _start: { @@ -19070,6 +19115,15 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_structInst___closed__18() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_structInst___closed__17; +x_2 = l_Lean_Parser_Term_structInstFields(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__19() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optIdent___closed__2; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; @@ -19077,56 +19131,56 @@ x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__19() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst___closed__18; +x_1 = l_Lean_Parser_Term_structInst___closed__19; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__20() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optEllipsis; -x_2 = l_Lean_Parser_Term_structInst___closed__19; +x_2 = l_Lean_Parser_Term_structInst___closed__20; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__21() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst___closed__17; -x_2 = l_Lean_Parser_Term_structInst___closed__20; +x_1 = l_Lean_Parser_Term_structInst___closed__18; +x_2 = l_Lean_Parser_Term_structInst___closed__21; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__22() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst___closed__11; -x_2 = l_Lean_Parser_Term_structInst___closed__21; +x_2 = l_Lean_Parser_Term_structInst___closed__22; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__23() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__15; -x_2 = l_Lean_Parser_Term_structInst___closed__22; +x_2 = l_Lean_Parser_Term_structInst___closed__23; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__24() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__25() { _start: { lean_object* x_1; @@ -19134,62 +19188,62 @@ x_1 = lean_mk_string_unchecked(" }", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__25() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst___closed__24; +x_1 = l_Lean_Parser_Term_structInst___closed__25; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__26() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst___closed__23; -x_2 = l_Lean_Parser_Term_structInst___closed__25; +x_1 = l_Lean_Parser_Term_structInst___closed__24; +x_2 = l_Lean_Parser_Term_structInst___closed__26; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__27() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst___closed__5; -x_2 = l_Lean_Parser_Term_structInst___closed__26; +x_2 = l_Lean_Parser_Term_structInst___closed__27; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__28() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_structInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_structInst___closed__27; +x_3 = l_Lean_Parser_Term_structInst___closed__28; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__29() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst___closed__3; -x_2 = l_Lean_Parser_Term_structInst___closed__28; +x_2 = l_Lean_Parser_Term_structInst___closed__29; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst___closed__30() { +static lean_object* _init_l_Lean_Parser_Term_structInst___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst___closed__2; -x_2 = l_Lean_Parser_Term_structInst___closed__29; +x_2 = l_Lean_Parser_Term_structInst___closed__30; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -19198,7 +19252,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_structInst___closed__30; +x_1 = l_Lean_Parser_Term_structInst___closed__31; return x_1; } } @@ -19335,6 +19389,15 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields_formatter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_Parser_Term_structInstFields___closed__2; +x_8 = l_Lean_PrettyPrinter_Formatter_node_formatter(x_7, x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} static lean_object* _init_l_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed__1() { _start: { @@ -20121,6 +20184,16 @@ return x_4; static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_structInst_formatter___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; @@ -20130,103 +20203,103 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst_formatter___closed__10; +x_1 = l_Lean_Parser_Term_structInst_formatter___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_optEllipsis_formatter__1___closed__2; -x_2 = l_Lean_Parser_Term_structInst_formatter___closed__11; +x_2 = l_Lean_Parser_Term_structInst_formatter___closed__12; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_formatter___closed__9; -x_2 = l_Lean_Parser_Term_structInst_formatter___closed__12; +x_1 = l_Lean_Parser_Term_structInst_formatter___closed__10; +x_2 = l_Lean_Parser_Term_structInst_formatter___closed__13; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst_formatter___closed__7; -x_2 = l_Lean_Parser_Term_structInst_formatter___closed__13; +x_2 = l_Lean_Parser_Term_structInst_formatter___closed__14; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst_formatter___closed__14; +x_1 = l_Lean_Parser_Term_structInst_formatter___closed__15; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst___closed__24; +x_1 = l_Lean_Parser_Term_structInst___closed__25; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__17() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_formatter___closed__15; -x_2 = l_Lean_Parser_Term_structInst_formatter___closed__16; +x_1 = l_Lean_Parser_Term_structInst_formatter___closed__16; +x_2 = l_Lean_Parser_Term_structInst_formatter___closed__17; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__18() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst_formatter___closed__2; -x_2 = l_Lean_Parser_Term_structInst_formatter___closed__17; +x_2 = l_Lean_Parser_Term_structInst_formatter___closed__18; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__19() { +static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_structInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_structInst_formatter___closed__18; +x_3 = l_Lean_Parser_Term_structInst_formatter___closed__19; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20239,7 +20312,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInst_formatter(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_structInst_formatter___closed__1; -x_7 = l_Lean_Parser_Term_structInst_formatter___closed__19; +x_7 = l_Lean_Parser_Term_structInst_formatter___closed__20; x_8 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -20277,6 +20350,15 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } +LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstFields_parenthesizer(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_Parser_Term_structInstFields___closed__2; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_node_parenthesizer(x_7, x_1, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} static lean_object* _init_l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__1() { _start: { @@ -21045,6 +21127,16 @@ return x_6; static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; @@ -21054,103 +21146,103 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__10; +x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_optEllipsis_parenthesizer__1___closed__2; -x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__12; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__9; -x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__12; +x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__13; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__7; -x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__13; +x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__14; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__14; +x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__15; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInst___closed__24; +x_1 = l_Lean_Parser_Term_structInst___closed__25; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__17() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__15; -x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__16; +x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__16; +x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__17; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__18() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__17; +x_2 = l_Lean_Parser_Term_structInst_parenthesizer___closed__18; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__19() { +static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_structInst___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_structInst_parenthesizer___closed__18; +x_3 = l_Lean_Parser_Term_structInst_parenthesizer___closed__19; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -21163,7 +21255,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInst_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_structInst_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_structInst_parenthesizer___closed__19; +x_7 = l_Lean_Parser_Term_structInst_parenthesizer___closed__20; x_8 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; } @@ -23421,7 +23513,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_formatter(uint8_t x_1, le if (x_1 == 0) { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Parser_Term_structInst_formatter___closed__10; +x_7 = l_Lean_Parser_Term_structInst_formatter___closed__11; x_8 = l_Lean_Parser_optional_formatter(x_7, x_2, x_3, x_4, x_5, x_6); return x_8; } @@ -23429,7 +23521,7 @@ else { lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = l_Lean_Parser_Term_binderType___closed__3; -x_10 = l_Lean_Parser_Term_structInst_formatter___closed__10; +x_10 = l_Lean_Parser_Term_structInst_formatter___closed__11; x_11 = l_Lean_PrettyPrinter_Formatter_node_formatter(x_9, x_10, x_2, x_3, x_4, x_5, x_6); return x_11; } @@ -25622,7 +25714,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeSpec_formatter___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_typeSpec___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_structInst_formatter___closed__10; +x_3 = l_Lean_Parser_Term_structInst_formatter___closed__11; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -69052,7 +69144,7 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69062,7 +69154,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -69072,7 +69164,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -69082,7 +69174,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -69092,7 +69184,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -69102,7 +69194,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69112,7 +69204,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -69122,7 +69214,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -69132,7 +69224,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -69142,7 +69234,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10() { _start: { lean_object* x_1; lean_object* x_2; @@ -69152,7 +69244,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69162,7 +69254,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -69172,7 +69264,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -69182,7 +69274,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -69192,7 +69284,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -69202,7 +69294,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69212,7 +69304,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -69222,7 +69314,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -69232,7 +69324,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -69242,7 +69334,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -69252,7 +69344,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69262,7 +69354,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -69272,7 +69364,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -69282,7 +69374,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -69292,7 +69384,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -69302,7 +69394,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69312,7 +69404,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -69322,7 +69414,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -69332,7 +69424,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -69342,7 +69434,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -69352,7 +69444,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69362,7 +69454,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -69372,7 +69464,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -69382,7 +69474,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -69392,7 +69484,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -69402,7 +69494,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69412,7 +69504,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -69422,7 +69514,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -69432,7 +69524,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -69442,7 +69534,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -69452,7 +69544,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -69462,7 +69554,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -69472,7 +69564,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -69482,7 +69574,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -69492,7 +69584,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -69502,7 +69594,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46() { _start: { lean_object* x_1; @@ -69510,29 +69602,29 @@ x_1 = lean_mk_string_unchecked("optSemicolon", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_Command_docComment___closed__1; x_2 = l_Lean_Parser_Command_docComment___closed__2; x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_71____closed__14; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49() { _start: { lean_object* x_1; @@ -69540,27 +69632,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52() { _start: { lean_object* x_1; @@ -69568,17 +69660,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon_formatter), 6, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54() { _start: { lean_object* x_1; @@ -69586,24 +69678,98 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Term_structInstFields___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_structInstFields___closed__2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_formatter), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_structInstFields_parenthesizer), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1; x_3 = l___regBuiltin_Lean_Parser_Term_letDecl_docString__1___closed__2; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3; x_6 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_185____closed__6; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) @@ -69613,7 +69779,7 @@ x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); x_9 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_185____closed__9; -x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4; +x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -69622,7 +69788,7 @@ x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); x_13 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_185____closed__12; -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -69630,10 +69796,10 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6; x_18 = l___regBuiltin_Lean_Parser_Term_haveDecl_docString__1___closed__2; -x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7; -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7; +x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8; x_21 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_20, x_6, x_16); if (lean_obj_tag(x_21) == 0) { @@ -69641,7 +69807,7 @@ lean_object* x_22; lean_object* x_23; lean_object* x_24; x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); lean_dec(x_21); -x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9; +x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9; x_24 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_23, x_22); if (lean_obj_tag(x_24) == 0) { @@ -69649,7 +69815,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; x_25 = lean_ctor_get(x_24, 1); lean_inc(x_25); lean_dec(x_24); -x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10; +x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10; x_27 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_26, x_25); if (lean_obj_tag(x_27) == 0) { @@ -69657,10 +69823,10 @@ lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean x_28 = lean_ctor_get(x_27, 1); lean_inc(x_28); lean_dec(x_27); -x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11; x_30 = l___regBuiltin_Lean_Parser_Term_sufficesDecl_docString__1___closed__2; -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12; -x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13; +x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12; +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13; x_33 = l_Lean_Parser_registerAlias(x_29, x_30, x_31, x_32, x_6, x_28); if (lean_obj_tag(x_33) == 0) { @@ -69668,7 +69834,7 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_33, 1); lean_inc(x_34); lean_dec(x_33); -x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14; +x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14; x_36 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_29, x_35, x_34); if (lean_obj_tag(x_36) == 0) { @@ -69676,7 +69842,7 @@ lean_object* x_37; lean_object* x_38; lean_object* x_39; x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); -x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15; +x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15; x_39 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_29, x_38, x_37); if (lean_obj_tag(x_39) == 0) { @@ -69684,10 +69850,10 @@ lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); lean_dec(x_39); -x_41 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16; +x_41 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16; x_42 = l_Lean_Parser_Term_letRecDecls___closed__2; -x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17; -x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18; +x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17; +x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18; x_45 = l_Lean_Parser_registerAlias(x_41, x_42, x_43, x_44, x_6, x_40); if (lean_obj_tag(x_45) == 0) { @@ -69695,7 +69861,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); lean_dec(x_45); -x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19; +x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19; x_48 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_41, x_47, x_46); if (lean_obj_tag(x_48) == 0) { @@ -69703,7 +69869,7 @@ lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20; x_51 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_41, x_50, x_49); if (lean_obj_tag(x_51) == 0) { @@ -69711,10 +69877,10 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); lean_dec(x_51); -x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21; +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21; x_54 = l_Lean_Parser_Term_hole___closed__2; -x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22; -x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23; +x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22; +x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23; x_57 = l_Lean_Parser_registerAlias(x_53, x_54, x_55, x_56, x_6, x_52); if (lean_obj_tag(x_57) == 0) { @@ -69722,7 +69888,7 @@ lean_object* x_58; lean_object* x_59; lean_object* x_60; x_58 = lean_ctor_get(x_57, 1); lean_inc(x_58); lean_dec(x_57); -x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24; +x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24; x_60 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_53, x_59, x_58); if (lean_obj_tag(x_60) == 0) { @@ -69730,7 +69896,7 @@ lean_object* x_61; lean_object* x_62; lean_object* x_63; x_61 = lean_ctor_get(x_60, 1); lean_inc(x_61); lean_dec(x_60); -x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25; +x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25; x_63 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_53, x_62, x_61); if (lean_obj_tag(x_63) == 0) { @@ -69738,10 +69904,10 @@ lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean x_64 = lean_ctor_get(x_63, 1); lean_inc(x_64); lean_dec(x_63); -x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26; +x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26; x_66 = l_Lean_Parser_Term_syntheticHole___closed__2; -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27; -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28; +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28; x_69 = l_Lean_Parser_registerAlias(x_65, x_66, x_67, x_68, x_6, x_64); if (lean_obj_tag(x_69) == 0) { @@ -69749,7 +69915,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; x_70 = lean_ctor_get(x_69, 1); lean_inc(x_70); lean_dec(x_69); -x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29; +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29; x_72 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_65, x_71, x_70); if (lean_obj_tag(x_72) == 0) { @@ -69757,7 +69923,7 @@ lean_object* x_73; lean_object* x_74; lean_object* x_75; x_73 = lean_ctor_get(x_72, 1); lean_inc(x_73); lean_dec(x_72); -x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30; +x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30; x_75 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_65, x_74, x_73); if (lean_obj_tag(x_75) == 0) { @@ -69765,10 +69931,10 @@ lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean x_76 = lean_ctor_get(x_75, 1); lean_inc(x_76); lean_dec(x_75); -x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31; +x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31; x_78 = l___regBuiltin_Lean_Parser_Term_matchDiscr_docString__1___closed__2; -x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32; -x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33; +x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32; +x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33; x_81 = l_Lean_Parser_registerAlias(x_77, x_78, x_79, x_80, x_6, x_76); if (lean_obj_tag(x_81) == 0) { @@ -69776,7 +69942,7 @@ lean_object* x_82; lean_object* x_83; lean_object* x_84; x_82 = lean_ctor_get(x_81, 1); lean_inc(x_82); lean_dec(x_81); -x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34; x_84 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_77, x_83, x_82); if (lean_obj_tag(x_84) == 0) { @@ -69784,7 +69950,7 @@ lean_object* x_85; lean_object* x_86; lean_object* x_87; x_85 = lean_ctor_get(x_84, 1); lean_inc(x_85); lean_dec(x_84); -x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35; +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35; x_87 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_77, x_86, x_85); if (lean_obj_tag(x_87) == 0) { @@ -69792,10 +69958,10 @@ lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean x_88 = lean_ctor_get(x_87, 1); lean_inc(x_88); lean_dec(x_87); -x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36; +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36; x_90 = l___regBuiltin_Lean_Parser_Term_bracketedBinder_docString__1___closed__2; -x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37; -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38; +x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37; +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38; x_93 = l_Lean_Parser_registerAlias(x_89, x_90, x_91, x_92, x_6, x_88); if (lean_obj_tag(x_93) == 0) { @@ -69803,7 +69969,7 @@ lean_object* x_94; lean_object* x_95; lean_object* x_96; x_94 = lean_ctor_get(x_93, 1); lean_inc(x_94); lean_dec(x_93); -x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39; +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39; x_96 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_89, x_95, x_94); if (lean_obj_tag(x_96) == 0) { @@ -69811,7 +69977,7 @@ lean_object* x_97; lean_object* x_98; lean_object* x_99; x_97 = lean_ctor_get(x_96, 1); lean_inc(x_97); lean_dec(x_96); -x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40; +x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40; x_99 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_89, x_98, x_97); if (lean_obj_tag(x_99) == 0) { @@ -69819,10 +69985,10 @@ lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; x_100 = lean_ctor_get(x_99, 1); lean_inc(x_100); lean_dec(x_99); -x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41; +x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41; x_102 = l___regBuiltin_Lean_Parser_Term_attrKind_docString__1___closed__2; -x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42; -x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43; +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42; +x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43; x_105 = l_Lean_Parser_registerAlias(x_101, x_102, x_103, x_104, x_6, x_100); if (lean_obj_tag(x_105) == 0) { @@ -69830,7 +69996,7 @@ lean_object* x_106; lean_object* x_107; lean_object* x_108; x_106 = lean_ctor_get(x_105, 1); lean_inc(x_106); lean_dec(x_105); -x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44; +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44; x_108 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_101, x_107, x_106); if (lean_obj_tag(x_108) == 0) { @@ -69838,7 +70004,7 @@ lean_object* x_109; lean_object* x_110; lean_object* x_111; x_109 = lean_ctor_get(x_108, 1); lean_inc(x_109); lean_dec(x_108); -x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45; +x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45; x_111 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_101, x_110, x_109); if (lean_obj_tag(x_111) == 0) { @@ -69846,10 +70012,10 @@ lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; x_112 = lean_ctor_get(x_111, 1); lean_inc(x_112); lean_dec(x_111); -x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47; -x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48; -x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50; -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51; +x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47; +x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48; +x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50; +x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51; x_117 = l_Lean_Parser_registerAlias(x_113, x_114, x_115, x_116, x_6, x_112); if (lean_obj_tag(x_117) == 0) { @@ -69857,7 +70023,7 @@ lean_object* x_118; lean_object* x_119; lean_object* x_120; x_118 = lean_ctor_get(x_117, 1); lean_inc(x_118); lean_dec(x_117); -x_119 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53; +x_119 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53; x_120 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_113, x_119, x_118); if (lean_obj_tag(x_120) == 0) { @@ -69865,95 +70031,53 @@ lean_object* x_121; lean_object* x_122; lean_object* x_123; x_121 = lean_ctor_get(x_120, 1); lean_inc(x_121); lean_dec(x_120); -x_122 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55; +x_122 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55; x_123 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_113, x_122, x_121); -return x_123; -} -else -{ -uint8_t x_124; -x_124 = !lean_is_exclusive(x_120); -if (x_124 == 0) -{ -return x_120; -} -else -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = lean_ctor_get(x_120, 0); -x_126 = lean_ctor_get(x_120, 1); -lean_inc(x_126); -lean_inc(x_125); -lean_dec(x_120); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -return x_127; -} -} -} -else -{ -uint8_t x_128; -x_128 = !lean_is_exclusive(x_117); -if (x_128 == 0) -{ -return x_117; -} -else -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_129 = lean_ctor_get(x_117, 0); -x_130 = lean_ctor_get(x_117, 1); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_124 = lean_ctor_get(x_123, 1); +lean_inc(x_124); +lean_dec(x_123); +x_125 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56; +x_126 = l_Lean_Parser_Term_structInstFields___closed__2; +x_127 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58; +x_128 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59; +x_129 = l_Lean_Parser_registerAlias(x_125, x_126, x_127, x_128, x_6, x_124); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_129, 1); lean_inc(x_130); -lean_inc(x_129); -lean_dec(x_117); -x_131 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_131, 0, x_129); -lean_ctor_set(x_131, 1, x_130); -return x_131; -} -} -} -else -{ -uint8_t x_132; -x_132 = !lean_is_exclusive(x_111); -if (x_132 == 0) -{ -return x_111; -} -else +lean_dec(x_129); +x_131 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61; +x_132 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_125, x_131, x_130); +if (lean_obj_tag(x_132) == 0) { lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_111, 0); -x_134 = lean_ctor_get(x_111, 1); -lean_inc(x_134); +x_133 = lean_ctor_get(x_132, 1); lean_inc(x_133); -lean_dec(x_111); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); +lean_dec(x_132); +x_134 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63; +x_135 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_125, x_134, x_133); return x_135; } -} -} else { uint8_t x_136; -x_136 = !lean_is_exclusive(x_108); +x_136 = !lean_is_exclusive(x_132); if (x_136 == 0) { -return x_108; +return x_132; } else { lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_108, 0); -x_138 = lean_ctor_get(x_108, 1); +x_137 = lean_ctor_get(x_132, 0); +x_138 = lean_ctor_get(x_132, 1); lean_inc(x_138); lean_inc(x_137); -lean_dec(x_108); +lean_dec(x_132); x_139 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); @@ -69964,19 +70088,19 @@ return x_139; else { uint8_t x_140; -x_140 = !lean_is_exclusive(x_105); +x_140 = !lean_is_exclusive(x_129); if (x_140 == 0) { -return x_105; +return x_129; } else { lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_141 = lean_ctor_get(x_105, 0); -x_142 = lean_ctor_get(x_105, 1); +x_141 = lean_ctor_get(x_129, 0); +x_142 = lean_ctor_get(x_129, 1); lean_inc(x_142); lean_inc(x_141); -lean_dec(x_105); +lean_dec(x_129); x_143 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_143, 0, x_141); lean_ctor_set(x_143, 1, x_142); @@ -69987,19 +70111,19 @@ return x_143; else { uint8_t x_144; -x_144 = !lean_is_exclusive(x_99); +x_144 = !lean_is_exclusive(x_123); if (x_144 == 0) { -return x_99; +return x_123; } else { lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_145 = lean_ctor_get(x_99, 0); -x_146 = lean_ctor_get(x_99, 1); +x_145 = lean_ctor_get(x_123, 0); +x_146 = lean_ctor_get(x_123, 1); lean_inc(x_146); lean_inc(x_145); -lean_dec(x_99); +lean_dec(x_123); x_147 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_147, 0, x_145); lean_ctor_set(x_147, 1, x_146); @@ -70010,19 +70134,19 @@ return x_147; else { uint8_t x_148; -x_148 = !lean_is_exclusive(x_96); +x_148 = !lean_is_exclusive(x_120); if (x_148 == 0) { -return x_96; +return x_120; } else { lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_96, 0); -x_150 = lean_ctor_get(x_96, 1); +x_149 = lean_ctor_get(x_120, 0); +x_150 = lean_ctor_get(x_120, 1); lean_inc(x_150); lean_inc(x_149); -lean_dec(x_96); +lean_dec(x_120); x_151 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_151, 0, x_149); lean_ctor_set(x_151, 1, x_150); @@ -70033,19 +70157,19 @@ return x_151; else { uint8_t x_152; -x_152 = !lean_is_exclusive(x_93); +x_152 = !lean_is_exclusive(x_117); if (x_152 == 0) { -return x_93; +return x_117; } else { lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_153 = lean_ctor_get(x_93, 0); -x_154 = lean_ctor_get(x_93, 1); +x_153 = lean_ctor_get(x_117, 0); +x_154 = lean_ctor_get(x_117, 1); lean_inc(x_154); lean_inc(x_153); -lean_dec(x_93); +lean_dec(x_117); x_155 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); @@ -70056,19 +70180,19 @@ return x_155; else { uint8_t x_156; -x_156 = !lean_is_exclusive(x_87); +x_156 = !lean_is_exclusive(x_111); if (x_156 == 0) { -return x_87; +return x_111; } else { lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_157 = lean_ctor_get(x_87, 0); -x_158 = lean_ctor_get(x_87, 1); +x_157 = lean_ctor_get(x_111, 0); +x_158 = lean_ctor_get(x_111, 1); lean_inc(x_158); lean_inc(x_157); -lean_dec(x_87); +lean_dec(x_111); x_159 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_159, 0, x_157); lean_ctor_set(x_159, 1, x_158); @@ -70079,19 +70203,19 @@ return x_159; else { uint8_t x_160; -x_160 = !lean_is_exclusive(x_84); +x_160 = !lean_is_exclusive(x_108); if (x_160 == 0) { -return x_84; +return x_108; } else { lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_161 = lean_ctor_get(x_84, 0); -x_162 = lean_ctor_get(x_84, 1); +x_161 = lean_ctor_get(x_108, 0); +x_162 = lean_ctor_get(x_108, 1); lean_inc(x_162); lean_inc(x_161); -lean_dec(x_84); +lean_dec(x_108); x_163 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_163, 0, x_161); lean_ctor_set(x_163, 1, x_162); @@ -70102,19 +70226,19 @@ return x_163; else { uint8_t x_164; -x_164 = !lean_is_exclusive(x_81); +x_164 = !lean_is_exclusive(x_105); if (x_164 == 0) { -return x_81; +return x_105; } else { lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_81, 0); -x_166 = lean_ctor_get(x_81, 1); +x_165 = lean_ctor_get(x_105, 0); +x_166 = lean_ctor_get(x_105, 1); lean_inc(x_166); lean_inc(x_165); -lean_dec(x_81); +lean_dec(x_105); x_167 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_167, 0, x_165); lean_ctor_set(x_167, 1, x_166); @@ -70125,19 +70249,19 @@ return x_167; else { uint8_t x_168; -x_168 = !lean_is_exclusive(x_75); +x_168 = !lean_is_exclusive(x_99); if (x_168 == 0) { -return x_75; +return x_99; } else { lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_75, 0); -x_170 = lean_ctor_get(x_75, 1); +x_169 = lean_ctor_get(x_99, 0); +x_170 = lean_ctor_get(x_99, 1); lean_inc(x_170); lean_inc(x_169); -lean_dec(x_75); +lean_dec(x_99); x_171 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_171, 0, x_169); lean_ctor_set(x_171, 1, x_170); @@ -70148,19 +70272,19 @@ return x_171; else { uint8_t x_172; -x_172 = !lean_is_exclusive(x_72); +x_172 = !lean_is_exclusive(x_96); if (x_172 == 0) { -return x_72; +return x_96; } else { lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_72, 0); -x_174 = lean_ctor_get(x_72, 1); +x_173 = lean_ctor_get(x_96, 0); +x_174 = lean_ctor_get(x_96, 1); lean_inc(x_174); lean_inc(x_173); -lean_dec(x_72); +lean_dec(x_96); x_175 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_175, 0, x_173); lean_ctor_set(x_175, 1, x_174); @@ -70171,19 +70295,19 @@ return x_175; else { uint8_t x_176; -x_176 = !lean_is_exclusive(x_69); +x_176 = !lean_is_exclusive(x_93); if (x_176 == 0) { -return x_69; +return x_93; } else { lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_177 = lean_ctor_get(x_69, 0); -x_178 = lean_ctor_get(x_69, 1); +x_177 = lean_ctor_get(x_93, 0); +x_178 = lean_ctor_get(x_93, 1); lean_inc(x_178); lean_inc(x_177); -lean_dec(x_69); +lean_dec(x_93); x_179 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_179, 0, x_177); lean_ctor_set(x_179, 1, x_178); @@ -70194,19 +70318,19 @@ return x_179; else { uint8_t x_180; -x_180 = !lean_is_exclusive(x_63); +x_180 = !lean_is_exclusive(x_87); if (x_180 == 0) { -return x_63; +return x_87; } else { lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_181 = lean_ctor_get(x_63, 0); -x_182 = lean_ctor_get(x_63, 1); +x_181 = lean_ctor_get(x_87, 0); +x_182 = lean_ctor_get(x_87, 1); lean_inc(x_182); lean_inc(x_181); -lean_dec(x_63); +lean_dec(x_87); x_183 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_183, 0, x_181); lean_ctor_set(x_183, 1, x_182); @@ -70217,19 +70341,19 @@ return x_183; else { uint8_t x_184; -x_184 = !lean_is_exclusive(x_60); +x_184 = !lean_is_exclusive(x_84); if (x_184 == 0) { -return x_60; +return x_84; } else { lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_185 = lean_ctor_get(x_60, 0); -x_186 = lean_ctor_get(x_60, 1); +x_185 = lean_ctor_get(x_84, 0); +x_186 = lean_ctor_get(x_84, 1); lean_inc(x_186); lean_inc(x_185); -lean_dec(x_60); +lean_dec(x_84); x_187 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_187, 0, x_185); lean_ctor_set(x_187, 1, x_186); @@ -70240,19 +70364,19 @@ return x_187; else { uint8_t x_188; -x_188 = !lean_is_exclusive(x_57); +x_188 = !lean_is_exclusive(x_81); if (x_188 == 0) { -return x_57; +return x_81; } else { lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_189 = lean_ctor_get(x_57, 0); -x_190 = lean_ctor_get(x_57, 1); +x_189 = lean_ctor_get(x_81, 0); +x_190 = lean_ctor_get(x_81, 1); lean_inc(x_190); lean_inc(x_189); -lean_dec(x_57); +lean_dec(x_81); x_191 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_191, 0, x_189); lean_ctor_set(x_191, 1, x_190); @@ -70263,19 +70387,19 @@ return x_191; else { uint8_t x_192; -x_192 = !lean_is_exclusive(x_51); +x_192 = !lean_is_exclusive(x_75); if (x_192 == 0) { -return x_51; +return x_75; } else { lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_51, 0); -x_194 = lean_ctor_get(x_51, 1); +x_193 = lean_ctor_get(x_75, 0); +x_194 = lean_ctor_get(x_75, 1); lean_inc(x_194); lean_inc(x_193); -lean_dec(x_51); +lean_dec(x_75); x_195 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_195, 0, x_193); lean_ctor_set(x_195, 1, x_194); @@ -70286,19 +70410,19 @@ return x_195; else { uint8_t x_196; -x_196 = !lean_is_exclusive(x_48); +x_196 = !lean_is_exclusive(x_72); if (x_196 == 0) { -return x_48; +return x_72; } else { lean_object* x_197; lean_object* x_198; lean_object* x_199; -x_197 = lean_ctor_get(x_48, 0); -x_198 = lean_ctor_get(x_48, 1); +x_197 = lean_ctor_get(x_72, 0); +x_198 = lean_ctor_get(x_72, 1); lean_inc(x_198); lean_inc(x_197); -lean_dec(x_48); +lean_dec(x_72); x_199 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_199, 0, x_197); lean_ctor_set(x_199, 1, x_198); @@ -70309,19 +70433,19 @@ return x_199; else { uint8_t x_200; -x_200 = !lean_is_exclusive(x_45); +x_200 = !lean_is_exclusive(x_69); if (x_200 == 0) { -return x_45; +return x_69; } else { lean_object* x_201; lean_object* x_202; lean_object* x_203; -x_201 = lean_ctor_get(x_45, 0); -x_202 = lean_ctor_get(x_45, 1); +x_201 = lean_ctor_get(x_69, 0); +x_202 = lean_ctor_get(x_69, 1); lean_inc(x_202); lean_inc(x_201); -lean_dec(x_45); +lean_dec(x_69); x_203 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_203, 0, x_201); lean_ctor_set(x_203, 1, x_202); @@ -70332,19 +70456,19 @@ return x_203; else { uint8_t x_204; -x_204 = !lean_is_exclusive(x_39); +x_204 = !lean_is_exclusive(x_63); if (x_204 == 0) { -return x_39; +return x_63; } else { lean_object* x_205; lean_object* x_206; lean_object* x_207; -x_205 = lean_ctor_get(x_39, 0); -x_206 = lean_ctor_get(x_39, 1); +x_205 = lean_ctor_get(x_63, 0); +x_206 = lean_ctor_get(x_63, 1); lean_inc(x_206); lean_inc(x_205); -lean_dec(x_39); +lean_dec(x_63); x_207 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_207, 0, x_205); lean_ctor_set(x_207, 1, x_206); @@ -70355,19 +70479,19 @@ return x_207; else { uint8_t x_208; -x_208 = !lean_is_exclusive(x_36); +x_208 = !lean_is_exclusive(x_60); if (x_208 == 0) { -return x_36; +return x_60; } else { lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_209 = lean_ctor_get(x_36, 0); -x_210 = lean_ctor_get(x_36, 1); +x_209 = lean_ctor_get(x_60, 0); +x_210 = lean_ctor_get(x_60, 1); lean_inc(x_210); lean_inc(x_209); -lean_dec(x_36); +lean_dec(x_60); x_211 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_211, 0, x_209); lean_ctor_set(x_211, 1, x_210); @@ -70378,19 +70502,19 @@ return x_211; else { uint8_t x_212; -x_212 = !lean_is_exclusive(x_33); +x_212 = !lean_is_exclusive(x_57); if (x_212 == 0) { -return x_33; +return x_57; } else { lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_213 = lean_ctor_get(x_33, 0); -x_214 = lean_ctor_get(x_33, 1); +x_213 = lean_ctor_get(x_57, 0); +x_214 = lean_ctor_get(x_57, 1); lean_inc(x_214); lean_inc(x_213); -lean_dec(x_33); +lean_dec(x_57); x_215 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_215, 0, x_213); lean_ctor_set(x_215, 1, x_214); @@ -70401,19 +70525,19 @@ return x_215; else { uint8_t x_216; -x_216 = !lean_is_exclusive(x_27); +x_216 = !lean_is_exclusive(x_51); if (x_216 == 0) { -return x_27; +return x_51; } else { lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_217 = lean_ctor_get(x_27, 0); -x_218 = lean_ctor_get(x_27, 1); +x_217 = lean_ctor_get(x_51, 0); +x_218 = lean_ctor_get(x_51, 1); lean_inc(x_218); lean_inc(x_217); -lean_dec(x_27); +lean_dec(x_51); x_219 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_219, 0, x_217); lean_ctor_set(x_219, 1, x_218); @@ -70424,19 +70548,19 @@ return x_219; else { uint8_t x_220; -x_220 = !lean_is_exclusive(x_24); +x_220 = !lean_is_exclusive(x_48); if (x_220 == 0) { -return x_24; +return x_48; } else { lean_object* x_221; lean_object* x_222; lean_object* x_223; -x_221 = lean_ctor_get(x_24, 0); -x_222 = lean_ctor_get(x_24, 1); +x_221 = lean_ctor_get(x_48, 0); +x_222 = lean_ctor_get(x_48, 1); lean_inc(x_222); lean_inc(x_221); -lean_dec(x_24); +lean_dec(x_48); x_223 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_223, 0, x_221); lean_ctor_set(x_223, 1, x_222); @@ -70447,19 +70571,19 @@ return x_223; else { uint8_t x_224; -x_224 = !lean_is_exclusive(x_21); +x_224 = !lean_is_exclusive(x_45); if (x_224 == 0) { -return x_21; +return x_45; } else { lean_object* x_225; lean_object* x_226; lean_object* x_227; -x_225 = lean_ctor_get(x_21, 0); -x_226 = lean_ctor_get(x_21, 1); +x_225 = lean_ctor_get(x_45, 0); +x_226 = lean_ctor_get(x_45, 1); lean_inc(x_226); lean_inc(x_225); -lean_dec(x_21); +lean_dec(x_45); x_227 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_227, 0, x_225); lean_ctor_set(x_227, 1, x_226); @@ -70470,19 +70594,19 @@ return x_227; else { uint8_t x_228; -x_228 = !lean_is_exclusive(x_15); +x_228 = !lean_is_exclusive(x_39); if (x_228 == 0) { -return x_15; +return x_39; } else { lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_229 = lean_ctor_get(x_15, 0); -x_230 = lean_ctor_get(x_15, 1); +x_229 = lean_ctor_get(x_39, 0); +x_230 = lean_ctor_get(x_39, 1); lean_inc(x_230); lean_inc(x_229); -lean_dec(x_15); +lean_dec(x_39); x_231 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_231, 0, x_229); lean_ctor_set(x_231, 1, x_230); @@ -70493,19 +70617,19 @@ return x_231; else { uint8_t x_232; -x_232 = !lean_is_exclusive(x_11); +x_232 = !lean_is_exclusive(x_36); if (x_232 == 0) { -return x_11; +return x_36; } else { lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_233 = lean_ctor_get(x_11, 0); -x_234 = lean_ctor_get(x_11, 1); +x_233 = lean_ctor_get(x_36, 0); +x_234 = lean_ctor_get(x_36, 1); lean_inc(x_234); lean_inc(x_233); -lean_dec(x_11); +lean_dec(x_36); x_235 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_235, 0, x_233); lean_ctor_set(x_235, 1, x_234); @@ -70516,19 +70640,19 @@ return x_235; else { uint8_t x_236; -x_236 = !lean_is_exclusive(x_7); +x_236 = !lean_is_exclusive(x_33); if (x_236 == 0) { -return x_7; +return x_33; } else { lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_237 = lean_ctor_get(x_7, 0); -x_238 = lean_ctor_get(x_7, 1); +x_237 = lean_ctor_get(x_33, 0); +x_238 = lean_ctor_get(x_33, 1); lean_inc(x_238); lean_inc(x_237); -lean_dec(x_7); +lean_dec(x_33); x_239 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_239, 0, x_237); lean_ctor_set(x_239, 1, x_238); @@ -70536,6 +70660,144 @@ return x_239; } } } +else +{ +uint8_t x_240; +x_240 = !lean_is_exclusive(x_27); +if (x_240 == 0) +{ +return x_27; +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_241 = lean_ctor_get(x_27, 0); +x_242 = lean_ctor_get(x_27, 1); +lean_inc(x_242); +lean_inc(x_241); +lean_dec(x_27); +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_241); +lean_ctor_set(x_243, 1, x_242); +return x_243; +} +} +} +else +{ +uint8_t x_244; +x_244 = !lean_is_exclusive(x_24); +if (x_244 == 0) +{ +return x_24; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_245 = lean_ctor_get(x_24, 0); +x_246 = lean_ctor_get(x_24, 1); +lean_inc(x_246); +lean_inc(x_245); +lean_dec(x_24); +x_247 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_246); +return x_247; +} +} +} +else +{ +uint8_t x_248; +x_248 = !lean_is_exclusive(x_21); +if (x_248 == 0) +{ +return x_21; +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; +x_249 = lean_ctor_get(x_21, 0); +x_250 = lean_ctor_get(x_21, 1); +lean_inc(x_250); +lean_inc(x_249); +lean_dec(x_21); +x_251 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_251, 0, x_249); +lean_ctor_set(x_251, 1, x_250); +return x_251; +} +} +} +else +{ +uint8_t x_252; +x_252 = !lean_is_exclusive(x_15); +if (x_252 == 0) +{ +return x_15; +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; +x_253 = lean_ctor_get(x_15, 0); +x_254 = lean_ctor_get(x_15, 1); +lean_inc(x_254); +lean_inc(x_253); +lean_dec(x_15); +x_255 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_255, 0, x_253); +lean_ctor_set(x_255, 1, x_254); +return x_255; +} +} +} +else +{ +uint8_t x_256; +x_256 = !lean_is_exclusive(x_11); +if (x_256 == 0) +{ +return x_11; +} +else +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; +x_257 = lean_ctor_get(x_11, 0); +x_258 = lean_ctor_get(x_11, 1); +lean_inc(x_258); +lean_inc(x_257); +lean_dec(x_11); +x_259 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_259, 0, x_257); +lean_ctor_set(x_259, 1, x_258); +return x_259; +} +} +} +else +{ +uint8_t x_260; +x_260 = !lean_is_exclusive(x_7); +if (x_260 == 0) +{ +return x_7; +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_261 = lean_ctor_get(x_7, 0); +x_262 = lean_ctor_get(x_7, 1); +lean_inc(x_262); +lean_inc(x_261); +lean_dec(x_7); +x_263 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_262); +return x_263; +} +} +} } lean_object* initialize_Lean_Parser_Attr(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Parser_Level(uint8_t builtin, lean_object*); @@ -72861,6 +73123,10 @@ l_Lean_Parser_Term_optEllipsis___closed__9 = _init_l_Lean_Parser_Term_optEllipsi lean_mark_persistent(l_Lean_Parser_Term_optEllipsis___closed__9); l_Lean_Parser_Term_optEllipsis = _init_l_Lean_Parser_Term_optEllipsis(); lean_mark_persistent(l_Lean_Parser_Term_optEllipsis); +l_Lean_Parser_Term_structInstFields___closed__1 = _init_l_Lean_Parser_Term_structInstFields___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_structInstFields___closed__1); +l_Lean_Parser_Term_structInstFields___closed__2 = _init_l_Lean_Parser_Term_structInstFields___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_structInstFields___closed__2); l_Lean_Parser_Term_structInst___closed__1 = _init_l_Lean_Parser_Term_structInst___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_structInst___closed__1); l_Lean_Parser_Term_structInst___closed__2 = _init_l_Lean_Parser_Term_structInst___closed__2(); @@ -72921,6 +73187,8 @@ l_Lean_Parser_Term_structInst___closed__29 = _init_l_Lean_Parser_Term_structInst lean_mark_persistent(l_Lean_Parser_Term_structInst___closed__29); l_Lean_Parser_Term_structInst___closed__30 = _init_l_Lean_Parser_Term_structInst___closed__30(); lean_mark_persistent(l_Lean_Parser_Term_structInst___closed__30); +l_Lean_Parser_Term_structInst___closed__31 = _init_l_Lean_Parser_Term_structInst___closed__31(); +lean_mark_persistent(l_Lean_Parser_Term_structInst___closed__31); l_Lean_Parser_Term_structInst = _init_l_Lean_Parser_Term_structInst(); lean_mark_persistent(l_Lean_Parser_Term_structInst); if (builtin) {res = l___regBuiltin_Lean_Parser_Term_structInst__1(lean_io_mk_world()); @@ -73095,6 +73363,8 @@ l_Lean_Parser_Term_structInst_formatter___closed__18 = _init_l_Lean_Parser_Term_ lean_mark_persistent(l_Lean_Parser_Term_structInst_formatter___closed__18); l_Lean_Parser_Term_structInst_formatter___closed__19 = _init_l_Lean_Parser_Term_structInst_formatter___closed__19(); lean_mark_persistent(l_Lean_Parser_Term_structInst_formatter___closed__19); +l_Lean_Parser_Term_structInst_formatter___closed__20 = _init_l_Lean_Parser_Term_structInst_formatter___closed__20(); +lean_mark_persistent(l_Lean_Parser_Term_structInst_formatter___closed__20); l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__1); l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_structInst_formatter__1___closed__2(); @@ -73245,6 +73515,8 @@ l_Lean_Parser_Term_structInst_parenthesizer___closed__18 = _init_l_Lean_Parser_T lean_mark_persistent(l_Lean_Parser_Term_structInst_parenthesizer___closed__18); l_Lean_Parser_Term_structInst_parenthesizer___closed__19 = _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__19(); lean_mark_persistent(l_Lean_Parser_Term_structInst_parenthesizer___closed__19); +l_Lean_Parser_Term_structInst_parenthesizer___closed__20 = _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__20(); +lean_mark_persistent(l_Lean_Parser_Term_structInst_parenthesizer___closed__20); l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer__1___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer__1___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer__1___closed__1); l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer__1___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer__1___closed__2(); @@ -81436,117 +81708,133 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Tactic_quotSeq_parenthesizer__1_ if (builtin) {res = l___regBuiltin_Lean_Parser_Tactic_quotSeq_parenthesizer__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__45); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__46); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__47); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__48); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__49); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__50); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__51); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__52); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__53); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__54); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369____closed__55); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5369_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__45); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__46); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__47); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__48); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__49); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__50); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__51); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__52); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__53); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__54); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__55); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__56); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__57); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__58); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__59); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__60); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__61); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__62); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379____closed__63); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5379_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index e12840e239ce..9aa3186dc102 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -6591,7 +6591,7 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = lean_alloc_ctor(5, 1, 0); +x_16 = lean_alloc_ctor(6, 1, 0); lean_ctor_set(x_16, 0, x_14); x_17 = lean_st_ref_take(x_7, x_15); x_18 = lean_ctor_get(x_17, 0); @@ -7868,7 +7868,7 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_alloc_ctor(11, 1, 0); +x_15 = lean_alloc_ctor(12, 1, 0); lean_ctor_set(x_15, 0, x_13); x_16 = lean_st_ref_take(x_6, x_14); x_17 = lean_ctor_get(x_16, 0); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 499d6a155133..f05627c6777a 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar__1___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__10; @@ -26,7 +25,6 @@ static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSeqLeft__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1___lambda__1___closed__1; @@ -48,7 +46,8 @@ lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_obj static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId___closed__1; @@ -92,15 +91,17 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___c static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun__1___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHAndThen__1___closed__4; -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1; lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabPProdMk___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMProdMk__1___closed__1; LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9(lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1___closed__3; @@ -141,7 +142,6 @@ lean_object* l_Lean_Expr_name_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__5___closed__2; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSeqRight___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -321,6 +321,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__4___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum__1___closed__1; LEAN_EXPORT lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -340,7 +341,6 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLazyBinop___closed__2; lean_object* l_Lean_PrettyPrinter_Delaborator_isNonConstFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__9___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -396,6 +396,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___lambda__2__ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__2(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -436,6 +437,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPPro uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__1; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; @@ -452,6 +454,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId___closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte__1___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam__1___closed__3; lean_object* l_Lean_getPPUniverses___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData__1___closed__2; @@ -459,6 +462,8 @@ uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLazyBinop___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__4; +uint8_t l_Lean_Meta_instDecidableEqCoeFnType(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar__1___closed__2; @@ -551,6 +556,7 @@ static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__6___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabHAndThen(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch__1___closed__1; @@ -565,6 +571,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___clos static lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar__1(lean_object*); extern lean_object* l_Lean_Parser_Command_declSig; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -605,6 +612,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHOrElse__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabForallBinders(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific__1___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__9___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object*, lean_object*, lean_object*); @@ -696,7 +704,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delabora extern lean_object* l_Lean_SubExpr_Pos_typeCoord; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__8; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__3___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__5; @@ -768,7 +775,6 @@ lean_object* l_Lean_getPPUnicodeFun___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabPProdMkCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma__1___closed__1; LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -838,6 +844,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___la LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabHOrElse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__4; lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__1; @@ -935,7 +942,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplic static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__6___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -989,6 +996,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__2; lean_object* lean_mk_syntax_ident(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1074,6 +1082,7 @@ lean_object* l_Lean_Syntax_mkCharLit(uint32_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; static lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; uint8_t l_Lean_getPPMotivesNonConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_shouldShowMotive___closed__1; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; @@ -1082,12 +1091,12 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHAndThen_ lean_object* l_Lean_getPPMVarsWithType___boxed(lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMProdMk__1(lean_object*); uint8_t l_Lean_Expr_isAutoParam(lean_object*); lean_object* l_Lean_getPPMVarsAnonymous___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__12___closed__1; -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__5; @@ -1134,6 +1143,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___c LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__6___boxed(lean_object**); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_hasMacroScopes(lean_object*); @@ -1155,7 +1165,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambd LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__3; @@ -1206,7 +1216,6 @@ lean_object* l_Lean_RBNode_find___at_Lean_PrettyPrinter_Delaborator_OptionsPerPo static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___closed__1; lean_object* l_Lean_getPPPiBinderTypes___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLazyBinop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__1___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__12; @@ -1308,7 +1317,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDelayedAssignedMVar___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabDIte___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar__1___closed__4; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1316,7 +1324,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVarAux___lambda__1___ uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Level_dec(lean_object*); uint8_t l_Lean_Syntax_hasIdent(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar__1___closed__3; @@ -1325,7 +1332,6 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_ static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1___closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar__1___closed__5; lean_object* l_Lean_getPPNumericTypes___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__4; @@ -1369,7 +1375,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNamedPattern(lean LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVarAux___lambda__1___closed__4; @@ -1383,7 +1388,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSeq__1___ static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__2; lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_processSpine___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHAndThen__1___closed__1; @@ -1404,6 +1408,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLit_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_processSpine___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_maxPrec; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__1; lean_object* l_Array_back_x21___rarg(lean_object*, lean_object*); @@ -1457,7 +1462,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr static lean_object* l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma__1(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__3; -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSeqLeft__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1501,6 +1505,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDatasOptions___ra LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__9(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVarAux___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__6; @@ -1534,7 +1539,6 @@ lean_object* l_Lean_getPPExplicit___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1552,6 +1556,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSeq___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getValues___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__6(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond__1(lean_object*); @@ -1637,7 +1643,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withOverApp___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2; lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__10; @@ -1701,6 +1706,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSeqLeft___lambda__1___ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__3___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr__5___closed__1; +static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSeqRight___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__1; @@ -8668,7 +8674,7 @@ static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Del _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -8684,6 +8690,26 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__2; +x_2 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__3; +x_3 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__4; +x_4 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20, lean_object* x_21) { _start: { @@ -8798,84 +8824,103 @@ goto block_44; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; x_66 = lean_unsigned_to_nat(2u); x_67 = l_Lean_Syntax_getArg(x_57, x_66); -x_68 = lean_unsigned_to_nat(3u); -x_69 = l_Lean_Syntax_getArg(x_57, x_68); -x_70 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; -lean_inc(x_69); -x_71 = l_Lean_Syntax_isOfKind(x_69, x_70); -if (x_71 == 0) +x_68 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_67); +x_69 = l_Lean_Syntax_isOfKind(x_67, x_68); +if (x_69 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_69); +lean_object* x_70; lean_object* x_71; lean_dec(x_67); lean_dec(x_57); -x_72 = lean_box(0); +x_70 = lean_box(0); lean_inc(x_3); -x_73 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_72, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_73; +x_71 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_70, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_71; goto block_44; } else { -lean_object* x_74; uint8_t x_75; -x_74 = l_Lean_Syntax_getArg(x_69, x_24); -lean_dec(x_69); -x_75 = l_Lean_Syntax_matchesNull(x_74, x_24); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_72 = l_Lean_Syntax_getArg(x_67, x_24); lean_dec(x_67); +x_73 = lean_unsigned_to_nat(3u); +x_74 = l_Lean_Syntax_getArg(x_57, x_73); +x_75 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; +lean_inc(x_74); +x_76 = l_Lean_Syntax_isOfKind(x_74, x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_74); +lean_dec(x_72); lean_dec(x_57); -x_76 = lean_box(0); +x_77 = lean_box(0); lean_inc(x_3); -x_77 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_76, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_77; +x_78 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_77, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_78; goto block_44; } else { -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = lean_unsigned_to_nat(4u); -x_79 = l_Lean_Syntax_getArg(x_57, x_78); -lean_dec(x_57); -x_80 = l_Lean_Syntax_isNone(x_79); +lean_object* x_79; uint8_t x_80; +x_79 = l_Lean_Syntax_getArg(x_74, x_24); +lean_dec(x_74); +x_80 = l_Lean_Syntax_matchesNull(x_79, x_24); if (x_80 == 0) { -uint8_t x_81; -x_81 = l_Lean_Syntax_matchesNull(x_79, x_66); -if (x_81 == 0) +lean_object* x_81; lean_object* x_82; +lean_dec(x_72); +lean_dec(x_57); +x_81 = lean_box(0); +lean_inc(x_3); +x_82 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_81, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_82; +goto block_44; +} +else { -lean_object* x_82; lean_object* x_83; -lean_dec(x_67); -x_82 = lean_box(0); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_unsigned_to_nat(4u); +x_84 = l_Lean_Syntax_getArg(x_57, x_83); +lean_dec(x_57); +x_85 = l_Lean_Syntax_isNone(x_84); +if (x_85 == 0) +{ +uint8_t x_86; +x_86 = l_Lean_Syntax_matchesNull(x_84, x_66); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +lean_dec(x_72); +x_87 = lean_box(0); lean_inc(x_3); -x_83 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_82, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_83; +x_88 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_87, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_88; goto block_44; } else { -lean_object* x_84; lean_object* x_85; +lean_object* x_89; lean_object* x_90; lean_dec(x_45); -x_84 = lean_box(0); -x_85 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_67, x_14, x_84, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -lean_dec(x_67); -x_28 = x_85; +x_89 = lean_box(0); +x_90 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_72, x_14, x_89, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +lean_dec(x_72); +x_28 = x_90; goto block_44; } } else { -lean_object* x_86; lean_object* x_87; -lean_dec(x_79); +lean_object* x_91; lean_object* x_92; +lean_dec(x_84); lean_dec(x_45); -x_86 = lean_box(0); -x_87 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_67, x_14, x_86, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -lean_dec(x_67); -x_28 = x_87; +x_91 = lean_box(0); +x_92 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_72, x_14, x_91, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +lean_dec(x_72); +x_28 = x_92; goto block_44; } } @@ -8885,29 +8930,30 @@ goto block_44; } } } +} else { -uint8_t x_88; +uint8_t x_93; lean_dec(x_45); lean_dec(x_14); -x_88 = !lean_is_exclusive(x_47); -if (x_88 == 0) +x_93 = !lean_is_exclusive(x_47); +if (x_93 == 0) { x_28 = x_47; goto block_44; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_47, 0); -x_90 = lean_ctor_get(x_47, 1); -lean_inc(x_90); -lean_inc(x_89); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_47, 0); +x_95 = lean_ctor_get(x_47, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_47); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_28 = x_91; +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_28 = x_96; goto block_44; } } @@ -9011,7 +9057,7 @@ return x_43; } else { -lean_object* x_92; +lean_object* x_97; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -9021,10 +9067,10 @@ lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); lean_dec(x_3); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_14); -lean_ctor_set(x_92, 1, x_21); -return x_92; +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_14); +lean_ctor_set(x_97, 1, x_21); +return x_97; } } } @@ -9373,84 +9419,103 @@ goto block_44; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; x_66 = lean_unsigned_to_nat(2u); x_67 = l_Lean_Syntax_getArg(x_57, x_66); -x_68 = lean_unsigned_to_nat(3u); -x_69 = l_Lean_Syntax_getArg(x_57, x_68); -x_70 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; -lean_inc(x_69); -x_71 = l_Lean_Syntax_isOfKind(x_69, x_70); -if (x_71 == 0) +x_68 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_67); +x_69 = l_Lean_Syntax_isOfKind(x_67, x_68); +if (x_69 == 0) { -lean_object* x_72; lean_object* x_73; -lean_dec(x_69); +lean_object* x_70; lean_object* x_71; lean_dec(x_67); lean_dec(x_57); -x_72 = lean_box(0); +x_70 = lean_box(0); lean_inc(x_3); -x_73 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_72, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_73; +x_71 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_70, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_71; goto block_44; } else { -lean_object* x_74; uint8_t x_75; -x_74 = l_Lean_Syntax_getArg(x_69, x_24); -lean_dec(x_69); -x_75 = l_Lean_Syntax_matchesNull(x_74, x_24); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_72 = l_Lean_Syntax_getArg(x_67, x_24); lean_dec(x_67); +x_73 = lean_unsigned_to_nat(3u); +x_74 = l_Lean_Syntax_getArg(x_57, x_73); +x_75 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; +lean_inc(x_74); +x_76 = l_Lean_Syntax_isOfKind(x_74, x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_74); +lean_dec(x_72); lean_dec(x_57); -x_76 = lean_box(0); +x_77 = lean_box(0); lean_inc(x_3); -x_77 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_76, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_77; +x_78 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_77, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_78; goto block_44; } else { -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = lean_unsigned_to_nat(4u); -x_79 = l_Lean_Syntax_getArg(x_57, x_78); -lean_dec(x_57); -x_80 = l_Lean_Syntax_isNone(x_79); +lean_object* x_79; uint8_t x_80; +x_79 = l_Lean_Syntax_getArg(x_74, x_24); +lean_dec(x_74); +x_80 = l_Lean_Syntax_matchesNull(x_79, x_24); if (x_80 == 0) { -uint8_t x_81; -x_81 = l_Lean_Syntax_matchesNull(x_79, x_66); -if (x_81 == 0) +lean_object* x_81; lean_object* x_82; +lean_dec(x_72); +lean_dec(x_57); +x_81 = lean_box(0); +lean_inc(x_3); +x_82 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_81, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_82; +goto block_44; +} +else { -lean_object* x_82; lean_object* x_83; -lean_dec(x_67); -x_82 = lean_box(0); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_unsigned_to_nat(4u); +x_84 = l_Lean_Syntax_getArg(x_57, x_83); +lean_dec(x_57); +x_85 = l_Lean_Syntax_isNone(x_84); +if (x_85 == 0) +{ +uint8_t x_86; +x_86 = l_Lean_Syntax_matchesNull(x_84, x_66); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +lean_dec(x_72); +x_87 = lean_box(0); lean_inc(x_3); -x_83 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_82, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -x_28 = x_83; +x_88 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(x_45, x_3, x_5, x_12, x_6, x_14, x_87, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +x_28 = x_88; goto block_44; } else { -lean_object* x_84; lean_object* x_85; +lean_object* x_89; lean_object* x_90; lean_dec(x_45); -x_84 = lean_box(0); -x_85 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_67, x_14, x_84, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -lean_dec(x_67); -x_28 = x_85; +x_89 = lean_box(0); +x_90 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_72, x_14, x_89, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +lean_dec(x_72); +x_28 = x_90; goto block_44; } } else { -lean_object* x_86; lean_object* x_87; -lean_dec(x_79); +lean_object* x_91; lean_object* x_92; +lean_dec(x_84); lean_dec(x_45); -x_86 = lean_box(0); -x_87 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_67, x_14, x_86, x_15, x_16, x_17, x_18, x_19, x_20, x_53); -lean_dec(x_67); -x_28 = x_87; +x_91 = lean_box(0); +x_92 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2(x_72, x_14, x_91, x_15, x_16, x_17, x_18, x_19, x_20, x_53); +lean_dec(x_72); +x_28 = x_92; goto block_44; } } @@ -9460,29 +9525,30 @@ goto block_44; } } } +} else { -uint8_t x_88; +uint8_t x_93; lean_dec(x_45); lean_dec(x_14); -x_88 = !lean_is_exclusive(x_47); -if (x_88 == 0) +x_93 = !lean_is_exclusive(x_47); +if (x_93 == 0) { x_28 = x_47; goto block_44; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_47, 0); -x_90 = lean_ctor_get(x_47, 1); -lean_inc(x_90); -lean_inc(x_89); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_47, 0); +x_95 = lean_ctor_get(x_47, 1); +lean_inc(x_95); +lean_inc(x_94); lean_dec(x_47); -x_91 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_28 = x_91; +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_28 = x_96; goto block_44; } } @@ -9586,7 +9652,7 @@ return x_43; } else { -lean_object* x_92; +lean_object* x_97; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -9596,10 +9662,10 @@ lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); lean_dec(x_3); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_14); -lean_ctor_set(x_92, 1, x_21); -return x_92; +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_14); +lean_ctor_set(x_97, 1, x_21); +return x_97; } } } @@ -9831,7 +9897,7 @@ uint8_t x_33; x_33 = !lean_is_exclusive(x_32); if (x_33 == 0) { -lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_34 = lean_ctor_get(x_32, 0); x_35 = lean_ctor_get(x_14, 5); lean_inc(x_35); @@ -9862,193 +9928,199 @@ lean_ctor_set(x_46, 0, x_37); lean_ctor_set(x_46, 1, x_40); lean_ctor_set(x_46, 2, x_45); x_47 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_37); +x_48 = l_Lean_Syntax_node1(x_37, x_47, x_46); +x_49 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; lean_inc(x_42); lean_inc(x_37); -x_48 = l_Lean_Syntax_node1(x_37, x_47, x_42); -x_49 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; +x_50 = l_Lean_Syntax_node1(x_37, x_49, x_42); +x_51 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; lean_inc(x_37); -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_37); -lean_ctor_set(x_50, 1, x_49); +x_52 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_52, 0, x_37); +lean_ctor_set(x_52, 1, x_51); if (lean_obj_tag(x_34) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; lean_inc(x_37); -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_37); -lean_ctor_set(x_52, 1, x_40); -lean_ctor_set(x_52, 2, x_51); -x_53 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_54 = l_Lean_Syntax_node6(x_37, x_53, x_39, x_42, x_46, x_48, x_52, x_50); -lean_ctor_set(x_32, 0, x_54); +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_37); +lean_ctor_set(x_54, 1, x_40); +lean_ctor_set(x_54, 2, x_53); +x_55 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_56 = l_Lean_Syntax_node6(x_37, x_55, x_39, x_42, x_48, x_50, x_54, x_52); +lean_ctor_set(x_32, 0, x_56); return x_32; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_55 = lean_ctor_get(x_34, 0); -lean_inc(x_55); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_57 = lean_ctor_get(x_34, 0); +lean_inc(x_57); lean_dec(x_34); -x_56 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +x_58 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; lean_inc(x_37); -x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_37); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Array_mkArray2___rarg(x_57, x_55); -x_59 = l_Array_append___rarg(x_41, x_58); -lean_dec(x_58); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_37); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Array_mkArray2___rarg(x_59, x_57); +x_61 = l_Array_append___rarg(x_41, x_60); +lean_dec(x_60); lean_inc(x_37); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_37); -lean_ctor_set(x_60, 1, x_40); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_62 = l_Lean_Syntax_node6(x_37, x_61, x_39, x_42, x_46, x_48, x_60, x_50); -lean_ctor_set(x_32, 0, x_62); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_37); +lean_ctor_set(x_62, 1, x_40); +lean_ctor_set(x_62, 2, x_61); +x_63 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_64 = l_Lean_Syntax_node6(x_37, x_63, x_39, x_42, x_48, x_50, x_62, x_52); +lean_ctor_set(x_32, 0, x_64); return x_32; } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_63 = lean_ctor_get(x_32, 0); -x_64 = lean_ctor_get(x_32, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_32); -x_65 = lean_ctor_get(x_14, 5); +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_65 = lean_ctor_get(x_32, 0); +x_66 = lean_ctor_get(x_32, 1); +lean_inc(x_66); lean_inc(x_65); -lean_dec(x_14); -x_66 = 0; -x_67 = l_Lean_SourceInfo_fromRef(x_65, x_66); -lean_dec(x_65); -x_68 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__6; -lean_inc(x_67); -x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_71 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_dec(x_32); +x_67 = lean_ctor_get(x_14, 5); lean_inc(x_67); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_70); -lean_ctor_set(x_72, 2, x_71); -x_73 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_74 = l_Lean_Syntax_SepArray_ofElems(x_73, x_29); +lean_dec(x_14); +x_68 = 0; +x_69 = l_Lean_SourceInfo_fromRef(x_67, x_68); +lean_dec(x_67); +x_70 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__6; +lean_inc(x_69); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_73 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_inc(x_69); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; +x_76 = l_Lean_Syntax_SepArray_ofElems(x_75, x_29); lean_dec(x_29); -x_75 = l_Array_append___rarg(x_71, x_74); -lean_dec(x_74); -lean_inc(x_67); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_67); -lean_ctor_set(x_76, 1, x_70); -lean_ctor_set(x_76, 2, x_75); -x_77 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; -lean_inc(x_72); -lean_inc(x_67); -x_78 = l_Lean_Syntax_node1(x_67, x_77, x_72); -x_79 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; -lean_inc(x_67); -x_80 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_80, 0, x_67); -lean_ctor_set(x_80, 1, x_79); -if (lean_obj_tag(x_63) == 0) +x_77 = l_Array_append___rarg(x_73, x_76); +lean_dec(x_76); +lean_inc(x_69); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_69); +lean_ctor_set(x_78, 1, x_72); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_69); +x_80 = l_Lean_Syntax_node1(x_69, x_79, x_78); +x_81 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; +lean_inc(x_74); +lean_inc(x_69); +x_82 = l_Lean_Syntax_node1(x_69, x_81, x_74); +x_83 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; +lean_inc(x_69); +x_84 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_84, 0, x_69); +lean_ctor_set(x_84, 1, x_83); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_81 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; -lean_inc(x_67); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_67); -lean_ctor_set(x_82, 1, x_70); -lean_ctor_set(x_82, 2, x_81); -x_83 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_84 = l_Lean_Syntax_node6(x_67, x_83, x_69, x_72, x_76, x_78, x_82, x_80); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_64); -return x_85; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_85 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; +lean_inc(x_69); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_69); +lean_ctor_set(x_86, 1, x_72); +lean_ctor_set(x_86, 2, x_85); +x_87 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_88 = l_Lean_Syntax_node6(x_69, x_87, x_71, x_74, x_80, x_82, x_86, x_84); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_66); +return x_89; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_86 = lean_ctor_get(x_63, 0); -lean_inc(x_86); -lean_dec(x_63); -x_87 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_67); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_67); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Array_mkArray2___rarg(x_88, x_86); -x_90 = l_Array_append___rarg(x_71, x_89); -lean_dec(x_89); -lean_inc(x_67); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_67); -lean_ctor_set(x_91, 1, x_70); -lean_ctor_set(x_91, 2, x_90); -x_92 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_93 = l_Lean_Syntax_node6(x_67, x_92, x_69, x_72, x_76, x_78, x_91, x_80); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_64); -return x_94; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_90 = lean_ctor_get(x_65, 0); +lean_inc(x_90); +lean_dec(x_65); +x_91 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_69); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_69); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Array_mkArray2___rarg(x_92, x_90); +x_94 = l_Array_append___rarg(x_73, x_93); +lean_dec(x_93); +lean_inc(x_69); +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_69); +lean_ctor_set(x_95, 1, x_72); +lean_ctor_set(x_95, 2, x_94); +x_96 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_97 = l_Lean_Syntax_node6(x_69, x_96, x_71, x_74, x_80, x_82, x_95, x_84); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_66); +return x_98; } } } else { -uint8_t x_95; +uint8_t x_99; lean_dec(x_29); lean_dec(x_14); -x_95 = !lean_is_exclusive(x_32); -if (x_95 == 0) +x_99 = !lean_is_exclusive(x_32); +if (x_99 == 0) { return x_32; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_32, 0); -x_97 = lean_ctor_get(x_32, 1); -lean_inc(x_97); -lean_inc(x_96); +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_32, 0); +x_101 = lean_ctor_get(x_32, 1); +lean_inc(x_101); +lean_inc(x_100); lean_dec(x_32); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; } } } else { -uint8_t x_99; +uint8_t x_103; lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -x_99 = !lean_is_exclusive(x_28); -if (x_99 == 0) +x_103 = !lean_is_exclusive(x_28); +if (x_103 == 0) { return x_28; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_28, 0); -x_101 = lean_ctor_get(x_28, 1); -lean_inc(x_101); -lean_inc(x_100); +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_28, 0); +x_105 = lean_ctor_get(x_28, 1); +lean_inc(x_105); +lean_inc(x_104); lean_dec(x_28); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } } @@ -10204,7 +10276,7 @@ uint8_t x_33; x_33 = !lean_is_exclusive(x_32); if (x_33 == 0) { -lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_34 = lean_ctor_get(x_32, 0); x_35 = lean_ctor_get(x_14, 5); lean_inc(x_35); @@ -10235,193 +10307,199 @@ lean_ctor_set(x_46, 0, x_37); lean_ctor_set(x_46, 1, x_40); lean_ctor_set(x_46, 2, x_45); x_47 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_37); +x_48 = l_Lean_Syntax_node1(x_37, x_47, x_46); +x_49 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; lean_inc(x_42); lean_inc(x_37); -x_48 = l_Lean_Syntax_node1(x_37, x_47, x_42); -x_49 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; +x_50 = l_Lean_Syntax_node1(x_37, x_49, x_42); +x_51 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; lean_inc(x_37); -x_50 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_50, 0, x_37); -lean_ctor_set(x_50, 1, x_49); +x_52 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_52, 0, x_37); +lean_ctor_set(x_52, 1, x_51); if (lean_obj_tag(x_34) == 0) { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; lean_inc(x_37); -x_52 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_52, 0, x_37); -lean_ctor_set(x_52, 1, x_40); -lean_ctor_set(x_52, 2, x_51); -x_53 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_54 = l_Lean_Syntax_node6(x_37, x_53, x_39, x_42, x_46, x_48, x_52, x_50); -lean_ctor_set(x_32, 0, x_54); +x_54 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_54, 0, x_37); +lean_ctor_set(x_54, 1, x_40); +lean_ctor_set(x_54, 2, x_53); +x_55 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_56 = l_Lean_Syntax_node6(x_37, x_55, x_39, x_42, x_48, x_50, x_54, x_52); +lean_ctor_set(x_32, 0, x_56); return x_32; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_55 = lean_ctor_get(x_34, 0); -lean_inc(x_55); +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_57 = lean_ctor_get(x_34, 0); +lean_inc(x_57); lean_dec(x_34); -x_56 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +x_58 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; lean_inc(x_37); -x_57 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_57, 0, x_37); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Array_mkArray2___rarg(x_57, x_55); -x_59 = l_Array_append___rarg(x_41, x_58); -lean_dec(x_58); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_37); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Array_mkArray2___rarg(x_59, x_57); +x_61 = l_Array_append___rarg(x_41, x_60); +lean_dec(x_60); lean_inc(x_37); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_37); -lean_ctor_set(x_60, 1, x_40); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_62 = l_Lean_Syntax_node6(x_37, x_61, x_39, x_42, x_46, x_48, x_60, x_50); -lean_ctor_set(x_32, 0, x_62); +x_62 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_62, 0, x_37); +lean_ctor_set(x_62, 1, x_40); +lean_ctor_set(x_62, 2, x_61); +x_63 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_64 = l_Lean_Syntax_node6(x_37, x_63, x_39, x_42, x_48, x_50, x_62, x_52); +lean_ctor_set(x_32, 0, x_64); return x_32; } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_63 = lean_ctor_get(x_32, 0); -x_64 = lean_ctor_get(x_32, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_32); -x_65 = lean_ctor_get(x_14, 5); +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_65 = lean_ctor_get(x_32, 0); +x_66 = lean_ctor_get(x_32, 1); +lean_inc(x_66); lean_inc(x_65); -lean_dec(x_14); -x_66 = 0; -x_67 = l_Lean_SourceInfo_fromRef(x_65, x_66); -lean_dec(x_65); -x_68 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__6; -lean_inc(x_67); -x_69 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -x_70 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_71 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_dec(x_32); +x_67 = lean_ctor_get(x_14, 5); lean_inc(x_67); -x_72 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_72, 0, x_67); -lean_ctor_set(x_72, 1, x_70); -lean_ctor_set(x_72, 2, x_71); -x_73 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_74 = l_Lean_Syntax_SepArray_ofElems(x_73, x_29); +lean_dec(x_14); +x_68 = 0; +x_69 = l_Lean_SourceInfo_fromRef(x_67, x_68); +lean_dec(x_67); +x_70 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__6; +lean_inc(x_69); +x_71 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_73 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_inc(x_69); +x_74 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_72); +lean_ctor_set(x_74, 2, x_73); +x_75 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; +x_76 = l_Lean_Syntax_SepArray_ofElems(x_75, x_29); lean_dec(x_29); -x_75 = l_Array_append___rarg(x_71, x_74); -lean_dec(x_74); -lean_inc(x_67); -x_76 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_76, 0, x_67); -lean_ctor_set(x_76, 1, x_70); -lean_ctor_set(x_76, 2, x_75); -x_77 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; -lean_inc(x_72); -lean_inc(x_67); -x_78 = l_Lean_Syntax_node1(x_67, x_77, x_72); -x_79 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; -lean_inc(x_67); -x_80 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_80, 0, x_67); -lean_ctor_set(x_80, 1, x_79); -if (lean_obj_tag(x_63) == 0) +x_77 = l_Array_append___rarg(x_73, x_76); +lean_dec(x_76); +lean_inc(x_69); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_69); +lean_ctor_set(x_78, 1, x_72); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; +lean_inc(x_69); +x_80 = l_Lean_Syntax_node1(x_69, x_79, x_78); +x_81 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7; +lean_inc(x_74); +lean_inc(x_69); +x_82 = l_Lean_Syntax_node1(x_69, x_81, x_74); +x_83 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__13; +lean_inc(x_69); +x_84 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_84, 0, x_69); +lean_ctor_set(x_84, 1, x_83); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_81 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; -lean_inc(x_67); -x_82 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_82, 0, x_67); -lean_ctor_set(x_82, 1, x_70); -lean_ctor_set(x_82, 2, x_81); -x_83 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_84 = l_Lean_Syntax_node6(x_67, x_83, x_69, x_72, x_76, x_78, x_82, x_80); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_64); -return x_85; +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_85 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__7; +lean_inc(x_69); +x_86 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_86, 0, x_69); +lean_ctor_set(x_86, 1, x_72); +lean_ctor_set(x_86, 2, x_85); +x_87 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_88 = l_Lean_Syntax_node6(x_69, x_87, x_71, x_74, x_80, x_82, x_86, x_84); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_66); +return x_89; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_86 = lean_ctor_get(x_63, 0); -lean_inc(x_86); -lean_dec(x_63); -x_87 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_67); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_67); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Array_mkArray2___rarg(x_88, x_86); -x_90 = l_Array_append___rarg(x_71, x_89); -lean_dec(x_89); -lean_inc(x_67); -x_91 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_91, 0, x_67); -lean_ctor_set(x_91, 1, x_70); -lean_ctor_set(x_91, 2, x_90); -x_92 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; -x_93 = l_Lean_Syntax_node6(x_67, x_92, x_69, x_72, x_76, x_78, x_91, x_80); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_64); -return x_94; +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_90 = lean_ctor_get(x_65, 0); +lean_inc(x_90); +lean_dec(x_65); +x_91 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_69); +x_92 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_92, 0, x_69); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Array_mkArray2___rarg(x_92, x_90); +x_94 = l_Array_append___rarg(x_73, x_93); +lean_dec(x_93); +lean_inc(x_69); +x_95 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_95, 0, x_69); +lean_ctor_set(x_95, 1, x_72); +lean_ctor_set(x_95, 2, x_94); +x_96 = l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; +x_97 = l_Lean_Syntax_node6(x_69, x_96, x_71, x_74, x_80, x_82, x_95, x_84); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_66); +return x_98; } } } else { -uint8_t x_95; +uint8_t x_99; lean_dec(x_29); lean_dec(x_14); -x_95 = !lean_is_exclusive(x_32); -if (x_95 == 0) +x_99 = !lean_is_exclusive(x_32); +if (x_99 == 0) { return x_32; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_32, 0); -x_97 = lean_ctor_get(x_32, 1); -lean_inc(x_97); -lean_inc(x_96); +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_32, 0); +x_101 = lean_ctor_get(x_32, 1); +lean_inc(x_101); +lean_inc(x_100); lean_dec(x_32); -x_98 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -return x_98; +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +return x_102; } } } else { -uint8_t x_99; +uint8_t x_103; lean_dec(x_15); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -x_99 = !lean_is_exclusive(x_28); -if (x_99 == 0) +x_103 = !lean_is_exclusive(x_28); +if (x_103 == 0) { return x_28; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_28, 0); -x_101 = lean_ctor_get(x_28, 1); -lean_inc(x_101); -lean_inc(x_100); +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_28, 0); +x_105 = lean_ctor_get(x_28, 1); +lean_inc(x_105); +lean_inc(x_104); lean_dec(x_28); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; } } } @@ -39715,7 +39793,7 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -39723,17 +39801,17 @@ x_1 = lean_mk_string_unchecked("coeNotation", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -39741,19 +39819,19 @@ x_1 = lean_mk_string_unchecked("↑", 3, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_9 = lean_ctor_get(x_6, 5); x_10 = 0; x_11 = l_Lean_SourceInfo_fromRef(x_9, x_10); -x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3; +x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3; lean_inc(x_11); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2; +x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2; x_15 = l_Lean_Syntax_node2(x_11, x_14, x_13, x_1); x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -39761,7 +39839,7 @@ lean_ctor_set(x_16, 1, x_8); return x_16; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -39769,17 +39847,17 @@ x_1 = lean_mk_string_unchecked("coeFunNotation", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -39787,19 +39865,19 @@ x_1 = lean_mk_string_unchecked("⇑", 3, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_9 = lean_ctor_get(x_6, 5); x_10 = 0; x_11 = l_Lean_SourceInfo_fromRef(x_9, x_10); -x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3; +x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3; lean_inc(x_11); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2; +x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2; x_15 = l_Lean_Syntax_node2(x_11, x_14, x_13, x_1); x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -39807,7 +39885,7 @@ lean_ctor_set(x_16, 1, x_8); return x_16; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -39815,17 +39893,17 @@ x_1 = lean_mk_string_unchecked("coeSortNotation", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3() { _start: { lean_object* x_1; @@ -39833,19 +39911,19 @@ x_1 = lean_mk_string_unchecked("↥", 3, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_9 = lean_ctor_get(x_6, 5); x_10 = 0; x_11 = l_Lean_SourceInfo_fromRef(x_9, x_10); -x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3; +x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3; lean_inc(x_11); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2; +x_14 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2; x_15 = l_Lean_Syntax_node2(x_11, x_14, x_13, x_1); x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -39853,122 +39931,257 @@ lean_ctor_set(x_16, 1, x_8); return x_16; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___boxed), 8, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___boxed), 8, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___boxed), 8, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4() { _start: { -uint8_t x_11; -x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); -switch (x_11) { -case 0: +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo), 8, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_13); +lean_object* x_11; +if (x_3 == 0) +{ +x_11 = x_10; +goto block_42; +} +else +{ +lean_object* x_43; uint8_t x_44; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_1); -x_14 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); -lean_closure_set(x_15, 0, x_13); -lean_closure_set(x_15, 1, x_14); -x_16 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1; -x_17 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); -lean_closure_set(x_17, 0, x_15); -lean_closure_set(x_17, 1, x_16); -x_18 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_12, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_12); -return x_18; +x_43 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_10); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +return x_43; } -case 1: +else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_1, 0); -lean_inc(x_19); -x_20 = lean_unsigned_to_nat(0u); -x_21 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_2, x_20); -x_22 = lean_nat_dec_eq(x_21, x_19); -lean_dec(x_21); -if (x_22 == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_43); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; +} +} +block_42: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_1, 1); -lean_inc(x_23); -lean_dec(x_1); -x_24 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_25 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -x_26 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_19, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_19); -return x_26; +uint8_t x_12; lean_object* x_13; uint8_t x_33; uint8_t x_34; +x_12 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); +x_33 = 1; +x_34 = l_Lean_Meta_instDecidableEqCoeFnType(x_12, x_33); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = lean_box(0); +x_13 = x_35; +goto block_32; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_1, 1); -lean_inc(x_27); +lean_object* x_36; uint8_t x_37; +x_36 = lean_unsigned_to_nat(0u); +x_37 = lean_nat_dec_lt(x_36, x_2); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = lean_box(0); +x_13 = x_38; +goto block_32; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_1, 1); +lean_inc(x_39); lean_dec(x_1); -x_28 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_29 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); -lean_closure_set(x_29, 0, x_27); -lean_closure_set(x_29, 1, x_28); -x_30 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2; -x_31 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); -lean_closure_set(x_31, 0, x_29); -lean_closure_set(x_31, 1, x_30); -x_32 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_19, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_19); -return x_32; +x_40 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4; +x_41 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_39, x_40, x_4, x_5, x_6, x_7, x_8, x_9, x_11); +lean_dec(x_39); +return x_41; } } +block_32: +{ +lean_dec(x_13); +switch (x_12) { +case 0: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +lean_dec(x_1); +x_15 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1; +x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); +lean_closure_set(x_18, 0, x_16); +lean_closure_set(x_18, 1, x_17); +x_19 = l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_11); +return x_19; +} +case 1: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; +x_22 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +x_23 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2; +x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_23); +x_25 = l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo(x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_11); +return x_25; +} default: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_33 = lean_ctor_get(x_1, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_1, 1); -lean_inc(x_34); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_26 = lean_ctor_get(x_1, 1); +lean_inc(x_26); lean_dec(x_1); -x_35 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_36 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); -lean_closure_set(x_36, 0, x_34); -lean_closure_set(x_36, 1, x_35); -x_37 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3; -x_38 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); -lean_closure_set(x_38, 0, x_36); -lean_closure_set(x_38, 1, x_37); -x_39 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_33, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_33); -return x_39; +x_27 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; +x_28 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); +lean_closure_set(x_28, 0, x_26); +lean_closure_set(x_28, 1, x_27); +x_29 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3; +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVarAux___spec__1___rarg), 9, 2); +lean_closure_set(x_30, 0, x_28); +lean_closure_set(x_30, 1, x_29); +x_31 = l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo(x_30, x_4, x_5, x_6, x_7, x_8, x_9, x_11); +return x_31; +} } } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_nat_dec_eq(x_1, x_2); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_13 = lean_nat_sub(x_1, x_2); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___boxed), 10, 2); +lean_closure_set(x_14, 0, x_3); +lean_closure_set(x_14, 1, x_13); +x_15 = 0; +x_16 = l_Lean_PrettyPrinter_Delaborator_delabAppCore(x_13, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_unsigned_to_nat(0u); +x_18 = 0; +x_19 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead(x_3, x_17, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; @@ -40000,52 +40213,95 @@ return x_14; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; x_15 = lean_ctor_get(x_11, 1); lean_inc(x_15); lean_dec(x_11); x_16 = lean_ctor_get(x_12, 0); lean_inc(x_16); lean_dec(x_12); -x_17 = l_Lean_PrettyPrinter_Delaborator_useAppExplicit___closed__1; +x_17 = lean_unsigned_to_nat(0u); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_17); +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +x_20 = lean_nat_dec_le(x_19, x_18); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_21 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_15); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +return x_21; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_21); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = l_Lean_PrettyPrinter_Delaborator_useAppExplicit___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_18 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -if (lean_obj_tag(x_18) == 0) +x_27 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_unbox(x_19); -lean_dec(x_19); -if (x_20 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_unbox(x_28); +lean_dec(x_28); +if (x_29 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_box(0); +x_32 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(x_18, x_19, x_16, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_30); +lean_dec(x_19); lean_dec(x_18); -x_22 = lean_box(0); -x_23 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4(x_16, x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_21); -return x_23; +return x_32; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); -lean_dec(x_18); -x_25 = lean_ctor_get(x_16, 1); -lean_inc(x_25); -x_26 = lean_unsigned_to_nat(0u); -x_27 = lean_nat_dec_eq(x_25, x_26); -lean_dec(x_25); -if (x_27 == 0) +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_27, 1); +lean_inc(x_33); +lean_dec(x_27); +x_34 = lean_ctor_get(x_16, 1); +lean_inc(x_34); +x_35 = lean_nat_dec_eq(x_34, x_17); +lean_dec(x_34); +if (x_35 == 0) { -lean_object* x_28; uint8_t x_29; +lean_object* x_36; uint8_t x_37; +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_16); lean_dec(x_7); lean_dec(x_6); @@ -40053,38 +40309,42 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_28 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_24); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +x_36 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_33); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) { -return x_28; +return x_36; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_36, 0); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_36); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } else { -lean_object* x_33; lean_object* x_34; -x_33 = lean_box(0); -x_34 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4(x_16, x_1, x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_24); -return x_34; +lean_object* x_41; lean_object* x_42; +x_41 = lean_box(0); +x_42 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(x_18, x_19, x_16, x_41, x_2, x_3, x_4, x_5, x_6, x_7, x_33); +lean_dec(x_19); +lean_dec(x_18); +return x_42; } } } else { -uint8_t x_35; +uint8_t x_43; +lean_dec(x_19); +lean_dec(x_18); lean_dec(x_16); lean_dec(x_7); lean_dec(x_6); @@ -40092,30 +40352,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_35 = !lean_is_exclusive(x_18); -if (x_35 == 0) +x_43 = !lean_is_exclusive(x_27); +if (x_43 == 0) { -return x_18; +return x_27; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_18, 0); -x_37 = lean_ctor_get(x_18, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_18); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_27, 0); +x_45 = lean_ctor_get(x_27, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_27); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} } } } } else { -lean_object* x_39; +lean_object* x_47; lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); @@ -40123,8 +40384,8 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_39 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_39; +x_47 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); +return x_47; } } } @@ -40132,7 +40393,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___boxed), 8, 0); return x_1; } } @@ -40158,18 +40419,15 @@ x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_ return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_12; +x_12 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); -return x_9; +lean_dec(x_1); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -40177,44 +40435,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambd { lean_object* x_9; x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); return x_9; } @@ -44349,7 +44569,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; -x_3 = lean_unsigned_to_nat(1202u); +x_3 = lean_unsigned_to_nat(1212u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -52613,7 +52833,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed__1; -x_3 = lean_unsigned_to_nat(1327u); +x_3 = lean_unsigned_to_nat(1337u); x_4 = lean_unsigned_to_nat(42u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -53870,6 +54090,10 @@ l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructure lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__4); l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5 = _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5(); lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5); +l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6 = _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__6); +l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7 = _init_l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__7); l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1); l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__2(); @@ -54515,30 +54739,32 @@ lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj__1_ if (builtin) {res = l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj__1(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__2); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__1); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__2); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__2___closed__3); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__1); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__2); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__3___closed__3); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2); -l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3); +}l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__1); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__2); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__1___closed__3); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__1); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__2); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__2___closed__3); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__1); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__2); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___lambda__3___closed__3); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__1); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__2); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__3); +l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator_delabHead___closed__4); l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__1); l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__2(); diff --git a/stage0/stdlib/Lean/Server/CodeActions/Provider.c b/stage0/stdlib/Lean/Server/CodeActions/Provider.c index 4c0e47b21e3e..4fc4a3d3dabc 100644 --- a/stage0/stdlib/Lean/Server/CodeActions/Provider.c +++ b/stage0/stdlib/Lean/Server/CodeActions/Provider.c @@ -4388,7 +4388,7 @@ return x_48; } } } -case 2: +case 3: { lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; x_49 = lean_ctor_get(x_21, 0); @@ -4689,7 +4689,7 @@ if (lean_obj_tag(x_4) == 1) lean_object* x_6; x_6 = lean_ctor_get(x_4, 0); lean_inc(x_6); -if (lean_obj_tag(x_6) == 2) +if (lean_obj_tag(x_6) == 3) { lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; x_7 = lean_ctor_get(x_6, 0); diff --git a/stage0/stdlib/Lean/Server/Completion.c b/stage0/stdlib/Lean/Server/Completion.c index ead8fe9c1df7..4ba430568216 100644 --- a/stage0/stdlib/Lean/Server/Completion.c +++ b/stage0/stdlib/Lean/Server/Completion.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Server.Completion -// Imports: Lean.Environment Lean.Parser.Term Lean.Data.FuzzyMatching Lean.Data.Lsp.LanguageFeatures Lean.Data.Lsp.Capabilities Lean.Data.Lsp.Utf16 Lean.Meta.CompletionName Lean.Meta.Tactic.Apply Lean.Meta.Match.MatcherInfo Lean.Elab.Tactic.Doc Lean.Server.InfoUtils Lean.Parser.Extension Lean.Server.FileSource Lean.Server.CompletionItemData +// Imports: Lean.Server.Completion.CompletionCollectors #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,34564 +13,3413 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter(lean_object*); -lean_object* l_Lean_Name_getNumParts(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_ppExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2___boxed__const__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1; -lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(lean_object*, lean_object*, size_t, size_t); -lean_object* l_Lean_getStructureParentInfo(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; -static lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10; -static lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(lean_object*, lean_object*, size_t, size_t); +lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; uint32_t lean_string_utf8_get(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__3(lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2; -LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_private_to_user_name(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_mkObj(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(lean_object*); +LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12(lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_ConstantInfo_type(lean_object*); -lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); -double lean_float_div(double, double); -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Exception_isInterrupt(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31; +extern lean_object* l_Lean_Lsp_instHashableMarkupContent; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3; static lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_matchNamespace___closed__1; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; -lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getId(lean_object*); -uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_idCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21; +lean_object* l_Lean_Server_Completion_optionCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1___boxed(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20; +uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(lean_object*); lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_foldlMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13; -uint8_t l_Lean_Name_isAnonymous(lean_object*); -lean_object* l_Lean_Syntax_getArgs(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_eligibleHeaderDeclsRef; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2; lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); -lean_object* l_Array_insertAt_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_ConstantInfo_name(lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f(lean_object*); -lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_environment_find(lean_object*, lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Lsp_instDecidableEqCompletionItemKind___boxed(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__2(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1; -static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2; -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed(lean_object*); -static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1; -lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_List_head_x3f___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1; -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___boxed(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136_(lean_object*); -extern lean_object* l_Lean_LocalContext_empty; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_instInhabitedJson; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(lean_object*, lean_object*, size_t, size_t); -uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_string_hash(lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_set(lean_object*, lean_object*, uint32_t); -lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(lean_object*); -lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1; -uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_List_any___rarg(lean_object*, lean_object*); -lean_object* l_Lean_getAliasState(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1(lean_object*); +lean_object* l_Std_DHashMap_Internal_AssocList_map_go___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_InductiveVal_numTypeFormers(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1; -uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); -static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1; -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8; -static double l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -extern lean_object* l_instInhabitedNat; -LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Parser_getTokenTable(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_projectionFnInfoExt; -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3; -uint8_t l_Lean_Expr_hasMVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getStr_x3f(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6; -lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Linter_instInhabitedDeprecationEntry; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2(lean_object*); -uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19; +uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(uint8_t); +static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16(lean_object*, size_t, size_t, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemDataWithId; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Char_toLower(uint32_t); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_find_x3f___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Info_tailPos_x3f(lean_object*); -uint8_t l_List_elem___at_Lean_addAliasEntry___spec__16(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1; -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_ConstantInfo_isTheorem(lean_object*); -lean_object* lean_local_ctx_find(lean_object*, lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1; LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2; -lean_object* lean_st_ref_take(lean_object*, lean_object*); -lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_next(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f(lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; uint64_t lean_uint64_shift_right(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +extern lean_object* l_Lean_Lsp_instHashableInsertReplaceEdit; LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Elab_Info_updateContext_x3f(lean_object*, lean_object*); -lean_object* l_Lean_Data_Trie_findPrefix_go___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix(lean_object*); -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount(lean_object*, lean_object*); -lean_object* l_Array_erase___at_Lean_getAllParentStructures___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Name_getPrefix(lean_object*); -lean_object* l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_ConstantInfo_isCtor(lean_object*); -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -lean_object* l_Lean_Syntax_getKind(lean_object*); -static lean_object* l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12___boxed(lean_object*); lean_object* lean_string_mk(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_CompletionInfo_lctx(lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_outOfBounds___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace(lean_object*, lean_object*); -lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionIdentifier; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11; -lean_object* l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___boxed(lean_object*); +extern lean_object* l_instHashableString; LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_List_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_NameSSet_instInhabited; -uint8_t l_Lean_isPrivatePrefix(lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getOptionDecls(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_find_x3f___closed__1; -uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); -lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); +static lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1; +static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3; +lean_object* l_Lean_Server_Completion_tacticCompletion(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(lean_object*); -lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(lean_object*); -lean_object* l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(lean_object*, lean_object*, double); +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6(lean_object*, size_t, size_t, uint64_t); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10; lean_object* l_panic___at_String_fromUTF8_x21___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(lean_object*, lean_object*); lean_object* l_List_replicateTR___rarg(lean_object*, lean_object*); -lean_object* l_Lean_FileMap_lineStart(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -static lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33; +static lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(lean_object*, lean_object*); +lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17; -uint8_t lean_name_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3(lean_object*, lean_object*); -lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29; extern lean_object* l_Lean_Lsp_instInhabitedCompletionItem; -lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3___boxed(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__1(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isIdent(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_LocalDecl_fvarId(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8; -static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemDataWithId; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___boxed(lean_object**); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(lean_object*); -extern lean_object* l_Std_Format_defWidth; -static lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3; -lean_object* l_Lean_Elab_Tactic_Doc_allTacticDocs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(size_t, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate(lean_object*, lean_object*); -uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1; -lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, uint8_t); -static lean_object* l_Lean_Lsp_CompletionItem_resolve___closed__1; -static lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_structureResolutionExt; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___boxed(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose(lean_object*); -lean_object* l_Lean_LocalDecl_userName(lean_object*); -static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_Linter_deprecatedAttr; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_isToken(lean_object*, lean_object*); -lean_object* l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(lean_object*, lean_object*); -lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Name_hasMacroScopes(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1(lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3; -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_searchAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1; -uint8_t l_Lean_Name_isAtomic(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Meta_allowCompletion(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__14(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems(lean_object*); +lean_object* l_Lean_Lsp_instDecidableEqCompletionItemTag___boxed(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26; +lean_object* l_Array_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(size_t, size_t, lean_object*); +static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1; +extern lean_object* l_Lean_Lsp_instHashableCompletionItemKind; +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_instHashableOption___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19(lean_object*, size_t, size_t, lean_object*); lean_object* lean_string_length(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Syntax_hasArgs(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_findStack_x3f(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8; +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1; -lean_object* lean_nat_mod(lean_object*, lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2; -uint8_t l_Lean_ConstantInfo_isInductive(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18; -uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3; -LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1; -lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1; -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___boxed(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3; -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; -lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7(lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint64_t l_Lean_Name_hash___override(lean_object*); +static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(lean_object*, lean_object*); uint8_t lean_float_beq(double, double); -lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -lean_object* l_String_toName(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__2___boxed(lean_object*, lean_object*); -lean_object* l_Lean_LocalDecl_type(lean_object*); uint64_t lean_uint64_xor(uint64_t, uint64_t); -extern lean_object* l_Id_instMonad; lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object*, lean_object*, lean_object*); -lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Completion_find_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_isStructure(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Expr_isProp(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18; lean_object* lean_nat_mul(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1; -lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); -lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(lean_object*, lean_object*, size_t, size_t); -lean_object* l_Lean_Name_getString_x21(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1; -lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(lean_object*, lean_object*, size_t, size_t); -lean_object* lean_erase_macro_scopes(lean_object*); -lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___boxed(lean_object**); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(lean_object*); -uint8_t l_Lean_Elab_Info_isCompletion(lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14; +extern lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; -extern lean_object* l_Lean_namespacesExt; -lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2; -lean_object* l_List_dropWhile___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4; -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___boxed(lean_object**); -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8(lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_findPrioritizedCompletionPartitionsAt(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_dotCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Elab_Info_isSmaller(lean_object*, lean_object*); -extern lean_object* l_Lean_instInhabitedName; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion___hyg_743_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__15(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -lean_object* l_Lean_Elab_Info_lctx(lean_object*); size_t lean_array_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_completeNamespaces___closed__1; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Server_Completion_cmpModPrivate(lean_object*, lean_object*); -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4; -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5; -lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); -lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(lean_object*, size_t, size_t); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3; +lean_object* l_Lean_Server_Completion_dotIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__5(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Lsp_instDecidableEqMarkupContent___boxed(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1; -uint8_t l_Lean_Name_isSuffixOf(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_cmpModPrivate___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1___boxed(lean_object*); -static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8; -lean_object* l_Char_utf8Size(uint32_t); -uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_404_(uint8_t, uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(lean_object*); +lean_object* l_Lean_Server_Completion_fieldIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -lean_object* l_Lean_Elab_Info_stx(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Exception_isRuntime(lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4; -static lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4; -lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); -uint8_t l_Lean_Expr_isForall(lean_object*); -uint8_t l_Lean_Elab_Info_occursInOrOnBoundary(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_is_class(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2; -lean_object* l_Lean_Expr_consumeMData(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Name_isInternal(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____spec__1(lean_object*, lean_object*); +extern lean_object* l_Lean_Lsp_instBEqInsertReplaceEdit; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4; +extern lean_object* l_Lean_Lsp_instHashableCompletionItemTag; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionIdentifier; -uint8_t lean_local_ctx_is_empty(lean_object*); -lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_CompletionItem_resolve___closed__2; -lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3; -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instDecidableEqString___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1(lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1; -static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3; -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_data_value_to_string(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_expr_instantiate1(x_1, x_3); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg(x_9, x_2, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -if (lean_obj_tag(x_1) == 7) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); -x_12 = 1; -x_13 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_404_(x_11, x_12); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_14 = lean_apply_6(x_2, x_1, x_3, x_4, x_5, x_6, x_7); -return x_14; -} -else -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; -lean_dec(x_1); -x_15 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1___boxed), 8, 2); -lean_closure_set(x_15, 0, x_10); -lean_closure_set(x_15, 1, x_2); -x_16 = 0; -x_17 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(x_8, x_11, x_9, x_15, x_16, x_3, x_4, x_5, x_6, x_7); -return x_17; -} -} -else -{ -lean_object* x_18; -x_18 = lean_apply_6(x_2, x_1, x_3, x_4, x_5, x_6, x_7); -return x_18; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg), 7, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_1); -return x_9; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("no inductive constructor matched", 32, 32); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2; -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("declName", 8, 8); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("const", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("[anonymous]", 11, 11); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("expected a `Name`, got '", 24, 24); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("'", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_3 = l_Except_orElseLazy___rarg(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_4 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2; -x_5 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_1); -x_6 = lean_array_mk(x_5); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -x_8 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3; -x_9 = lean_unsigned_to_nat(1u); -x_10 = l_Lean_Json_parseTagged(x_2, x_8, x_9, x_7); -lean_dec(x_7); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_58; -x_58 = !lean_is_exclusive(x_10); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; -x_59 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_60 = l_Except_orElseLazy___rarg(x_10, x_59); -return x_60; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_10, 0); -lean_inc(x_61); -lean_dec(x_10); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -x_63 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_64 = l_Except_orElseLazy___rarg(x_62, x_63); -return x_64; -} -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = lean_ctor_get(x_10, 0); -lean_inc(x_65); -lean_dec(x_10); -x_66 = lean_array_get_size(x_65); -x_67 = lean_unsigned_to_nat(0u); -x_68 = lean_nat_dec_lt(x_67, x_66); -lean_dec(x_66); -if (x_68 == 0) -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_65); -x_69 = l_Lean_instInhabitedJson; -x_70 = l_outOfBounds___rarg(x_69); -x_11 = x_70; -goto block_57; -} -else -{ -lean_object* x_71; -x_71 = lean_array_fget(x_65, x_67); -lean_dec(x_65); -x_11 = x_71; -goto block_57; -} -} -block_57: -{ -lean_object* x_12; -lean_inc(x_11); -x_12 = l_Lean_Json_getStr_x3f(x_11); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -lean_dec(x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_15 = l_Except_orElseLazy___rarg(x_12, x_14); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_12, 0); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -x_18 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_19 = l_Except_orElseLazy___rarg(x_17, x_18); -return x_19; -} -} -else -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_12, 0); -x_22 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5; -x_23 = lean_string_dec_eq(x_21, x_22); -if (x_23 == 0) -{ -lean_object* x_24; uint8_t x_25; -x_24 = l_String_toName(x_21); -x_25 = l_Lean_Name_isAnonymous(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_11); -x_26 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_12, 0, x_26); -x_27 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_28 = l_Except_orElseLazy___rarg(x_12, x_27); -return x_28; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_24); -x_29 = lean_unsigned_to_nat(80u); -x_30 = l_Lean_Json_pretty(x_11, x_29); -x_31 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6; -x_32 = lean_string_append(x_31, x_30); -lean_dec(x_30); -x_33 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -x_34 = lean_string_append(x_32, x_33); -lean_ctor_set_tag(x_12, 0); -lean_ctor_set(x_12, 0, x_34); -x_35 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_36 = l_Except_orElseLazy___rarg(x_12, x_35); -return x_36; -} -} -else -{ -lean_object* x_37; -lean_free_object(x_12); -lean_dec(x_21); -lean_dec(x_11); -x_37 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10; -return x_37; -} -} -else -{ -lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_38 = lean_ctor_get(x_12, 0); -lean_inc(x_38); -lean_dec(x_12); -x_39 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5; -x_40 = lean_string_dec_eq(x_38, x_39); -if (x_40 == 0) -{ -lean_object* x_41; uint8_t x_42; -x_41 = l_String_toName(x_38); -x_42 = l_Lean_Name_isAnonymous(x_41); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -lean_dec(x_11); -x_43 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_43, 0, x_41); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -x_45 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_46 = l_Except_orElseLazy___rarg(x_44, x_45); -return x_46; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_41); -x_47 = lean_unsigned_to_nat(80u); -x_48 = l_Lean_Json_pretty(x_11, x_47); -x_49 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6; -x_50 = lean_string_append(x_49, x_48); -lean_dec(x_48); -x_51 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -x_52 = lean_string_append(x_50, x_51); -x_53 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_53, 0, x_52); -x_54 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4; -x_55 = l_Except_orElseLazy___rarg(x_53, x_54); -return x_55; -} -} -else -{ -lean_object* x_56; -lean_dec(x_38); -lean_dec(x_11); -x_56 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10; -return x_56; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("id", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3; -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("fvar", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_box(0); -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6; -x_4 = lean_unsigned_to_nat(1u); -x_5 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5; -lean_inc(x_1); -x_6 = l_Lean_Json_parseTagged(x_1, x_3, x_4, x_5); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___boxed), 3, 2); -lean_closure_set(x_7, 0, x_2); -lean_closure_set(x_7, 1, x_1); -if (lean_obj_tag(x_6) == 0) -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_6); -if (x_51 == 0) -{ -lean_object* x_52; -x_52 = l_Except_orElseLazy___rarg(x_6, x_7); -return x_52; -} -else -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_6, 0); -lean_inc(x_53); -lean_dec(x_6); -x_54 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_54, 0, x_53); -x_55 = l_Except_orElseLazy___rarg(x_54, x_7); -return x_55; -} -} -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_6, 0); -lean_inc(x_56); -lean_dec(x_6); -x_57 = lean_array_get_size(x_56); -x_58 = lean_unsigned_to_nat(0u); -x_59 = lean_nat_dec_lt(x_58, x_57); -lean_dec(x_57); -if (x_59 == 0) -{ -lean_object* x_60; lean_object* x_61; -lean_dec(x_56); -x_60 = l_Lean_instInhabitedJson; -x_61 = l_outOfBounds___rarg(x_60); -x_8 = x_61; -goto block_50; -} -else -{ -lean_object* x_62; -x_62 = lean_array_fget(x_56, x_58); -lean_dec(x_56); -x_8 = x_62; -goto block_50; -} -} -block_50: -{ -lean_object* x_9; -lean_inc(x_8); -x_9 = l_Lean_Json_getStr_x3f(x_8); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_10; -lean_dec(x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; -x_11 = l_Except_orElseLazy___rarg(x_9, x_7); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = l_Except_orElseLazy___rarg(x_13, x_7); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_ctor_get(x_9, 0); -x_17 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5; -x_18 = lean_string_dec_eq(x_16, x_17); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = l_String_toName(x_16); -x_20 = l_Lean_Name_isAnonymous(x_19); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -lean_dec(x_8); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_9, 0, x_21); -x_22 = l_Except_orElseLazy___rarg(x_9, x_7); -return x_22; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_19); -x_23 = lean_unsigned_to_nat(80u); -x_24 = l_Lean_Json_pretty(x_8, x_23); -x_25 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6; -x_26 = lean_string_append(x_25, x_24); -lean_dec(x_24); -x_27 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -x_28 = lean_string_append(x_26, x_27); -lean_ctor_set_tag(x_9, 0); -lean_ctor_set(x_9, 0, x_28); -x_29 = l_Except_orElseLazy___rarg(x_9, x_7); -return x_29; -} -} -else -{ -lean_object* x_30; lean_object* x_31; -lean_free_object(x_9); -lean_dec(x_16); -lean_dec(x_8); -x_30 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8; -x_31 = l_Except_orElseLazy___rarg(x_30, x_7); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_ctor_get(x_9, 0); -lean_inc(x_32); -lean_dec(x_9); -x_33 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5; -x_34 = lean_string_dec_eq(x_32, x_33); -if (x_34 == 0) -{ -lean_object* x_35; uint8_t x_36; -x_35 = l_String_toName(x_32); -x_36 = l_Lean_Name_isAnonymous(x_35); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_8); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_35); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = l_Except_orElseLazy___rarg(x_38, x_7); -return x_39; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -lean_dec(x_35); -x_40 = lean_unsigned_to_nat(80u); -x_41 = l_Lean_Json_pretty(x_8, x_40); -x_42 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6; -x_43 = lean_string_append(x_42, x_41); -lean_dec(x_41); -x_44 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -x_45 = lean_string_append(x_43, x_44); -x_46 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_46, 0, x_45); -x_47 = l_Except_orElseLazy___rarg(x_46, x_7); -return x_47; -} -} -else -{ -lean_object* x_48; lean_object* x_49; -lean_dec(x_32); -lean_dec(x_8); -x_48 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8; -x_49 = l_Except_orElseLazy___rarg(x_48, x_7); -return x_49; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__2(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionIdentifier() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1; -return x_1; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = 0; -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1___boxed), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270_(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_2; -x_2 = !lean_is_exclusive(x_1); -if (x_2 == 0) -{ -lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_3 = lean_ctor_get(x_1, 0); -x_4 = 1; -x_5 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_6 = l_Lean_Name_toString(x_3, x_4, x_5); -lean_ctor_set_tag(x_1, 3); -lean_ctor_set(x_1, 0, x_6); -x_7 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_1); -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_8); -lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Json_mkObj(x_10); -x_12 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3; -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -x_14 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_9); -x_15 = l_Lean_Json_mkObj(x_14); -return x_15; -} -else -{ -lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -lean_dec(x_1); -x_17 = 1; -x_18 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_19 = l_Lean_Name_toString(x_16, x_17, x_18); -x_20 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1; -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_Json_mkObj(x_24); -x_26 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3; -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_23); -x_29 = l_Lean_Json_mkObj(x_28); -return x_29; -} -} -else -{ -uint8_t x_30; -x_30 = !lean_is_exclusive(x_1); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_31 = lean_ctor_get(x_1, 0); -x_32 = 1; -x_33 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_34 = l_Lean_Name_toString(x_31, x_32, x_33); -lean_ctor_set_tag(x_1, 3); -lean_ctor_set(x_1, 0, x_34); -x_35 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_1); -x_37 = lean_box(0); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Json_mkObj(x_38); -x_40 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6; -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_37); -x_43 = l_Lean_Json_mkObj(x_42); -return x_43; -} -else -{ -lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_44 = lean_ctor_get(x_1, 0); -lean_inc(x_44); -lean_dec(x_1); -x_45 = 1; -x_46 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_47 = l_Lean_Name_toString(x_44, x_45, x_46); -x_48 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_48); -x_51 = lean_box(0); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_Json_mkObj(x_52); -x_54 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6; -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -x_56 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_51); -x_57 = l_Lean_Json_mkObj(x_56); -return x_57; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionIdentifier() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1; -return x_1; -} -} -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Json_getObjValD(x_1, x_2); -switch (lean_obj_tag(x_3)) { -case 0: -{ -lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1; -return x_4; -} -case 1: -{ -lean_object* x_5; -x_5 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136_(x_3); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -return x_5; -} -else -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; -} -} -else -{ -uint8_t x_9; -x_9 = !lean_is_exclusive(x_5); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_5, 0, x_11); -return x_5; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_5, 0); -lean_inc(x_12); -lean_dec(x_5); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -} -default: -{ -lean_object* x_15; uint8_t x_16; -lean_inc(x_3); -x_15 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136_(x_3); -x_16 = !lean_is_exclusive(x_3); -if (x_16 == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_3, 0); -lean_dec(x_17); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_18; -lean_free_object(x_3); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_15, 0); -lean_inc(x_19); -lean_dec(x_15); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_19); -return x_20; -} -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -lean_object* x_22; -x_22 = lean_ctor_get(x_15, 0); -lean_ctor_set_tag(x_3, 1); -lean_ctor_set(x_3, 0, x_22); -lean_ctor_set(x_15, 0, x_3); -return x_15; -} -else -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -lean_ctor_set_tag(x_3, 1); -lean_ctor_set(x_3, 0, x_23); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_3); -return x_24; -} -} -} -else -{ -lean_dec(x_3); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_15, 0); -lean_inc(x_25); -if (lean_is_exclusive(x_15)) { - lean_ctor_release(x_15, 0); - x_26 = x_15; -} else { - lean_dec_ref(x_15); - x_26 = lean_box(0); -} -if (lean_is_scalar(x_26)) { - x_27 = lean_alloc_ctor(0, 1, 0); -} else { - x_27 = x_26; -} -lean_ctor_set(x_27, 0, x_25); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_15, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_15)) { - lean_ctor_release(x_15, 0); - x_29 = x_15; -} else { - lean_dec_ref(x_15); - x_29 = lean_box(0); -} -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_28); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(1, 1, 0); -} else { - x_31 = x_29; -} -lean_ctor_set(x_31, 0, x_30); -return x_31; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("params", 6, 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lsp", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("CompletionItemDataWithId", 24, 24); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3; -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5; -x_2 = 1; -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(".", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9; -x_2 = 1; -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("id\?", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15; -x_2 = 1; -x_3 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1; -lean_inc(x_1); -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_4; -lean_dec(x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13; -x_7 = lean_string_append(x_6, x_5); -lean_dec(x_5); -lean_ctor_set(x_3, 0, x_7); -return x_3; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_3, 0); -lean_inc(x_8); -lean_dec(x_3); -x_9 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13; -x_10 = lean_string_append(x_9, x_8); -lean_dec(x_8); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_3, 0); -lean_inc(x_12); -lean_dec(x_3); -x_13 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1(x_1, x_13); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -lean_dec(x_12); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18; -x_18 = lean_string_append(x_17, x_16); -lean_dec(x_16); -lean_ctor_set(x_14, 0, x_18); -return x_14; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_14, 0); -lean_inc(x_19); -lean_dec(x_14); -x_20 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18; -x_21 = lean_string_append(x_20, x_19); -lean_dec(x_19); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_14); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_14, 0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_12); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_14, 0, x_25); -return x_14; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_14, 0); -lean_inc(x_26); -lean_dec(x_14); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_12); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -return x_28; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemDataWithId() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; -lean_dec(x_1); -x_3 = lean_box(0); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); -x_5 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270_(x_4); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_5); -x_7 = lean_box(0); -x_8 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_8, 0, x_6); -lean_ctor_set(x_8, 1, x_7); -return x_8; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(x_2); -x_4 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -x_6 = lean_box(0); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -lean_dec(x_1); -x_9 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1; -x_10 = l_Lean_Json_opt___at___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____spec__1(x_9, x_8); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_6); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_7); -lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_14 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_12, x_13); -x_15 = l_Lean_Json_mkObj(x_14); -return x_15; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemDataWithId() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_ppExpr(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_7, 0); -x_10 = l_Std_Format_defWidth; -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_format_pretty(x_9, x_10, x_11, x_11); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_13 = lean_ctor_get(x_7, 0); -x_14 = lean_ctor_get(x_7, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_7); -x_15 = l_Std_Format_defWidth; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_format_pretty(x_13, x_15, x_16, x_16); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_14); -return x_18; -} -} -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_7); -if (x_19 == 0) -{ -return x_7; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_7, 0); -x_21 = lean_ctor_get(x_7, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_7); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; -} -} -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_resolve___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed), 7, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_resolve___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_76; -x_8 = lean_st_ref_get(x_6, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_3, 1); -lean_inc(x_12); -x_13 = l_Lean_Lsp_CompletionItem_resolve___closed__1; -x_76 = lean_ctor_get(x_1, 1); -lean_inc(x_76); -if (lean_obj_tag(x_76) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_77; lean_object* x_78; -lean_dec(x_12); -x_77 = lean_ctor_get(x_2, 0); -lean_inc(x_77); -lean_dec(x_2); -x_78 = lean_environment_find(x_11, x_77); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -x_79 = lean_box(0); -x_14 = x_79; -goto block_75; -} -else -{ -uint8_t x_80; -x_80 = !lean_is_exclusive(x_78); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_78, 0); -x_82 = l_Lean_ConstantInfo_type(x_81); -lean_dec(x_81); -lean_ctor_set(x_78, 0, x_82); -x_14 = x_78; -goto block_75; -} -else -{ -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_78, 0); -lean_inc(x_83); -lean_dec(x_78); -x_84 = l_Lean_ConstantInfo_type(x_83); -lean_dec(x_83); -x_85 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_85, 0, x_84); -x_14 = x_85; -goto block_75; -} -} -} -else -{ -lean_object* x_86; lean_object* x_87; -lean_dec(x_11); -x_86 = lean_ctor_get(x_2, 0); -lean_inc(x_86); -lean_dec(x_2); -x_87 = lean_local_ctx_find(x_12, x_86); -if (lean_obj_tag(x_87) == 0) -{ -lean_object* x_88; -x_88 = lean_box(0); -x_14 = x_88; -goto block_75; -} -else -{ -uint8_t x_89; -x_89 = !lean_is_exclusive(x_87); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; -x_90 = lean_ctor_get(x_87, 0); -x_91 = l_Lean_LocalDecl_type(x_90); -lean_dec(x_90); -lean_ctor_set(x_87, 0, x_91); -x_14 = x_87; -goto block_75; -} -else -{ -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_87, 0); -lean_inc(x_92); -lean_dec(x_87); -x_93 = l_Lean_LocalDecl_type(x_92); -lean_dec(x_92); -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_93); -x_14 = x_94; -goto block_75; -} -} -} -} -else -{ -lean_object* x_95; lean_object* x_96; -lean_dec(x_76); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_2); -x_95 = lean_box(0); -x_96 = lean_apply_7(x_13, x_1, x_95, x_3, x_4, x_5, x_6, x_10); -return x_96; -} -block_75: -{ -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_box(0); -x_16 = !lean_is_exclusive(x_1); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_1, 1); -lean_dec(x_17); -lean_ctor_set(x_1, 1, x_15); -x_18 = lean_box(0); -x_19 = lean_apply_7(x_13, x_1, x_18, x_3, x_4, x_5, x_6, x_10); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_20 = lean_ctor_get(x_1, 0); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); -x_23 = lean_ctor_get(x_1, 4); -x_24 = lean_ctor_get(x_1, 5); -x_25 = lean_ctor_get(x_1, 6); -x_26 = lean_ctor_get(x_1, 7); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_1); -x_27 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_15); -lean_ctor_set(x_27, 2, x_21); -lean_ctor_set(x_27, 3, x_22); -lean_ctor_set(x_27, 4, x_23); -lean_ctor_set(x_27, 5, x_24); -lean_ctor_set(x_27, 6, x_25); -lean_ctor_set(x_27, 7, x_26); -x_28 = lean_box(0); -x_29 = lean_apply_7(x_13, x_27, x_28, x_3, x_4, x_5, x_6, x_10); -return x_29; -} -} -else -{ -uint8_t x_30; -x_30 = !lean_is_exclusive(x_14); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_14, 0); -x_32 = l_Lean_Lsp_CompletionItem_resolve___closed__2; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_33 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg(x_31, x_32, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -lean_ctor_set(x_14, 0, x_34); -x_36 = !lean_is_exclusive(x_1); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_1, 1); -lean_dec(x_37); -lean_ctor_set(x_1, 1, x_14); -x_38 = lean_box(0); -x_39 = lean_apply_7(x_13, x_1, x_38, x_3, x_4, x_5, x_6, x_35); -return x_39; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_40 = lean_ctor_get(x_1, 0); -x_41 = lean_ctor_get(x_1, 2); -x_42 = lean_ctor_get(x_1, 3); -x_43 = lean_ctor_get(x_1, 4); -x_44 = lean_ctor_get(x_1, 5); -x_45 = lean_ctor_get(x_1, 6); -x_46 = lean_ctor_get(x_1, 7); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_1); -x_47 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_14); -lean_ctor_set(x_47, 2, x_41); -lean_ctor_set(x_47, 3, x_42); -lean_ctor_set(x_47, 4, x_43); -lean_ctor_set(x_47, 5, x_44); -lean_ctor_set(x_47, 6, x_45); -lean_ctor_set(x_47, 7, x_46); -x_48 = lean_box(0); -x_49 = lean_apply_7(x_13, x_47, x_48, x_3, x_4, x_5, x_6, x_35); -return x_49; -} -} -else -{ -uint8_t x_50; -lean_free_object(x_14); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_50 = !lean_is_exclusive(x_33); -if (x_50 == 0) -{ -return x_33; -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_33, 0); -x_52 = lean_ctor_get(x_33, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_33); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -return x_53; -} -} -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_14, 0); -lean_inc(x_54); -lean_dec(x_14); -x_55 = l_Lean_Lsp_CompletionItem_resolve___closed__2; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_56 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_consumeImplicitPrefix___rarg(x_54, x_55, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_57); -x_60 = lean_ctor_get(x_1, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_1, 2); -lean_inc(x_61); -x_62 = lean_ctor_get(x_1, 3); -lean_inc(x_62); -x_63 = lean_ctor_get(x_1, 4); -lean_inc(x_63); -x_64 = lean_ctor_get(x_1, 5); -lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 6); -lean_inc(x_65); -x_66 = lean_ctor_get(x_1, 7); -lean_inc(x_66); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - lean_ctor_release(x_1, 2); - lean_ctor_release(x_1, 3); - lean_ctor_release(x_1, 4); - lean_ctor_release(x_1, 5); - lean_ctor_release(x_1, 6); - lean_ctor_release(x_1, 7); - x_67 = x_1; -} else { - lean_dec_ref(x_1); - x_67 = lean_box(0); -} -if (lean_is_scalar(x_67)) { - x_68 = lean_alloc_ctor(0, 8, 0); -} else { - x_68 = x_67; -} -lean_ctor_set(x_68, 0, x_60); -lean_ctor_set(x_68, 1, x_59); -lean_ctor_set(x_68, 2, x_61); -lean_ctor_set(x_68, 3, x_62); -lean_ctor_set(x_68, 4, x_63); -lean_ctor_set(x_68, 5, x_64); -lean_ctor_set(x_68, 6, x_65); -lean_ctor_set(x_68, 7, x_66); -x_69 = lean_box(0); -x_70 = lean_apply_7(x_13, x_68, x_69, x_3, x_4, x_5, x_6, x_58); -return x_70; -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_71 = lean_ctor_get(x_56, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_56, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_73 = x_56; -} else { - lean_dec_ref(x_56); - x_73 = lean_box(0); -} -if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(1, 2, 0); -} else { - x_74 = x_73; -} -lean_ctor_set(x_74, 0, x_71); -lean_ctor_set(x_74, 1, x_72); -return x_74; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_Lsp_CompletionItem_resolve___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Lsp_CompletionItem_resolve___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion___hyg_743_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_box(0); -x_3 = lean_st_mk_ref(x_2, x_1); -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -return x_3; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_dec(x_3); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_5); -lean_ctor_set(x_7, 1, x_6); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_5; -lean_dec(x_1); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_4); -return x_5; -} -else -{ -uint8_t x_6; -lean_dec(x_2); -x_6 = !lean_is_exclusive(x_3); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_ctor_get(x_3, 0); -x_8 = lean_ctor_get(x_3, 1); -x_9 = lean_ctor_get(x_3, 2); -lean_inc(x_7); -lean_inc(x_1); -x_10 = l_Lean_Meta_allowCompletion(x_1, x_7); -if (x_10 == 0) -{ -lean_object* x_11; -lean_free_object(x_3); -lean_dec(x_8); -lean_dec(x_7); -x_11 = lean_box(0); -x_2 = x_11; -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_4); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; lean_object* x_29; uint8_t x_30; -x_14 = lean_ctor_get(x_4, 0); -x_15 = lean_ctor_get(x_4, 1); -x_16 = lean_array_get_size(x_15); -x_17 = l_Lean_Name_hash___override(x_7); -x_18 = 32; -x_19 = lean_uint64_shift_right(x_17, x_18); -x_20 = lean_uint64_xor(x_17, x_19); -x_21 = 16; -x_22 = lean_uint64_shift_right(x_20, x_21); -x_23 = lean_uint64_xor(x_20, x_22); -x_24 = lean_uint64_to_usize(x_23); -x_25 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_26 = 1; -x_27 = lean_usize_sub(x_25, x_26); -x_28 = lean_usize_land(x_24, x_27); -x_29 = lean_array_uget(x_15, x_28); -x_30 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_29); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_31 = lean_unsigned_to_nat(1u); -x_32 = lean_nat_add(x_14, x_31); -lean_dec(x_14); -lean_ctor_set(x_3, 2, x_29); -x_33 = lean_array_uset(x_15, x_28, x_3); -x_34 = lean_unsigned_to_nat(4u); -x_35 = lean_nat_mul(x_32, x_34); -x_36 = lean_unsigned_to_nat(3u); -x_37 = lean_nat_div(x_35, x_36); -lean_dec(x_35); -x_38 = lean_array_get_size(x_33); -x_39 = lean_nat_dec_le(x_37, x_38); -lean_dec(x_38); -lean_dec(x_37); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; -x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_33); -lean_ctor_set(x_4, 1, x_40); -lean_ctor_set(x_4, 0, x_32); -x_41 = lean_box(0); -x_2 = x_41; -x_3 = x_9; -goto _start; -} -else -{ -lean_object* x_43; -lean_ctor_set(x_4, 1, x_33); -lean_ctor_set(x_4, 0, x_32); -x_43 = lean_box(0); -x_2 = x_43; -x_3 = x_9; -goto _start; -} -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -lean_free_object(x_3); -x_45 = lean_box(0); -x_46 = lean_array_uset(x_15, x_28, x_45); -x_47 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_29); -x_48 = lean_array_uset(x_46, x_28, x_47); -lean_ctor_set(x_4, 1, x_48); -x_49 = lean_box(0); -x_2 = x_49; -x_3 = x_9; -goto _start; -} -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; lean_object* x_66; uint8_t x_67; -x_51 = lean_ctor_get(x_4, 0); -x_52 = lean_ctor_get(x_4, 1); -lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_4); -x_53 = lean_array_get_size(x_52); -x_54 = l_Lean_Name_hash___override(x_7); -x_55 = 32; -x_56 = lean_uint64_shift_right(x_54, x_55); -x_57 = lean_uint64_xor(x_54, x_56); -x_58 = 16; -x_59 = lean_uint64_shift_right(x_57, x_58); -x_60 = lean_uint64_xor(x_57, x_59); -x_61 = lean_uint64_to_usize(x_60); -x_62 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_63 = 1; -x_64 = lean_usize_sub(x_62, x_63); -x_65 = lean_usize_land(x_61, x_64); -x_66 = lean_array_uget(x_52, x_65); -x_67 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_66); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -x_68 = lean_unsigned_to_nat(1u); -x_69 = lean_nat_add(x_51, x_68); -lean_dec(x_51); -lean_ctor_set(x_3, 2, x_66); -x_70 = lean_array_uset(x_52, x_65, x_3); -x_71 = lean_unsigned_to_nat(4u); -x_72 = lean_nat_mul(x_69, x_71); -x_73 = lean_unsigned_to_nat(3u); -x_74 = lean_nat_div(x_72, x_73); -lean_dec(x_72); -x_75 = lean_array_get_size(x_70); -x_76 = lean_nat_dec_le(x_74, x_75); -lean_dec(x_75); -lean_dec(x_74); -if (x_76 == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_70); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_69); -lean_ctor_set(x_78, 1, x_77); -x_79 = lean_box(0); -x_2 = x_79; -x_3 = x_9; -x_4 = x_78; -goto _start; -} -else -{ -lean_object* x_81; lean_object* x_82; -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_69); -lean_ctor_set(x_81, 1, x_70); -x_82 = lean_box(0); -x_2 = x_82; -x_3 = x_9; -x_4 = x_81; -goto _start; -} -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_free_object(x_3); -x_84 = lean_box(0); -x_85 = lean_array_uset(x_52, x_65, x_84); -x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_66); -x_87 = lean_array_uset(x_85, x_65, x_86); -x_88 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_88, 0, x_51); -lean_ctor_set(x_88, 1, x_87); -x_89 = lean_box(0); -x_2 = x_89; -x_3 = x_9; -x_4 = x_88; -goto _start; -} -} -} -} -else -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; -x_91 = lean_ctor_get(x_3, 0); -x_92 = lean_ctor_get(x_3, 1); -x_93 = lean_ctor_get(x_3, 2); -lean_inc(x_93); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_3); -lean_inc(x_91); -lean_inc(x_1); -x_94 = l_Lean_Meta_allowCompletion(x_1, x_91); -if (x_94 == 0) -{ -lean_object* x_95; -lean_dec(x_92); -lean_dec(x_91); -x_95 = lean_box(0); -x_2 = x_95; -x_3 = x_93; -goto _start; -} -else -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; uint64_t x_105; uint64_t x_106; uint64_t x_107; size_t x_108; size_t x_109; size_t x_110; size_t x_111; size_t x_112; lean_object* x_113; uint8_t x_114; -x_97 = lean_ctor_get(x_4, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_4, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_4)) { - lean_ctor_release(x_4, 0); - lean_ctor_release(x_4, 1); - x_99 = x_4; -} else { - lean_dec_ref(x_4); - x_99 = lean_box(0); -} -x_100 = lean_array_get_size(x_98); -x_101 = l_Lean_Name_hash___override(x_91); -x_102 = 32; -x_103 = lean_uint64_shift_right(x_101, x_102); -x_104 = lean_uint64_xor(x_101, x_103); -x_105 = 16; -x_106 = lean_uint64_shift_right(x_104, x_105); -x_107 = lean_uint64_xor(x_104, x_106); -x_108 = lean_uint64_to_usize(x_107); -x_109 = lean_usize_of_nat(x_100); -lean_dec(x_100); -x_110 = 1; -x_111 = lean_usize_sub(x_109, x_110); -x_112 = lean_usize_land(x_108, x_111); -x_113 = lean_array_uget(x_98, x_112); -x_114 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_91, x_113); -if (x_114 == 0) -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_115 = lean_unsigned_to_nat(1u); -x_116 = lean_nat_add(x_97, x_115); -lean_dec(x_97); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_91); -lean_ctor_set(x_117, 1, x_92); -lean_ctor_set(x_117, 2, x_113); -x_118 = lean_array_uset(x_98, x_112, x_117); -x_119 = lean_unsigned_to_nat(4u); -x_120 = lean_nat_mul(x_116, x_119); -x_121 = lean_unsigned_to_nat(3u); -x_122 = lean_nat_div(x_120, x_121); -lean_dec(x_120); -x_123 = lean_array_get_size(x_118); -x_124 = lean_nat_dec_le(x_122, x_123); -lean_dec(x_123); -lean_dec(x_122); -if (x_124 == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; -x_125 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_118); -if (lean_is_scalar(x_99)) { - x_126 = lean_alloc_ctor(0, 2, 0); -} else { - x_126 = x_99; -} -lean_ctor_set(x_126, 0, x_116); -lean_ctor_set(x_126, 1, x_125); -x_127 = lean_box(0); -x_2 = x_127; -x_3 = x_93; -x_4 = x_126; -goto _start; -} -else -{ -lean_object* x_129; lean_object* x_130; -if (lean_is_scalar(x_99)) { - x_129 = lean_alloc_ctor(0, 2, 0); -} else { - x_129 = x_99; -} -lean_ctor_set(x_129, 0, x_116); -lean_ctor_set(x_129, 1, x_118); -x_130 = lean_box(0); -x_2 = x_130; -x_3 = x_93; -x_4 = x_129; -goto _start; -} -} -else -{ -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_132 = lean_box(0); -x_133 = lean_array_uset(x_98, x_112, x_132); -x_134 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_91, x_92, x_113); -x_135 = lean_array_uset(x_133, x_112, x_134); -if (lean_is_scalar(x_99)) { - x_136 = lean_alloc_ctor(0, 2, 0); -} else { - x_136 = x_99; -} -lean_ctor_set(x_136, 0, x_97); -lean_ctor_set(x_136, 1, x_135); -x_137 = lean_box(0); -x_2 = x_137; -x_3 = x_93; -x_4 = x_136; -goto _start; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_3, x_4); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; -lean_dec(x_5); -x_8 = lean_array_uget(x_2, x_3); -x_9 = lean_box(0); -lean_inc(x_1); -x_10 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(x_1, x_9, x_8, x_6); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = 1; -x_14 = lean_usize_add(x_3, x_13); -x_3 = x_14; -x_5 = x_11; -x_6 = x_12; -goto _start; -} -else -{ -lean_object* x_16; -lean_dec(x_1); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_5); -lean_ctor_set(x_16, 1, x_6); -return x_16; -} -} -} -static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Server_Completion_eligibleHeaderDeclsRef; -return x_1; -} -} -static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_unsigned_to_nat(1u); -x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2; -x_3 = lean_mk_array(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1; -x_4 = lean_st_ref_take(x_3, x_2); -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_6 = lean_ctor_get(x_4, 1); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -lean_dec(x_8); -x_10 = lean_array_get_size(x_9); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_nat_dec_lt(x_11, x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_1); -x_13 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; -x_14 = lean_st_ref_set(x_3, x_13, x_6); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -x_17 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -lean_ctor_set(x_14, 0, x_17); -return x_14; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_14, 1); -lean_inc(x_18); -lean_dec(x_14); -x_19 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -uint8_t x_21; -x_21 = lean_nat_dec_le(x_10, x_10); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_1); -x_22 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; -x_23 = lean_st_ref_set(x_3, x_22, x_6); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -lean_dec(x_25); -x_26 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -lean_ctor_set(x_23, 0, x_26); -return x_23; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_23, 1); -lean_inc(x_27); -lean_dec(x_23); -x_28 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -else -{ -size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_30 = 0; -x_31 = lean_usize_of_nat(x_10); -lean_dec(x_10); -x_32 = lean_box(0); -x_33 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; -x_34 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(x_1, x_9, x_30, x_31, x_32, x_33); -lean_dec(x_9); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -lean_inc(x_35); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_st_ref_set(x_3, x_36, x_6); -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) -{ -lean_object* x_39; -x_39 = lean_ctor_get(x_37, 0); -lean_dec(x_39); -lean_ctor_set(x_37, 0, x_35); -return x_37; -} -else -{ -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_35); -lean_ctor_set(x_41, 1, x_40); -return x_41; -} -} -} -} -else -{ -lean_object* x_42; uint8_t x_43; -lean_dec(x_1); -x_42 = lean_ctor_get(x_4, 1); -lean_inc(x_42); -lean_dec(x_4); -x_43 = !lean_is_exclusive(x_5); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_44 = lean_ctor_get(x_5, 0); -lean_inc(x_44); -x_45 = lean_st_ref_set(x_3, x_5, x_42); -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) -{ -lean_object* x_47; -x_47 = lean_ctor_get(x_45, 0); -lean_dec(x_47); -lean_ctor_set(x_45, 0, x_44); -return x_45; -} -else -{ -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_45, 1); -lean_inc(x_48); -lean_dec(x_45); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_44); -lean_ctor_set(x_49, 1, x_48); -return x_49; -} -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_50 = lean_ctor_get(x_5, 0); -lean_inc(x_50); -lean_dec(x_5); -lean_inc(x_50); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -x_52 = lean_st_ref_set(x_3, x_51, x_42); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_54 = x_52; -} else { - lean_dec_ref(x_52); - x_54 = lean_box(0); -} -if (lean_is_scalar(x_54)) { - x_55 = lean_alloc_ctor(0, 2, 0); -} else { - x_55 = x_54; -} -lean_ctor_set(x_55, 0, x_50); -lean_ctor_set(x_55, 1, x_53); -return x_55; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(x_1, x_2, x_7, x_8, x_5, x_6); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(x_1, x_2, x_4, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -lean_dec(x_2); -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_apply_2(x_6, lean_box(0), x_3); -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_3); -x_8 = lean_ctor_get(x_4, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 2); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -lean_inc(x_2); -x_12 = lean_apply_2(x_2, x_8, x_9); -x_13 = lean_alloc_closure((void*)(l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1), 4, 3); -lean_closure_set(x_13, 0, x_1); -lean_closure_set(x_13, 1, x_2); -lean_closure_set(x_13, 2, x_10); -x_14 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_12, x_13); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg), 4, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_PersistentHashMap_foldlMAux___rarg(x_1, lean_box(0), lean_box(0), lean_box(0), x_3, x_2, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__3(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg), 4, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_apply_2(x_1, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed), 4, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = lean_box(0); -x_6 = l_Lean_PersistentHashMap_foldlMAux___rarg(x_1, lean_box(0), lean_box(0), lean_box(0), x_4, x_2, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = 1; -x_8 = lean_usize_add(x_1, x_7); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_2, x_3, x_4, x_8, x_5, x_6); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_6); -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -x_9 = lean_array_uget(x_3, x_4); -x_10 = lean_box(0); -lean_inc(x_2); -lean_inc(x_1); -x_11 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(x_1, x_2, x_10, x_9); -x_12 = lean_box_usize(x_4); -x_13 = lean_box_usize(x_5); -x_14 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed), 6, 5); -lean_closure_set(x_14, 0, x_12); -lean_closure_set(x_14, 1, x_1); -lean_closure_set(x_14, 2, x_2); -lean_closure_set(x_14, 3, x_3); -lean_closure_set(x_14, 4, x_13); -x_15 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_11, x_14); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_3); -lean_dec(x_2); -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -lean_dec(x_1); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_apply_2(x_17, lean_box(0), x_6); -return x_18; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed), 6, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -lean_inc(x_4); -x_6 = l_Lean_Meta_allowCompletion(x_1, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_7 = lean_ctor_get(x_2, 0); -lean_inc(x_7); -lean_dec(x_2); -x_8 = lean_ctor_get(x_7, 1); -lean_inc(x_8); -lean_dec(x_7); -x_9 = lean_box(0); -x_10 = lean_apply_2(x_8, lean_box(0), x_9); -return x_10; -} -else -{ -lean_object* x_11; -lean_dec(x_2); -x_11 = lean_apply_2(x_3, x_4, x_5); -return x_11; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -lean_inc(x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1), 5, 3); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_3); -x_8 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(x_2, x_6, x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_7); -lean_inc(x_3); -lean_inc(x_2); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_10, 0, x_1); -lean_closure_set(x_10, 1, x_2); -lean_closure_set(x_10, 2, x_3); -if (x_9 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -lean_dec(x_2); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_box(0); -x_14 = lean_apply_2(x_12, lean_box(0), x_13); -x_15 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_14, x_10); -return x_15; -} -else -{ -uint8_t x_16; -x_16 = lean_nat_dec_le(x_7, x_7); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); -x_17 = lean_ctor_get(x_2, 0); -lean_inc(x_17); -lean_dec(x_2); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_box(0); -x_20 = lean_apply_2(x_18, lean_box(0), x_19); -x_21 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_20, x_10); -return x_21; -} -else -{ -size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = 0; -x_23 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_24 = lean_box(0); -x_25 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_2, x_3, x_6, x_22, x_23, x_24); -x_26 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_25, x_10); -return x_26; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_Server_Completion_getEligibleHeaderDecls), 2, 1); -lean_closure_set(x_6, 0, x_5); -x_7 = lean_apply_2(x_1, lean_box(0), x_6); -lean_inc(x_4); -x_8 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3), 5, 4); -lean_closure_set(x_8, 0, x_5); -lean_closure_set(x_8, 1, x_2); -lean_closure_set(x_8, 2, x_3); -lean_closure_set(x_8, 3, x_4); -x_9 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_ctor_get(x_2, 0); -lean_inc(x_7); -lean_dec(x_2); -lean_inc(x_6); -x_8 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4), 5, 4); -lean_closure_set(x_8, 0, x_4); -lean_closure_set(x_8, 1, x_1); -lean_closure_set(x_8, 2, x_5); -lean_closure_set(x_8, 3, x_6); -x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___boxed), 5, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(x_7, x_2, x_3, x_4, x_8, x_6); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_1, x_2, x_3, x_7, x_8, x_6); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(x_1, x_2, x_3, x_4); -lean_dec(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -return x_6; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; uint8_t x_19; -x_4 = lean_ctor_get(x_1, 1); -x_5 = lean_array_get_size(x_4); -x_6 = l_Lean_Name_hash___override(x_3); -x_7 = 32; -x_8 = lean_uint64_shift_right(x_6, x_7); -x_9 = lean_uint64_xor(x_6, x_8); -x_10 = 16; -x_11 = lean_uint64_shift_right(x_9, x_10); -x_12 = lean_uint64_xor(x_9, x_11); -x_13 = lean_uint64_to_usize(x_12); -x_14 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_15 = 1; -x_16 = lean_usize_sub(x_14, x_15); -x_17 = lean_usize_land(x_13, x_16); -x_18 = lean_array_uget(x_4, x_17); -x_19 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_3, x_18); -lean_dec(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_2, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_21, x_3); -if (x_22 == 0) -{ -uint8_t x_23; -lean_dec(x_3); -lean_dec(x_2); -x_23 = 0; -return x_23; -} -else -{ -uint8_t x_24; -x_24 = l_Lean_Meta_allowCompletion(x_2, x_3); -return x_24; -} -} -else -{ -uint8_t x_25; -lean_dec(x_3); -lean_dec(x_2); -x_25 = 1; -return x_25; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(x_1, x_2, x_3); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = lean_string_utf8_at_end(x_2, x_1); -if (x_3 == 0) -{ -uint32_t x_4; lean_object* x_5; uint32_t x_6; lean_object* x_7; lean_object* x_8; -x_4 = lean_string_utf8_get(x_2, x_1); -x_5 = l_Char_toLower(x_4); -x_6 = lean_unbox_uint32(x_5); -lean_dec(x_5); -x_7 = lean_string_utf8_set(x_2, x_1, x_6); -x_8 = lean_string_utf8_next(x_7, x_1); -lean_dec(x_1); -x_1 = x_8; -x_2 = x_7; -goto _start; -} -else -{ -lean_dec(x_1); -return x_2; -} -} -} -LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; double x_7; double x_8; uint8_t x_9; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_unbox_float(x_4); -x_8 = lean_unbox_float(x_6); -x_9 = lean_float_beq(x_7, x_8); -if (x_9 == 0) -{ -double x_10; double x_11; uint8_t x_12; -lean_dec(x_5); -lean_dec(x_3); -x_10 = lean_unbox_float(x_6); -lean_dec(x_6); -x_11 = lean_unbox_float(x_4); -lean_dec(x_4); -x_12 = lean_float_decLt(x_10, x_11); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_6); -lean_dec(x_4); -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -lean_dec(x_3); -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(x_14, x_13); -x_16 = lean_ctor_get(x_5, 0); -lean_inc(x_16); -lean_dec(x_5); -x_17 = l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(x_14, x_16); -x_18 = lean_string_dec_lt(x_15, x_17); -lean_dec(x_17); -lean_dec(x_15); -return x_18; -} -} -} -static lean_object* _init_l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1___boxed), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_nat_dec_lt(x_2, x_3); -if (x_4 == 0) -{ -lean_dec(x_2); -return x_1; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1; -lean_inc(x_2); -x_6 = l_Array_qpartition___rarg(x_1, x_5, x_2, x_3); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = lean_nat_dec_le(x_3, x_7); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_8, x_2, x_7); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_7, x_11); -lean_dec(x_7); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -else -{ -lean_dec(x_7); -lean_dec(x_2); -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = lean_usize_dec_lt(x_2, x_1); -if (x_4 == 0) -{ -return x_3; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; -x_5 = lean_array_uget(x_3, x_2); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = lean_ctor_get(x_5, 0); -lean_inc(x_8); -lean_dec(x_5); -x_9 = 1; -x_10 = lean_usize_add(x_2, x_9); -x_11 = lean_array_uset(x_7, x_2, x_8); -x_2 = x_10; -x_3 = x_11; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; -x_2 = lean_array_get_size(x_1); -x_3 = lean_unsigned_to_nat(1u); -x_4 = lean_nat_sub(x_2, x_3); -lean_dec(x_2); -x_5 = lean_unsigned_to_nat(0u); -x_6 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_1, x_5, x_4); -lean_dec(x_4); -x_7 = lean_array_size(x_6); -x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(x_7, x_8, x_6); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; size_t x_5; lean_object* x_6; -x_4 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(x_4, x_5, x_3); -return x_6; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_4); -lean_ctor_set(x_11, 1, x_3); -x_12 = !lean_is_exclusive(x_1); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_13 = lean_ctor_get(x_1, 6); -lean_dec(x_13); -x_14 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(x_11); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_1, 6, x_15); -x_16 = lean_st_ref_take(x_5, x_10); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -x_20 = lean_box_float(x_2); -lean_ctor_set(x_16, 1, x_20); -lean_ctor_set(x_16, 0, x_1); -x_21 = lean_array_push(x_18, x_16); -x_22 = lean_st_ref_set(x_5, x_21, x_19); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_22, 0); -lean_dec(x_24); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_22, 0, x_25); -return x_22; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_dec(x_22); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_26); -return x_28; -} -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_29 = lean_ctor_get(x_16, 0); -x_30 = lean_ctor_get(x_16, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_16); -x_31 = lean_box_float(x_2); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_array_push(x_29, x_32); -x_34 = lean_st_ref_set(x_5, x_33, x_30); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_36 = x_34; -} else { - lean_dec_ref(x_34); - x_36 = lean_box(0); -} -x_37 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -if (lean_is_scalar(x_36)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_36; -} -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_35); -return x_38; -} -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_39 = lean_ctor_get(x_1, 0); -x_40 = lean_ctor_get(x_1, 1); -x_41 = lean_ctor_get(x_1, 2); -x_42 = lean_ctor_get(x_1, 3); -x_43 = lean_ctor_get(x_1, 4); -x_44 = lean_ctor_get(x_1, 5); -x_45 = lean_ctor_get(x_1, 7); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_1); -x_46 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(x_11); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_48, 0, x_39); -lean_ctor_set(x_48, 1, x_40); -lean_ctor_set(x_48, 2, x_41); -lean_ctor_set(x_48, 3, x_42); -lean_ctor_set(x_48, 4, x_43); -lean_ctor_set(x_48, 5, x_44); -lean_ctor_set(x_48, 6, x_47); -lean_ctor_set(x_48, 7, x_45); -x_49 = lean_st_ref_take(x_5, x_10); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_52 = x_49; -} else { - lean_dec_ref(x_49); - x_52 = lean_box(0); -} -x_53 = lean_box_float(x_2); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_52; -} -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_53); -x_55 = lean_array_push(x_50, x_54); -x_56 = lean_st_ref_set(x_5, x_55, x_51); -x_57 = lean_ctor_get(x_56, 1); -lean_inc(x_57); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_58 = x_56; -} else { - lean_dec_ref(x_56); - x_58 = lean_box(0); -} -x_59 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -if (lean_is_scalar(x_58)) { - x_60 = lean_alloc_ctor(0, 2, 0); -} else { - x_60 = x_58; -} -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_57); -return x_60; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -double x_11; lean_object* x_12; -x_11 = lean_unbox_float(x_2); -lean_dec(x_2); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_2 = 1; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); -x_4 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_4, 0, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("\n\n", 2, 2); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, double x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1; -x_16 = 1; -x_17 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_18 = l_Lean_Name_toString(x_1, x_16, x_17); -x_19 = lean_box(0); -x_20 = lean_box(x_2); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_3); -if (lean_obj_tag(x_4) == 0) -{ -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_23, 0, x_18); -lean_ctor_set(x_23, 1, x_19); -lean_ctor_set(x_23, 2, x_19); -lean_ctor_set(x_23, 3, x_21); -lean_ctor_set(x_23, 4, x_19); -lean_ctor_set(x_23, 5, x_19); -lean_ctor_set(x_23, 6, x_19); -lean_ctor_set(x_23, 7, x_5); -x_24 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_23, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_24; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_7, 0); -lean_inc(x_25); -lean_dec(x_7); -x_26 = lean_apply_1(x_15, x_25); -x_27 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_27, 0, x_18); -lean_ctor_set(x_27, 1, x_19); -lean_ctor_set(x_27, 2, x_26); -lean_ctor_set(x_27, 3, x_21); -lean_ctor_set(x_27, 4, x_19); -lean_ctor_set(x_27, 5, x_19); -lean_ctor_set(x_27, 6, x_19); -lean_ctor_set(x_27, 7, x_5); -x_28 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_27, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_28; -} -} -else -{ -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_4, 0); -lean_inc(x_29); -lean_dec(x_4); -x_30 = lean_apply_1(x_15, x_29); -x_31 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_31, 0, x_18); -lean_ctor_set(x_31, 1, x_19); -lean_ctor_set(x_31, 2, x_30); -lean_ctor_set(x_31, 3, x_21); -lean_ctor_set(x_31, 4, x_19); -lean_ctor_set(x_31, 5, x_19); -lean_ctor_set(x_31, 6, x_19); -lean_ctor_set(x_31, 7, x_5); -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_31, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_32; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_33 = lean_ctor_get(x_4, 0); -lean_inc(x_33); -lean_dec(x_4); -x_34 = lean_ctor_get(x_7, 0); -lean_inc(x_34); -lean_dec(x_7); -x_35 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_36 = lean_string_append(x_35, x_33); -lean_dec(x_33); -x_37 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3; -x_38 = lean_string_append(x_36, x_37); -x_39 = lean_string_append(x_38, x_34); -lean_dec(x_34); -x_40 = lean_string_append(x_39, x_35); -x_41 = lean_apply_1(x_15, x_40); -x_42 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_42, 0, x_18); -lean_ctor_set(x_42, 1, x_19); -lean_ctor_set(x_42, 2, x_41); -lean_ctor_set(x_42, 3, x_21); -lean_ctor_set(x_42, 4, x_19); -lean_ctor_set(x_42, 5, x_19); -lean_ctor_set(x_42, 6, x_19); -lean_ctor_set(x_42, 7, x_5); -x_43 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_42, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_43; -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Linter_deprecatedAttr; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2; -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("`", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` has been deprecated.", 22, 22); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` has been deprecated, use `", 28, 28); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("` instead.", 10, 10); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(lean_object* x_1, lean_object* x_2, uint8_t x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_12 = lean_st_ref_get(x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_27 = lean_ctor_get(x_2, 0); -lean_inc(x_27); -x_28 = l_Lean_Linter_instInhabitedDeprecationEntry; -x_29 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1; -lean_inc(x_27); -x_30 = l_Lean_ParametricAttribute_getParam_x3f___rarg(x_28, x_29, x_15, x_27); -if (lean_obj_tag(x_30) == 0) -{ -lean_object* x_31; -lean_dec(x_27); -x_31 = lean_box(0); -x_16 = x_31; -x_17 = x_31; -goto block_26; -} -else -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_30, 0); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_35); -lean_dec(x_33); -if (lean_obj_tag(x_35) == 0) -{ -uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_36 = 1; -x_37 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_38 = l_Lean_Name_toString(x_27, x_36, x_37); -x_39 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -x_40 = lean_string_append(x_39, x_38); -lean_dec(x_38); -x_41 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; -x_42 = lean_string_append(x_40, x_41); -lean_ctor_set(x_30, 0, x_42); -x_43 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_30; -x_17 = x_43; -goto block_26; -} -else -{ -uint8_t x_44; -lean_free_object(x_30); -x_44 = !lean_is_exclusive(x_35); -if (x_44 == 0) -{ -lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_45 = lean_ctor_get(x_35, 0); -x_46 = 1; -x_47 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_48 = l_Lean_Name_toString(x_27, x_46, x_47); -x_49 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -x_50 = lean_string_append(x_49, x_48); -lean_dec(x_48); -x_51 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; -x_52 = lean_string_append(x_50, x_51); -x_53 = l_Lean_Name_toString(x_45, x_46, x_47); -x_54 = lean_string_append(x_52, x_53); -lean_dec(x_53); -x_55 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; -x_56 = lean_string_append(x_54, x_55); -lean_ctor_set(x_35, 0, x_56); -x_57 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_35; -x_17 = x_57; -goto block_26; -} -else -{ -lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_58 = lean_ctor_get(x_35, 0); -lean_inc(x_58); -lean_dec(x_35); -x_59 = 1; -x_60 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_61 = l_Lean_Name_toString(x_27, x_59, x_60); -x_62 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -x_63 = lean_string_append(x_62, x_61); -lean_dec(x_61); -x_64 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; -x_65 = lean_string_append(x_63, x_64); -x_66 = l_Lean_Name_toString(x_58, x_59, x_60); -x_67 = lean_string_append(x_65, x_66); -lean_dec(x_66); -x_68 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; -x_69 = lean_string_append(x_67, x_68); -x_70 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_70, 0, x_69); -x_71 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_70; -x_17 = x_71; -goto block_26; -} -} -} -else -{ -uint8_t x_72; -lean_free_object(x_30); -lean_dec(x_33); -lean_dec(x_27); -x_72 = !lean_is_exclusive(x_34); -if (x_72 == 0) -{ -lean_object* x_73; -x_73 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_34; -x_17 = x_73; -goto block_26; -} -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_34, 0); -lean_inc(x_74); -lean_dec(x_34); -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_74); -x_76 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_75; -x_17 = x_76; -goto block_26; -} -} -} -else -{ -lean_object* x_77; lean_object* x_78; -x_77 = lean_ctor_get(x_30, 0); -lean_inc(x_77); -lean_dec(x_30); -x_78 = lean_ctor_get(x_77, 1); -lean_inc(x_78); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -x_79 = lean_ctor_get(x_77, 0); -lean_inc(x_79); -lean_dec(x_77); -if (lean_obj_tag(x_79) == 0) -{ -uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_80 = 1; -x_81 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_82 = l_Lean_Name_toString(x_27, x_80, x_81); -x_83 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -x_84 = lean_string_append(x_83, x_82); -lean_dec(x_82); -x_85 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; -x_86 = lean_string_append(x_84, x_85); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -x_88 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_87; -x_17 = x_88; -goto block_26; -} -else -{ -lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_89 = lean_ctor_get(x_79, 0); -lean_inc(x_89); -if (lean_is_exclusive(x_79)) { - lean_ctor_release(x_79, 0); - x_90 = x_79; -} else { - lean_dec_ref(x_79); - x_90 = lean_box(0); -} -x_91 = 1; -x_92 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_93 = l_Lean_Name_toString(x_27, x_91, x_92); -x_94 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; -x_95 = lean_string_append(x_94, x_93); -lean_dec(x_93); -x_96 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; -x_97 = lean_string_append(x_95, x_96); -x_98 = l_Lean_Name_toString(x_89, x_91, x_92); -x_99 = lean_string_append(x_97, x_98); -lean_dec(x_98); -x_100 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; -x_101 = lean_string_append(x_99, x_100); -if (lean_is_scalar(x_90)) { - x_102 = lean_alloc_ctor(1, 1, 0); -} else { - x_102 = x_90; -} -lean_ctor_set(x_102, 0, x_101); -x_103 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_102; -x_17 = x_103; -goto block_26; -} -} -else -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_77); -lean_dec(x_27); -x_104 = lean_ctor_get(x_78, 0); -lean_inc(x_104); -if (lean_is_exclusive(x_78)) { - lean_ctor_release(x_78, 0); - x_105 = x_78; -} else { - lean_dec_ref(x_78); - x_105 = lean_box(0); -} -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 1, 0); -} else { - x_106 = x_105; -} -lean_ctor_set(x_106, 0, x_104); -x_107 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; -x_16 = x_106; -x_17 = x_107; -goto block_26; -} -} -} -} -else -{ -lean_object* x_108; -x_108 = lean_box(0); -x_16 = x_108; -x_17 = x_108; -goto block_26; -} -block_26: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_18 = lean_ctor_get(x_2, 0); -lean_inc(x_18); -x_19 = 1; -x_20 = l_Lean_findDocString_x3f(x_15, x_18, x_19, x_14); -lean_dec(x_15); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(x_1, x_3, x_2, x_16, x_17, x_4, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -return x_23; -} -else -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_15); -x_24 = lean_box(0); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(x_1, x_3, x_2, x_16, x_17, x_4, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_25; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; double x_16; lean_object* x_17; -x_15 = lean_unbox(x_2); -lean_dec(x_2); -x_16 = lean_unbox_float(x_6); -lean_dec(x_6); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(x_1, x_15, x_3, x_4, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -return x_17; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; double x_13; lean_object* x_14; -x_12 = lean_unbox(x_3); -lean_dec(x_3); -x_13 = lean_unbox_float(x_4); -lean_dec(x_4); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_14; -} -} -static lean_object* _init_l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_projectionFnInfoExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_instInhabitedProjectionFunctionInfo; -x_14 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; -x_15 = l_Lean_MapDeclarationExtension_contains___rarg(x_13, x_14, x_12, x_1); -lean_dec(x_12); -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_9, 0, x_17); -return x_9; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_18 = lean_ctor_get(x_9, 0); -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_9); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_instInhabitedProjectionFunctionInfo; -x_22 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; -x_23 = l_Lean_MapDeclarationExtension_contains___rarg(x_21, x_22, x_20, x_1); -lean_dec(x_20); -x_24 = lean_box(x_23); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_19); -return x_26; -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_6, 5); -x_10 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_9); -lean_ctor_set(x_13, 1, x_12); -lean_ctor_set_tag(x_10, 1); -lean_ctor_set(x_10, 0, x_13); -return x_10; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_10, 0); -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_10); -lean_inc(x_9); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_14); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -} -} -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unknown constant '", 18, 18); -return x_1; -} -} -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_free_object(x_9); -x_15 = 0; -x_16 = l_Lean_MessageData_ofConstName(x_1, x_15); -x_17 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -return x_21; -} -else -{ -uint8_t x_22; -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_14); -if (x_22 == 0) -{ -lean_ctor_set(x_9, 0, x_14); -return x_9; -} -else -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_14, 0); -lean_inc(x_23); -lean_dec(x_14); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_9, 0, x_24); -return x_9; -} -} -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_9, 0); -x_26 = lean_ctor_get(x_9, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_9); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -lean_inc(x_1); -x_28 = lean_environment_find(x_27, x_1); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_29 = 0; -x_30 = l_Lean_MessageData_ofConstName(x_1, x_29); -x_31 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_26); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_1); -x_36 = lean_ctor_get(x_28, 0); -lean_inc(x_36); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - x_37 = x_28; -} else { - lean_dec_ref(x_28); - x_37 = lean_box(0); -} -if (lean_is_scalar(x_37)) { - x_38 = lean_alloc_ctor(1, 1, 0); -} else { - x_38 = x_37; -} -lean_ctor_set(x_38, 0, x_36); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_26); -return x_39; -} -} -} -} -static lean_object* _init_l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 1; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_22; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - x_13 = x_1; -} else { - lean_dec_ref(x_1); - x_13 = lean_box(0); -} -x_22 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) -{ -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); -if (lean_obj_tag(x_25) == 6) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_26 = lean_ctor_get(x_22, 1); -lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_27, 4); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_unsigned_to_nat(0u); -x_30 = lean_nat_dec_eq(x_28, x_29); -lean_dec(x_28); -x_31 = lean_box(x_30); -lean_ctor_set(x_23, 0, x_31); -x_14 = x_23; -x_15 = x_26; -goto block_21; -} -else -{ -lean_object* x_32; lean_object* x_33; -lean_free_object(x_23); -lean_dec(x_25); -x_32 = lean_ctor_get(x_22, 1); -lean_inc(x_32); -lean_dec(x_22); -x_33 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_14 = x_33; -x_15 = x_32; -goto block_21; -} -} -else -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_23, 0); -lean_inc(x_34); -lean_dec(x_23); -if (lean_obj_tag(x_34) == 6) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; -x_35 = lean_ctor_get(x_22, 1); -lean_inc(x_35); -lean_dec(x_22); -x_36 = lean_ctor_get(x_34, 0); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_36, 4); -lean_inc(x_37); -lean_dec(x_36); -x_38 = lean_unsigned_to_nat(0u); -x_39 = lean_nat_dec_eq(x_37, x_38); -lean_dec(x_37); -x_40 = lean_box(x_39); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_40); -x_14 = x_41; -x_15 = x_35; -goto block_21; -} -else -{ -lean_object* x_42; lean_object* x_43; -lean_dec(x_34); -x_42 = lean_ctor_get(x_22, 1); -lean_inc(x_42); -lean_dec(x_22); -x_43 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_14 = x_43; -x_15 = x_42; -goto block_21; -} -} -} -else -{ -uint8_t x_44; -lean_dec(x_13); -lean_dec(x_12); -x_44 = !lean_is_exclusive(x_22); -if (x_44 == 0) -{ -return x_22; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_22, 0); -x_46 = lean_ctor_get(x_22, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_22); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} -} -block_21: -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_12); -x_18 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -if (lean_is_scalar(x_13)) { - x_19 = lean_alloc_ctor(0, 2, 0); -} else { - x_19 = x_13; - lean_ctor_set_tag(x_19, 0); -} -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_15); -return x_19; -} -else -{ -lean_dec(x_13); -x_1 = x_12; -x_8 = x_15; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -if (lean_obj_tag(x_11) == 5) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_13 = lean_ctor_get(x_9, 1); -x_14 = lean_ctor_get(x_9, 0); -lean_dec(x_14); -x_15 = lean_ctor_get(x_11, 0); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_16, 2); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_Expr_isProp(x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = l_Lean_InductiveVal_numTypeFormers(x_15); -x_20 = lean_unsigned_to_nat(1u); -x_21 = lean_nat_dec_eq(x_19, x_20); -lean_dec(x_19); -if (x_21 == 0) -{ -lean_object* x_22; -lean_dec(x_15); -x_22 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_22); -return x_9; -} -else -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_15, 2); -lean_inc(x_23); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_nat_dec_eq(x_23, x_24); -lean_dec(x_23); -if (x_25 == 0) -{ -lean_object* x_26; -lean_dec(x_15); -x_26 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_26); -return x_9; -} -else -{ -lean_object* x_27; uint8_t x_28; -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -x_28 = lean_nat_dec_eq(x_27, x_24); -lean_dec(x_27); -if (x_28 == 0) -{ -lean_object* x_29; -lean_dec(x_15); -x_29 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_29); -return x_9; -} -else -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_15, 4); -lean_inc(x_30); -x_31 = l_List_isEmpty___rarg(x_30); -if (x_31 == 0) -{ -uint8_t x_32; -x_32 = lean_ctor_get_uint8(x_15, sizeof(void*)*6); -if (x_32 == 0) -{ -uint8_t x_33; -x_33 = lean_ctor_get_uint8(x_15, sizeof(void*)*6 + 1); -lean_dec(x_15); -if (x_33 == 0) -{ -lean_object* x_34; -lean_free_object(x_9); -x_34 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_13); -return x_34; -} -else -{ -lean_object* x_35; -lean_dec(x_30); -x_35 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_35); -return x_9; -} -} -else -{ -lean_object* x_36; -lean_dec(x_30); -lean_dec(x_15); -x_36 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_36); -return x_9; -} -} -else -{ -lean_object* x_37; -lean_dec(x_30); -lean_dec(x_15); -x_37 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_37); -return x_9; -} -} -} -} -} -else -{ -lean_object* x_38; -lean_dec(x_15); -x_38 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_38); -return x_9; -} -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_39 = lean_ctor_get(x_9, 1); -lean_inc(x_39); -lean_dec(x_9); -x_40 = lean_ctor_get(x_11, 0); -lean_inc(x_40); -lean_dec(x_11); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_41, 2); -lean_inc(x_42); -lean_dec(x_41); -x_43 = l_Lean_Expr_isProp(x_42); -lean_dec(x_42); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_44 = l_Lean_InductiveVal_numTypeFormers(x_40); -x_45 = lean_unsigned_to_nat(1u); -x_46 = lean_nat_dec_eq(x_44, x_45); -lean_dec(x_44); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; -lean_dec(x_40); -x_47 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_39); -return x_48; -} -else -{ -lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_49 = lean_ctor_get(x_40, 2); -lean_inc(x_49); -x_50 = lean_unsigned_to_nat(0u); -x_51 = lean_nat_dec_eq(x_49, x_50); -lean_dec(x_49); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_40); -x_52 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_39); -return x_53; -} -else -{ -lean_object* x_54; uint8_t x_55; -x_54 = lean_ctor_get(x_40, 1); -lean_inc(x_54); -x_55 = lean_nat_dec_eq(x_54, x_50); -lean_dec(x_54); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; -lean_dec(x_40); -x_56 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_39); -return x_57; -} -else -{ -lean_object* x_58; uint8_t x_59; -x_58 = lean_ctor_get(x_40, 4); -lean_inc(x_58); -x_59 = l_List_isEmpty___rarg(x_58); -if (x_59 == 0) -{ -uint8_t x_60; -x_60 = lean_ctor_get_uint8(x_40, sizeof(void*)*6); -if (x_60 == 0) -{ -uint8_t x_61; -x_61 = lean_ctor_get_uint8(x_40, sizeof(void*)*6 + 1); -lean_dec(x_40); -if (x_61 == 0) -{ -lean_object* x_62; -x_62 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_58, x_2, x_3, x_4, x_5, x_6, x_7, x_39); -return x_62; -} -else -{ -lean_object* x_63; lean_object* x_64; -lean_dec(x_58); -x_63 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_39); -return x_64; -} -} -else -{ -lean_object* x_65; lean_object* x_66; -lean_dec(x_58); -lean_dec(x_40); -x_65 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_39); -return x_66; -} -} -else -{ -lean_object* x_67; lean_object* x_68; -lean_dec(x_58); -lean_dec(x_40); -x_67 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_39); -return x_68; -} -} -} -} -} -else -{ -lean_object* x_69; lean_object* x_70; -lean_dec(x_40); -x_69 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_39); -return x_70; -} -} -} -else -{ -uint8_t x_71; -lean_dec(x_11); -x_71 = !lean_is_exclusive(x_9); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_9, 0); -lean_dec(x_72); -x_73 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -lean_ctor_set(x_9, 0, x_73); -return x_9; -} -else -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_74 = lean_ctor_get(x_9, 1); -lean_inc(x_74); -lean_dec(x_9); -x_75 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2; -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -return x_76; -} -} -} -else -{ -uint8_t x_77; -x_77 = !lean_is_exclusive(x_9); -if (x_77 == 0) -{ -return x_9; -} -else -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_9, 0); -x_79 = lean_ctor_get(x_9, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_9); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 20; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 2; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 4; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 22; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 21; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 12; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 6; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 3; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_ConstantInfo_isCtor(x_1); -if (x_14 == 0) -{ -uint8_t x_15; -x_15 = l_Lean_ConstantInfo_isInductive(x_1); -if (x_15 == 0) -{ -uint8_t x_16; -lean_dec(x_13); -x_16 = l_Lean_ConstantInfo_isTheorem(x_1); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -lean_free_object(x_9); -x_17 = l_Lean_ConstantInfo_name(x_1); -x_18 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = l_Lean_ConstantInfo_type(x_1); -x_24 = !lean_is_exclusive(x_6); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_6, 9); -lean_dec(x_25); -x_26 = lean_unsigned_to_nat(0u); -lean_ctor_set(x_6, 9, x_26); -x_27 = lean_whnf(x_23, x_4, x_5, x_6, x_7, x_22); -if (lean_obj_tag(x_27) == 0) -{ -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; uint8_t x_30; -x_29 = lean_ctor_get(x_27, 0); -x_30 = l_Lean_Expr_isForall(x_29); -lean_dec(x_29); -if (x_30 == 0) -{ -lean_object* x_31; -x_31 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1; -lean_ctor_set(x_27, 0, x_31); -return x_27; -} -else -{ -lean_object* x_32; -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2; -lean_ctor_set(x_27, 0, x_32); -return x_27; -} -} -else -{ -lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = l_Lean_Expr_isForall(x_33); -lean_dec(x_33); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -x_36 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1; -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_34); -return x_37; -} -else -{ -lean_object* x_38; lean_object* x_39; -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_34); -return x_39; -} -} -} -else -{ -uint8_t x_40; -x_40 = !lean_is_exclusive(x_27); -if (x_40 == 0) -{ -return x_27; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_27, 0); -x_42 = lean_ctor_get(x_27, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_27); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_44 = lean_ctor_get(x_6, 0); -x_45 = lean_ctor_get(x_6, 1); -x_46 = lean_ctor_get(x_6, 2); -x_47 = lean_ctor_get(x_6, 3); -x_48 = lean_ctor_get(x_6, 4); -x_49 = lean_ctor_get(x_6, 5); -x_50 = lean_ctor_get(x_6, 6); -x_51 = lean_ctor_get(x_6, 7); -x_52 = lean_ctor_get(x_6, 8); -x_53 = lean_ctor_get(x_6, 10); -x_54 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); -x_55 = lean_ctor_get(x_6, 11); -x_56 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); -lean_inc(x_55); -lean_inc(x_53); -lean_inc(x_52); -lean_inc(x_51); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_6); -x_57 = lean_unsigned_to_nat(0u); -x_58 = lean_alloc_ctor(0, 12, 2); -lean_ctor_set(x_58, 0, x_44); -lean_ctor_set(x_58, 1, x_45); -lean_ctor_set(x_58, 2, x_46); -lean_ctor_set(x_58, 3, x_47); -lean_ctor_set(x_58, 4, x_48); -lean_ctor_set(x_58, 5, x_49); -lean_ctor_set(x_58, 6, x_50); -lean_ctor_set(x_58, 7, x_51); -lean_ctor_set(x_58, 8, x_52); -lean_ctor_set(x_58, 9, x_57); -lean_ctor_set(x_58, 10, x_53); -lean_ctor_set(x_58, 11, x_55); -lean_ctor_set_uint8(x_58, sizeof(void*)*12, x_54); -lean_ctor_set_uint8(x_58, sizeof(void*)*12 + 1, x_56); -x_59 = lean_whnf(x_23, x_4, x_5, x_58, x_7, x_22); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_62 = x_59; -} else { - lean_dec_ref(x_59); - x_62 = lean_box(0); -} -x_63 = l_Lean_Expr_isForall(x_60); -lean_dec(x_60); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; -x_64 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1; -if (lean_is_scalar(x_62)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_62; -} -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_61); -return x_65; -} -else -{ -lean_object* x_66; lean_object* x_67; -x_66 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2; -if (lean_is_scalar(x_62)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_62; -} -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_61); -return x_67; -} -} -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_59, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_59, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_70 = x_59; -} else { - lean_dec_ref(x_59); - x_70 = lean_box(0); -} -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(1, 2, 0); -} else { - x_71 = x_70; -} -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; -} -} -} -else -{ -uint8_t x_72; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_72 = !lean_is_exclusive(x_18); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_18, 0); -lean_dec(x_73); -x_74 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3; -lean_ctor_set(x_18, 0, x_74); -return x_18; -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_18, 1); -lean_inc(x_75); -lean_dec(x_18); -x_76 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3; -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_75); -return x_77; -} -} -} -else -{ -lean_object* x_78; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_78 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4; -lean_ctor_set(x_9, 0, x_78); -return x_9; -} -} -else -{ -lean_object* x_79; uint8_t x_80; -x_79 = l_Lean_ConstantInfo_name(x_1); -lean_inc(x_79); -x_80 = lean_is_class(x_13, x_79); -if (x_80 == 0) -{ -lean_object* x_81; -lean_free_object(x_9); -x_81 = l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_79, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -if (lean_obj_tag(x_81) == 0) -{ -lean_object* x_82; lean_object* x_83; uint8_t x_84; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -lean_dec(x_82); -x_84 = lean_unbox(x_83); -lean_dec(x_83); -if (x_84 == 0) -{ -uint8_t x_85; -x_85 = !lean_is_exclusive(x_81); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_81, 0); -lean_dec(x_86); -x_87 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5; -lean_ctor_set(x_81, 0, x_87); -return x_81; -} -else -{ -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_ctor_get(x_81, 1); -lean_inc(x_88); -lean_dec(x_81); -x_89 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5; -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_88); -return x_90; -} -} -else -{ -uint8_t x_91; -x_91 = !lean_is_exclusive(x_81); -if (x_91 == 0) -{ -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_81, 0); -lean_dec(x_92); -x_93 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6; -lean_ctor_set(x_81, 0, x_93); -return x_81; -} -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_94 = lean_ctor_get(x_81, 1); -lean_inc(x_94); -lean_dec(x_81); -x_95 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6; -x_96 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_94); -return x_96; -} -} -} -else -{ -uint8_t x_97; -x_97 = !lean_is_exclusive(x_81); -if (x_97 == 0) -{ -return x_81; -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_81, 0); -x_99 = lean_ctor_get(x_81, 1); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_81); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_99); -return x_100; -} -} -} -else -{ -lean_object* x_101; -lean_dec(x_79); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_101 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7; -lean_ctor_set(x_9, 0, x_101); -return x_9; -} -} -} -else -{ -lean_object* x_102; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_102 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8; -lean_ctor_set(x_9, 0, x_102); -return x_9; -} -} -else -{ -lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_103 = lean_ctor_get(x_9, 0); -x_104 = lean_ctor_get(x_9, 1); -lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_9); -x_105 = lean_ctor_get(x_103, 0); -lean_inc(x_105); -lean_dec(x_103); -x_106 = l_Lean_ConstantInfo_isCtor(x_1); -if (x_106 == 0) -{ -uint8_t x_107; -x_107 = l_Lean_ConstantInfo_isInductive(x_1); -if (x_107 == 0) -{ -uint8_t x_108; -lean_dec(x_105); -x_108 = l_Lean_ConstantInfo_isTheorem(x_1); -if (x_108 == 0) -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; -x_109 = l_Lean_ConstantInfo_name(x_1); -x_110 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_109, x_2, x_3, x_4, x_5, x_6, x_7, x_104); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -lean_dec(x_111); -x_113 = lean_unbox(x_112); -lean_dec(x_112); -if (x_113 == 0) -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_114 = lean_ctor_get(x_110, 1); -lean_inc(x_114); -lean_dec(x_110); -x_115 = l_Lean_ConstantInfo_type(x_1); -x_116 = lean_ctor_get(x_6, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_6, 1); -lean_inc(x_117); -x_118 = lean_ctor_get(x_6, 2); -lean_inc(x_118); -x_119 = lean_ctor_get(x_6, 3); -lean_inc(x_119); -x_120 = lean_ctor_get(x_6, 4); -lean_inc(x_120); -x_121 = lean_ctor_get(x_6, 5); -lean_inc(x_121); -x_122 = lean_ctor_get(x_6, 6); -lean_inc(x_122); -x_123 = lean_ctor_get(x_6, 7); -lean_inc(x_123); -x_124 = lean_ctor_get(x_6, 8); -lean_inc(x_124); -x_125 = lean_ctor_get(x_6, 10); -lean_inc(x_125); -x_126 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); -x_127 = lean_ctor_get(x_6, 11); -lean_inc(x_127); -x_128 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - lean_ctor_release(x_6, 1); - lean_ctor_release(x_6, 2); - lean_ctor_release(x_6, 3); - lean_ctor_release(x_6, 4); - lean_ctor_release(x_6, 5); - lean_ctor_release(x_6, 6); - lean_ctor_release(x_6, 7); - lean_ctor_release(x_6, 8); - lean_ctor_release(x_6, 9); - lean_ctor_release(x_6, 10); - lean_ctor_release(x_6, 11); - x_129 = x_6; -} else { - lean_dec_ref(x_6); - x_129 = lean_box(0); -} -x_130 = lean_unsigned_to_nat(0u); -if (lean_is_scalar(x_129)) { - x_131 = lean_alloc_ctor(0, 12, 2); -} else { - x_131 = x_129; -} -lean_ctor_set(x_131, 0, x_116); -lean_ctor_set(x_131, 1, x_117); -lean_ctor_set(x_131, 2, x_118); -lean_ctor_set(x_131, 3, x_119); -lean_ctor_set(x_131, 4, x_120); -lean_ctor_set(x_131, 5, x_121); -lean_ctor_set(x_131, 6, x_122); -lean_ctor_set(x_131, 7, x_123); -lean_ctor_set(x_131, 8, x_124); -lean_ctor_set(x_131, 9, x_130); -lean_ctor_set(x_131, 10, x_125); -lean_ctor_set(x_131, 11, x_127); -lean_ctor_set_uint8(x_131, sizeof(void*)*12, x_126); -lean_ctor_set_uint8(x_131, sizeof(void*)*12 + 1, x_128); -x_132 = lean_whnf(x_115, x_4, x_5, x_131, x_7, x_114); -if (lean_obj_tag(x_132) == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); -lean_inc(x_134); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_135 = x_132; -} else { - lean_dec_ref(x_132); - x_135 = lean_box(0); -} -x_136 = l_Lean_Expr_isForall(x_133); -lean_dec(x_133); -if (x_136 == 0) -{ -lean_object* x_137; lean_object* x_138; -x_137 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1; -if (lean_is_scalar(x_135)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_135; -} -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_134); -return x_138; -} -else -{ -lean_object* x_139; lean_object* x_140; -x_139 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2; -if (lean_is_scalar(x_135)) { - x_140 = lean_alloc_ctor(0, 2, 0); -} else { - x_140 = x_135; -} -lean_ctor_set(x_140, 0, x_139); -lean_ctor_set(x_140, 1, x_134); -return x_140; -} -} -else -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_141 = lean_ctor_get(x_132, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_132, 1); -lean_inc(x_142); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_143 = x_132; -} else { - lean_dec_ref(x_132); - x_143 = lean_box(0); -} -if (lean_is_scalar(x_143)) { - x_144 = lean_alloc_ctor(1, 2, 0); -} else { - x_144 = x_143; -} -lean_ctor_set(x_144, 0, x_141); -lean_ctor_set(x_144, 1, x_142); -return x_144; -} -} -else -{ -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_145 = lean_ctor_get(x_110, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_146 = x_110; -} else { - lean_dec_ref(x_110); - x_146 = lean_box(0); -} -x_147 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3; -if (lean_is_scalar(x_146)) { - x_148 = lean_alloc_ctor(0, 2, 0); -} else { - x_148 = x_146; -} -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_145); -return x_148; -} -} -else -{ -lean_object* x_149; lean_object* x_150; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_149 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4; -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_149); -lean_ctor_set(x_150, 1, x_104); -return x_150; -} -} -else -{ -lean_object* x_151; uint8_t x_152; -x_151 = l_Lean_ConstantInfo_name(x_1); -lean_inc(x_151); -x_152 = lean_is_class(x_105, x_151); -if (x_152 == 0) -{ -lean_object* x_153; -x_153 = l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_151, x_2, x_3, x_4, x_5, x_6, x_7, x_104); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -if (lean_obj_tag(x_153) == 0) -{ -lean_object* x_154; lean_object* x_155; uint8_t x_156; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -lean_dec(x_154); -x_156 = lean_unbox(x_155); -lean_dec(x_155); -if (x_156 == 0) -{ -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_157 = lean_ctor_get(x_153, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_158 = x_153; -} else { - lean_dec_ref(x_153); - x_158 = lean_box(0); -} -x_159 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5; -if (lean_is_scalar(x_158)) { - x_160 = lean_alloc_ctor(0, 2, 0); -} else { - x_160 = x_158; -} -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_157); -return x_160; -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_161 = lean_ctor_get(x_153, 1); -lean_inc(x_161); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_162 = x_153; -} else { - lean_dec_ref(x_153); - x_162 = lean_box(0); -} -x_163 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6; -if (lean_is_scalar(x_162)) { - x_164 = lean_alloc_ctor(0, 2, 0); -} else { - x_164 = x_162; -} -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_161); -return x_164; -} -} -else -{ -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -x_165 = lean_ctor_get(x_153, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_153, 1); -lean_inc(x_166); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_167 = x_153; -} else { - lean_dec_ref(x_153); - x_167 = lean_box(0); -} -if (lean_is_scalar(x_167)) { - x_168 = lean_alloc_ctor(1, 2, 0); -} else { - x_168 = x_167; -} -lean_ctor_set(x_168, 0, x_165); -lean_ctor_set(x_168, 1, x_166); -return x_168; -} -} -else -{ -lean_object* x_169; lean_object* x_170; -lean_dec(x_151); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_169 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7; -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_104); -return x_170; -} -} -} -else -{ -lean_object* x_171; lean_object* x_172; -lean_dec(x_105); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_171 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8; -x_172 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_172, 0, x_171); -lean_ctor_set(x_172, 1, x_104); -return x_172; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_throwError___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_isEnumType___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(lean_object* x_1, lean_object* x_2, double x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -lean_inc(x_2); -x_16 = lean_environment_find(x_15, x_2); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_11, 0, x_17); -return x_11; -} -else -{ -lean_object* x_18; lean_object* x_19; -lean_free_object(x_11); -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) -{ -uint8_t x_21; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -x_23 = lean_box(0); -lean_ctor_set(x_19, 0, x_23); -return x_19; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_box(0); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} -} -else -{ -lean_object* x_27; uint8_t x_28; -x_27 = lean_ctor_get(x_19, 1); -lean_inc(x_27); -lean_dec(x_19); -x_28 = !lean_is_exclusive(x_20); -if (x_28 == 0) -{ -lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_20, 0); -lean_ctor_set_tag(x_20, 0); -lean_ctor_set(x_20, 0, x_2); -x_30 = lean_unbox(x_29); -lean_dec(x_29); -x_31 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_20, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; -x_32 = lean_ctor_get(x_20, 0); -lean_inc(x_32); -lean_dec(x_20); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_2); -x_34 = lean_unbox(x_32); -lean_dec(x_32); -x_35 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_33, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_35; -} -} -} -else -{ -uint8_t x_36; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_19); -if (x_36 == 0) -{ -return x_19; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_19, 0); -x_38 = lean_ctor_get(x_19, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_19); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} -} -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_11, 0); -x_41 = lean_ctor_get(x_11, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_11); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -lean_dec(x_40); -lean_inc(x_2); -x_43 = lean_environment_find(x_42, x_2); -if (lean_obj_tag(x_43) == 0) -{ -lean_object* x_44; lean_object* x_45; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_44 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_41); -return x_45; -} -else -{ -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_43, 0); -lean_inc(x_46); -lean_dec(x_43); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_47 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_41); -lean_dec(x_46); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_50 = x_47; -} else { - lean_dec_ref(x_47); - x_50 = lean_box(0); -} -x_51 = lean_box(0); -if (lean_is_scalar(x_50)) { - x_52 = lean_alloc_ctor(0, 2, 0); -} else { - x_52 = x_50; -} -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_49); -return x_52; -} -else -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; -x_53 = lean_ctor_get(x_47, 1); -lean_inc(x_53); -lean_dec(x_47); -x_54 = lean_ctor_get(x_48, 0); -lean_inc(x_54); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - x_55 = x_48; -} else { - lean_dec_ref(x_48); - x_55 = lean_box(0); -} -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(0, 1, 0); -} else { - x_56 = x_55; - lean_ctor_set_tag(x_56, 0); -} -lean_ctor_set(x_56, 0, x_2); -x_57 = lean_unbox(x_54); -lean_dec(x_54); -x_58 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_56, x_57, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_53); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_58; -} -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_59 = lean_ctor_get(x_47, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_47, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_61 = x_47; -} else { - lean_dec_ref(x_47); - x_61 = lean_box(0); -} -if (lean_is_scalar(x_61)) { - x_62 = lean_alloc_ctor(1, 2, 0); -} else { - x_62 = x_61; -} -lean_ctor_set(x_62, 0, x_59); -lean_ctor_set(x_62, 1, x_60); -return x_62; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -double x_11; lean_object* x_12; -x_11 = lean_unbox_float(x_3); -lean_dec(x_3); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -return x_12; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("keyword", 7, 7); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 13; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_box(0); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2; -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; -x_13 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_13, 0, x_1); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set(x_13, 2, x_10); -lean_ctor_set(x_13, 3, x_12); -lean_ctor_set(x_13, 4, x_10); -lean_ctor_set(x_13, 5, x_10); -lean_ctor_set(x_13, 6, x_10); -lean_ctor_set(x_13, 7, x_10); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_13, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -double x_10; lean_object* x_11; -x_10 = lean_unbox_float(x_2); -lean_dec(x_2); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("namespace", 9, 9); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 8; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_10 = 1; -x_11 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_12 = l_Lean_Name_toString(x_1, x_10, x_11); -x_13 = lean_box(0); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2; -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3; -x_16 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_16, 0, x_12); -lean_ctor_set(x_16, 1, x_14); -lean_ctor_set(x_16, 2, x_13); -lean_ctor_set(x_16, 3, x_15); -lean_ctor_set(x_16, 4, x_13); -lean_ctor_set(x_16, 5, x_13); -lean_ctor_set(x_16, 6, x_13); -lean_ctor_set(x_16, 7, x_13); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_16, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_17; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -double x_10; lean_object* x_11; -x_10 = lean_unbox_float(x_2); -lean_dec(x_2); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_st_mk_ref(x_1, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -lean_inc(x_10); -x_12 = lean_apply_7(x_2, x_3, x_10, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_st_ref_get(x_10, x_14); -lean_dec(x_10); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = lean_box(0); -lean_ctor_set(x_15, 0, x_18); -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_15, 1); -lean_inc(x_19); -lean_dec(x_15); -x_20 = lean_box(0); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -else -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_13); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_13, 0); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_15); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_15, 0); -x_26 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_25); -x_27 = 1; -x_28 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); -lean_ctor_set(x_13, 0, x_28); -lean_ctor_set(x_15, 0, x_13); -return x_15; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_29 = lean_ctor_get(x_15, 0); -x_30 = lean_ctor_get(x_15, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_15); -x_31 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_29); -x_32 = 1; -x_33 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); -lean_ctor_set(x_13, 0, x_33); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_13); -lean_ctor_set(x_34, 1, x_30); -return x_34; -} -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_13); -x_35 = lean_ctor_get(x_15, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_15, 1); -lean_inc(x_36); -if (lean_is_exclusive(x_15)) { - lean_ctor_release(x_15, 0); - lean_ctor_release(x_15, 1); - x_37 = x_15; -} else { - lean_dec_ref(x_15); - x_37 = lean_box(0); -} -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_35); -x_39 = 1; -x_40 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set_uint8(x_40, sizeof(void*)*1, x_39); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_40); -if (lean_is_scalar(x_37)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_37; -} -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_36); -return x_42; -} -} -} -else -{ -uint8_t x_43; -lean_dec(x_10); -x_43 = !lean_is_exclusive(x_12); -if (x_43 == 0) -{ -return x_12; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_12, 0); -x_45 = lean_ctor_get(x_12, 1); -lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_12); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM___lambda__1), 8, 3); -lean_closure_set(x_7, 0, x_6); -lean_closure_set(x_7, 1, x_4); -lean_closure_set(x_7, 2, x_1); -x_8 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_3, x_7, x_5); -return x_8; -} -} -static double _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; double x_3; -x_1 = lean_unsigned_to_nat(1u); -x_2 = 1; -x_3 = l_Float_ofScientific(x_1, x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 1) -{ -lean_object* x_3; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) -{ -if (lean_obj_tag(x_2) == 1) -{ -lean_object* x_4; -x_4 = lean_ctor_get(x_2, 0); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_5; lean_object* x_6; double x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_ctor_get(x_2, 1); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -x_8 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_5, x_6, x_7); -return x_8; -} -else -{ -lean_object* x_9; -lean_dec(x_1); -x_9 = lean_box(0); -return x_9; -} -} -else -{ -lean_object* x_10; -lean_dec(x_1); -x_10 = lean_box(0); -return x_10; -} -} -else -{ -lean_object* x_11; -lean_dec(x_3); -lean_dec(x_1); -x_11 = lean_box(0); -return x_11; -} -} -else -{ -lean_object* x_12; -lean_dec(x_1); -x_12 = lean_box(0); -return x_12; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -lean_inc(x_1); -x_7 = lean_private_to_user_name(x_1); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_1); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; -} -else -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_st_ref_get(x_5, x_6); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -lean_inc(x_11); -x_16 = l_Lean_mkPrivateName(x_15, x_11); -x_17 = lean_name_eq(x_16, x_1); -lean_dec(x_1); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; -lean_free_object(x_7); -lean_dec(x_11); -x_18 = lean_box(0); -lean_ctor_set(x_12, 0, x_18); -return x_12; -} -else -{ -lean_ctor_set(x_12, 0, x_7); -return x_12; -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_12, 0); -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_12); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -lean_inc(x_11); -x_22 = l_Lean_mkPrivateName(x_21, x_11); -x_23 = lean_name_eq(x_22, x_1); -lean_dec(x_1); -lean_dec(x_22); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -lean_free_object(x_7); -lean_dec(x_11); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_20); -return x_25; -} -else -{ -lean_object* x_26; -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_7); -lean_ctor_set(x_26, 1, x_20); -return x_26; -} -} -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_27 = lean_ctor_get(x_7, 0); -lean_inc(x_27); -lean_dec(x_7); -x_28 = lean_st_ref_get(x_5, x_6); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_31 = x_28; -} else { - lean_dec_ref(x_28); - x_31 = lean_box(0); -} -x_32 = lean_ctor_get(x_29, 0); -lean_inc(x_32); -lean_dec(x_29); -lean_inc(x_27); -x_33 = l_Lean_mkPrivateName(x_32, x_27); -x_34 = lean_name_eq(x_33, x_1); -lean_dec(x_1); -lean_dec(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -lean_dec(x_27); -x_35 = lean_box(0); -if (lean_is_scalar(x_31)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_31; -} -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_30); -return x_36; -} -else -{ -lean_object* x_37; lean_object* x_38; -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_27); -if (lean_is_scalar(x_31)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_31; -} -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_30); -return x_38; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_box(0); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed), 6, 0); -return x_1; -} -} -static double _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1() { -_start: -{ -double x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_2 = lean_box_float(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_box(0); -x_12 = l_Lean_Name_replacePrefix(x_1, x_2, x_11); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1; -if (x_3 == 0) -{ -if (lean_obj_tag(x_4) == 1) -{ -if (lean_obj_tag(x_12) == 1) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_ctor_get(x_4, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 1); -lean_inc(x_15); -lean_dec(x_4); -x_16 = lean_ctor_get(x_12, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -x_18 = lean_name_eq(x_14, x_16); -if (x_18 == 0) -{ -uint8_t x_19; -x_19 = l_Lean_Name_isAnonymous(x_14); -lean_dec(x_14); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_12); -x_20 = lean_box(0); -x_21 = lean_apply_6(x_13, x_20, x_6, x_7, x_8, x_9, x_10); -return x_21; -} -else -{ -double x_22; lean_object* x_23; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -x_23 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_15, x_17, x_22); -lean_dec(x_17); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_16); -lean_dec(x_12); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_10); -return x_25; -} -else -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_27 = lean_ctor_get(x_23, 0); -x_28 = l_Lean_Name_getNumParts(x_16); -lean_dec(x_16); -x_29 = lean_unsigned_to_nat(1u); -x_30 = lean_nat_add(x_28, x_29); -lean_dec(x_28); -x_31 = 0; -x_32 = lean_unsigned_to_nat(0u); -x_33 = l_Float_ofScientific(x_30, x_31, x_32); -lean_dec(x_30); -x_34 = lean_unbox_float(x_27); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_33); -x_36 = lean_box_float(x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_12); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_23, 0, x_37); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_23); -lean_ctor_set(x_38, 1, x_10); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; double x_45; double x_46; double x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_39 = lean_ctor_get(x_23, 0); -lean_inc(x_39); -lean_dec(x_23); -x_40 = l_Lean_Name_getNumParts(x_16); -lean_dec(x_16); -x_41 = lean_unsigned_to_nat(1u); -x_42 = lean_nat_add(x_40, x_41); -lean_dec(x_40); -x_43 = 0; -x_44 = lean_unsigned_to_nat(0u); -x_45 = l_Float_ofScientific(x_42, x_43, x_44); -lean_dec(x_42); -x_46 = lean_unbox_float(x_39); -lean_dec(x_39); -x_47 = lean_float_div(x_46, x_45); -x_48 = lean_box_float(x_47); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_12); -lean_ctor_set(x_49, 1, x_48); -x_50 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_10); -return x_51; -} -} -} -} -else -{ -double x_52; lean_object* x_53; -lean_dec(x_16); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_52 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -x_53 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_15, x_17, x_52); -if (lean_obj_tag(x_53) == 0) -{ -lean_object* x_54; lean_object* x_55; -lean_dec(x_17); -x_54 = lean_box(0); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_10); -return x_55; -} -else -{ -uint8_t x_56; -x_56 = !lean_is_exclusive(x_53); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_57 = lean_ctor_get(x_53, 0); -x_58 = l_Lean_Name_str___override(x_11, x_17); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_57); -lean_ctor_set(x_53, 0, x_59); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_53); -lean_ctor_set(x_60, 1, x_10); -return x_60; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_61 = lean_ctor_get(x_53, 0); -lean_inc(x_61); -lean_dec(x_53); -x_62 = l_Lean_Name_str___override(x_11, x_17); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_61); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_10); -return x_65; -} -} -} -} -else -{ -lean_object* x_66; lean_object* x_67; -lean_dec(x_12); -lean_dec(x_4); -x_66 = lean_box(0); -x_67 = lean_apply_6(x_13, x_66, x_6, x_7, x_8, x_9, x_10); -return x_67; -} -} -else -{ -lean_object* x_68; lean_object* x_69; -lean_dec(x_12); -lean_dec(x_4); -x_68 = lean_box(0); -x_69 = lean_apply_6(x_13, x_68, x_6, x_7, x_8, x_9, x_10); -return x_69; -} -} -else -{ -uint8_t x_70; -x_70 = l_Lean_Name_isPrefixOf(x_4, x_12); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; -lean_dec(x_12); -lean_dec(x_4); -x_71 = lean_box(0); -x_72 = lean_apply_6(x_13, x_71, x_6, x_7, x_8, x_9, x_10); -return x_72; -} -else -{ -lean_object* x_73; uint8_t x_74; -x_73 = l_Lean_Name_replacePrefix(x_12, x_4, x_11); -lean_dec(x_4); -x_74 = l_Lean_Name_isAtomic(x_73); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; -lean_dec(x_73); -x_75 = lean_box(0); -x_76 = lean_apply_6(x_13, x_75, x_6, x_7, x_8, x_9, x_10); -return x_76; -} -else -{ -uint8_t x_77; -x_77 = l_Lean_Name_isAnonymous(x_73); -if (x_77 == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_78 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_73); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_80, 0, x_79); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_10); -return x_81; -} -else -{ -lean_object* x_82; lean_object* x_83; -lean_dec(x_73); -x_82 = lean_box(0); -x_83 = lean_apply_6(x_13, x_82, x_6, x_7, x_8, x_9, x_10); -return x_83; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_10, 0, x_14); -return x_10; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -} -else -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_10); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_10, 1); -x_20 = lean_ctor_get(x_10, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_11, 0); -lean_inc(x_21); -lean_dec(x_11); -x_22 = l_Lean_Name_isPrefixOf(x_1, x_21); -if (x_22 == 0) -{ -lean_object* x_23; -lean_dec(x_21); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_23 = lean_box(0); -lean_ctor_set(x_10, 0, x_23); -return x_10; -} -else -{ -lean_object* x_24; lean_object* x_25; -lean_free_object(x_10); -x_24 = lean_box(0); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_21, x_1, x_3, x_2, x_24, x_5, x_6, x_7, x_8, x_19); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_10, 1); -lean_inc(x_26); -lean_dec(x_10); -x_27 = lean_ctor_get(x_11, 0); -lean_inc(x_27); -lean_dec(x_11); -x_28 = l_Lean_Name_isPrefixOf(x_1, x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_27); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_26); -return x_30; -} -else -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_box(0); -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_27, x_1, x_3, x_2, x_31, x_5, x_6, x_7, x_8, x_26); -return x_32; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_3); -lean_dec(x_3); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_3); -lean_dec(x_3); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_1); -return x_11; -} -} -static lean_object* _init_l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedName; -x_2 = l_instInhabitedNat; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Server.Completion", 22, 22); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("_private.Lean.Server.Completion.0.Lean.Server.Completion.truncate.go", 68, 68); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2; -x_3 = lean_unsigned_to_nat(293u); -x_4 = lean_unsigned_to_nat(25u); -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go(lean_object* x_1, lean_object* x_2) { -_start: -{ -switch (lean_obj_tag(x_2)) { -case 0: -{ -lean_object* x_3; lean_object* x_4; -lean_dec(x_1); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -return x_4; -} -case 1: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_2, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_1); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go(x_1, x_5); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_7, 1); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_10, x_11); -x_13 = lean_nat_dec_le(x_1, x_12); -if (x_13 == 0) -{ -uint8_t x_14; lean_object* x_15; -lean_dec(x_9); -x_14 = l_Lean_Name_isAnonymous(x_5); -x_15 = lean_string_length(x_6); -if (x_14 == 0) -{ -lean_object* x_16; uint8_t x_17; -x_16 = lean_nat_add(x_12, x_15); -lean_dec(x_15); -lean_dec(x_12); -x_17 = lean_nat_dec_le(x_16, x_1); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec(x_16); -lean_dec(x_2); -x_18 = lean_nat_sub(x_1, x_11); -x_19 = lean_nat_sub(x_18, x_10); -lean_dec(x_10); -lean_dec(x_18); -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_string_utf8_extract(x_6, x_20, x_19); -lean_dec(x_19); -lean_dec(x_6); -x_22 = l_Lean_Name_str___override(x_5, x_21); -lean_ctor_set(x_7, 1, x_1); -lean_ctor_set(x_7, 0, x_22); -return x_7; -} -else -{ -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -lean_ctor_set(x_7, 1, x_16); -lean_ctor_set(x_7, 0, x_2); -return x_7; -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; -lean_dec(x_12); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_add(x_10, x_23); -x_25 = lean_nat_add(x_24, x_15); -lean_dec(x_15); -lean_dec(x_24); -x_26 = lean_nat_dec_le(x_25, x_1); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_25); -lean_dec(x_2); -x_27 = lean_nat_sub(x_1, x_23); -x_28 = lean_nat_sub(x_27, x_10); -lean_dec(x_10); -lean_dec(x_27); -x_29 = lean_string_utf8_extract(x_6, x_23, x_28); -lean_dec(x_28); -lean_dec(x_6); -x_30 = l_Lean_Name_str___override(x_5, x_29); -lean_ctor_set(x_7, 1, x_1); -lean_ctor_set(x_7, 0, x_30); -return x_7; -} -else -{ -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -lean_ctor_set(x_7, 1, x_25); -lean_ctor_set(x_7, 0, x_2); -return x_7; -} -} -} -else -{ -lean_dec(x_12); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_7; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_31 = lean_ctor_get(x_7, 0); -x_32 = lean_ctor_get(x_7, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_7); -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_32, x_33); -x_35 = lean_nat_dec_le(x_1, x_34); -if (x_35 == 0) -{ -uint8_t x_36; lean_object* x_37; -lean_dec(x_31); -x_36 = l_Lean_Name_isAnonymous(x_5); -x_37 = lean_string_length(x_6); -if (x_36 == 0) -{ -lean_object* x_38; uint8_t x_39; -x_38 = lean_nat_add(x_34, x_37); -lean_dec(x_37); -lean_dec(x_34); -x_39 = lean_nat_dec_le(x_38, x_1); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -lean_dec(x_38); -lean_dec(x_2); -x_40 = lean_nat_sub(x_1, x_33); -x_41 = lean_nat_sub(x_40, x_32); -lean_dec(x_32); -lean_dec(x_40); -x_42 = lean_unsigned_to_nat(0u); -x_43 = lean_string_utf8_extract(x_6, x_42, x_41); -lean_dec(x_41); -lean_dec(x_6); -x_44 = l_Lean_Name_str___override(x_5, x_43); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_1); -return x_45; -} -else -{ -lean_object* x_46; -lean_dec(x_32); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_2); -lean_ctor_set(x_46, 1, x_38); -return x_46; -} -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -lean_dec(x_34); -x_47 = lean_unsigned_to_nat(0u); -x_48 = lean_nat_add(x_32, x_47); -x_49 = lean_nat_add(x_48, x_37); -lean_dec(x_37); -lean_dec(x_48); -x_50 = lean_nat_dec_le(x_49, x_1); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -lean_dec(x_49); -lean_dec(x_2); -x_51 = lean_nat_sub(x_1, x_47); -x_52 = lean_nat_sub(x_51, x_32); -lean_dec(x_32); -lean_dec(x_51); -x_53 = lean_string_utf8_extract(x_6, x_47, x_52); -lean_dec(x_52); -lean_dec(x_6); -x_54 = l_Lean_Name_str___override(x_5, x_53); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_1); -return x_55; -} -else -{ -lean_object* x_56; -lean_dec(x_32); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_2); -lean_ctor_set(x_56, 1, x_49); -return x_56; -} -} -} -else -{ -lean_object* x_57; -lean_dec(x_34); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_31); -lean_ctor_set(x_57, 1, x_32); -return x_57; -} -} -} -default: -{ -lean_object* x_58; lean_object* x_59; -lean_dec(x_2); -lean_dec(x_1); -x_58 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4; -x_59 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1(x_58); -return x_59; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go(x_2, x_1); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -lean_dec(x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1() { -_start: -{ -double x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_2 = lean_box_float(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_Completion_matchNamespace___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace(lean_object* x_1, lean_object* x_2, uint8_t x_3) { -_start: -{ -if (x_3 == 0) -{ -if (lean_obj_tag(x_1) == 1) -{ -if (lean_obj_tag(x_2) == 1) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 1); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_2, 1); -lean_inc(x_7); -lean_dec(x_2); -x_8 = lean_name_eq(x_4, x_6); -lean_dec(x_6); -if (x_8 == 0) -{ -lean_object* x_9; -lean_dec(x_7); -x_9 = lean_box(0); -return x_9; -} -else -{ -double x_10; lean_object* x_11; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -x_11 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_7, x_5, x_10); -return x_11; -} -} -else -{ -lean_object* x_12; -lean_dec(x_2); -x_12 = lean_box(0); -return x_12; -} -} -else -{ -lean_object* x_13; -lean_dec(x_2); -x_13 = lean_box(0); -return x_13; -} -} -else -{ -uint8_t x_14; -x_14 = lean_name_eq(x_2, x_1); -if (x_14 == 0) -{ -uint8_t x_15; -x_15 = l_Lean_Name_isPrefixOf(x_2, x_1); -lean_dec(x_2); -if (x_15 == 0) -{ -lean_object* x_16; -x_16 = lean_box(0); -return x_16; -} -else -{ -lean_object* x_17; -x_17 = l_Lean_Server_Completion_matchNamespace___closed__1; -return x_17; -} -} -else -{ -lean_object* x_18; -lean_dec(x_2); -x_18 = lean_box(0); -return x_18; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = lean_unbox(x_3); -lean_dec(x_3); -x_5 = l_Lean_Server_Completion_matchNamespace(x_1, x_2, x_4); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; -lean_inc(x_1); -lean_inc(x_5); -x_13 = l_Lean_Name_append(x_5, x_1); -x_14 = l_Lean_Server_Completion_matchNamespace(x_4, x_13, x_2); -if (lean_obj_tag(x_14) == 0) -{ -if (lean_obj_tag(x_5) == 1) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_5, 0); -lean_inc(x_15); -lean_dec(x_5); -x_5 = x_15; -goto _start; -} -else -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_12); -return x_18; -} -} -else -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_1); -x_19 = lean_ctor_get(x_14, 0); -lean_inc(x_19); -lean_dec(x_14); -x_20 = lean_apply_10(x_3, x_4, x_5, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_20; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_2); -lean_dec(x_2); -x_14 = l_Lean_Server_Completion_completeNamespaces_visitNamespaces(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_9); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_9); -x_20 = lean_ctor_get(x_8, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_8, 1); -lean_inc(x_21); -if (lean_is_exclusive(x_8)) { - lean_ctor_release(x_8, 0); - lean_ctor_release(x_8, 1); - x_22 = x_8; -} else { - lean_dec_ref(x_8); - x_22 = lean_box(0); -} -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_20, 0); -lean_inc(x_38); -lean_dec(x_20); -lean_inc(x_1); -lean_inc(x_38); -x_39 = l_Lean_Name_append(x_38, x_1); -lean_inc(x_39); -x_40 = l_Lean_Server_Completion_matchNamespace(x_3, x_39, x_2); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; -lean_dec(x_39); -lean_dec(x_38); -lean_inc(x_6); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_6); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_23 = x_42; -x_24 = x_17; -goto block_37; -} -else -{ -if (x_2 == 0) -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; double x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -lean_dec(x_39); -x_43 = lean_ctor_get(x_40, 0); -lean_inc(x_43); -lean_dec(x_40); -x_44 = lean_box(0); -lean_inc(x_3); -x_45 = l_Lean_Name_replacePrefix(x_3, x_38, x_44); -lean_dec(x_38); -x_46 = lean_unbox_float(x_43); -lean_dec(x_43); -lean_inc(x_11); -x_47 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(x_45, x_46, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3; -x_23 = x_49; -x_24 = x_48; -goto block_37; -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; double x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -lean_dec(x_38); -x_50 = lean_ctor_get(x_40, 0); -lean_inc(x_50); -lean_dec(x_40); -x_51 = lean_box(0); -lean_inc(x_3); -x_52 = l_Lean_Name_replacePrefix(x_3, x_39, x_51); -lean_dec(x_39); -x_53 = lean_unbox_float(x_50); -lean_dec(x_50); -lean_inc(x_11); -x_54 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(x_52, x_53, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3; -x_23 = x_56; -x_24 = x_55; -goto block_37; -} -} -} -else -{ -lean_object* x_57; lean_object* x_58; -lean_dec(x_20); -lean_inc(x_6); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_6); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_23 = x_58; -x_24 = x_17; -goto block_37; -} -block_37: -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; -x_26 = lean_ctor_get(x_23, 0); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; lean_object* x_28; -lean_dec(x_21); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec(x_26); -lean_ctor_set(x_23, 0, x_27); -if (lean_is_scalar(x_22)) { - x_28 = lean_alloc_ctor(0, 2, 0); -} else { - x_28 = x_22; - lean_ctor_set_tag(x_28, 0); -} -lean_ctor_set(x_28, 0, x_23); -lean_ctor_set(x_28, 1, x_24); -return x_28; -} -else -{ -lean_object* x_29; -lean_free_object(x_23); -lean_dec(x_22); -x_29 = lean_ctor_get(x_26, 0); -lean_inc(x_29); -lean_dec(x_26); -x_8 = x_21; -x_9 = x_29; -x_10 = lean_box(0); -x_17 = x_24; -goto _start; -} -} -else -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_23, 0); -lean_inc(x_31); -lean_dec(x_23); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_21); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_1); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -if (lean_is_scalar(x_22)) { - x_34 = lean_alloc_ctor(0, 2, 0); -} else { - x_34 = x_22; - lean_ctor_set_tag(x_34, 0); -} -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_24); -return x_34; -} -else -{ -lean_object* x_35; -lean_dec(x_22); -x_35 = lean_ctor_get(x_31, 0); -lean_inc(x_35); -lean_dec(x_31); -x_8 = x_21; -x_9 = x_35; -x_10 = lean_box(0); -x_17 = x_24; -goto _start; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_dec(x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_16 = lean_apply_9(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_2 = x_25; -x_3 = x_15; -x_10 = x_24; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) -{ -return x_16; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_array_uget(x_2, x_3); -switch (lean_obj_tag(x_14)) { -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -lean_inc(x_1); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_17 = lean_apply_10(x_1, x_5, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_17, 0, x_21); -return x_17; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -else -{ -lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_dec(x_17); -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -lean_dec(x_18); -x_27 = 1; -x_28 = lean_usize_add(x_3, x_27); -x_3 = x_28; -x_5 = x_26; -x_12 = x_25; -goto _start; -} -} -else -{ -uint8_t x_30; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) -{ -return x_17; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -case 1: -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_14, 0); -lean_inc(x_34); -lean_dec(x_14); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_35 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_1, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) -{ -uint8_t x_37; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_35, 0); -lean_dec(x_38); -x_39 = lean_box(0); -lean_ctor_set(x_35, 0, x_39); -return x_35; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = lean_box(0); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} -} -else -{ -lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_dec(x_35); -x_44 = lean_ctor_get(x_36, 0); -lean_inc(x_44); -lean_dec(x_36); -x_45 = 1; -x_46 = lean_usize_add(x_3, x_45); -x_3 = x_46; -x_5 = x_44; -x_12 = x_43; -goto _start; -} -} -else -{ -uint8_t x_48; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_35); -if (x_48 == 0) -{ -return x_35; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_35, 0); -x_50 = lean_ctor_get(x_35, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_35); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} -} -default: -{ -size_t x_52; size_t x_53; -x_52 = 1; -x_53 = lean_usize_add(x_3, x_52); -x_3 = x_53; -goto _start; -} -} -} -else -{ -lean_object* x_55; lean_object* x_56; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_5); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_12); -return x_56; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed), 12, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_array_get_size(x_2); -x_15 = lean_nat_dec_lt(x_5, x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_6); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_13); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_fget(x_2, x_5); -x_19 = lean_array_fget(x_3, x_5); -lean_inc(x_1); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_20 = lean_apply_10(x_1, x_6, x_18, x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_20, 0, x_24); -return x_20; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 1); -lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_dec(x_20); -x_29 = lean_ctor_get(x_21, 0); -lean_inc(x_29); -lean_dec(x_21); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_5, x_30); -lean_dec(x_5); -x_4 = lean_box(0); -x_5 = x_31; -x_6 = x_29; -x_13 = x_28; -goto _start; -} -} -else -{ -uint8_t x_33; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_20); -if (x_33 == 0) -{ -return x_20; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_20, 0); -x_35 = lean_ctor_get(x_20, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_20); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed), 13, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_array_get_size(x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_3); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_2); -lean_ctor_set(x_16, 1, x_10); -return x_16; -} -else -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_13, x_13); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_2); -lean_ctor_set(x_18, 1, x_10); -return x_18; -} -else -{ -size_t x_19; size_t x_20; lean_object* x_21; -lean_free_object(x_2); -x_19 = 0; -x_20 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_21 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(x_1, x_12, x_19, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_12); -return x_21; -} -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); -lean_dec(x_2); -x_23 = lean_array_get_size(x_22); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_nat_dec_lt(x_24, x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_3); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_10); -return x_27; -} -else -{ -uint8_t x_28; -x_28 = lean_nat_dec_le(x_23, x_23); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_3); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_10); -return x_30; -} -else -{ -size_t x_31; size_t x_32; lean_object* x_33; -x_31 = 0; -x_32 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_33 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(x_1, x_22, x_31, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_22); -return x_33; -} -} -} -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_2, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_2, 1); -lean_inc(x_35); -lean_dec(x_2); -x_36 = lean_unsigned_to_nat(0u); -x_37 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(x_1, x_34, x_35, lean_box(0), x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_35); -lean_dec(x_34); -return x_37; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg), 10, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = lean_box(0); -x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -x_15 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -x_5 = x_25; -x_12 = x_24; -goto _start; -} -} -else -{ -uint8_t x_29; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_12); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_array_get_size(x_11); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -lean_dec(x_11); -x_15 = lean_ctor_get(x_1, 1); -lean_inc(x_15); -lean_dec(x_1); -x_16 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_16; -} -else -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_12, x_12); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_12); -lean_dec(x_11); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); -lean_dec(x_1); -x_19 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_19; -} -else -{ -size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_20 = 0; -x_21 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_22 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_11); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_24); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_dec(x_23); -x_32 = lean_ctor_get(x_1, 1); -lean_inc(x_32); -lean_dec(x_1); -x_33 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_31); -return x_33; -} -} -else -{ -uint8_t x_34; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_23); -if (x_34 == 0) -{ -return x_23; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_23, 0); -x_36 = lean_ctor_get(x_23, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_23); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (x_1 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_2); -x_13 = lean_box(0); -x_14 = l_Lean_Name_replacePrefix(x_3, x_4, x_13); -lean_dec(x_4); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = l_Lean_Name_append(x_4, x_2); -x_17 = lean_box(0); -x_18 = l_Lean_Name_replacePrefix(x_3, x_16, x_17); -lean_dec(x_16); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 4); -lean_inc(x_14); -lean_dec(x_1); -x_15 = l_Lean_Server_Completion_completeNamespaces_visitNamespaces(x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_15; -} -} -static lean_object* _init_l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = l_Lean_Name_isInternal(x_6); -if (x_15 == 0) -{ -uint8_t x_16; -x_16 = l_Lean_Environment_contains(x_1, x_6); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_17 = lean_ctor_get(x_2, 0); -lean_inc(x_17); -lean_dec(x_2); -x_18 = lean_ctor_get(x_17, 5); -lean_inc(x_18); -x_19 = lean_box(0); -x_20 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -lean_inc(x_8); -lean_inc(x_18); -lean_inc(x_6); -lean_inc(x_3); -x_21 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(x_3, x_4, x_6, x_18, x_19, x_20, x_18, x_18, x_20, lean_box(0), x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_18); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_box(0); -x_27 = l_Lean_Server_Completion_completeNamespaces___lambda__2(x_17, x_3, x_4, x_5, x_6, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_25); -return x_27; -} -else -{ -uint8_t x_28; -lean_dec(x_17); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_28 = !lean_is_exclusive(x_21); -if (x_28 == 0) -{ -lean_object* x_29; uint8_t x_30; -x_29 = lean_ctor_get(x_21, 0); -lean_dec(x_29); -x_30 = !lean_is_exclusive(x_24); -if (x_30 == 0) -{ -lean_ctor_set(x_21, 0, x_24); -return x_21; -} -else -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_24, 0); -lean_inc(x_31); -lean_dec(x_24); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_21, 0, x_32); -return x_21; -} -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_33 = lean_ctor_get(x_21, 1); -lean_inc(x_33); -lean_dec(x_21); -x_34 = lean_ctor_get(x_24, 0); -lean_inc(x_34); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - x_35 = x_24; -} else { - lean_dec_ref(x_24); - x_35 = lean_box(0); -} -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(1, 1, 0); -} else { - x_36 = x_35; -} -lean_ctor_set(x_36, 0, x_34); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_33); -return x_37; -} -} -} -else -{ -lean_object* x_38; lean_object* x_39; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_14); -return x_39; -} -} -else -{ -lean_object* x_40; lean_object* x_41; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_40 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_14); -return x_41; -} -} -} -static lean_object* _init_l_Lean_Server_Completion_completeNamespaces___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_namespacesExt; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_11 = lean_st_ref_get(x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_box(x_3); -lean_inc(x_2); -x_16 = lean_alloc_closure((void*)(l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed), 12, 2); -lean_closure_set(x_16, 0, x_15); -lean_closure_set(x_16, 1, x_2); -x_17 = l_Lean_NameSSet_instInhabited; -x_18 = l_Lean_Server_Completion_completeNamespaces___closed__1; -x_19 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_17, x_18, x_14); -x_20 = lean_box(x_3); -x_21 = lean_alloc_closure((void*)(l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed), 14, 5); -lean_closure_set(x_21, 0, x_14); -lean_closure_set(x_21, 1, x_1); -lean_closure_set(x_21, 2, x_2); -lean_closure_set(x_21, 3, x_20); -lean_closure_set(x_21, 4, x_16); -x_22 = l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(x_19, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_22; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -uint8_t x_18; lean_object* x_19; -x_18 = lean_unbox(x_2); -lean_dec(x_2); -x_19 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; double x_14; lean_object* x_15; -x_13 = lean_unbox(x_1); -lean_dec(x_1); -x_14 = lean_unbox_float(x_5); -lean_dec(x_5); -x_15 = l_Lean_Server_Completion_completeNamespaces___lambda__1(x_13, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_3); -lean_dec(x_3); -x_15 = l_Lean_Server_Completion_completeNamespaces___lambda__2(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_6); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l_Lean_Server_Completion_completeNamespaces___lambda__3(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_7); -return x_16; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_3); -lean_dec(x_3); -x_12 = l_Lean_Server_Completion_completeNamespaces(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_searchAlias(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_1); -lean_inc(x_3); -lean_inc(x_5); -x_13 = lean_apply_2(x_1, x_5, x_3); -if (lean_obj_tag(x_13) == 0) -{ -if (lean_obj_tag(x_5) == 1) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_5, 0); -lean_inc(x_14); -lean_dec(x_5); -x_5 = x_14; -goto _start; -} -else -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_12); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_5); -lean_dec(x_1); -x_18 = lean_ctor_get(x_13, 0); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_apply_10(x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_2) == 1) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_12 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = lean_box(0); -lean_ctor_set(x_12, 0, x_16); -return x_12; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_13, 0); -lean_inc(x_20); -lean_dec(x_13); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_2 = x_11; -x_3 = x_22; -x_10 = x_21; -goto _start; -} -} -else -{ -uint8_t x_24; -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_12); -if (x_24 == 0) -{ -return x_12; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_12, 0); -x_26 = lean_ctor_get(x_12, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_12); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; -} -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_28 = lean_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_3); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_10); -return x_31; -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_34; -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_3, x_4, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_box(0); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_6); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_14 = x_39; -x_15 = x_36; -goto block_33; -} -else -{ -uint8_t x_40; -x_40 = !lean_is_exclusive(x_35); -if (x_40 == 0) -{ -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = lean_ctor_get(x_35, 0); -x_42 = lean_ctor_get(x_34, 1); -lean_inc(x_42); -lean_dec(x_34); -x_43 = !lean_is_exclusive(x_41); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_box(0); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_35); -x_46 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_46, 0, x_45); -x_14 = x_46; -x_15 = x_42; -goto block_33; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_47 = lean_ctor_get(x_41, 0); -x_48 = lean_ctor_get(x_41, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_41); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -lean_ctor_set(x_35, 0, x_49); -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_35); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_14 = x_52; -x_15 = x_42; -goto block_33; -} -} -else -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_6); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_54 = lean_ctor_get(x_35, 0); -x_55 = lean_ctor_get(x_6, 0); -x_56 = lean_ctor_get(x_34, 1); -lean_inc(x_56); -lean_dec(x_34); -x_57 = !lean_is_exclusive(x_54); -if (x_57 == 0) -{ -uint8_t x_58; -x_58 = !lean_is_exclusive(x_55); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = lean_ctor_get(x_54, 0); -x_60 = lean_ctor_get(x_54, 1); -x_61 = lean_ctor_get(x_55, 0); -x_62 = lean_ctor_get(x_55, 1); -x_63 = l_Lean_Name_isSuffixOf(x_59, x_61); -if (x_63 == 0) -{ -lean_object* x_64; -lean_dec(x_60); -lean_dec(x_59); -x_64 = lean_box(0); -lean_ctor_set(x_54, 1, x_6); -lean_ctor_set(x_54, 0, x_64); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -else -{ -lean_object* x_65; -lean_dec(x_62); -lean_dec(x_61); -lean_ctor_set(x_55, 1, x_60); -lean_ctor_set(x_55, 0, x_59); -x_65 = lean_box(0); -lean_ctor_set(x_54, 1, x_6); -lean_ctor_set(x_54, 0, x_65); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_66 = lean_ctor_get(x_54, 0); -x_67 = lean_ctor_get(x_54, 1); -x_68 = lean_ctor_get(x_55, 0); -x_69 = lean_ctor_get(x_55, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_55); -x_70 = l_Lean_Name_isSuffixOf(x_66, x_68); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; -lean_dec(x_67); -lean_dec(x_66); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -lean_ctor_set(x_6, 0, x_71); -x_72 = lean_box(0); -lean_ctor_set(x_54, 1, x_6); -lean_ctor_set(x_54, 0, x_72); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -else -{ -lean_object* x_73; lean_object* x_74; -lean_dec(x_69); -lean_dec(x_68); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_66); -lean_ctor_set(x_73, 1, x_67); -lean_ctor_set(x_6, 0, x_73); -x_74 = lean_box(0); -lean_ctor_set(x_54, 1, x_6); -lean_ctor_set(x_54, 0, x_74); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -} -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_75 = lean_ctor_get(x_54, 0); -x_76 = lean_ctor_get(x_54, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_54); -x_77 = lean_ctor_get(x_55, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_55, 1); -lean_inc(x_78); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_79 = x_55; -} else { - lean_dec_ref(x_55); - x_79 = lean_box(0); -} -x_80 = l_Lean_Name_isSuffixOf(x_75, x_77); -if (x_80 == 0) -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_76); -lean_dec(x_75); -if (lean_is_scalar(x_79)) { - x_81 = lean_alloc_ctor(0, 2, 0); -} else { - x_81 = x_79; -} -lean_ctor_set(x_81, 0, x_77); -lean_ctor_set(x_81, 1, x_78); -lean_ctor_set(x_6, 0, x_81); -x_82 = lean_box(0); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_6); -lean_ctor_set(x_35, 0, x_83); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_78); -lean_dec(x_77); -if (lean_is_scalar(x_79)) { - x_84 = lean_alloc_ctor(0, 2, 0); -} else { - x_84 = x_79; -} -lean_ctor_set(x_84, 0, x_75); -lean_ctor_set(x_84, 1, x_76); -lean_ctor_set(x_6, 0, x_84); -x_85 = lean_box(0); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_6); -lean_ctor_set(x_35, 0, x_86); -x_14 = x_35; -x_15 = x_56; -goto block_33; -} -} -} -else -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_87 = lean_ctor_get(x_35, 0); -x_88 = lean_ctor_get(x_6, 0); -lean_inc(x_88); -lean_dec(x_6); -x_89 = lean_ctor_get(x_34, 1); -lean_inc(x_89); -lean_dec(x_34); -x_90 = lean_ctor_get(x_87, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_87, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_92 = x_87; -} else { - lean_dec_ref(x_87); - x_92 = lean_box(0); -} -x_93 = lean_ctor_get(x_88, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_88, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_95 = x_88; -} else { - lean_dec_ref(x_88); - x_95 = lean_box(0); -} -x_96 = l_Lean_Name_isSuffixOf(x_90, x_93); -if (x_96 == 0) -{ -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -lean_dec(x_91); -lean_dec(x_90); -if (lean_is_scalar(x_95)) { - x_97 = lean_alloc_ctor(0, 2, 0); -} else { - x_97 = x_95; -} -lean_ctor_set(x_97, 0, x_93); -lean_ctor_set(x_97, 1, x_94); -x_98 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_98, 0, x_97); -x_99 = lean_box(0); -if (lean_is_scalar(x_92)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_92; -} -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_98); -lean_ctor_set(x_35, 0, x_100); -x_14 = x_35; -x_15 = x_89; -goto block_33; -} -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -lean_dec(x_94); -lean_dec(x_93); -if (lean_is_scalar(x_95)) { - x_101 = lean_alloc_ctor(0, 2, 0); -} else { - x_101 = x_95; -} -lean_ctor_set(x_101, 0, x_90); -lean_ctor_set(x_101, 1, x_91); -x_102 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = lean_box(0); -if (lean_is_scalar(x_92)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_92; -} -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_102); -lean_ctor_set(x_35, 0, x_104); -x_14 = x_35; -x_15 = x_89; -goto block_33; -} -} -} -} -else -{ -lean_object* x_105; -x_105 = lean_ctor_get(x_35, 0); -lean_inc(x_105); -lean_dec(x_35); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_106 = lean_ctor_get(x_34, 1); -lean_inc(x_106); -lean_dec(x_34); -x_107 = lean_ctor_get(x_105, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_105, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_109 = x_105; -} else { - lean_dec_ref(x_105); - x_109 = lean_box(0); -} -if (lean_is_scalar(x_109)) { - x_110 = lean_alloc_ctor(0, 2, 0); -} else { - x_110 = x_109; -} -lean_ctor_set(x_110, 0, x_107); -lean_ctor_set(x_110, 1, x_108); -x_111 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_111, 0, x_110); -x_112 = lean_box(0); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_114, 0, x_113); -x_14 = x_114; -x_15 = x_106; -goto block_33; -} -else -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_115 = lean_ctor_get(x_6, 0); -lean_inc(x_115); -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - x_116 = x_6; -} else { - lean_dec_ref(x_6); - x_116 = lean_box(0); -} -x_117 = lean_ctor_get(x_34, 1); -lean_inc(x_117); -lean_dec(x_34); -x_118 = lean_ctor_get(x_105, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_105, 1); -lean_inc(x_119); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_120 = x_105; -} else { - lean_dec_ref(x_105); - x_120 = lean_box(0); -} -x_121 = lean_ctor_get(x_115, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_115, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_123 = x_115; -} else { - lean_dec_ref(x_115); - x_123 = lean_box(0); -} -x_124 = l_Lean_Name_isSuffixOf(x_118, x_121); -if (x_124 == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_119); -lean_dec(x_118); -if (lean_is_scalar(x_123)) { - x_125 = lean_alloc_ctor(0, 2, 0); -} else { - x_125 = x_123; -} -lean_ctor_set(x_125, 0, x_121); -lean_ctor_set(x_125, 1, x_122); -if (lean_is_scalar(x_116)) { - x_126 = lean_alloc_ctor(1, 1, 0); -} else { - x_126 = x_116; -} -lean_ctor_set(x_126, 0, x_125); -x_127 = lean_box(0); -if (lean_is_scalar(x_120)) { - x_128 = lean_alloc_ctor(0, 2, 0); -} else { - x_128 = x_120; -} -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_126); -x_129 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_129, 0, x_128); -x_14 = x_129; -x_15 = x_117; -goto block_33; -} -else -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -lean_dec(x_122); -lean_dec(x_121); -if (lean_is_scalar(x_123)) { - x_130 = lean_alloc_ctor(0, 2, 0); -} else { - x_130 = x_123; -} -lean_ctor_set(x_130, 0, x_118); -lean_ctor_set(x_130, 1, x_119); -if (lean_is_scalar(x_116)) { - x_131 = lean_alloc_ctor(1, 1, 0); -} else { - x_131 = x_116; -} -lean_ctor_set(x_131, 0, x_130); -x_132 = lean_box(0); -if (lean_is_scalar(x_120)) { - x_133 = lean_alloc_ctor(0, 2, 0); -} else { - x_133 = x_120; -} -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_131); -x_134 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_134, 0, x_133); -x_14 = x_134; -x_15 = x_117; -goto block_33; -} -} -} -} -} -else -{ -uint8_t x_135; -lean_dec(x_6); -x_135 = !lean_is_exclusive(x_34); -if (x_135 == 0) -{ -return x_34; -} -else -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = lean_ctor_get(x_34, 0); -x_137 = lean_ctor_get(x_34, 1); -lean_inc(x_137); -lean_inc(x_136); -lean_dec(x_34); -x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_136); -lean_ctor_set(x_138, 1, x_137); -return x_138; -} -} -block_33: -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_14, 0); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -x_20 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; -lean_ctor_set(x_17, 0, x_20); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_14); -lean_ctor_set(x_21, 1, x_15); -return x_21; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_14, 0, x_24); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_15); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_26 = lean_ctor_get(x_14, 0); -lean_inc(x_26); -lean_dec(x_14); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_28 = x_26; -} else { - lean_dec_ref(x_26); - x_28 = lean_box(0); -} -x_29 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; -if (lean_is_scalar(x_28)) { - x_30 = lean_alloc_ctor(0, 2, 0); -} else { - x_30 = x_28; -} -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_27); -x_31 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_15); -return x_32; -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_3); -lean_dec(x_2); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_8); -lean_ctor_set(x_18, 1, x_10); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; -} -else -{ -lean_object* x_21; -lean_dec(x_8); -x_21 = lean_ctor_get(x_7, 0); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_7, 1); -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); -x_25 = l_List_elem___at_Lean_addAliasEntry___spec__16(x_3, x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_box(0); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_3); -lean_inc(x_2); -x_27 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(x_23, x_2, x_1, x_3, x_26, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_3); -lean_dec(x_2); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 0); -lean_dec(x_30); -x_31 = lean_box(0); -lean_ctor_set(x_27, 0, x_31); -return x_27; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_27, 1); -lean_inc(x_32); -lean_dec(x_27); -x_33 = lean_box(0); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -return x_34; -} -} -else -{ -uint8_t x_35; -x_35 = !lean_is_exclusive(x_28); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_28, 0); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) -{ -uint8_t x_38; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_3); -lean_dec(x_2); -x_38 = !lean_is_exclusive(x_27); -if (x_38 == 0) -{ -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_27, 0); -lean_dec(x_39); -x_40 = !lean_is_exclusive(x_36); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_36, 0); -lean_dec(x_41); -x_42 = lean_ctor_get(x_37, 0); -lean_inc(x_42); -lean_dec(x_37); -lean_ctor_set(x_36, 0, x_42); -return x_27; -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_36, 1); -lean_inc(x_43); -lean_dec(x_36); -x_44 = lean_ctor_get(x_37, 0); -lean_inc(x_44); -lean_dec(x_37); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); -lean_ctor_set(x_28, 0, x_45); -return x_27; -} -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = lean_ctor_get(x_27, 1); -lean_inc(x_46); -lean_dec(x_27); -x_47 = lean_ctor_get(x_36, 1); -lean_inc(x_47); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_48 = x_36; -} else { - lean_dec_ref(x_36); - x_48 = lean_box(0); -} -x_49 = lean_ctor_get(x_37, 0); -lean_inc(x_49); -lean_dec(x_37); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_48; -} -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_47); -lean_ctor_set(x_28, 0, x_50); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_28); -lean_ctor_set(x_51, 1, x_46); -return x_51; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_free_object(x_28); -x_52 = lean_ctor_get(x_27, 1); -lean_inc(x_52); -lean_dec(x_27); -x_53 = lean_ctor_get(x_36, 1); -lean_inc(x_53); -lean_dec(x_36); -x_54 = lean_ctor_get(x_37, 0); -lean_inc(x_54); -lean_dec(x_37); -x_7 = x_22; -x_8 = x_54; -x_9 = lean_box(0); -x_10 = x_53; -x_17 = x_52; -goto _start; -} -} -else -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_28, 0); -lean_inc(x_56); -lean_dec(x_28); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -if (lean_obj_tag(x_57) == 0) -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_3); -lean_dec(x_2); -x_58 = lean_ctor_get(x_27, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_59 = x_27; -} else { - lean_dec_ref(x_27); - x_59 = lean_box(0); -} -x_60 = lean_ctor_get(x_56, 1); -lean_inc(x_60); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_61 = x_56; -} else { - lean_dec_ref(x_56); - x_61 = lean_box(0); -} -x_62 = lean_ctor_get(x_57, 0); -lean_inc(x_62); -lean_dec(x_57); -if (lean_is_scalar(x_61)) { - x_63 = lean_alloc_ctor(0, 2, 0); -} else { - x_63 = x_61; -} -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -if (lean_is_scalar(x_59)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_59; -} -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_58); -return x_65; -} -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_27, 1); -lean_inc(x_66); -lean_dec(x_27); -x_67 = lean_ctor_get(x_56, 1); -lean_inc(x_67); -lean_dec(x_56); -x_68 = lean_ctor_get(x_57, 0); -lean_inc(x_68); -lean_dec(x_57); -x_7 = x_22; -x_8 = x_68; -x_9 = lean_box(0); -x_10 = x_67; -x_17 = x_66; -goto _start; -} -} -} -} -else -{ -uint8_t x_70; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_3); -lean_dec(x_2); -x_70 = !lean_is_exclusive(x_27); -if (x_70 == 0) -{ -return x_27; -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_27, 0); -x_72 = lean_ctor_get(x_27, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_27); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; -} -} -} -else -{ -lean_object* x_74; -x_74 = lean_box(0); -x_7 = x_22; -x_8 = x_74; -x_9 = lean_box(0); -goto _start; -} -} -else -{ -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_7, 1); -x_77 = lean_box(0); -x_7 = x_76; -x_8 = x_77; -x_9 = lean_box(0); -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_dec(x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_16 = lean_apply_9(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_2 = x_25; -x_3 = x_15; -x_10 = x_24; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) -{ -return x_16; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = lean_box(0); -x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -x_15 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__3(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -x_5 = x_25; -x_12 = x_24; -goto _start; -} -} -else -{ -uint8_t x_29; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_12); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -lean_inc(x_3); -x_12 = l_Lean_Meta_allowCompletion(x_1, x_3); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_11); -return x_14; -} -else -{ -lean_object* x_15; -x_15 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_15; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -lean_inc(x_12); -x_13 = l_Lean_Server_Completion_getEligibleHeaderDecls(x_12, x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_array_get_size(x_16); -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_nat_dec_lt(x_18, x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_17); -lean_dec(x_16); -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); -lean_closure_set(x_22, 0, x_12); -lean_closure_set(x_22, 1, x_1); -x_23 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__4(x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -return x_23; -} -else -{ -uint8_t x_24; -x_24 = lean_nat_dec_le(x_17, x_17); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_17); -lean_dec(x_16); -x_25 = lean_ctor_get(x_12, 1); -lean_inc(x_25); -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); -lean_closure_set(x_27, 0, x_12); -lean_closure_set(x_27, 1, x_1); -x_28 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__4(x_26, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -return x_28; -} -else -{ -size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_29 = 0; -x_30 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_31 = lean_box(0); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_32 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6(x_1, x_16, x_29, x_30, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -lean_dec(x_16); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -if (lean_obj_tag(x_33) == 0) -{ -uint8_t x_34; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_32); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_32, 0); -lean_dec(x_35); -x_36 = lean_box(0); -lean_ctor_set(x_32, 0, x_36); -return x_32; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_32, 1); -lean_inc(x_37); -lean_dec(x_32); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -return x_39; -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_33); -x_40 = lean_ctor_get(x_32, 1); -lean_inc(x_40); -lean_dec(x_32); -x_41 = lean_ctor_get(x_12, 1); -lean_inc(x_41); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -lean_dec(x_41); -x_43 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); -lean_closure_set(x_43, 0, x_12); -lean_closure_set(x_43, 1, x_1); -x_44 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__4(x_42, x_43, x_2, x_3, x_4, x_5, x_6, x_7, x_40); -return x_44; -} -} -else -{ -uint8_t x_45; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_45 = !lean_is_exclusive(x_32); -if (x_45 == 0) -{ -return x_32; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_32, 0); -x_47 = lean_ctor_get(x_32, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_32); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_5, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 1); -lean_inc(x_16); -lean_dec(x_5); -lean_inc(x_15); -lean_inc(x_1); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(x_2, x_1, x_15); -if (x_17 == 0) -{ -lean_dec(x_15); -x_5 = x_16; -goto _start; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Lean_Name_getString_x21(x_3); -x_20 = lean_box(0); -x_21 = l_Lean_Name_str___override(x_20, x_19); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_6); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_21, x_15, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -if (lean_obj_tag(x_23) == 0) -{ -uint8_t x_24; -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 0); -lean_dec(x_25); -x_26 = lean_box(0); -lean_ctor_set(x_22, 0, x_26); -return x_22; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_dec(x_22); -x_28 = lean_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -else -{ -lean_object* x_30; -lean_dec(x_23); -x_30 = lean_ctor_get(x_22, 1); -lean_inc(x_30); -lean_dec(x_22); -x_5 = x_16; -x_12 = x_30; -goto _start; -} -} -else -{ -uint8_t x_32; -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_22); -if (x_32 == 0) -{ -return x_22; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_22); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_5, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_5, 1); -lean_inc(x_16); -lean_dec(x_5); -lean_inc(x_15); -lean_inc(x_1); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(x_2, x_1, x_15); -if (x_17 == 0) -{ -lean_dec(x_15); -x_5 = x_16; -goto _start; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Lean_Name_getString_x21(x_3); -x_20 = lean_box(0); -x_21 = l_Lean_Name_str___override(x_20, x_19); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_6); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_21, x_15, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -if (lean_obj_tag(x_23) == 0) -{ -uint8_t x_24; -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 0); -lean_dec(x_25); -x_26 = lean_box(0); -lean_ctor_set(x_22, 0, x_26); -return x_22; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_22, 1); -lean_inc(x_27); -lean_dec(x_22); -x_28 = lean_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -else -{ -lean_object* x_30; -lean_dec(x_23); -x_30 = lean_ctor_get(x_22, 1); -lean_inc(x_30); -lean_dec(x_22); -x_5 = x_16; -x_12 = x_30; -goto _start; -} -} -else -{ -uint8_t x_32; -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_22); -if (x_32 == 0) -{ -return x_22; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_22, 0); -x_34 = lean_ctor_get(x_22, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_22); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_dec(x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_16 = lean_apply_9(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_2 = x_25; -x_3 = x_15; -x_10 = x_24; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) -{ -return x_16; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_array_uget(x_2, x_3); -switch (lean_obj_tag(x_14)) { -case 0: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -lean_inc(x_1); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_17 = lean_apply_10(x_1, x_5, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) -{ -uint8_t x_19; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_17, 0, x_21); -return x_17; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -else -{ -lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_dec(x_17); -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -lean_dec(x_18); -x_27 = 1; -x_28 = lean_usize_add(x_3, x_27); -x_3 = x_28; -x_5 = x_26; -x_12 = x_25; -goto _start; -} -} -else -{ -uint8_t x_30; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_17); -if (x_30 == 0) -{ -return x_17; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -case 1: -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_14, 0); -lean_inc(x_34); -lean_dec(x_14); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_35 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_1, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -if (lean_obj_tag(x_36) == 0) -{ -uint8_t x_37; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_37 = !lean_is_exclusive(x_35); -if (x_37 == 0) -{ -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_35, 0); -lean_dec(x_38); -x_39 = lean_box(0); -lean_ctor_set(x_35, 0, x_39); -return x_35; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_35, 1); -lean_inc(x_40); -lean_dec(x_35); -x_41 = lean_box(0); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -return x_42; -} -} -else -{ -lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; -x_43 = lean_ctor_get(x_35, 1); -lean_inc(x_43); -lean_dec(x_35); -x_44 = lean_ctor_get(x_36, 0); -lean_inc(x_44); -lean_dec(x_36); -x_45 = 1; -x_46 = lean_usize_add(x_3, x_45); -x_3 = x_46; -x_5 = x_44; -x_12 = x_43; -goto _start; -} -} -else -{ -uint8_t x_48; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_35); -if (x_48 == 0) -{ -return x_35; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_35, 0); -x_50 = lean_ctor_get(x_35, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_35); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; -} -} -} -default: -{ -size_t x_52; size_t x_53; -x_52 = 1; -x_53 = lean_usize_add(x_3, x_52); -x_3 = x_53; -goto _start; -} -} -} -else -{ -lean_object* x_55; lean_object* x_56; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_55 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_55, 0, x_5); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_12); -return x_56; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed), 12, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_array_get_size(x_2); -x_15 = lean_nat_dec_lt(x_5, x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_6); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_13); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_array_fget(x_2, x_5); -x_19 = lean_array_fget(x_3, x_5); -lean_inc(x_1); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_20 = lean_apply_10(x_1, x_6, x_18, x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) -{ -uint8_t x_22; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -x_24 = lean_box(0); -lean_ctor_set(x_20, 0, x_24); -return x_20; -} -else -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 1); -lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -return x_27; -} -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_dec(x_20); -x_29 = lean_ctor_get(x_21, 0); -lean_inc(x_29); -lean_dec(x_21); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_5, x_30); -lean_dec(x_5); -x_4 = lean_box(0); -x_5 = x_31; -x_6 = x_29; -x_13 = x_28; -goto _start; -} -} -else -{ -uint8_t x_33; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_20); -if (x_33 == 0) -{ -return x_20; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_20, 0); -x_35 = lean_ctor_get(x_20, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_20); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed), 13, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_array_get_size(x_12); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_3); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_2); -lean_ctor_set(x_16, 1, x_10); -return x_16; -} -else -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_13, x_13); -if (x_17 == 0) -{ -lean_object* x_18; -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_3); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_2); -lean_ctor_set(x_18, 1, x_10); -return x_18; -} -else -{ -size_t x_19; size_t x_20; lean_object* x_21; -lean_free_object(x_2); -x_19 = 0; -x_20 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_21 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(x_1, x_12, x_19, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_12); -return x_21; -} -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_22 = lean_ctor_get(x_2, 0); -lean_inc(x_22); -lean_dec(x_2); -x_23 = lean_array_get_size(x_22); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_nat_dec_lt(x_24, x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_3); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_10); -return x_27; -} -else -{ -uint8_t x_28; -x_28 = lean_nat_dec_le(x_23, x_23); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_3); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_10); -return x_30; -} -else -{ -size_t x_31; size_t x_32; lean_object* x_33; -x_31 = 0; -x_32 = lean_usize_of_nat(x_23); -lean_dec(x_23); -x_33 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(x_1, x_22, x_31, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_22); -return x_33; -} -} -} -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_2, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_2, 1); -lean_inc(x_35); -lean_dec(x_2); -x_36 = lean_unsigned_to_nat(0u); -x_37 = l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(x_1, x_34, x_35, lean_box(0), x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_35); -lean_dec(x_34); -return x_37; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg), 10, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = lean_box(0); -x_12 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -x_15 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__10(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -x_5 = x_25; -x_12 = x_24; -goto _start; -} -} -else -{ -uint8_t x_29; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_12); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_array_get_size(x_11); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -lean_dec(x_11); -x_15 = lean_ctor_get(x_1, 1); -lean_inc(x_15); -lean_dec(x_1); -x_16 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__11(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_16; -} -else -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_12, x_12); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_12); -lean_dec(x_11); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); -lean_dec(x_1); -x_19 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__11(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_19; -} -else -{ -size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_20 = 0; -x_21 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_22 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_23 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_11); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_24); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_dec(x_23); -x_32 = lean_ctor_get(x_1, 1); -lean_inc(x_32); -lean_dec(x_1); -x_33 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__11(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_31); -return x_33; -} -} -else -{ -uint8_t x_34; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_23); -if (x_34 == 0) -{ -return x_23; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_23, 0); -x_36 = lean_ctor_get(x_23, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_23); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; -x_14 = l_Lean_Name_isPrefixOf(x_1, x_5); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_13); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_box(0); -lean_inc(x_5); -x_18 = l_Lean_Name_replacePrefix(x_5, x_1, x_17); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_2, x_18); -lean_dec(x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_20 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_13); -return x_21; -} -else -{ -lean_object* x_22; double x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_unbox_float(x_22); -lean_dec(x_22); -x_24 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8(x_3, x_4, x_5, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_5); -return x_24; -} -} -} -} -static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_8); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_8); -x_19 = lean_ctor_get(x_7, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_7, 1); -lean_inc(x_20); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - x_21 = x_7; -} else { - lean_dec_ref(x_7); - x_21 = lean_box(0); -} -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_30 = lean_ctor_get(x_19, 0); -lean_inc(x_30); -lean_dec(x_19); -x_31 = l_Lean_getAliasState(x_2); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_32 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed), 13, 4); -lean_closure_set(x_32, 0, x_30); -lean_closure_set(x_32, 1, x_1); -lean_closure_set(x_32, 2, x_2); -lean_closure_set(x_32, 3, x_3); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_33 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__9(x_31, x_32, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -if (lean_obj_tag(x_34) == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_box(0); -x_22 = x_36; -x_23 = x_35; -goto block_29; -} -else -{ -lean_object* x_37; lean_object* x_38; -lean_dec(x_34); -x_37 = lean_ctor_get(x_33, 1); -lean_inc(x_37); -lean_dec(x_33); -x_38 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_22 = x_38; -x_23 = x_37; -goto block_29; -} -} -else -{ -uint8_t x_39; -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_39 = !lean_is_exclusive(x_33); -if (x_39 == 0) -{ -return x_33; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_33, 0); -x_41 = lean_ctor_get(x_33, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_33); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; -} -} -} -else -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_19, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_19, 1); -lean_inc(x_44); -lean_dec(x_19); -lean_inc(x_44); -lean_inc(x_2); -x_45 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_allowCompletion(x_3, x_2, x_44); -if (x_45 == 0) -{ -lean_object* x_46; -lean_dec(x_44); -lean_dec(x_43); -x_46 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_22 = x_46; -x_23 = x_16; -goto block_29; -} -else -{ -lean_object* x_47; -lean_inc(x_1); -x_47 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_1, x_43); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; -lean_dec(x_44); -lean_dec(x_43); -x_48 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_22 = x_48; -x_23 = x_16; -goto block_29; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; double x_53; lean_object* x_54; -x_49 = lean_ctor_get(x_47, 0); -lean_inc(x_49); -lean_dec(x_47); -x_50 = l_Lean_Name_getString_x21(x_43); -lean_dec(x_43); -x_51 = lean_box(0); -x_52 = l_Lean_Name_str___override(x_51, x_50); -x_53 = lean_unbox_float(x_49); -lean_dec(x_49); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_10); -x_54 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_52, x_44, x_53, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = lean_box(0); -x_22 = x_57; -x_23 = x_56; -goto block_29; -} -else -{ -lean_object* x_58; lean_object* x_59; -lean_dec(x_55); -x_58 = lean_ctor_get(x_54, 1); -lean_inc(x_58); -lean_dec(x_54); -x_59 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_22 = x_59; -x_23 = x_58; -goto block_29; -} -} -else -{ -uint8_t x_60; -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_60 = !lean_is_exclusive(x_54); -if (x_60 == 0) -{ -return x_54; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_54, 0); -x_62 = lean_ctor_get(x_54, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_54); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; -} -} -} -} -} -block_29: -{ -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_20); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_box(0); -if (lean_is_scalar(x_21)) { - x_25 = lean_alloc_ctor(0, 2, 0); -} else { - x_25 = x_21; - lean_ctor_set_tag(x_25, 0); -} -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; -} -else -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_21); -x_26 = lean_ctor_get(x_22, 0); -lean_inc(x_26); -lean_dec(x_22); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec(x_26); -x_7 = x_20; -x_8 = x_27; -x_9 = lean_box(0); -x_16 = x_23; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_2); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 2); -lean_inc(x_15); -lean_dec(x_3); -lean_inc(x_1); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_16 = lean_apply_9(x_1, x_13, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_2 = x_25; -x_3 = x_15; -x_10 = x_24; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_15); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) -{ -return x_16; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = lean_box(0); -x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -x_15 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__19(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 0, x_20); -return x_16; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_dec(x_16); -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -lean_dec(x_17); -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -x_5 = x_25; -x_12 = x_24; -goto _start; -} -} -else -{ -uint8_t x_29; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_16); -if (x_29 == 0) -{ -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -} -else -{ -lean_object* x_33; lean_object* x_34; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_5); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_12); -return x_34; -} -} -} -LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_array_get_size(x_11); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_nat_dec_lt(x_13, x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -lean_dec(x_11); -x_15 = lean_ctor_get(x_1, 1); -lean_inc(x_15); -lean_dec(x_1); -x_16 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__20(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_16; -} -else -{ -uint8_t x_17; -x_17 = lean_nat_dec_le(x_12, x_12); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_12); -lean_dec(x_11); -x_18 = lean_ctor_get(x_1, 1); -lean_inc(x_18); -lean_dec(x_1); -x_19 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__20(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_19; -} -else -{ -size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; -x_20 = 0; -x_21 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_22 = lean_box(0); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_23 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_11); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_dec(x_24); -x_31 = lean_ctor_get(x_23, 1); -lean_inc(x_31); -lean_dec(x_23); -x_32 = lean_ctor_get(x_1, 1); -lean_inc(x_32); -lean_dec(x_1); -x_33 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__20(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_31); -return x_33; -} -} -else -{ -uint8_t x_34; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_23); -if (x_34 == 0) -{ -return x_23; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_23, 0); -x_36 = lean_ctor_get(x_23, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_23); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_6, x_5); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_8); -lean_dec(x_1); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_7); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_14); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; double x_27; lean_object* x_28; -lean_dec(x_7); -x_18 = lean_array_uget(x_4, x_6); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -lean_inc(x_1); -x_28 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_18, x_27); -if (lean_obj_tag(x_28) == 0) -{ -lean_object* x_29; -lean_dec(x_18); -x_29 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_19 = x_29; -x_20 = x_14; -goto block_26; -} -else -{ -lean_object* x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_unbox_float(x_30); -lean_dec(x_30); -lean_inc(x_8); -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem(x_18, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_19 = x_34; -x_20 = x_33; -goto block_26; -} -block_26: -{ -lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec(x_21); -x_23 = 1; -x_24 = lean_usize_add(x_6, x_23); -x_6 = x_24; -x_7 = x_22; -x_14 = x_20; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -uint8_t x_17; -x_17 = lean_usize_dec_lt(x_8, x_7); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_1); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_9); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_16); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_40; -x_20 = lean_array_uget(x_6, x_8); -x_40 = !lean_is_exclusive(x_9); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_9, 1); -x_42 = lean_ctor_get(x_9, 0); -lean_dec(x_42); -lean_inc(x_10); -lean_inc(x_41); -lean_inc(x_1); -x_43 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_20, x_41, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_20); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) -{ -lean_object* x_46; -x_46 = lean_ctor_get(x_44, 0); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_43, 1); -lean_inc(x_47); -lean_dec(x_43); -lean_ctor_set(x_9, 0, x_44); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, x_9); -x_49 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_21 = x_49; -x_22 = x_47; -goto block_39; -} -else -{ -lean_object* x_50; uint8_t x_51; -lean_dec(x_41); -x_50 = lean_ctor_get(x_43, 1); -lean_inc(x_50); -lean_dec(x_43); -x_51 = !lean_is_exclusive(x_46); -if (x_51 == 0) -{ -lean_object* x_52; -x_52 = lean_ctor_get(x_46, 0); -lean_inc(x_5); -lean_ctor_set(x_9, 1, x_52); -lean_ctor_set(x_9, 0, x_5); -lean_ctor_set(x_46, 0, x_9); -x_21 = x_44; -x_22 = x_50; -goto block_39; -} -else -{ -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_46, 0); -lean_inc(x_53); -lean_dec(x_46); -lean_inc(x_5); -lean_ctor_set(x_9, 1, x_53); -lean_ctor_set(x_9, 0, x_5); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_9); -lean_ctor_set(x_44, 0, x_54); -x_21 = x_44; -x_22 = x_50; -goto block_39; -} -} -} -else -{ -lean_object* x_55; -x_55 = lean_ctor_get(x_44, 0); -lean_inc(x_55); -lean_dec(x_44); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_43, 1); -lean_inc(x_56); -lean_dec(x_43); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_9, 0, x_57); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_9); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_21 = x_59; -x_22 = x_56; -goto block_39; -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_41); -x_60 = lean_ctor_get(x_43, 1); -lean_inc(x_60); -lean_dec(x_43); -x_61 = lean_ctor_get(x_55, 0); -lean_inc(x_61); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - x_62 = x_55; -} else { - lean_dec_ref(x_55); - x_62 = lean_box(0); -} -lean_inc(x_5); -lean_ctor_set(x_9, 1, x_61); -lean_ctor_set(x_9, 0, x_5); -if (lean_is_scalar(x_62)) { - x_63 = lean_alloc_ctor(1, 1, 0); -} else { - x_63 = x_62; -} -lean_ctor_set(x_63, 0, x_9); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -x_21 = x_64; -x_22 = x_60; -goto block_39; -} -} -} -else -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_65 = lean_ctor_get(x_9, 1); -lean_inc(x_65); -lean_dec(x_9); -lean_inc(x_10); -lean_inc(x_65); -lean_inc(x_1); -x_66 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_20, x_65, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_20); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - x_69 = x_67; -} else { - lean_dec_ref(x_67); - x_69 = lean_box(0); -} -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_66, 1); -lean_inc(x_70); -lean_dec(x_66); -if (lean_is_scalar(x_69)) { - x_71 = lean_alloc_ctor(1, 1, 0); -} else { - x_71 = x_69; -} -lean_ctor_set(x_71, 0, x_68); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_65); -x_73 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_73, 0, x_72); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -x_21 = x_74; -x_22 = x_70; -goto block_39; -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_65); -x_75 = lean_ctor_get(x_66, 1); -lean_inc(x_75); -lean_dec(x_66); -x_76 = lean_ctor_get(x_68, 0); -lean_inc(x_76); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - x_77 = x_68; -} else { - lean_dec_ref(x_68); - x_77 = lean_box(0); -} -lean_inc(x_5); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_5); -lean_ctor_set(x_78, 1, x_76); -if (lean_is_scalar(x_77)) { - x_79 = lean_alloc_ctor(1, 1, 0); -} else { - x_79 = x_77; -} -lean_ctor_set(x_79, 0, x_78); -if (lean_is_scalar(x_69)) { - x_80 = lean_alloc_ctor(1, 1, 0); -} else { - x_80 = x_69; -} -lean_ctor_set(x_80, 0, x_79); -x_21 = x_80; -x_22 = x_75; -goto block_39; -} -} -block_39: -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_21); -if (x_23 == 0) -{ -lean_object* x_24; -x_24 = lean_ctor_get(x_21, 0); -if (lean_obj_tag(x_24) == 0) -{ -lean_object* x_25; lean_object* x_26; -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_1); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -lean_dec(x_24); -lean_ctor_set(x_21, 0, x_25); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_21); -lean_ctor_set(x_26, 1, x_22); -return x_26; -} -else -{ -lean_object* x_27; size_t x_28; size_t x_29; -lean_free_object(x_21); -x_27 = lean_ctor_get(x_24, 0); -lean_inc(x_27); -lean_dec(x_24); -x_28 = 1; -x_29 = lean_usize_add(x_8, x_28); -x_8 = x_29; -x_9 = x_27; -x_16 = x_22; -goto _start; -} -} -else -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_21, 0); -lean_inc(x_31); -lean_dec(x_21); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_1); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_22); -return x_34; -} -else -{ -lean_object* x_35; size_t x_36; size_t x_37; -x_35 = lean_ctor_get(x_31, 0); -lean_inc(x_35); -lean_dec(x_31); -x_36 = 1; -x_37 = lean_usize_add(x_8, x_36); -x_8 = x_37; -x_9 = x_35; -x_16 = x_22; -goto _start; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_7, x_6); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_1); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_8); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_15); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_19 = lean_array_uget(x_5, x_7); -x_28 = lean_ctor_get(x_8, 1); -lean_inc(x_28); -if (lean_is_exclusive(x_8)) { - lean_ctor_release(x_8, 0); - lean_ctor_release(x_8, 1); - x_29 = x_8; -} else { - lean_dec_ref(x_8); - x_29 = lean_box(0); -} -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_28); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_30 = x_48; -x_31 = x_15; -goto block_46; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_28); -x_49 = lean_ctor_get(x_19, 0); -lean_inc(x_49); -lean_dec(x_19); -x_50 = l_Lean_LocalDecl_userName(x_49); -lean_inc(x_1); -x_51 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_1, x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; -lean_dec(x_50); -lean_dec(x_49); -x_52 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_52; -x_31 = x_15; -goto block_46; -} -else -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_51); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; uint8_t x_56; double x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_54 = lean_ctor_get(x_51, 0); -x_55 = l_Lean_LocalDecl_fvarId(x_49); -lean_dec(x_49); -lean_ctor_set(x_51, 0, x_55); -x_56 = 5; -x_57 = lean_unbox_float(x_54); -lean_dec(x_54); -lean_inc(x_9); -x_58 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_50, x_51, x_56, x_57, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_60; -x_31 = x_59; -goto block_46; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; double x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_61 = lean_ctor_get(x_51, 0); -lean_inc(x_61); -lean_dec(x_51); -x_62 = l_Lean_LocalDecl_fvarId(x_49); -lean_dec(x_49); -x_63 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_63, 0, x_62); -x_64 = 5; -x_65 = lean_unbox_float(x_61); -lean_dec(x_61); -lean_inc(x_9); -x_66 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_50, x_63, x_64, x_65, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_68; -x_31 = x_67; -goto block_46; -} -} -} -block_27: -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec(x_22); -x_24 = 1; -x_25 = lean_usize_add(x_7, x_24); -x_7 = x_25; -x_8 = x_23; -x_15 = x_21; -goto _start; -} -block_46: -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_30, 0); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_29; -} -lean_ctor_set(x_36, 0, x_4); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_33, 0, x_36); -x_20 = x_30; -x_21 = x_31; -goto block_27; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 0); -lean_inc(x_37); -lean_dec(x_33); -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_29; -} -lean_ctor_set(x_38, 0, x_4); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_30, 0, x_39); -x_20 = x_30; -x_21 = x_31; -goto block_27; -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_30, 0); -lean_inc(x_40); -lean_dec(x_30); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - x_42 = x_40; -} else { - lean_dec_ref(x_40); - x_42 = lean_box(0); -} -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_29; -} -lean_ctor_set(x_43, 0, x_4); -lean_ctor_set(x_43, 1, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(1, 1, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_20 = x_45; -x_21 = x_31; -goto block_27; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_1); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_12 = lean_ctor_get(x_3, 0); -x_13 = lean_box(0); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_4); -x_16 = lean_array_size(x_12); -x_17 = 0; -lean_inc(x_5); -x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26(x_1, x_2, x_12, x_13, x_14, x_12, x_16, x_17, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_box(0); -x_25 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_23, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -lean_dec(x_5); -return x_25; -} -else -{ -uint8_t x_26; -lean_dec(x_20); -lean_dec(x_5); -x_26 = !lean_is_exclusive(x_18); -if (x_26 == 0) -{ -lean_object* x_27; uint8_t x_28; -x_27 = lean_ctor_get(x_18, 0); -lean_dec(x_27); -x_28 = !lean_is_exclusive(x_21); -if (x_28 == 0) -{ -lean_ctor_set(x_18, 0, x_21); -return x_18; -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_21, 0); -lean_inc(x_29); -lean_dec(x_21); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_18, 0, x_30); -return x_18; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_18, 1); -lean_inc(x_31); -lean_dec(x_18); -x_32 = lean_ctor_get(x_21, 0); -lean_inc(x_32); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - x_33 = x_21; -} else { - lean_dec_ref(x_21); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(1, 1, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_32); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_31); -return x_35; -} -} -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_36 = lean_ctor_get(x_3, 0); -x_37 = lean_box(0); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_4); -x_40 = lean_array_size(x_36); -x_41 = 0; -lean_inc(x_5); -x_42 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27(x_1, x_36, x_37, x_38, x_36, x_40, x_41, x_39, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -lean_dec(x_43); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_46 = lean_ctor_get(x_42, 1); -lean_inc(x_46); -lean_dec(x_42); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_box(0); -x_49 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_47, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_46); -lean_dec(x_5); -return x_49; -} -else -{ -uint8_t x_50; -lean_dec(x_44); -lean_dec(x_5); -x_50 = !lean_is_exclusive(x_42); -if (x_50 == 0) -{ -lean_object* x_51; uint8_t x_52; -x_51 = lean_ctor_get(x_42, 0); -lean_dec(x_51); -x_52 = !lean_is_exclusive(x_45); -if (x_52 == 0) -{ -lean_ctor_set(x_42, 0, x_45); -return x_42; -} -else -{ -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_45, 0); -lean_inc(x_53); -lean_dec(x_45); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_42, 0, x_54); -return x_42; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_55 = lean_ctor_get(x_42, 1); -lean_inc(x_55); -lean_dec(x_42); -x_56 = lean_ctor_get(x_45, 0); -lean_inc(x_56); -if (lean_is_exclusive(x_45)) { - lean_ctor_release(x_45, 0); - x_57 = x_45; -} else { - lean_dec_ref(x_45); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 1, 0); -} else { - x_58 = x_57; -} -lean_ctor_set(x_58, 0, x_56); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_55); -return x_59; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -uint8_t x_16; -x_16 = lean_usize_dec_lt(x_7, x_6); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_1); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_8); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_15); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_19 = lean_array_uget(x_5, x_7); -x_28 = lean_ctor_get(x_8, 1); -lean_inc(x_28); -if (lean_is_exclusive(x_8)) { - lean_ctor_release(x_8, 0); - lean_ctor_release(x_8, 1); - x_29 = x_8; -} else { - lean_dec_ref(x_8); - x_29 = lean_box(0); -} -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_28); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_30 = x_48; -x_31 = x_15; -goto block_46; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -lean_dec(x_28); -x_49 = lean_ctor_get(x_19, 0); -lean_inc(x_49); -lean_dec(x_19); -x_50 = l_Lean_LocalDecl_userName(x_49); -lean_inc(x_1); -x_51 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_1, x_50); -if (lean_obj_tag(x_51) == 0) -{ -lean_object* x_52; -lean_dec(x_50); -lean_dec(x_49); -x_52 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_52; -x_31 = x_15; -goto block_46; -} -else -{ -uint8_t x_53; -x_53 = !lean_is_exclusive(x_51); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; uint8_t x_56; double x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_54 = lean_ctor_get(x_51, 0); -x_55 = l_Lean_LocalDecl_fvarId(x_49); -lean_dec(x_49); -lean_ctor_set(x_51, 0, x_55); -x_56 = 5; -x_57 = lean_unbox_float(x_54); -lean_dec(x_54); -lean_inc(x_9); -x_58 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_50, x_51, x_56, x_57, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_60; -x_31 = x_59; -goto block_46; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; double x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_61 = lean_ctor_get(x_51, 0); -lean_inc(x_61); -lean_dec(x_51); -x_62 = l_Lean_LocalDecl_fvarId(x_49); -lean_dec(x_49); -x_63 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_63, 0, x_62); -x_64 = 5; -x_65 = lean_unbox_float(x_61); -lean_dec(x_61); -lean_inc(x_9); -x_66 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_50, x_63, x_64, x_65, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -x_67 = lean_ctor_get(x_66, 1); -lean_inc(x_67); -lean_dec(x_66); -x_68 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_30 = x_68; -x_31 = x_67; -goto block_46; -} -} -} -block_27: -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -lean_dec(x_22); -x_24 = 1; -x_25 = lean_usize_add(x_7, x_24); -x_7 = x_25; -x_8 = x_23; -x_15 = x_21; -goto _start; -} -block_46: -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_30); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_30, 0); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_29; -} -lean_ctor_set(x_36, 0, x_4); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_33, 0, x_36); -x_20 = x_30; -x_21 = x_31; -goto block_27; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 0); -lean_inc(x_37); -lean_dec(x_33); -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_38 = lean_alloc_ctor(0, 2, 0); -} else { - x_38 = x_29; -} -lean_ctor_set(x_38, 0, x_4); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_30, 0, x_39); -x_20 = x_30; -x_21 = x_31; -goto block_27; -} -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_30, 0); -lean_inc(x_40); -lean_dec(x_30); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - x_42 = x_40; -} else { - lean_dec_ref(x_40); - x_42 = lean_box(0); -} -lean_inc(x_4); -if (lean_is_scalar(x_29)) { - x_43 = lean_alloc_ctor(0, 2, 0); -} else { - x_43 = x_29; -} -lean_ctor_set(x_43, 0, x_4); -lean_ctor_set(x_43, 1, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(1, 1, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_43); -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_44); -x_20 = x_45; -x_21 = x_31; -goto block_27; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_1); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_12 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_3, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_3); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -lean_dec(x_4); -lean_dec(x_1); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_12, 0); -lean_dec(x_17); -x_18 = lean_ctor_get(x_15, 0); -lean_inc(x_18); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_18); -return x_12; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_12, 1); -lean_inc(x_19); -lean_dec(x_12); -x_20 = lean_ctor_get(x_15, 0); -lean_inc(x_20); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_20); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -lean_free_object(x_13); -x_22 = lean_ctor_get(x_12, 1); -lean_inc(x_22); -lean_dec(x_12); -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = lean_ctor_get(x_2, 1); -x_25 = lean_box(0); -x_26 = lean_box(0); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_23); -x_28 = lean_array_size(x_24); -x_29 = 0; -lean_inc(x_4); -x_30 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28(x_1, x_24, x_25, x_26, x_24, x_28, x_29, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_22); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -lean_dec(x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -if (lean_obj_tag(x_33) == 0) -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_box(0); -x_37 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(x_35, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_34); -lean_dec(x_4); -return x_37; -} -else -{ -uint8_t x_38; -lean_dec(x_32); -lean_dec(x_4); -x_38 = !lean_is_exclusive(x_30); -if (x_38 == 0) -{ -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_30, 0); -lean_dec(x_39); -x_40 = !lean_is_exclusive(x_33); -if (x_40 == 0) -{ -lean_ctor_set(x_30, 0, x_33); -return x_30; -} -else -{ -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_33, 0); -lean_inc(x_41); -lean_dec(x_33); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_30, 0, x_42); -return x_30; -} -} -else -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_43 = lean_ctor_get(x_30, 1); -lean_inc(x_43); -lean_dec(x_30); -x_44 = lean_ctor_get(x_33, 0); -lean_inc(x_44); -if (lean_is_exclusive(x_33)) { - lean_ctor_release(x_33, 0); - x_45 = x_33; -} else { - lean_dec_ref(x_33); - x_45 = lean_box(0); -} -if (lean_is_scalar(x_45)) { - x_46 = lean_alloc_ctor(1, 1, 0); -} else { - x_46 = x_45; -} -lean_ctor_set(x_46, 0, x_44); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_43); -return x_47; -} -} -} -} -else -{ -lean_object* x_48; -x_48 = lean_ctor_get(x_13, 0); -lean_inc(x_48); -lean_dec(x_13); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_4); -lean_dec(x_1); -x_49 = lean_ctor_get(x_12, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - lean_ctor_release(x_12, 1); - x_50 = x_12; -} else { - lean_dec_ref(x_12); - x_50 = lean_box(0); -} -x_51 = lean_ctor_get(x_48, 0); -lean_inc(x_51); -lean_dec(x_48); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_51); -if (lean_is_scalar(x_50)) { - x_53 = lean_alloc_ctor(0, 2, 0); -} else { - x_53 = x_50; -} -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_49); -return x_53; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; size_t x_60; size_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_54 = lean_ctor_get(x_12, 1); -lean_inc(x_54); -lean_dec(x_12); -x_55 = lean_ctor_get(x_48, 0); -lean_inc(x_55); -lean_dec(x_48); -x_56 = lean_ctor_get(x_2, 1); -x_57 = lean_box(0); -x_58 = lean_box(0); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_55); -x_60 = lean_array_size(x_56); -x_61 = 0; -lean_inc(x_4); -x_62 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28(x_1, x_56, x_57, x_58, x_56, x_60, x_61, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_54); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -lean_dec(x_63); -x_65 = lean_ctor_get(x_64, 0); -lean_inc(x_65); -if (lean_obj_tag(x_65) == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_62, 1); -lean_inc(x_66); -lean_dec(x_62); -x_67 = lean_ctor_get(x_64, 1); -lean_inc(x_67); -lean_dec(x_64); -x_68 = lean_box(0); -x_69 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(x_67, x_68, x_4, x_5, x_6, x_7, x_8, x_9, x_66); -lean_dec(x_4); -return x_69; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_64); -lean_dec(x_4); -x_70 = lean_ctor_get(x_62, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_71 = x_62; -} else { - lean_dec_ref(x_62); - x_71 = lean_box(0); -} -x_72 = lean_ctor_get(x_65, 0); -lean_inc(x_72); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - x_73 = x_65; -} else { - lean_dec_ref(x_65); - x_73 = lean_box(0); -} -if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(1, 1, 0); -} else { - x_74 = x_73; -} -lean_ctor_set(x_74, 0, x_72); -if (lean_is_scalar(x_71)) { - x_75 = lean_alloc_ctor(0, 2, 0); -} else { - x_75 = x_71; -} -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_70); -return x_75; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_4, x_1, x_2, x_3, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_5); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_dec(x_13); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_20); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_14); -if (x_25 == 0) -{ -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_14, 0); -x_28 = lean_ctor_get(x_13, 0); -lean_dec(x_28); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_box(0); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_14); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_13, 0, x_32); -return x_13; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_ctor_set(x_14, 0, x_35); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_14); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_13, 0, x_38); -return x_13; -} -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_39 = lean_ctor_get(x_14, 0); -x_40 = lean_ctor_get(x_13, 1); -lean_inc(x_40); -lean_dec(x_13); -x_41 = lean_ctor_get(x_39, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_39, 1); -lean_inc(x_42); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_43 = x_39; -} else { - lean_dec_ref(x_39); - x_43 = lean_box(0); -} -if (lean_is_scalar(x_43)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_43; -} -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_42); -lean_ctor_set(x_14, 0, x_44); -x_45 = lean_box(0); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_14); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_40); -return x_48; -} -} -else -{ -uint8_t x_49; -x_49 = !lean_is_exclusive(x_5); -if (x_49 == 0) -{ -uint8_t x_50; -x_50 = !lean_is_exclusive(x_13); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_51 = lean_ctor_get(x_14, 0); -x_52 = lean_ctor_get(x_5, 0); -x_53 = lean_ctor_get(x_13, 0); -lean_dec(x_53); -x_54 = !lean_is_exclusive(x_51); -if (x_54 == 0) -{ -uint8_t x_55; -x_55 = !lean_is_exclusive(x_52); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_56 = lean_ctor_get(x_51, 0); -x_57 = lean_ctor_get(x_51, 1); -x_58 = lean_ctor_get(x_52, 0); -x_59 = lean_ctor_get(x_52, 1); -x_60 = l_Lean_Name_isSuffixOf(x_56, x_58); -if (x_60 == 0) -{ -lean_object* x_61; -lean_dec(x_57); -lean_dec(x_56); -x_61 = lean_box(0); -lean_ctor_set(x_51, 1, x_5); -lean_ctor_set(x_51, 0, x_61); -return x_13; -} -else -{ -lean_object* x_62; -lean_dec(x_59); -lean_dec(x_58); -lean_ctor_set(x_52, 1, x_57); -lean_ctor_set(x_52, 0, x_56); -x_62 = lean_box(0); -lean_ctor_set(x_51, 1, x_5); -lean_ctor_set(x_51, 0, x_62); -return x_13; -} -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_63 = lean_ctor_get(x_51, 0); -x_64 = lean_ctor_get(x_51, 1); -x_65 = lean_ctor_get(x_52, 0); -x_66 = lean_ctor_get(x_52, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_52); -x_67 = l_Lean_Name_isSuffixOf(x_63, x_65); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; -lean_dec(x_64); -lean_dec(x_63); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_65); -lean_ctor_set(x_68, 1, x_66); -lean_ctor_set(x_5, 0, x_68); -x_69 = lean_box(0); -lean_ctor_set(x_51, 1, x_5); -lean_ctor_set(x_51, 0, x_69); -return x_13; -} -else -{ -lean_object* x_70; lean_object* x_71; -lean_dec(x_66); -lean_dec(x_65); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_63); -lean_ctor_set(x_70, 1, x_64); -lean_ctor_set(x_5, 0, x_70); -x_71 = lean_box(0); -lean_ctor_set(x_51, 1, x_5); -lean_ctor_set(x_51, 0, x_71); -return x_13; -} -} -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_72 = lean_ctor_get(x_51, 0); -x_73 = lean_ctor_get(x_51, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_51); -x_74 = lean_ctor_get(x_52, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_52, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_76 = x_52; -} else { - lean_dec_ref(x_52); - x_76 = lean_box(0); -} -x_77 = l_Lean_Name_isSuffixOf(x_72, x_74); -if (x_77 == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; -lean_dec(x_73); -lean_dec(x_72); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_74); -lean_ctor_set(x_78, 1, x_75); -lean_ctor_set(x_5, 0, x_78); -x_79 = lean_box(0); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_5); -lean_ctor_set(x_14, 0, x_80); -return x_13; -} -else -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_75); -lean_dec(x_74); -if (lean_is_scalar(x_76)) { - x_81 = lean_alloc_ctor(0, 2, 0); -} else { - x_81 = x_76; -} -lean_ctor_set(x_81, 0, x_72); -lean_ctor_set(x_81, 1, x_73); -lean_ctor_set(x_5, 0, x_81); -x_82 = lean_box(0); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_5); -lean_ctor_set(x_14, 0, x_83); -return x_13; -} -} -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_84 = lean_ctor_get(x_14, 0); -x_85 = lean_ctor_get(x_5, 0); -x_86 = lean_ctor_get(x_13, 1); -lean_inc(x_86); -lean_dec(x_13); -x_87 = lean_ctor_get(x_84, 0); -lean_inc(x_87); -x_88 = lean_ctor_get(x_84, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_89 = x_84; -} else { - lean_dec_ref(x_84); - x_89 = lean_box(0); -} -x_90 = lean_ctor_get(x_85, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_85, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_85)) { - lean_ctor_release(x_85, 0); - lean_ctor_release(x_85, 1); - x_92 = x_85; -} else { - lean_dec_ref(x_85); - x_92 = lean_box(0); -} -x_93 = l_Lean_Name_isSuffixOf(x_87, x_90); -if (x_93 == 0) -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -lean_dec(x_88); -lean_dec(x_87); -if (lean_is_scalar(x_92)) { - x_94 = lean_alloc_ctor(0, 2, 0); -} else { - x_94 = x_92; -} -lean_ctor_set(x_94, 0, x_90); -lean_ctor_set(x_94, 1, x_91); -lean_ctor_set(x_5, 0, x_94); -x_95 = lean_box(0); -if (lean_is_scalar(x_89)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_89; -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_5); -lean_ctor_set(x_14, 0, x_96); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_14); -lean_ctor_set(x_97, 1, x_86); -return x_97; -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_91); -lean_dec(x_90); -if (lean_is_scalar(x_92)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_92; -} -lean_ctor_set(x_98, 0, x_87); -lean_ctor_set(x_98, 1, x_88); -lean_ctor_set(x_5, 0, x_98); -x_99 = lean_box(0); -if (lean_is_scalar(x_89)) { - x_100 = lean_alloc_ctor(0, 2, 0); -} else { - x_100 = x_89; -} -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_5); -lean_ctor_set(x_14, 0, x_100); -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_14); -lean_ctor_set(x_101, 1, x_86); -return x_101; -} -} -} -else -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_102 = lean_ctor_get(x_14, 0); -x_103 = lean_ctor_get(x_5, 0); -lean_inc(x_103); -lean_dec(x_5); -x_104 = lean_ctor_get(x_13, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_13)) { - lean_ctor_release(x_13, 0); - lean_ctor_release(x_13, 1); - x_105 = x_13; -} else { - lean_dec_ref(x_13); - x_105 = lean_box(0); -} -x_106 = lean_ctor_get(x_102, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_102, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_108 = x_102; -} else { - lean_dec_ref(x_102); - x_108 = lean_box(0); -} -x_109 = lean_ctor_get(x_103, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_103, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_111 = x_103; -} else { - lean_dec_ref(x_103); - x_111 = lean_box(0); -} -x_112 = l_Lean_Name_isSuffixOf(x_106, x_109); -if (x_112 == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_107); -lean_dec(x_106); -if (lean_is_scalar(x_111)) { - x_113 = lean_alloc_ctor(0, 2, 0); -} else { - x_113 = x_111; -} -lean_ctor_set(x_113, 0, x_109); -lean_ctor_set(x_113, 1, x_110); -x_114 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_114, 0, x_113); -x_115 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_116 = lean_alloc_ctor(0, 2, 0); -} else { - x_116 = x_108; -} -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_114); -lean_ctor_set(x_14, 0, x_116); -if (lean_is_scalar(x_105)) { - x_117 = lean_alloc_ctor(0, 2, 0); -} else { - x_117 = x_105; -} -lean_ctor_set(x_117, 0, x_14); -lean_ctor_set(x_117, 1, x_104); -return x_117; -} -else -{ -lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_110); -lean_dec(x_109); -if (lean_is_scalar(x_111)) { - x_118 = lean_alloc_ctor(0, 2, 0); -} else { - x_118 = x_111; -} -lean_ctor_set(x_118, 0, x_106); -lean_ctor_set(x_118, 1, x_107); -x_119 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_119, 0, x_118); -x_120 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_121 = lean_alloc_ctor(0, 2, 0); -} else { - x_121 = x_108; -} -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_119); -lean_ctor_set(x_14, 0, x_121); -if (lean_is_scalar(x_105)) { - x_122 = lean_alloc_ctor(0, 2, 0); -} else { - x_122 = x_105; -} -lean_ctor_set(x_122, 0, x_14); -lean_ctor_set(x_122, 1, x_104); -return x_122; -} -} -} -} -else -{ -lean_object* x_123; -x_123 = lean_ctor_get(x_14, 0); -lean_inc(x_123); -lean_dec(x_14); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_124 = lean_ctor_get(x_13, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_13)) { - lean_ctor_release(x_13, 0); - lean_ctor_release(x_13, 1); - x_125 = x_13; -} else { - lean_dec_ref(x_13); - x_125 = lean_box(0); -} -x_126 = lean_ctor_get(x_123, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_123, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_128 = x_123; -} else { - lean_dec_ref(x_123); - x_128 = lean_box(0); -} -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(0, 2, 0); -} else { - x_129 = x_128; -} -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_127); -x_130 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_130, 0, x_129); -x_131 = lean_box(0); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_130); -x_133 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_133, 0, x_132); -if (lean_is_scalar(x_125)) { - x_134 = lean_alloc_ctor(0, 2, 0); -} else { - x_134 = x_125; -} -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_124); -return x_134; -} -else -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; -x_135 = lean_ctor_get(x_5, 0); -lean_inc(x_135); -if (lean_is_exclusive(x_5)) { - lean_ctor_release(x_5, 0); - x_136 = x_5; -} else { - lean_dec_ref(x_5); - x_136 = lean_box(0); -} -x_137 = lean_ctor_get(x_13, 1); -lean_inc(x_137); -if (lean_is_exclusive(x_13)) { - lean_ctor_release(x_13, 0); - lean_ctor_release(x_13, 1); - x_138 = x_13; -} else { - lean_dec_ref(x_13); - x_138 = lean_box(0); -} -x_139 = lean_ctor_get(x_123, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_123, 1); -lean_inc(x_140); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_141 = x_123; -} else { - lean_dec_ref(x_123); - x_141 = lean_box(0); -} -x_142 = lean_ctor_get(x_135, 0); -lean_inc(x_142); -x_143 = lean_ctor_get(x_135, 1); -lean_inc(x_143); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_144 = x_135; -} else { - lean_dec_ref(x_135); - x_144 = lean_box(0); -} -x_145 = l_Lean_Name_isSuffixOf(x_139, x_142); -if (x_145 == 0) -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; -lean_dec(x_140); -lean_dec(x_139); -if (lean_is_scalar(x_144)) { - x_146 = lean_alloc_ctor(0, 2, 0); -} else { - x_146 = x_144; -} -lean_ctor_set(x_146, 0, x_142); -lean_ctor_set(x_146, 1, x_143); -if (lean_is_scalar(x_136)) { - x_147 = lean_alloc_ctor(1, 1, 0); -} else { - x_147 = x_136; -} -lean_ctor_set(x_147, 0, x_146); -x_148 = lean_box(0); -if (lean_is_scalar(x_141)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_141; -} -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_147); -x_150 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_150, 0, x_149); -if (lean_is_scalar(x_138)) { - x_151 = lean_alloc_ctor(0, 2, 0); -} else { - x_151 = x_138; -} -lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_137); -return x_151; -} -else -{ -lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; -lean_dec(x_143); -lean_dec(x_142); -if (lean_is_scalar(x_144)) { - x_152 = lean_alloc_ctor(0, 2, 0); -} else { - x_152 = x_144; -} -lean_ctor_set(x_152, 0, x_139); -lean_ctor_set(x_152, 1, x_140); -if (lean_is_scalar(x_136)) { - x_153 = lean_alloc_ctor(1, 1, 0); -} else { - x_153 = x_136; -} -lean_ctor_set(x_153, 0, x_152); -x_154 = lean_box(0); -if (lean_is_scalar(x_141)) { - x_155 = lean_alloc_ctor(0, 2, 0); -} else { - x_155 = x_141; -} -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_153); -x_156 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_156, 0, x_155); -if (lean_is_scalar(x_138)) { - x_157 = lean_alloc_ctor(0, 2, 0); -} else { - x_157 = x_138; -} -lean_ctor_set(x_157, 0, x_156); -lean_ctor_set(x_157, 1, x_137); -return x_157; -} -} -} -} -} -else -{ -uint8_t x_158; -lean_dec(x_5); -x_158 = !lean_is_exclusive(x_13); -if (x_158 == 0) -{ -return x_13; -} -else -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_13, 0); -x_160 = lean_ctor_get(x_13, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_13); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set(x_161, 1, x_160); -return x_161; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_49; lean_object* x_50; lean_object* x_58; -x_13 = lean_box(x_2); -lean_inc(x_4); -lean_inc(x_1); -x_14 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed), 12, 3); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_13); -lean_closure_set(x_14, 2, x_4); -x_15 = lean_ctor_get(x_3, 0); -lean_inc(x_15); -lean_dec(x_3); -x_16 = lean_ctor_get(x_15, 4); -lean_inc(x_16); -x_17 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_58 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(x_14, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; -lean_dec(x_15); -lean_dec(x_1); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_49 = x_17; -x_50 = x_60; -goto block_57; -} -else -{ -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_61 = lean_ctor_get(x_59, 0); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_58, 1); -lean_inc(x_62); -lean_dec(x_58); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_ctor_get(x_15, 5); -lean_inc(x_64); -lean_dec(x_15); -x_65 = lean_box(0); -x_66 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_4); -lean_inc(x_1); -x_67 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1(x_2, x_1, x_4, x_64, x_65, x_64, x_64, x_66, lean_box(0), x_63, x_6, x_7, x_8, x_9, x_10, x_11, x_62); -lean_dec(x_64); -if (lean_obj_tag(x_67) == 0) -{ -lean_object* x_68; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; -lean_dec(x_1); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_49 = x_17; -x_50 = x_69; -goto block_57; -} -else -{ -uint8_t x_70; -x_70 = !lean_is_exclusive(x_68); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_71 = lean_ctor_get(x_68, 0); -x_72 = lean_ctor_get(x_67, 1); -lean_inc(x_72); -lean_dec(x_67); -x_73 = !lean_is_exclusive(x_71); -if (x_73 == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_74 = lean_ctor_get(x_71, 1); -x_75 = lean_ctor_get(x_71, 0); -lean_dec(x_75); -x_76 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_4); -x_77 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_76, x_1, x_2, x_4, x_8, x_9, x_10, x_11, x_72); -if (lean_obj_tag(x_77) == 0) -{ -lean_object* x_78; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -lean_ctor_set(x_71, 0, x_66); -x_49 = x_68; -x_50 = x_79; -goto block_57; -} -else -{ -uint8_t x_80; -x_80 = !lean_is_exclusive(x_78); -if (x_80 == 0) -{ -if (lean_obj_tag(x_74) == 0) -{ -lean_object* x_81; lean_object* x_82; uint8_t x_83; -x_81 = lean_ctor_get(x_78, 0); -x_82 = lean_ctor_get(x_77, 1); -lean_inc(x_82); -lean_dec(x_77); -x_83 = !lean_is_exclusive(x_81); -if (x_83 == 0) -{ -lean_ctor_set(x_71, 1, x_78); -lean_ctor_set(x_71, 0, x_66); -x_49 = x_68; -x_50 = x_82; -goto block_57; -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_81, 0); -x_85 = lean_ctor_get(x_81, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_81); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_84); -lean_ctor_set(x_86, 1, x_85); -lean_ctor_set(x_78, 0, x_86); -lean_ctor_set(x_71, 1, x_78); -lean_ctor_set(x_71, 0, x_66); -x_49 = x_68; -x_50 = x_82; -goto block_57; -} -} -else -{ -lean_object* x_87; uint8_t x_88; -x_87 = lean_ctor_get(x_78, 0); -lean_free_object(x_71); -lean_free_object(x_68); -x_88 = !lean_is_exclusive(x_74); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; uint8_t x_91; -x_89 = lean_ctor_get(x_74, 0); -x_90 = lean_ctor_get(x_77, 1); -lean_inc(x_90); -lean_dec(x_77); -x_91 = !lean_is_exclusive(x_87); -if (x_91 == 0) -{ -uint8_t x_92; -x_92 = !lean_is_exclusive(x_89); -if (x_92 == 0) -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_93 = lean_ctor_get(x_87, 0); -x_94 = lean_ctor_get(x_87, 1); -x_95 = lean_ctor_get(x_89, 0); -x_96 = lean_ctor_get(x_89, 1); -x_97 = l_Lean_Name_isSuffixOf(x_93, x_95); -if (x_97 == 0) -{ -lean_dec(x_94); -lean_dec(x_93); -lean_ctor_set(x_87, 1, x_74); -lean_ctor_set(x_87, 0, x_66); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -else -{ -lean_dec(x_96); -lean_dec(x_95); -lean_ctor_set(x_89, 1, x_94); -lean_ctor_set(x_89, 0, x_93); -lean_ctor_set(x_87, 1, x_74); -lean_ctor_set(x_87, 0, x_66); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; -x_98 = lean_ctor_get(x_87, 0); -x_99 = lean_ctor_get(x_87, 1); -x_100 = lean_ctor_get(x_89, 0); -x_101 = lean_ctor_get(x_89, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_89); -x_102 = l_Lean_Name_isSuffixOf(x_98, x_100); -if (x_102 == 0) -{ -lean_object* x_103; -lean_dec(x_99); -lean_dec(x_98); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_100); -lean_ctor_set(x_103, 1, x_101); -lean_ctor_set(x_74, 0, x_103); -lean_ctor_set(x_87, 1, x_74); -lean_ctor_set(x_87, 0, x_66); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -else -{ -lean_object* x_104; -lean_dec(x_101); -lean_dec(x_100); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_98); -lean_ctor_set(x_104, 1, x_99); -lean_ctor_set(x_74, 0, x_104); -lean_ctor_set(x_87, 1, x_74); -lean_ctor_set(x_87, 0, x_66); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; -x_105 = lean_ctor_get(x_87, 0); -x_106 = lean_ctor_get(x_87, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_87); -x_107 = lean_ctor_get(x_89, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_89, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_109 = x_89; -} else { - lean_dec_ref(x_89); - x_109 = lean_box(0); -} -x_110 = l_Lean_Name_isSuffixOf(x_105, x_107); -if (x_110 == 0) -{ -lean_object* x_111; lean_object* x_112; -lean_dec(x_106); -lean_dec(x_105); -if (lean_is_scalar(x_109)) { - x_111 = lean_alloc_ctor(0, 2, 0); -} else { - x_111 = x_109; -} -lean_ctor_set(x_111, 0, x_107); -lean_ctor_set(x_111, 1, x_108); -lean_ctor_set(x_74, 0, x_111); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_66); -lean_ctor_set(x_112, 1, x_74); -lean_ctor_set(x_78, 0, x_112); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -else -{ -lean_object* x_113; lean_object* x_114; -lean_dec(x_108); -lean_dec(x_107); -if (lean_is_scalar(x_109)) { - x_113 = lean_alloc_ctor(0, 2, 0); -} else { - x_113 = x_109; -} -lean_ctor_set(x_113, 0, x_105); -lean_ctor_set(x_113, 1, x_106); -lean_ctor_set(x_74, 0, x_113); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_66); -lean_ctor_set(x_114, 1, x_74); -lean_ctor_set(x_78, 0, x_114); -x_49 = x_78; -x_50 = x_90; -goto block_57; -} -} -} -else -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_115 = lean_ctor_get(x_74, 0); -lean_inc(x_115); -lean_dec(x_74); -x_116 = lean_ctor_get(x_77, 1); -lean_inc(x_116); -lean_dec(x_77); -x_117 = lean_ctor_get(x_87, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_87, 1); -lean_inc(x_118); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_119 = x_87; -} else { - lean_dec_ref(x_87); - x_119 = lean_box(0); -} -x_120 = lean_ctor_get(x_115, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_115, 1); -lean_inc(x_121); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_122 = x_115; -} else { - lean_dec_ref(x_115); - x_122 = lean_box(0); -} -x_123 = l_Lean_Name_isSuffixOf(x_117, x_120); -if (x_123 == 0) -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_118); -lean_dec(x_117); -if (lean_is_scalar(x_122)) { - x_124 = lean_alloc_ctor(0, 2, 0); -} else { - x_124 = x_122; -} -lean_ctor_set(x_124, 0, x_120); -lean_ctor_set(x_124, 1, x_121); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_124); -if (lean_is_scalar(x_119)) { - x_126 = lean_alloc_ctor(0, 2, 0); -} else { - x_126 = x_119; -} -lean_ctor_set(x_126, 0, x_66); -lean_ctor_set(x_126, 1, x_125); -lean_ctor_set(x_78, 0, x_126); -x_49 = x_78; -x_50 = x_116; -goto block_57; -} -else -{ -lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_121); -lean_dec(x_120); -if (lean_is_scalar(x_122)) { - x_127 = lean_alloc_ctor(0, 2, 0); -} else { - x_127 = x_122; -} -lean_ctor_set(x_127, 0, x_117); -lean_ctor_set(x_127, 1, x_118); -x_128 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_128, 0, x_127); -if (lean_is_scalar(x_119)) { - x_129 = lean_alloc_ctor(0, 2, 0); -} else { - x_129 = x_119; -} -lean_ctor_set(x_129, 0, x_66); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_78, 0, x_129); -x_49 = x_78; -x_50 = x_116; -goto block_57; -} -} -} -} -else -{ -lean_object* x_130; -x_130 = lean_ctor_get(x_78, 0); -lean_inc(x_130); -lean_dec(x_78); -if (lean_obj_tag(x_74) == 0) -{ -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_131 = lean_ctor_get(x_77, 1); -lean_inc(x_131); -lean_dec(x_77); -x_132 = lean_ctor_get(x_130, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_134 = x_130; -} else { - lean_dec_ref(x_130); - x_134 = lean_box(0); -} -if (lean_is_scalar(x_134)) { - x_135 = lean_alloc_ctor(0, 2, 0); -} else { - x_135 = x_134; -} -lean_ctor_set(x_135, 0, x_132); -lean_ctor_set(x_135, 1, x_133); -x_136 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_71, 1, x_136); -lean_ctor_set(x_71, 0, x_66); -x_49 = x_68; -x_50 = x_131; -goto block_57; -} -else -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; -lean_free_object(x_71); -lean_free_object(x_68); -x_137 = lean_ctor_get(x_74, 0); -lean_inc(x_137); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - x_138 = x_74; -} else { - lean_dec_ref(x_74); - x_138 = lean_box(0); -} -x_139 = lean_ctor_get(x_77, 1); -lean_inc(x_139); -lean_dec(x_77); -x_140 = lean_ctor_get(x_130, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_130, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_142 = x_130; -} else { - lean_dec_ref(x_130); - x_142 = lean_box(0); -} -x_143 = lean_ctor_get(x_137, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_137, 1); -lean_inc(x_144); -if (lean_is_exclusive(x_137)) { - lean_ctor_release(x_137, 0); - lean_ctor_release(x_137, 1); - x_145 = x_137; -} else { - lean_dec_ref(x_137); - x_145 = lean_box(0); -} -x_146 = l_Lean_Name_isSuffixOf(x_140, x_143); -if (x_146 == 0) -{ -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; -lean_dec(x_141); -lean_dec(x_140); -if (lean_is_scalar(x_145)) { - x_147 = lean_alloc_ctor(0, 2, 0); -} else { - x_147 = x_145; -} -lean_ctor_set(x_147, 0, x_143); -lean_ctor_set(x_147, 1, x_144); -if (lean_is_scalar(x_138)) { - x_148 = lean_alloc_ctor(1, 1, 0); -} else { - x_148 = x_138; -} -lean_ctor_set(x_148, 0, x_147); -if (lean_is_scalar(x_142)) { - x_149 = lean_alloc_ctor(0, 2, 0); -} else { - x_149 = x_142; -} -lean_ctor_set(x_149, 0, x_66); -lean_ctor_set(x_149, 1, x_148); -x_150 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_150, 0, x_149); -x_49 = x_150; -x_50 = x_139; -goto block_57; -} -else -{ -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_144); -lean_dec(x_143); -if (lean_is_scalar(x_145)) { - x_151 = lean_alloc_ctor(0, 2, 0); -} else { - x_151 = x_145; -} -lean_ctor_set(x_151, 0, x_140); -lean_ctor_set(x_151, 1, x_141); -if (lean_is_scalar(x_138)) { - x_152 = lean_alloc_ctor(1, 1, 0); -} else { - x_152 = x_138; -} -lean_ctor_set(x_152, 0, x_151); -if (lean_is_scalar(x_142)) { - x_153 = lean_alloc_ctor(0, 2, 0); -} else { - x_153 = x_142; -} -lean_ctor_set(x_153, 0, x_66); -lean_ctor_set(x_153, 1, x_152); -x_154 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_154, 0, x_153); -x_49 = x_154; -x_50 = x_139; -goto block_57; -} -} -} -} -} -else -{ -uint8_t x_155; -lean_free_object(x_71); -lean_dec(x_74); -lean_free_object(x_68); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_155 = !lean_is_exclusive(x_77); -if (x_155 == 0) -{ -return x_77; -} -else -{ -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_77, 0); -x_157 = lean_ctor_get(x_77, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_77); -x_158 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -return x_158; -} -} -} -else -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_71, 1); -lean_inc(x_159); -lean_dec(x_71); -x_160 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_4); -x_161 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_160, x_1, x_2, x_4, x_8, x_9, x_10, x_11, x_72); -if (lean_obj_tag(x_161) == 0) -{ -lean_object* x_162; -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -if (lean_obj_tag(x_162) == 0) -{ -lean_object* x_163; lean_object* x_164; -x_163 = lean_ctor_get(x_161, 1); -lean_inc(x_163); -lean_dec(x_161); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_66); -lean_ctor_set(x_164, 1, x_159); -lean_ctor_set(x_68, 0, x_164); -x_49 = x_68; -x_50 = x_163; -goto block_57; -} -else -{ -lean_object* x_165; lean_object* x_166; -x_165 = lean_ctor_get(x_162, 0); -lean_inc(x_165); -if (lean_is_exclusive(x_162)) { - lean_ctor_release(x_162, 0); - x_166 = x_162; -} else { - lean_dec_ref(x_162); - x_166 = lean_box(0); -} -if (lean_obj_tag(x_159) == 0) -{ -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_167 = lean_ctor_get(x_161, 1); -lean_inc(x_167); -lean_dec(x_161); -x_168 = lean_ctor_get(x_165, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_165, 1); -lean_inc(x_169); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - lean_ctor_release(x_165, 1); - x_170 = x_165; -} else { - lean_dec_ref(x_165); - x_170 = lean_box(0); -} -if (lean_is_scalar(x_170)) { - x_171 = lean_alloc_ctor(0, 2, 0); -} else { - x_171 = x_170; -} -lean_ctor_set(x_171, 0, x_168); -lean_ctor_set(x_171, 1, x_169); -if (lean_is_scalar(x_166)) { - x_172 = lean_alloc_ctor(1, 1, 0); -} else { - x_172 = x_166; -} -lean_ctor_set(x_172, 0, x_171); -x_173 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_173, 0, x_66); -lean_ctor_set(x_173, 1, x_172); -lean_ctor_set(x_68, 0, x_173); -x_49 = x_68; -x_50 = x_167; -goto block_57; -} -else -{ -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; -lean_free_object(x_68); -x_174 = lean_ctor_get(x_159, 0); -lean_inc(x_174); -if (lean_is_exclusive(x_159)) { - lean_ctor_release(x_159, 0); - x_175 = x_159; -} else { - lean_dec_ref(x_159); - x_175 = lean_box(0); -} -x_176 = lean_ctor_get(x_161, 1); -lean_inc(x_176); -lean_dec(x_161); -x_177 = lean_ctor_get(x_165, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_165, 1); -lean_inc(x_178); -if (lean_is_exclusive(x_165)) { - lean_ctor_release(x_165, 0); - lean_ctor_release(x_165, 1); - x_179 = x_165; -} else { - lean_dec_ref(x_165); - x_179 = lean_box(0); -} -x_180 = lean_ctor_get(x_174, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_174, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - x_182 = x_174; -} else { - lean_dec_ref(x_174); - x_182 = lean_box(0); -} -x_183 = l_Lean_Name_isSuffixOf(x_177, x_180); -if (x_183 == 0) -{ -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_178); -lean_dec(x_177); -if (lean_is_scalar(x_182)) { - x_184 = lean_alloc_ctor(0, 2, 0); -} else { - x_184 = x_182; -} -lean_ctor_set(x_184, 0, x_180); -lean_ctor_set(x_184, 1, x_181); -if (lean_is_scalar(x_175)) { - x_185 = lean_alloc_ctor(1, 1, 0); -} else { - x_185 = x_175; -} -lean_ctor_set(x_185, 0, x_184); -if (lean_is_scalar(x_179)) { - x_186 = lean_alloc_ctor(0, 2, 0); -} else { - x_186 = x_179; -} -lean_ctor_set(x_186, 0, x_66); -lean_ctor_set(x_186, 1, x_185); -if (lean_is_scalar(x_166)) { - x_187 = lean_alloc_ctor(1, 1, 0); -} else { - x_187 = x_166; -} -lean_ctor_set(x_187, 0, x_186); -x_49 = x_187; -x_50 = x_176; -goto block_57; -} -else -{ -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -lean_dec(x_181); -lean_dec(x_180); -if (lean_is_scalar(x_182)) { - x_188 = lean_alloc_ctor(0, 2, 0); -} else { - x_188 = x_182; -} -lean_ctor_set(x_188, 0, x_177); -lean_ctor_set(x_188, 1, x_178); -if (lean_is_scalar(x_175)) { - x_189 = lean_alloc_ctor(1, 1, 0); -} else { - x_189 = x_175; -} -lean_ctor_set(x_189, 0, x_188); -if (lean_is_scalar(x_179)) { - x_190 = lean_alloc_ctor(0, 2, 0); -} else { - x_190 = x_179; -} -lean_ctor_set(x_190, 0, x_66); -lean_ctor_set(x_190, 1, x_189); -if (lean_is_scalar(x_166)) { - x_191 = lean_alloc_ctor(1, 1, 0); -} else { - x_191 = x_166; -} -lean_ctor_set(x_191, 0, x_190); -x_49 = x_191; -x_50 = x_176; -goto block_57; -} -} -} -} -else -{ -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_159); -lean_free_object(x_68); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_192 = lean_ctor_get(x_161, 0); -lean_inc(x_192); -x_193 = lean_ctor_get(x_161, 1); -lean_inc(x_193); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - lean_ctor_release(x_161, 1); - x_194 = x_161; -} else { - lean_dec_ref(x_161); - x_194 = lean_box(0); -} -if (lean_is_scalar(x_194)) { - x_195 = lean_alloc_ctor(1, 2, 0); -} else { - x_195 = x_194; -} -lean_ctor_set(x_195, 0, x_192); -lean_ctor_set(x_195, 1, x_193); -return x_195; -} -} -} -else -{ -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -x_196 = lean_ctor_get(x_68, 0); -lean_inc(x_196); -lean_dec(x_68); -x_197 = lean_ctor_get(x_67, 1); -lean_inc(x_197); -lean_dec(x_67); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -if (lean_is_exclusive(x_196)) { - lean_ctor_release(x_196, 0); - lean_ctor_release(x_196, 1); - x_199 = x_196; -} else { - lean_dec_ref(x_196); - x_199 = lean_box(0); -} -x_200 = lean_box(0); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_4); -x_201 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_200, x_1, x_2, x_4, x_8, x_9, x_10, x_11, x_197); -if (lean_obj_tag(x_201) == 0) -{ -lean_object* x_202; -x_202 = lean_ctor_get(x_201, 0); -lean_inc(x_202); -if (lean_obj_tag(x_202) == 0) -{ -lean_object* x_203; lean_object* x_204; lean_object* x_205; -x_203 = lean_ctor_get(x_201, 1); -lean_inc(x_203); -lean_dec(x_201); -if (lean_is_scalar(x_199)) { - x_204 = lean_alloc_ctor(0, 2, 0); -} else { - x_204 = x_199; -} -lean_ctor_set(x_204, 0, x_66); -lean_ctor_set(x_204, 1, x_198); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_204); -x_49 = x_205; -x_50 = x_203; -goto block_57; -} -else -{ -lean_object* x_206; lean_object* x_207; -x_206 = lean_ctor_get(x_202, 0); -lean_inc(x_206); -if (lean_is_exclusive(x_202)) { - lean_ctor_release(x_202, 0); - x_207 = x_202; -} else { - lean_dec_ref(x_202); - x_207 = lean_box(0); -} -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; -x_208 = lean_ctor_get(x_201, 1); -lean_inc(x_208); -lean_dec(x_201); -x_209 = lean_ctor_get(x_206, 0); -lean_inc(x_209); -x_210 = lean_ctor_get(x_206, 1); -lean_inc(x_210); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - lean_ctor_release(x_206, 1); - x_211 = x_206; -} else { - lean_dec_ref(x_206); - x_211 = lean_box(0); -} -if (lean_is_scalar(x_211)) { - x_212 = lean_alloc_ctor(0, 2, 0); -} else { - x_212 = x_211; -} -lean_ctor_set(x_212, 0, x_209); -lean_ctor_set(x_212, 1, x_210); -if (lean_is_scalar(x_207)) { - x_213 = lean_alloc_ctor(1, 1, 0); -} else { - x_213 = x_207; -} -lean_ctor_set(x_213, 0, x_212); -if (lean_is_scalar(x_199)) { - x_214 = lean_alloc_ctor(0, 2, 0); -} else { - x_214 = x_199; -} -lean_ctor_set(x_214, 0, x_66); -lean_ctor_set(x_214, 1, x_213); -x_215 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_215, 0, x_214); -x_49 = x_215; -x_50 = x_208; -goto block_57; -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; -lean_dec(x_199); -x_216 = lean_ctor_get(x_198, 0); -lean_inc(x_216); -if (lean_is_exclusive(x_198)) { - lean_ctor_release(x_198, 0); - x_217 = x_198; -} else { - lean_dec_ref(x_198); - x_217 = lean_box(0); -} -x_218 = lean_ctor_get(x_201, 1); -lean_inc(x_218); -lean_dec(x_201); -x_219 = lean_ctor_get(x_206, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_206, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - lean_ctor_release(x_206, 1); - x_221 = x_206; -} else { - lean_dec_ref(x_206); - x_221 = lean_box(0); -} -x_222 = lean_ctor_get(x_216, 0); -lean_inc(x_222); -x_223 = lean_ctor_get(x_216, 1); -lean_inc(x_223); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_224 = x_216; -} else { - lean_dec_ref(x_216); - x_224 = lean_box(0); -} -x_225 = l_Lean_Name_isSuffixOf(x_219, x_222); -if (x_225 == 0) -{ -lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -lean_dec(x_220); -lean_dec(x_219); -if (lean_is_scalar(x_224)) { - x_226 = lean_alloc_ctor(0, 2, 0); -} else { - x_226 = x_224; -} -lean_ctor_set(x_226, 0, x_222); -lean_ctor_set(x_226, 1, x_223); -if (lean_is_scalar(x_217)) { - x_227 = lean_alloc_ctor(1, 1, 0); -} else { - x_227 = x_217; -} -lean_ctor_set(x_227, 0, x_226); -if (lean_is_scalar(x_221)) { - x_228 = lean_alloc_ctor(0, 2, 0); -} else { - x_228 = x_221; -} -lean_ctor_set(x_228, 0, x_66); -lean_ctor_set(x_228, 1, x_227); -if (lean_is_scalar(x_207)) { - x_229 = lean_alloc_ctor(1, 1, 0); -} else { - x_229 = x_207; -} -lean_ctor_set(x_229, 0, x_228); -x_49 = x_229; -x_50 = x_218; -goto block_57; -} -else -{ -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -lean_dec(x_223); -lean_dec(x_222); -if (lean_is_scalar(x_224)) { - x_230 = lean_alloc_ctor(0, 2, 0); -} else { - x_230 = x_224; -} -lean_ctor_set(x_230, 0, x_219); -lean_ctor_set(x_230, 1, x_220); -if (lean_is_scalar(x_217)) { - x_231 = lean_alloc_ctor(1, 1, 0); -} else { - x_231 = x_217; -} -lean_ctor_set(x_231, 0, x_230); -if (lean_is_scalar(x_221)) { - x_232 = lean_alloc_ctor(0, 2, 0); -} else { - x_232 = x_221; -} -lean_ctor_set(x_232, 0, x_66); -lean_ctor_set(x_232, 1, x_231); -if (lean_is_scalar(x_207)) { - x_233 = lean_alloc_ctor(1, 1, 0); -} else { - x_233 = x_207; -} -lean_ctor_set(x_233, 0, x_232); -x_49 = x_233; -x_50 = x_218; -goto block_57; -} -} -} -} -else -{ -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; -lean_dec(x_199); -lean_dec(x_198); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_234 = lean_ctor_get(x_201, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_201, 1); -lean_inc(x_235); -if (lean_is_exclusive(x_201)) { - lean_ctor_release(x_201, 0); - lean_ctor_release(x_201, 1); - x_236 = x_201; -} else { - lean_dec_ref(x_201); - x_236 = lean_box(0); -} -if (lean_is_scalar(x_236)) { - x_237 = lean_alloc_ctor(1, 2, 0); -} else { - x_237 = x_236; -} -lean_ctor_set(x_237, 0, x_234); -lean_ctor_set(x_237, 1, x_235); -return x_237; -} -} -} -} -else -{ -uint8_t x_238; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_1); -x_238 = !lean_is_exclusive(x_67); -if (x_238 == 0) -{ -return x_67; -} -else -{ -lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_239 = lean_ctor_get(x_67, 0); -x_240 = lean_ctor_get(x_67, 1); -lean_inc(x_240); -lean_inc(x_239); -lean_dec(x_67); -x_241 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_241, 0, x_239); -lean_ctor_set(x_241, 1, x_240); -return x_241; -} -} -} -} -else -{ -uint8_t x_242; -lean_dec(x_15); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_1); -x_242 = !lean_is_exclusive(x_58); -if (x_242 == 0) -{ -return x_58; -} -else -{ -lean_object* x_243; lean_object* x_244; lean_object* x_245; -x_243 = lean_ctor_get(x_58, 0); -x_244 = lean_ctor_get(x_58, 1); -lean_inc(x_244); -lean_inc(x_243); -lean_dec(x_58); -x_245 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_245, 0, x_243); -lean_ctor_set(x_245, 1, x_244); -return x_245; -} -} -block_48: -{ -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_20; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_17); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -else -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_18, 0); -lean_inc(x_21); -lean_dec(x_18); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; lean_object* x_23; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_19); -return x_23; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_24 = lean_ctor_get(x_21, 0); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_29 = !lean_is_exclusive(x_27); -if (x_29 == 0) -{ -lean_object* x_30; -x_30 = lean_ctor_get(x_27, 0); -lean_dec(x_30); -lean_ctor_set(x_27, 0, x_17); -return x_27; -} -else -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_dec(x_27); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_17); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} -} -else -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_27, 1); -lean_inc(x_33); -lean_dec(x_27); -x_34 = !lean_is_exclusive(x_28); -if (x_34 == 0) -{ -lean_object* x_35; uint8_t x_36; double x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_28, 0); -lean_ctor_set_tag(x_28, 0); -lean_ctor_set(x_28, 0, x_4); -x_36 = lean_unbox(x_35); -lean_dec(x_35); -x_37 = lean_unbox_float(x_26); -lean_dec(x_26); -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_25, x_28, x_36, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; uint8_t x_41; double x_42; lean_object* x_43; -x_39 = lean_ctor_get(x_28, 0); -lean_inc(x_39); -lean_dec(x_28); -x_40 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_40, 0, x_4); -x_41 = lean_unbox(x_39); -lean_dec(x_39); -x_42 = lean_unbox_float(x_26); -lean_dec(x_26); -x_43 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_25, x_40, x_41, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_43; -} -} -} -else -{ -uint8_t x_44; -lean_dec(x_26); -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_44 = !lean_is_exclusive(x_27); -if (x_44 == 0) -{ -return x_27; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_27, 0); -x_46 = lean_ctor_get(x_27, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_27); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} -} -} -} -} -block_57: -{ -if (lean_obj_tag(x_49) == 0) -{ -x_18 = x_17; -x_19 = x_50; -goto block_48; -} -else -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_49); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_49, 0); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -lean_ctor_set(x_49, 0, x_53); -x_18 = x_49; -x_19 = x_50; -goto block_48; -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_49, 0); -lean_inc(x_54); -lean_dec(x_49); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); -lean_dec(x_54); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_18 = x_56; -x_19 = x_50; -goto block_48; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = l_Lean_Name_isPrefixOf(x_2, x_3); -if (x_4 == 0) -{ -lean_object* x_5; -lean_dec(x_3); -lean_dec(x_1); -x_5 = lean_box(0); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_box(0); -x_7 = l_Lean_Name_replacePrefix(x_3, x_2, x_6); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic(x_1, x_7); -lean_dec(x_7); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7(x_1, x_2, x_3, x_5, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_1, 4); -lean_inc(x_13); -lean_dec(x_1); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore_searchAlias(x_2, x_3, x_4, x_5, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_Server_Completion_completeNamespaces(x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_12 = lean_st_ref_get(x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_box(x_2); -lean_inc(x_3); -lean_inc(x_1); -x_17 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed), 12, 3); -lean_closure_set(x_17, 0, x_1); -lean_closure_set(x_17, 1, x_16); -lean_closure_set(x_17, 2, x_3); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_18 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -if (lean_obj_tag(x_19) == 0) -{ -uint8_t x_20; -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_20 = !lean_is_exclusive(x_18); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_18, 0); -lean_dec(x_21); -x_22 = lean_box(0); -lean_ctor_set(x_18, 0, x_22); -return x_18; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_18, 1); -lean_inc(x_23); -lean_dec(x_18); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; -} -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_19); -x_26 = lean_ctor_get(x_18, 1); -lean_inc(x_26); -lean_dec(x_18); -lean_inc(x_1); -x_27 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed), 3, 1); -lean_closure_set(x_27, 0, x_1); -lean_inc(x_15); -x_28 = l_Lean_Server_Completion_getEligibleHeaderDecls(x_15, x_26); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_29); -lean_inc(x_15); -x_31 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed), 12, 2); -lean_closure_set(x_31, 0, x_15); -lean_closure_set(x_31, 1, x_29); -x_32 = lean_ctor_get(x_3, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_32, 5); -lean_inc(x_33); -x_34 = lean_box(0); -x_35 = lean_box(0); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_33); -lean_inc(x_15); -lean_inc(x_1); -x_36 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17(x_1, x_15, x_29, x_33, x_34, x_33, x_33, x_35, lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_33); -if (lean_obj_tag(x_36) == 0) -{ -lean_object* x_37; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) -{ -uint8_t x_38; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_27); -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_38 = !lean_is_exclusive(x_36); -if (x_38 == 0) -{ -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_36, 0); -lean_dec(x_39); -x_40 = lean_box(0); -lean_ctor_set(x_36, 0, x_40); -return x_36; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_36, 1); -lean_inc(x_41); -lean_dec(x_36); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; -} -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -lean_dec(x_37); -x_44 = lean_ctor_get(x_36, 1); -lean_inc(x_44); -lean_dec(x_36); -x_45 = l_Lean_getAliasState(x_15); -x_46 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__5), 12, 3); -lean_closure_set(x_46, 0, x_32); -lean_closure_set(x_46, 1, x_27); -lean_closure_set(x_46, 2, x_31); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_47 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__18(x_45, x_46, x_5, x_6, x_7, x_8, x_9, x_10, x_44); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -if (lean_obj_tag(x_48) == 0) -{ -uint8_t x_49; -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_49 = !lean_is_exclusive(x_47); -if (x_49 == 0) -{ -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_47, 0); -lean_dec(x_50); -x_51 = lean_box(0); -lean_ctor_set(x_47, 0, x_51); -return x_47; -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_47, 1); -lean_inc(x_52); -lean_dec(x_47); -x_53 = lean_box(0); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -return x_54; -} -} -else -{ -lean_dec(x_48); -if (lean_obj_tag(x_1) == 1) -{ -lean_object* x_55; -x_55 = lean_ctor_get(x_1, 0); -lean_inc(x_55); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_56 = lean_ctor_get(x_47, 1); -lean_inc(x_56); -lean_dec(x_47); -x_57 = lean_ctor_get(x_1, 1); -lean_inc(x_57); -x_58 = l_Lean_Parser_getTokenTable(x_15); -lean_dec(x_15); -x_59 = lean_unsigned_to_nat(0u); -x_60 = l_Lean_Data_Trie_findPrefix_go___rarg(x_57, x_58, x_59); -x_61 = lean_array_size(x_60); -x_62 = 0; -lean_inc(x_5); -x_63 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23(x_57, x_60, x_34, x_60, x_61, x_62, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_56); -lean_dec(x_60); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -x_65 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_64); -return x_65; -} -else -{ -lean_object* x_66; lean_object* x_67; -lean_dec(x_55); -lean_dec(x_15); -x_66 = lean_ctor_get(x_47, 1); -lean_inc(x_66); -lean_dec(x_47); -x_67 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_66); -return x_67; -} -} -else -{ -lean_object* x_68; lean_object* x_69; -lean_dec(x_15); -x_68 = lean_ctor_get(x_47, 1); -lean_inc(x_68); -lean_dec(x_47); -x_69 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_68); -return x_69; -} -} -} -else -{ -uint8_t x_70; -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_70 = !lean_is_exclusive(x_47); -if (x_70 == 0) -{ -return x_47; -} -else -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_47, 0); -x_72 = lean_ctor_get(x_47, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_47); -x_73 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -return x_73; -} -} -} -} -else -{ -uint8_t x_74; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_27); -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_74 = !lean_is_exclusive(x_36); -if (x_74 == 0) -{ -return x_36; -} -else -{ -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_36, 0); -x_76 = lean_ctor_get(x_36, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_36); -x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -return x_77; -} -} -} -} -else -{ -uint8_t x_78; -lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_78 = !lean_is_exclusive(x_18); -if (x_78 == 0) -{ -return x_18; -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_18, 0); -x_80 = lean_ctor_get(x_18, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_18); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -return x_81; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; -x_12 = l_Lean_Name_isAtomic(x_3); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_box(0); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_3, x_2, x_1, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_15 = lean_ctor_get(x_7, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_box(0); -lean_inc(x_5); -lean_inc(x_3); -x_18 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24(x_3, x_16, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_16); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_3, x_2, x_1, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -return x_20; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_box(0); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_2, 0); -lean_inc(x_15); -lean_dec(x_2); -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate(x_4, x_15); -x_17 = 0; -x_18 = lean_box(0); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_17, x_16, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = l_Lean_Name_hasMacroScopes(x_3); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_box(0); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_4, x_5, x_3, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_15; -} -else -{ -lean_object* x_16; -x_16 = l_Lean_Syntax_getHeadInfo(x_2); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_16); -x_17 = lean_erase_macro_scopes(x_3); -x_18 = lean_box(0); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_4, x_5, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; -lean_dec(x_16); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_20 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_12); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_3); -lean_dec(x_3); -x_15 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_1); -return x_15; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -uint8_t x_18; lean_object* x_19; -x_18 = lean_unbox(x_1); -lean_dec(x_1); -x_19 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__6(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -double x_13; lean_object* x_14; -x_13 = lean_unbox_float(x_4); -lean_dec(x_4); -x_14 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__7(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -double x_13; lean_object* x_14; -x_13 = lean_unbox_float(x_4); -lean_dec(x_4); -x_14 = l_List_forM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__8(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__16(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_1); -return x_14; -} -} -LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -lean_object* x_17; -x_17 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_14 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__22(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__23(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { -_start: -{ -size_t x_17; size_t x_18; lean_object* x_19; -x_17 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_18 = lean_unbox_usize(x_8); -lean_dec(x_8); -x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__26(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_17 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__27(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_17 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__28(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__24(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_2); -lean_dec(x_2); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__1(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_2); -lean_dec(x_2); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__2(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_5); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__3(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -double x_13; lean_object* x_14; -x_13 = lean_unbox_float(x_5); -lean_dec(x_5); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__4(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_3); -lean_dec(x_3); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__6(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_4); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_2); -lean_dec(x_2); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_4); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_2); -lean_dec(x_2); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_4); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_3); -lean_dec(x_3); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_5); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_5); -lean_dec(x_5); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_box(x_7); -lean_inc(x_2); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___boxed), 12, 5); -lean_closure_set(x_10, 0, x_2); -lean_closure_set(x_10, 1, x_4); -lean_closure_set(x_10, 2, x_5); -lean_closure_set(x_10, 3, x_6); -lean_closure_set(x_10, 4, x_9); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_10, x_8); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_7); -lean_dec(x_7); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_unfoldDefinition_x3f(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -return x_7; -} -else -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = l_Lean_Exception_isInterrupt(x_9); -if (x_10 == 0) -{ -uint8_t x_11; -x_11 = l_Lean_Exception_isRuntime(x_9); -if (x_11 == 0) -{ -lean_object* x_12; -lean_dec(x_9); -x_12 = lean_box(0); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -return x_7; -} -} -else -{ -return x_7; -} -} -else -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_7, 0); -x_14 = lean_ctor_get(x_7, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_7); -x_15 = l_Lean_Exception_isInterrupt(x_13); -if (x_15 == 0) -{ -uint8_t x_16; -x_16 = l_Lean_Exception_isRuntime(x_13); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_13); -x_17 = lean_box(0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_14); -return x_18; -} -else -{ -lean_object* x_19; -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_13); -lean_ctor_set(x_19, 1, x_14); -return x_19; -} -} -else -{ -lean_object* x_20; -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_14); -return x_20; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_11 = !lean_is_exclusive(x_9); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_9, 0); -lean_dec(x_12); -x_13 = 0; -x_14 = lean_box(x_13); -lean_ctor_set(x_9, 0, x_14); -return x_9; -} -else -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = 0; -x_17 = lean_box(x_16); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_15); -return x_18; -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_dec(x_9); -x_20 = lean_ctor_get(x_10, 0); -lean_inc(x_20); -lean_dec(x_10); -x_21 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(x_20, x_2, x_4, x_5, x_6, x_7, x_19); -return x_21; -} -} -else -{ -uint8_t x_22; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_22 = !lean_is_exclusive(x_9); -if (x_22 == 0) -{ -return x_9; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_9, 0); -x_24 = lean_ctor_get(x_9, 1); -lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_9); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_Expr_getAppFn(x_1); -if (lean_obj_tag(x_8) == 4) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -lean_dec(x_8); -lean_inc(x_9); -x_10 = lean_private_to_user_name(x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = lean_name_eq(x_9, x_2); -lean_dec(x_9); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_12, x_3, x_4, x_5, x_6, x_7); -return x_13; -} -else -{ -uint8_t x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_14 = 1; -x_15 = lean_box(x_14); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_7); -return x_16; -} -} -else -{ -lean_object* x_17; uint8_t x_18; -lean_dec(x_9); -x_17 = lean_ctor_get(x_10, 0); -lean_inc(x_17); -lean_dec(x_10); -x_18 = lean_name_eq(x_17, x_2); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_box(0); -x_20 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_19, x_3, x_4, x_5, x_6, x_7); -return x_20; -} -else -{ -uint8_t x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_21 = 1; -x_22 = lean_box(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_7); -return x_23; -} -} -} -else -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_8); -x_24 = lean_box(0); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_7); -return x_25; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; -x_14 = lean_usize_dec_lt(x_7, x_6); -if (x_14 == 0) -{ -lean_object* x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_8); -x_16 = lean_array_uget(x_5, x_7); -x_17 = l_Lean_Expr_fvarId_x21(x_16); -lean_dec(x_16); -lean_inc(x_9); -x_18 = l_Lean_FVarId_getDecl(x_17, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_LocalDecl_type(x_19); -lean_dec(x_19); -x_22 = l_Lean_Expr_consumeMData(x_21); -lean_dec(x_21); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(x_22, x_1, x_9, x_10, x_11, x_12, x_20); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_unbox(x_24); -lean_dec(x_24); -if (x_25 == 0) -{ -lean_object* x_26; size_t x_27; size_t x_28; -x_26 = lean_ctor_get(x_23, 1); -lean_inc(x_26); -lean_dec(x_23); -x_27 = 1; -x_28 = lean_usize_add(x_7, x_27); -lean_inc(x_4); -{ -size_t _tmp_6 = x_28; -lean_object* _tmp_7 = x_4; -lean_object* _tmp_12 = x_26; -x_7 = _tmp_6; -x_8 = _tmp_7; -x_13 = _tmp_12; -} -goto _start; -} -else -{ -uint8_t x_30; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -x_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_23, 0); -lean_dec(x_31); -x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1; -lean_ctor_set(x_23, 0, x_32); -return x_23; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_23, 1); -lean_inc(x_33); -lean_dec(x_23); -x_34 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_33); -return x_35; -} -} -} -else -{ -uint8_t x_36; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -x_36 = !lean_is_exclusive(x_23); -if (x_36 == 0) -{ -return x_23; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_23, 0); -x_38 = lean_ctor_get(x_23, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_23); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; -} -} -} -else -{ -uint8_t x_40; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -x_40 = !lean_is_exclusive(x_18); -if (x_40 == 0) -{ -return x_18; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_18, 0); -x_42 = lean_ctor_get(x_18, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_18); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; lean_object* x_9; -x_7 = 0; -x_8 = lean_box(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_6); -return x_9; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed), 6, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_box(0); -x_10 = lean_array_size(x_2); -x_11 = 0; -x_12 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(x_1, x_2, x_9, x_12, x_2, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1; -x_18 = lean_box(0); -x_19 = lean_apply_6(x_17, x_18, x_4, x_5, x_6, x_7, x_16); -return x_19; -} -else -{ -uint8_t x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_20 = !lean_is_exclusive(x_13); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_13, 0); -lean_dec(x_21); -x_22 = lean_ctor_get(x_15, 0); -lean_inc(x_22); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_22); -return x_13; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_13, 1); -lean_inc(x_23); -lean_dec(x_13); -x_24 = lean_ctor_get(x_15, 0); -lean_inc(x_24); -lean_dec(x_15); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; -} -} -} -else -{ -uint8_t x_26; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_26 = !lean_is_exclusive(x_13); -if (x_26 == 0) -{ -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_13, 0); -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_13); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_8 = l_Lean_ConstantInfo_type(x_2); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed), 8, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = 0; -x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -size_t x_14; size_t x_15; lean_object* x_16; -x_14 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_15 = lean_unbox_usize(x_7); -lean_dec(x_7); -x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_16; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_Expr_consumeMData(x_3); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDefEqToAppOf(x_9, x_1, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_8 = l_Lean_ConstantInfo_type(x_2); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed), 8, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = 0; -x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 2) -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_ctor_get(x_1, 1); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_nat_dec_eq(x_2, x_3); -if (x_4 == 0) -{ -lean_inc(x_1); -return x_1; -} -else -{ -uint8_t x_5; -x_5 = l_Lean_isPrivatePrefix(x_1); -if (x_5 == 0) -{ -lean_inc(x_1); -return x_1; -} -else -{ -lean_object* x_6; -x_6 = lean_box(0); -return x_6; -} -} -} -else -{ -lean_inc(x_1); -return x_1; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT uint8_t l_Lean_Server_Completion_cmpModPrivate(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix(x_1); -lean_dec(x_1); -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_stripPrivatePrefix(x_2); -lean_dec(x_2); -switch (lean_obj_tag(x_3)) { -case 0: -{ -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -x_5 = 1; -return x_5; -} -else -{ -uint8_t x_6; -lean_dec(x_4); -x_6 = 0; -return x_6; -} -} -case 1: -{ -if (lean_obj_tag(x_4) == 1) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_ctor_get(x_3, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_3, 1); -lean_inc(x_8); -lean_dec(x_3); -x_9 = lean_ctor_get(x_4, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 1); -lean_inc(x_10); -lean_dec(x_4); -x_11 = lean_string_dec_lt(x_8, x_10); -if (x_11 == 0) -{ -uint8_t x_12; -x_12 = lean_string_dec_eq(x_8, x_10); -lean_dec(x_10); -lean_dec(x_8); -if (x_12 == 0) -{ -uint8_t x_13; -lean_dec(x_9); -lean_dec(x_7); -x_13 = 2; -return x_13; -} -else -{ -x_1 = x_7; -x_2 = x_9; -goto _start; -} -} -else -{ -uint8_t x_15; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_15 = 0; -return x_15; -} -} -else -{ -uint8_t x_16; -lean_dec(x_4); -lean_dec(x_3); -x_16 = 2; -return x_16; -} -} -default: -{ -switch (lean_obj_tag(x_4)) { -case 0: -{ -uint8_t x_17; -lean_dec(x_3); -x_17 = 2; -return x_17; -} -case 1: -{ -uint8_t x_18; -lean_dec(x_4); -lean_dec(x_3); -x_18 = 0; -return x_18; -} -default: -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_3, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_3, 1); -lean_inc(x_20); -lean_dec(x_3); -x_21 = lean_ctor_get(x_4, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_4, 1); -lean_inc(x_22); -lean_dec(x_4); -x_23 = lean_nat_dec_lt(x_20, x_22); -if (x_23 == 0) -{ -uint8_t x_24; -x_24 = lean_nat_dec_eq(x_20, x_22); -lean_dec(x_22); -lean_dec(x_20); -if (x_24 == 0) -{ -uint8_t x_25; -lean_dec(x_21); -lean_dec(x_19); -x_25 = 2; -return x_25; -} -else -{ -x_1 = x_19; -x_2 = x_21; -goto _start; -} -} -else -{ -uint8_t x_27; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -x_27 = 0; -return x_27; -} -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Server_Completion_cmpModPrivate___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Server_Completion_cmpModPrivate(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; uint8_t x_5; lean_object* x_6; -x_4 = lean_box(0); -x_5 = 0; -x_6 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_6, 0, x_4); -lean_ctor_set(x_6, 1, x_2); -lean_ctor_set(x_6, 2, x_3); -lean_ctor_set(x_6, 3, x_4); -lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); -return x_6; -} -else -{ -uint8_t x_7; -x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -if (x_7 == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_1); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -x_11 = lean_ctor_get(x_1, 2); -x_12 = lean_ctor_get(x_1, 3); -lean_inc(x_10); -lean_inc(x_2); -x_13 = l_Lean_Server_Completion_cmpModPrivate(x_2, x_10); -switch (x_13) { -case 0: -{ -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_9, x_2, x_3); -x_15 = 0; -lean_ctor_set(x_1, 0, x_14); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); -return x_1; -} -case 1: -{ -uint8_t x_16; -lean_dec(x_11); -lean_dec(x_10); -x_16 = 0; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); -return x_1; -} -default: -{ -lean_object* x_17; uint8_t x_18; -x_17 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_12, x_2, x_3); -x_18 = 0; -lean_ctor_set(x_1, 3, x_17); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); -return x_1; -} -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_1, 0); -x_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_1); -lean_inc(x_20); -lean_inc(x_2); -x_23 = l_Lean_Server_Completion_cmpModPrivate(x_2, x_20); -switch (x_23) { -case 0: -{ -lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_24 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_19, x_2, x_3); -x_25 = 0; -x_26 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_20); -lean_ctor_set(x_26, 2, x_21); -lean_ctor_set(x_26, 3, x_22); -lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); -return x_26; -} -case 1: -{ -uint8_t x_27; lean_object* x_28; -lean_dec(x_21); -lean_dec(x_20); -x_27 = 0; -x_28 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_28, 0, x_19); -lean_ctor_set(x_28, 1, x_2); -lean_ctor_set(x_28, 2, x_3); -lean_ctor_set(x_28, 3, x_22); -lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); -return x_28; -} -default: -{ -lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_29 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_22, x_2, x_3); -x_30 = 0; -x_31 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -lean_ctor_set(x_31, 2, x_21); -lean_ctor_set(x_31, 3, x_29); -lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); -return x_31; -} -} -} -} -else -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_1); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_33 = lean_ctor_get(x_1, 0); -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 2); -x_36 = lean_ctor_get(x_1, 3); -lean_inc(x_34); -lean_inc(x_2); -x_37 = l_Lean_Server_Completion_cmpModPrivate(x_2, x_34); -switch (x_37) { -case 0: -{ -lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_33, x_2, x_3); -x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); -if (x_39 == 0) -{ -lean_object* x_40; -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; -x_41 = lean_ctor_get(x_38, 3); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) -{ -uint8_t x_42; -x_42 = !lean_is_exclusive(x_38); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_38, 3); -lean_dec(x_43); -x_44 = lean_ctor_get(x_38, 0); -lean_dec(x_44); -lean_ctor_set(x_38, 0, x_41); -x_45 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); -return x_1; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_38, 1); -x_47 = lean_ctor_get(x_38, 2); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_38); -x_48 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_48, 0, x_41); -lean_ctor_set(x_48, 1, x_46); -lean_ctor_set(x_48, 2, x_47); -lean_ctor_set(x_48, 3, x_41); -lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); -x_49 = 1; -lean_ctor_set(x_1, 0, x_48); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); -return x_1; -} -} -else -{ -uint8_t x_50; -x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); -if (x_50 == 0) -{ -uint8_t x_51; -x_51 = !lean_is_exclusive(x_38); -if (x_51 == 0) -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_52 = lean_ctor_get(x_38, 1); -x_53 = lean_ctor_get(x_38, 2); -x_54 = lean_ctor_get(x_38, 3); -lean_dec(x_54); -x_55 = lean_ctor_get(x_38, 0); -lean_dec(x_55); -x_56 = !lean_is_exclusive(x_41); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; -x_57 = lean_ctor_get(x_41, 0); -x_58 = lean_ctor_get(x_41, 1); -x_59 = lean_ctor_get(x_41, 2); -x_60 = lean_ctor_get(x_41, 3); -x_61 = 1; -lean_ctor_set(x_41, 3, x_57); -lean_ctor_set(x_41, 2, x_53); -lean_ctor_set(x_41, 1, x_52); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_60); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); -x_62 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_59); -lean_ctor_set(x_1, 1, x_58); -lean_ctor_set(x_1, 0, x_41); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); -return x_1; -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; -x_63 = lean_ctor_get(x_41, 0); -x_64 = lean_ctor_get(x_41, 1); -x_65 = lean_ctor_get(x_41, 2); -x_66 = lean_ctor_get(x_41, 3); -lean_inc(x_66); -lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_41); -x_67 = 1; -x_68 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_68, 0, x_40); -lean_ctor_set(x_68, 1, x_52); -lean_ctor_set(x_68, 2, x_53); -lean_ctor_set(x_68, 3, x_63); -lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_66); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); -x_69 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_65); -lean_ctor_set(x_1, 1, x_64); -lean_ctor_set(x_1, 0, x_68); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); -return x_1; -} -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_70 = lean_ctor_get(x_38, 1); -x_71 = lean_ctor_get(x_38, 2); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_38); -x_72 = lean_ctor_get(x_41, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_41, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_41, 2); -lean_inc(x_74); -x_75 = lean_ctor_get(x_41, 3); -lean_inc(x_75); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - lean_ctor_release(x_41, 2); - lean_ctor_release(x_41, 3); - x_76 = x_41; -} else { - lean_dec_ref(x_41); - x_76 = lean_box(0); -} -x_77 = 1; -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(1, 4, 1); -} else { - x_78 = x_76; -} -lean_ctor_set(x_78, 0, x_40); -lean_ctor_set(x_78, 1, x_70); -lean_ctor_set(x_78, 2, x_71); -lean_ctor_set(x_78, 3, x_72); -lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); -x_79 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_79, 0, x_75); -lean_ctor_set(x_79, 1, x_34); -lean_ctor_set(x_79, 2, x_35); -lean_ctor_set(x_79, 3, x_36); -lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); -x_80 = 0; -lean_ctor_set(x_1, 3, x_79); -lean_ctor_set(x_1, 2, x_74); -lean_ctor_set(x_1, 1, x_73); -lean_ctor_set(x_1, 0, x_78); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); -return x_1; -} -} -else -{ -uint8_t x_81; -lean_free_object(x_1); -x_81 = !lean_is_exclusive(x_41); -if (x_81 == 0) -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_82 = lean_ctor_get(x_41, 3); -lean_dec(x_82); -x_83 = lean_ctor_get(x_41, 2); -lean_dec(x_83); -x_84 = lean_ctor_get(x_41, 1); -lean_dec(x_84); -x_85 = lean_ctor_get(x_41, 0); -lean_dec(x_85); -x_86 = 1; -lean_ctor_set(x_41, 3, x_36); -lean_ctor_set(x_41, 2, x_35); -lean_ctor_set(x_41, 1, x_34); -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); -return x_41; -} -else -{ -uint8_t x_87; lean_object* x_88; -lean_dec(x_41); -x_87 = 1; -x_88 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_88, 0, x_38); -lean_ctor_set(x_88, 1, x_34); -lean_ctor_set(x_88, 2, x_35); -lean_ctor_set(x_88, 3, x_36); -lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); -return x_88; -} -} -} -} -else -{ -uint8_t x_89; -x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); -if (x_89 == 0) -{ -uint8_t x_90; -x_90 = !lean_is_exclusive(x_38); -if (x_90 == 0) -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_91 = lean_ctor_get(x_38, 1); -x_92 = lean_ctor_get(x_38, 2); -x_93 = lean_ctor_get(x_38, 3); -x_94 = lean_ctor_get(x_38, 0); -lean_dec(x_94); -x_95 = !lean_is_exclusive(x_40); -if (x_95 == 0) -{ -uint8_t x_96; uint8_t x_97; -x_96 = 1; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); -x_97 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_40); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); -return x_1; -} -else -{ -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; -x_98 = lean_ctor_get(x_40, 0); -x_99 = lean_ctor_get(x_40, 1); -x_100 = lean_ctor_get(x_40, 2); -x_101 = lean_ctor_get(x_40, 3); -lean_inc(x_101); -lean_inc(x_100); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_40); -x_102 = 1; -x_103 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_103, 0, x_98); -lean_ctor_set(x_103, 1, x_99); -lean_ctor_set(x_103, 2, x_100); -lean_ctor_set(x_103, 3, x_101); -lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); -lean_ctor_set(x_38, 3, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 1, x_34); -lean_ctor_set(x_38, 0, x_93); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); -x_104 = 0; -lean_ctor_set(x_1, 3, x_38); -lean_ctor_set(x_1, 2, x_92); -lean_ctor_set(x_1, 1, x_91); -lean_ctor_set(x_1, 0, x_103); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); -return x_1; -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_105 = lean_ctor_get(x_38, 1); -x_106 = lean_ctor_get(x_38, 2); -x_107 = lean_ctor_get(x_38, 3); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_38); -x_108 = lean_ctor_get(x_40, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_40, 1); -lean_inc(x_109); -x_110 = lean_ctor_get(x_40, 2); -lean_inc(x_110); -x_111 = lean_ctor_get(x_40, 3); -lean_inc(x_111); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_112 = x_40; -} else { - lean_dec_ref(x_40); - x_112 = lean_box(0); -} -x_113 = 1; -if (lean_is_scalar(x_112)) { - x_114 = lean_alloc_ctor(1, 4, 1); -} else { - x_114 = x_112; -} -lean_ctor_set(x_114, 0, x_108); -lean_ctor_set(x_114, 1, x_109); -lean_ctor_set(x_114, 2, x_110); -lean_ctor_set(x_114, 3, x_111); -lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); -x_115 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_115, 0, x_107); -lean_ctor_set(x_115, 1, x_34); -lean_ctor_set(x_115, 2, x_35); -lean_ctor_set(x_115, 3, x_36); -lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); -x_116 = 0; -lean_ctor_set(x_1, 3, x_115); -lean_ctor_set(x_1, 2, x_106); -lean_ctor_set(x_1, 1, x_105); -lean_ctor_set(x_1, 0, x_114); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); -return x_1; -} -} -else -{ -lean_object* x_117; -x_117 = lean_ctor_get(x_38, 3); -lean_inc(x_117); -if (lean_obj_tag(x_117) == 0) -{ -uint8_t x_118; -lean_free_object(x_1); -x_118 = !lean_is_exclusive(x_40); -if (x_118 == 0) -{ -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_119 = lean_ctor_get(x_40, 3); -lean_dec(x_119); -x_120 = lean_ctor_get(x_40, 2); -lean_dec(x_120); -x_121 = lean_ctor_get(x_40, 1); -lean_dec(x_121); -x_122 = lean_ctor_get(x_40, 0); -lean_dec(x_122); -x_123 = 1; -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); -return x_40; -} -else -{ -uint8_t x_124; lean_object* x_125; -lean_dec(x_40); -x_124 = 1; -x_125 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_125, 0, x_38); -lean_ctor_set(x_125, 1, x_34); -lean_ctor_set(x_125, 2, x_35); -lean_ctor_set(x_125, 3, x_36); -lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); -return x_125; -} -} -else -{ -uint8_t x_126; -x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); -if (x_126 == 0) -{ -uint8_t x_127; -lean_free_object(x_1); -x_127 = !lean_is_exclusive(x_38); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_128 = lean_ctor_get(x_38, 1); -x_129 = lean_ctor_get(x_38, 2); -x_130 = lean_ctor_get(x_38, 3); -lean_dec(x_130); -x_131 = lean_ctor_get(x_38, 0); -lean_dec(x_131); -x_132 = !lean_is_exclusive(x_117); -if (x_132 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; -x_133 = lean_ctor_get(x_117, 0); -x_134 = lean_ctor_get(x_117, 1); -x_135 = lean_ctor_get(x_117, 2); -x_136 = lean_ctor_get(x_117, 3); -x_137 = 1; -lean_inc(x_40); -lean_ctor_set(x_117, 3, x_133); -lean_ctor_set(x_117, 2, x_129); -lean_ctor_set(x_117, 1, x_128); -lean_ctor_set(x_117, 0, x_40); -x_138 = !lean_is_exclusive(x_40); -if (x_138 == 0) -{ -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; -x_139 = lean_ctor_get(x_40, 3); -lean_dec(x_139); -x_140 = lean_ctor_get(x_40, 2); -lean_dec(x_140); -x_141 = lean_ctor_get(x_40, 1); -lean_dec(x_141); -x_142 = lean_ctor_get(x_40, 0); -lean_dec(x_142); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -lean_ctor_set(x_40, 3, x_36); -lean_ctor_set(x_40, 2, x_35); -lean_ctor_set(x_40, 1, x_34); -lean_ctor_set(x_40, 0, x_136); -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); -x_143 = 0; -lean_ctor_set(x_38, 3, x_40); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); -return x_38; -} -else -{ -lean_object* x_144; uint8_t x_145; -lean_dec(x_40); -lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); -x_144 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_144, 0, x_136); -lean_ctor_set(x_144, 1, x_34); -lean_ctor_set(x_144, 2, x_35); -lean_ctor_set(x_144, 3, x_36); -lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); -x_145 = 0; -lean_ctor_set(x_38, 3, x_144); -lean_ctor_set(x_38, 2, x_135); -lean_ctor_set(x_38, 1, x_134); -lean_ctor_set(x_38, 0, x_117); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); -return x_38; -} -} -else -{ -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; -x_146 = lean_ctor_get(x_117, 0); -x_147 = lean_ctor_get(x_117, 1); -x_148 = lean_ctor_get(x_117, 2); -x_149 = lean_ctor_get(x_117, 3); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_inc(x_146); -lean_dec(x_117); -x_150 = 1; -lean_inc(x_40); -x_151 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_151, 0, x_40); -lean_ctor_set(x_151, 1, x_128); -lean_ctor_set(x_151, 2, x_129); -lean_ctor_set(x_151, 3, x_146); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_152 = x_40; -} else { - lean_dec_ref(x_40); - x_152 = lean_box(0); -} -lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); -if (lean_is_scalar(x_152)) { - x_153 = lean_alloc_ctor(1, 4, 1); -} else { - x_153 = x_152; -} -lean_ctor_set(x_153, 0, x_149); -lean_ctor_set(x_153, 1, x_34); -lean_ctor_set(x_153, 2, x_35); -lean_ctor_set(x_153, 3, x_36); -lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); -x_154 = 0; -lean_ctor_set(x_38, 3, x_153); -lean_ctor_set(x_38, 2, x_148); -lean_ctor_set(x_38, 1, x_147); -lean_ctor_set(x_38, 0, x_151); -lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); -return x_38; -} -} -else -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; -x_155 = lean_ctor_get(x_38, 1); -x_156 = lean_ctor_get(x_38, 2); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_38); -x_157 = lean_ctor_get(x_117, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_117, 1); -lean_inc(x_158); -x_159 = lean_ctor_get(x_117, 2); -lean_inc(x_159); -x_160 = lean_ctor_get(x_117, 3); -lean_inc(x_160); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - lean_ctor_release(x_117, 2); - lean_ctor_release(x_117, 3); - x_161 = x_117; -} else { - lean_dec_ref(x_117); - x_161 = lean_box(0); -} -x_162 = 1; -lean_inc(x_40); -if (lean_is_scalar(x_161)) { - x_163 = lean_alloc_ctor(1, 4, 1); -} else { - x_163 = x_161; -} -lean_ctor_set(x_163, 0, x_40); -lean_ctor_set(x_163, 1, x_155); -lean_ctor_set(x_163, 2, x_156); -lean_ctor_set(x_163, 3, x_157); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_164 = x_40; -} else { - lean_dec_ref(x_40); - x_164 = lean_box(0); -} -lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 4, 1); -} else { - x_165 = x_164; -} -lean_ctor_set(x_165, 0, x_160); -lean_ctor_set(x_165, 1, x_34); -lean_ctor_set(x_165, 2, x_35); -lean_ctor_set(x_165, 3, x_36); -lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); -x_166 = 0; -x_167 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_167, 0, x_163); -lean_ctor_set(x_167, 1, x_158); -lean_ctor_set(x_167, 2, x_159); -lean_ctor_set(x_167, 3, x_165); -lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); -return x_167; -} -} -else -{ -uint8_t x_168; -x_168 = !lean_is_exclusive(x_38); -if (x_168 == 0) -{ -lean_object* x_169; lean_object* x_170; uint8_t x_171; -x_169 = lean_ctor_get(x_38, 3); -lean_dec(x_169); -x_170 = lean_ctor_get(x_38, 0); -lean_dec(x_170); -x_171 = !lean_is_exclusive(x_40); -if (x_171 == 0) -{ -uint8_t x_172; -lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); -x_172 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); -return x_1; -} -else -{ -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; -x_173 = lean_ctor_get(x_40, 0); -x_174 = lean_ctor_get(x_40, 1); -x_175 = lean_ctor_get(x_40, 2); -x_176 = lean_ctor_get(x_40, 3); -lean_inc(x_176); -lean_inc(x_175); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_40); -x_177 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_177, 0, x_173); -lean_ctor_set(x_177, 1, x_174); -lean_ctor_set(x_177, 2, x_175); -lean_ctor_set(x_177, 3, x_176); -lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); -lean_ctor_set(x_38, 0, x_177); -x_178 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); -return x_1; -} -} -else -{ -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; -x_179 = lean_ctor_get(x_38, 1); -x_180 = lean_ctor_get(x_38, 2); -lean_inc(x_180); -lean_inc(x_179); -lean_dec(x_38); -x_181 = lean_ctor_get(x_40, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_40, 1); -lean_inc(x_182); -x_183 = lean_ctor_get(x_40, 2); -lean_inc(x_183); -x_184 = lean_ctor_get(x_40, 3); -lean_inc(x_184); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - lean_ctor_release(x_40, 2); - lean_ctor_release(x_40, 3); - x_185 = x_40; -} else { - lean_dec_ref(x_40); - x_185 = lean_box(0); -} -if (lean_is_scalar(x_185)) { - x_186 = lean_alloc_ctor(1, 4, 1); -} else { - x_186 = x_185; -} -lean_ctor_set(x_186, 0, x_181); -lean_ctor_set(x_186, 1, x_182); -lean_ctor_set(x_186, 2, x_183); -lean_ctor_set(x_186, 3, x_184); -lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); -x_187 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_179); -lean_ctor_set(x_187, 2, x_180); -lean_ctor_set(x_187, 3, x_117); -lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); -x_188 = 1; -lean_ctor_set(x_1, 0, x_187); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); -return x_1; -} -} -} -} -} -} -else -{ -uint8_t x_189; -x_189 = 1; -lean_ctor_set(x_1, 0, x_38); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); -return x_1; -} -} -case 1: -{ -uint8_t x_190; -lean_dec(x_35); -lean_dec(x_34); -x_190 = 1; -lean_ctor_set(x_1, 2, x_3); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); -return x_1; -} -default: -{ -lean_object* x_191; uint8_t x_192; -x_191 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_36, x_2, x_3); -x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); -if (x_192 == 0) -{ -lean_object* x_193; -x_193 = lean_ctor_get(x_191, 0); -lean_inc(x_193); -if (lean_obj_tag(x_193) == 0) -{ -lean_object* x_194; -x_194 = lean_ctor_get(x_191, 3); -lean_inc(x_194); -if (lean_obj_tag(x_194) == 0) -{ -uint8_t x_195; -x_195 = !lean_is_exclusive(x_191); -if (x_195 == 0) -{ -lean_object* x_196; lean_object* x_197; uint8_t x_198; -x_196 = lean_ctor_get(x_191, 3); -lean_dec(x_196); -x_197 = lean_ctor_get(x_191, 0); -lean_dec(x_197); -lean_ctor_set(x_191, 0, x_194); -x_198 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); -return x_1; -} -else -{ -lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; -x_199 = lean_ctor_get(x_191, 1); -x_200 = lean_ctor_get(x_191, 2); -lean_inc(x_200); -lean_inc(x_199); -lean_dec(x_191); -x_201 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_201, 0, x_194); -lean_ctor_set(x_201, 1, x_199); -lean_ctor_set(x_201, 2, x_200); -lean_ctor_set(x_201, 3, x_194); -lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); -x_202 = 1; -lean_ctor_set(x_1, 3, x_201); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); -return x_1; -} -} -else -{ -uint8_t x_203; -x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); -if (x_203 == 0) -{ -uint8_t x_204; -x_204 = !lean_is_exclusive(x_191); -if (x_204 == 0) -{ -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_205 = lean_ctor_get(x_191, 1); -x_206 = lean_ctor_get(x_191, 2); -x_207 = lean_ctor_get(x_191, 3); -lean_dec(x_207); -x_208 = lean_ctor_get(x_191, 0); -lean_dec(x_208); -x_209 = !lean_is_exclusive(x_194); -if (x_209 == 0) -{ -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; -x_210 = lean_ctor_get(x_194, 0); -x_211 = lean_ctor_get(x_194, 1); -x_212 = lean_ctor_get(x_194, 2); -x_213 = lean_ctor_get(x_194, 3); -x_214 = 1; -lean_ctor_set(x_194, 3, x_193); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); -lean_ctor_set(x_191, 3, x_213); -lean_ctor_set(x_191, 2, x_212); -lean_ctor_set(x_191, 1, x_211); -lean_ctor_set(x_191, 0, x_210); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); -x_215 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_194); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); -return x_1; -} -else -{ -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; -x_216 = lean_ctor_get(x_194, 0); -x_217 = lean_ctor_get(x_194, 1); -x_218 = lean_ctor_get(x_194, 2); -x_219 = lean_ctor_get(x_194, 3); -lean_inc(x_219); -lean_inc(x_218); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_194); -x_220 = 1; -x_221 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_221, 0, x_33); -lean_ctor_set(x_221, 1, x_34); -lean_ctor_set(x_221, 2, x_35); -lean_ctor_set(x_221, 3, x_193); -lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); -lean_ctor_set(x_191, 3, x_219); -lean_ctor_set(x_191, 2, x_218); -lean_ctor_set(x_191, 1, x_217); -lean_ctor_set(x_191, 0, x_216); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); -x_222 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_206); -lean_ctor_set(x_1, 1, x_205); -lean_ctor_set(x_1, 0, x_221); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); -return x_1; -} -} -else -{ -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; -x_223 = lean_ctor_get(x_191, 1); -x_224 = lean_ctor_get(x_191, 2); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_191); -x_225 = lean_ctor_get(x_194, 0); -lean_inc(x_225); -x_226 = lean_ctor_get(x_194, 1); -lean_inc(x_226); -x_227 = lean_ctor_get(x_194, 2); -lean_inc(x_227); -x_228 = lean_ctor_get(x_194, 3); -lean_inc(x_228); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - lean_ctor_release(x_194, 2); - lean_ctor_release(x_194, 3); - x_229 = x_194; -} else { - lean_dec_ref(x_194); - x_229 = lean_box(0); -} -x_230 = 1; -if (lean_is_scalar(x_229)) { - x_231 = lean_alloc_ctor(1, 4, 1); -} else { - x_231 = x_229; -} -lean_ctor_set(x_231, 0, x_33); -lean_ctor_set(x_231, 1, x_34); -lean_ctor_set(x_231, 2, x_35); -lean_ctor_set(x_231, 3, x_193); -lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); -x_232 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_232, 0, x_225); -lean_ctor_set(x_232, 1, x_226); -lean_ctor_set(x_232, 2, x_227); -lean_ctor_set(x_232, 3, x_228); -lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); -x_233 = 0; -lean_ctor_set(x_1, 3, x_232); -lean_ctor_set(x_1, 2, x_224); -lean_ctor_set(x_1, 1, x_223); -lean_ctor_set(x_1, 0, x_231); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); -return x_1; -} -} -else -{ -uint8_t x_234; -lean_free_object(x_1); -x_234 = !lean_is_exclusive(x_194); -if (x_234 == 0) -{ -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; -x_235 = lean_ctor_get(x_194, 3); -lean_dec(x_235); -x_236 = lean_ctor_get(x_194, 2); -lean_dec(x_236); -x_237 = lean_ctor_get(x_194, 1); -lean_dec(x_237); -x_238 = lean_ctor_get(x_194, 0); -lean_dec(x_238); -x_239 = 1; -lean_ctor_set(x_194, 3, x_191); -lean_ctor_set(x_194, 2, x_35); -lean_ctor_set(x_194, 1, x_34); -lean_ctor_set(x_194, 0, x_33); -lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); -return x_194; -} -else -{ -uint8_t x_240; lean_object* x_241; -lean_dec(x_194); -x_240 = 1; -x_241 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_241, 0, x_33); -lean_ctor_set(x_241, 1, x_34); -lean_ctor_set(x_241, 2, x_35); -lean_ctor_set(x_241, 3, x_191); -lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); -return x_241; -} -} -} -} -else -{ -uint8_t x_242; -x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); -if (x_242 == 0) -{ -uint8_t x_243; -x_243 = !lean_is_exclusive(x_191); -if (x_243 == 0) -{ -lean_object* x_244; uint8_t x_245; -x_244 = lean_ctor_get(x_191, 0); -lean_dec(x_244); -x_245 = !lean_is_exclusive(x_193); -if (x_245 == 0) -{ -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; -x_246 = lean_ctor_get(x_193, 0); -x_247 = lean_ctor_get(x_193, 1); -x_248 = lean_ctor_get(x_193, 2); -x_249 = lean_ctor_get(x_193, 3); -x_250 = 1; -lean_ctor_set(x_193, 3, x_246); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); -lean_ctor_set(x_191, 0, x_249); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); -x_251 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_248); -lean_ctor_set(x_1, 1, x_247); -lean_ctor_set(x_1, 0, x_193); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); -return x_1; -} -else -{ -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; -x_252 = lean_ctor_get(x_193, 0); -x_253 = lean_ctor_get(x_193, 1); -x_254 = lean_ctor_get(x_193, 2); -x_255 = lean_ctor_get(x_193, 3); -lean_inc(x_255); -lean_inc(x_254); -lean_inc(x_253); -lean_inc(x_252); -lean_dec(x_193); -x_256 = 1; -x_257 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_257, 0, x_33); -lean_ctor_set(x_257, 1, x_34); -lean_ctor_set(x_257, 2, x_35); -lean_ctor_set(x_257, 3, x_252); -lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); -lean_ctor_set(x_191, 0, x_255); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); -x_258 = 0; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set(x_1, 2, x_254); -lean_ctor_set(x_1, 1, x_253); -lean_ctor_set(x_1, 0, x_257); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); -return x_1; -} -} -else -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; -x_259 = lean_ctor_get(x_191, 1); -x_260 = lean_ctor_get(x_191, 2); -x_261 = lean_ctor_get(x_191, 3); -lean_inc(x_261); -lean_inc(x_260); -lean_inc(x_259); -lean_dec(x_191); -x_262 = lean_ctor_get(x_193, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_193, 1); -lean_inc(x_263); -x_264 = lean_ctor_get(x_193, 2); -lean_inc(x_264); -x_265 = lean_ctor_get(x_193, 3); -lean_inc(x_265); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_266 = x_193; -} else { - lean_dec_ref(x_193); - x_266 = lean_box(0); -} -x_267 = 1; -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(1, 4, 1); -} else { - x_268 = x_266; -} -lean_ctor_set(x_268, 0, x_33); -lean_ctor_set(x_268, 1, x_34); -lean_ctor_set(x_268, 2, x_35); -lean_ctor_set(x_268, 3, x_262); -lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); -x_269 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_269, 0, x_265); -lean_ctor_set(x_269, 1, x_259); -lean_ctor_set(x_269, 2, x_260); -lean_ctor_set(x_269, 3, x_261); -lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); -x_270 = 0; -lean_ctor_set(x_1, 3, x_269); -lean_ctor_set(x_1, 2, x_264); -lean_ctor_set(x_1, 1, x_263); -lean_ctor_set(x_1, 0, x_268); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); -return x_1; -} -} -else -{ -lean_object* x_271; -x_271 = lean_ctor_get(x_191, 3); -lean_inc(x_271); -if (lean_obj_tag(x_271) == 0) -{ -uint8_t x_272; -lean_free_object(x_1); -x_272 = !lean_is_exclusive(x_193); -if (x_272 == 0) -{ -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; -x_273 = lean_ctor_get(x_193, 3); -lean_dec(x_273); -x_274 = lean_ctor_get(x_193, 2); -lean_dec(x_274); -x_275 = lean_ctor_get(x_193, 1); -lean_dec(x_275); -x_276 = lean_ctor_get(x_193, 0); -lean_dec(x_276); -x_277 = 1; -lean_ctor_set(x_193, 3, x_191); -lean_ctor_set(x_193, 2, x_35); -lean_ctor_set(x_193, 1, x_34); -lean_ctor_set(x_193, 0, x_33); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); -return x_193; -} -else -{ -uint8_t x_278; lean_object* x_279; -lean_dec(x_193); -x_278 = 1; -x_279 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_279, 0, x_33); -lean_ctor_set(x_279, 1, x_34); -lean_ctor_set(x_279, 2, x_35); -lean_ctor_set(x_279, 3, x_191); -lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); -return x_279; -} -} -else -{ -uint8_t x_280; -x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); -if (x_280 == 0) -{ -uint8_t x_281; -lean_free_object(x_1); -x_281 = !lean_is_exclusive(x_191); -if (x_281 == 0) -{ -lean_object* x_282; lean_object* x_283; uint8_t x_284; -x_282 = lean_ctor_get(x_191, 3); -lean_dec(x_282); -x_283 = lean_ctor_get(x_191, 0); -lean_dec(x_283); -x_284 = !lean_is_exclusive(x_271); -if (x_284 == 0) -{ -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; -x_285 = lean_ctor_get(x_271, 0); -x_286 = lean_ctor_get(x_271, 1); -x_287 = lean_ctor_get(x_271, 2); -x_288 = lean_ctor_get(x_271, 3); -x_289 = 1; -lean_inc(x_193); -lean_ctor_set(x_271, 3, x_193); -lean_ctor_set(x_271, 2, x_35); -lean_ctor_set(x_271, 1, x_34); -lean_ctor_set(x_271, 0, x_33); -x_290 = !lean_is_exclusive(x_193); -if (x_290 == 0) -{ -lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; -x_291 = lean_ctor_get(x_193, 3); -lean_dec(x_291); -x_292 = lean_ctor_get(x_193, 2); -lean_dec(x_292); -x_293 = lean_ctor_get(x_193, 1); -lean_dec(x_293); -x_294 = lean_ctor_get(x_193, 0); -lean_dec(x_294); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -lean_ctor_set(x_193, 3, x_288); -lean_ctor_set(x_193, 2, x_287); -lean_ctor_set(x_193, 1, x_286); -lean_ctor_set(x_193, 0, x_285); -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); -x_295 = 0; -lean_ctor_set(x_191, 3, x_193); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); -return x_191; -} -else -{ -lean_object* x_296; uint8_t x_297; -lean_dec(x_193); -lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); -x_296 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_296, 0, x_285); -lean_ctor_set(x_296, 1, x_286); -lean_ctor_set(x_296, 2, x_287); -lean_ctor_set(x_296, 3, x_288); -lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); -x_297 = 0; -lean_ctor_set(x_191, 3, x_296); -lean_ctor_set(x_191, 0, x_271); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); -return x_191; -} -} -else -{ -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; -x_298 = lean_ctor_get(x_271, 0); -x_299 = lean_ctor_get(x_271, 1); -x_300 = lean_ctor_get(x_271, 2); -x_301 = lean_ctor_get(x_271, 3); -lean_inc(x_301); -lean_inc(x_300); -lean_inc(x_299); -lean_inc(x_298); -lean_dec(x_271); -x_302 = 1; -lean_inc(x_193); -x_303 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_303, 0, x_33); -lean_ctor_set(x_303, 1, x_34); -lean_ctor_set(x_303, 2, x_35); -lean_ctor_set(x_303, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_304 = x_193; -} else { - lean_dec_ref(x_193); - x_304 = lean_box(0); -} -lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); -if (lean_is_scalar(x_304)) { - x_305 = lean_alloc_ctor(1, 4, 1); -} else { - x_305 = x_304; -} -lean_ctor_set(x_305, 0, x_298); -lean_ctor_set(x_305, 1, x_299); -lean_ctor_set(x_305, 2, x_300); -lean_ctor_set(x_305, 3, x_301); -lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); -x_306 = 0; -lean_ctor_set(x_191, 3, x_305); -lean_ctor_set(x_191, 0, x_303); -lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); -return x_191; -} -} -else -{ -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; -x_307 = lean_ctor_get(x_191, 1); -x_308 = lean_ctor_get(x_191, 2); -lean_inc(x_308); -lean_inc(x_307); -lean_dec(x_191); -x_309 = lean_ctor_get(x_271, 0); -lean_inc(x_309); -x_310 = lean_ctor_get(x_271, 1); -lean_inc(x_310); -x_311 = lean_ctor_get(x_271, 2); -lean_inc(x_311); -x_312 = lean_ctor_get(x_271, 3); -lean_inc(x_312); -if (lean_is_exclusive(x_271)) { - lean_ctor_release(x_271, 0); - lean_ctor_release(x_271, 1); - lean_ctor_release(x_271, 2); - lean_ctor_release(x_271, 3); - x_313 = x_271; -} else { - lean_dec_ref(x_271); - x_313 = lean_box(0); -} -x_314 = 1; -lean_inc(x_193); -if (lean_is_scalar(x_313)) { - x_315 = lean_alloc_ctor(1, 4, 1); -} else { - x_315 = x_313; -} -lean_ctor_set(x_315, 0, x_33); -lean_ctor_set(x_315, 1, x_34); -lean_ctor_set(x_315, 2, x_35); -lean_ctor_set(x_315, 3, x_193); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_316 = x_193; -} else { - lean_dec_ref(x_193); - x_316 = lean_box(0); -} -lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); -if (lean_is_scalar(x_316)) { - x_317 = lean_alloc_ctor(1, 4, 1); -} else { - x_317 = x_316; -} -lean_ctor_set(x_317, 0, x_309); -lean_ctor_set(x_317, 1, x_310); -lean_ctor_set(x_317, 2, x_311); -lean_ctor_set(x_317, 3, x_312); -lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); -x_318 = 0; -x_319 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_319, 0, x_315); -lean_ctor_set(x_319, 1, x_307); -lean_ctor_set(x_319, 2, x_308); -lean_ctor_set(x_319, 3, x_317); -lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); -return x_319; -} -} -else -{ -uint8_t x_320; -x_320 = !lean_is_exclusive(x_191); -if (x_320 == 0) -{ -lean_object* x_321; lean_object* x_322; uint8_t x_323; -x_321 = lean_ctor_get(x_191, 3); -lean_dec(x_321); -x_322 = lean_ctor_get(x_191, 0); -lean_dec(x_322); -x_323 = !lean_is_exclusive(x_193); -if (x_323 == 0) -{ -uint8_t x_324; -lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); -x_324 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); -return x_1; -} -else -{ -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; -x_325 = lean_ctor_get(x_193, 0); -x_326 = lean_ctor_get(x_193, 1); -x_327 = lean_ctor_get(x_193, 2); -x_328 = lean_ctor_get(x_193, 3); -lean_inc(x_328); -lean_inc(x_327); -lean_inc(x_326); -lean_inc(x_325); -lean_dec(x_193); -x_329 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_329, 0, x_325); -lean_ctor_set(x_329, 1, x_326); -lean_ctor_set(x_329, 2, x_327); -lean_ctor_set(x_329, 3, x_328); -lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); -lean_ctor_set(x_191, 0, x_329); -x_330 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); -return x_1; -} -} -else -{ -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; -x_331 = lean_ctor_get(x_191, 1); -x_332 = lean_ctor_get(x_191, 2); -lean_inc(x_332); -lean_inc(x_331); -lean_dec(x_191); -x_333 = lean_ctor_get(x_193, 0); -lean_inc(x_333); -x_334 = lean_ctor_get(x_193, 1); -lean_inc(x_334); -x_335 = lean_ctor_get(x_193, 2); -lean_inc(x_335); -x_336 = lean_ctor_get(x_193, 3); -lean_inc(x_336); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - lean_ctor_release(x_193, 2); - lean_ctor_release(x_193, 3); - x_337 = x_193; -} else { - lean_dec_ref(x_193); - x_337 = lean_box(0); -} -if (lean_is_scalar(x_337)) { - x_338 = lean_alloc_ctor(1, 4, 1); -} else { - x_338 = x_337; -} -lean_ctor_set(x_338, 0, x_333); -lean_ctor_set(x_338, 1, x_334); -lean_ctor_set(x_338, 2, x_335); -lean_ctor_set(x_338, 3, x_336); -lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); -x_339 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_339, 0, x_338); -lean_ctor_set(x_339, 1, x_331); -lean_ctor_set(x_339, 2, x_332); -lean_ctor_set(x_339, 3, x_271); -lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); -x_340 = 1; -lean_ctor_set(x_1, 3, x_339); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); -return x_1; -} -} -} -} -} -} -else -{ -uint8_t x_341; -x_341 = 1; -lean_ctor_set(x_1, 3, x_191); -lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); -return x_1; -} -} -} -} -else -{ -lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_342 = lean_ctor_get(x_1, 0); -x_343 = lean_ctor_get(x_1, 1); -x_344 = lean_ctor_get(x_1, 2); -x_345 = lean_ctor_get(x_1, 3); -lean_inc(x_345); -lean_inc(x_344); -lean_inc(x_343); -lean_inc(x_342); -lean_dec(x_1); -lean_inc(x_343); -lean_inc(x_2); -x_346 = l_Lean_Server_Completion_cmpModPrivate(x_2, x_343); -switch (x_346) { -case 0: -{ -lean_object* x_347; uint8_t x_348; -x_347 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_342, x_2, x_3); -x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); -if (x_348 == 0) -{ -lean_object* x_349; -x_349 = lean_ctor_get(x_347, 0); -lean_inc(x_349); -if (lean_obj_tag(x_349) == 0) -{ -lean_object* x_350; -x_350 = lean_ctor_get(x_347, 3); -lean_inc(x_350); -if (lean_obj_tag(x_350) == 0) -{ -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; -x_351 = lean_ctor_get(x_347, 1); -lean_inc(x_351); -x_352 = lean_ctor_get(x_347, 2); -lean_inc(x_352); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_353 = x_347; -} else { - lean_dec_ref(x_347); - x_353 = lean_box(0); -} -if (lean_is_scalar(x_353)) { - x_354 = lean_alloc_ctor(1, 4, 1); -} else { - x_354 = x_353; -} -lean_ctor_set(x_354, 0, x_350); -lean_ctor_set(x_354, 1, x_351); -lean_ctor_set(x_354, 2, x_352); -lean_ctor_set(x_354, 3, x_350); -lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); -x_355 = 1; -x_356 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_356, 0, x_354); -lean_ctor_set(x_356, 1, x_343); -lean_ctor_set(x_356, 2, x_344); -lean_ctor_set(x_356, 3, x_345); -lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); -return x_356; -} -else -{ -uint8_t x_357; -x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); -if (x_357 == 0) -{ -lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; -x_358 = lean_ctor_get(x_347, 1); -lean_inc(x_358); -x_359 = lean_ctor_get(x_347, 2); -lean_inc(x_359); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_360 = x_347; -} else { - lean_dec_ref(x_347); - x_360 = lean_box(0); -} -x_361 = lean_ctor_get(x_350, 0); -lean_inc(x_361); -x_362 = lean_ctor_get(x_350, 1); -lean_inc(x_362); -x_363 = lean_ctor_get(x_350, 2); -lean_inc(x_363); -x_364 = lean_ctor_get(x_350, 3); -lean_inc(x_364); -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_365 = x_350; -} else { - lean_dec_ref(x_350); - x_365 = lean_box(0); -} -x_366 = 1; -if (lean_is_scalar(x_365)) { - x_367 = lean_alloc_ctor(1, 4, 1); -} else { - x_367 = x_365; -} -lean_ctor_set(x_367, 0, x_349); -lean_ctor_set(x_367, 1, x_358); -lean_ctor_set(x_367, 2, x_359); -lean_ctor_set(x_367, 3, x_361); -lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); -if (lean_is_scalar(x_360)) { - x_368 = lean_alloc_ctor(1, 4, 1); -} else { - x_368 = x_360; -} -lean_ctor_set(x_368, 0, x_364); -lean_ctor_set(x_368, 1, x_343); -lean_ctor_set(x_368, 2, x_344); -lean_ctor_set(x_368, 3, x_345); -lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); -x_369 = 0; -x_370 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_370, 0, x_367); -lean_ctor_set(x_370, 1, x_362); -lean_ctor_set(x_370, 2, x_363); -lean_ctor_set(x_370, 3, x_368); -lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); -return x_370; -} -else -{ -lean_object* x_371; uint8_t x_372; lean_object* x_373; -if (lean_is_exclusive(x_350)) { - lean_ctor_release(x_350, 0); - lean_ctor_release(x_350, 1); - lean_ctor_release(x_350, 2); - lean_ctor_release(x_350, 3); - x_371 = x_350; -} else { - lean_dec_ref(x_350); - x_371 = lean_box(0); -} -x_372 = 1; -if (lean_is_scalar(x_371)) { - x_373 = lean_alloc_ctor(1, 4, 1); -} else { - x_373 = x_371; -} -lean_ctor_set(x_373, 0, x_347); -lean_ctor_set(x_373, 1, x_343); -lean_ctor_set(x_373, 2, x_344); -lean_ctor_set(x_373, 3, x_345); -lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); -return x_373; -} -} -} -else -{ -uint8_t x_374; -x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); -if (x_374 == 0) -{ -lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; -x_375 = lean_ctor_get(x_347, 1); -lean_inc(x_375); -x_376 = lean_ctor_get(x_347, 2); -lean_inc(x_376); -x_377 = lean_ctor_get(x_347, 3); -lean_inc(x_377); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_378 = x_347; -} else { - lean_dec_ref(x_347); - x_378 = lean_box(0); -} -x_379 = lean_ctor_get(x_349, 0); -lean_inc(x_379); -x_380 = lean_ctor_get(x_349, 1); -lean_inc(x_380); -x_381 = lean_ctor_get(x_349, 2); -lean_inc(x_381); -x_382 = lean_ctor_get(x_349, 3); -lean_inc(x_382); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_383 = x_349; -} else { - lean_dec_ref(x_349); - x_383 = lean_box(0); -} -x_384 = 1; -if (lean_is_scalar(x_383)) { - x_385 = lean_alloc_ctor(1, 4, 1); -} else { - x_385 = x_383; -} -lean_ctor_set(x_385, 0, x_379); -lean_ctor_set(x_385, 1, x_380); -lean_ctor_set(x_385, 2, x_381); -lean_ctor_set(x_385, 3, x_382); -lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); -if (lean_is_scalar(x_378)) { - x_386 = lean_alloc_ctor(1, 4, 1); -} else { - x_386 = x_378; -} -lean_ctor_set(x_386, 0, x_377); -lean_ctor_set(x_386, 1, x_343); -lean_ctor_set(x_386, 2, x_344); -lean_ctor_set(x_386, 3, x_345); -lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); -x_387 = 0; -x_388 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_388, 0, x_385); -lean_ctor_set(x_388, 1, x_375); -lean_ctor_set(x_388, 2, x_376); -lean_ctor_set(x_388, 3, x_386); -lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); -return x_388; -} -else -{ -lean_object* x_389; -x_389 = lean_ctor_get(x_347, 3); -lean_inc(x_389); -if (lean_obj_tag(x_389) == 0) -{ -lean_object* x_390; uint8_t x_391; lean_object* x_392; -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_390 = x_349; -} else { - lean_dec_ref(x_349); - x_390 = lean_box(0); -} -x_391 = 1; -if (lean_is_scalar(x_390)) { - x_392 = lean_alloc_ctor(1, 4, 1); -} else { - x_392 = x_390; -} -lean_ctor_set(x_392, 0, x_347); -lean_ctor_set(x_392, 1, x_343); -lean_ctor_set(x_392, 2, x_344); -lean_ctor_set(x_392, 3, x_345); -lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); -return x_392; -} -else -{ -uint8_t x_393; -x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); -if (x_393 == 0) -{ -lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; -x_394 = lean_ctor_get(x_347, 1); -lean_inc(x_394); -x_395 = lean_ctor_get(x_347, 2); -lean_inc(x_395); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_396 = x_347; -} else { - lean_dec_ref(x_347); - x_396 = lean_box(0); -} -x_397 = lean_ctor_get(x_389, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_389, 1); -lean_inc(x_398); -x_399 = lean_ctor_get(x_389, 2); -lean_inc(x_399); -x_400 = lean_ctor_get(x_389, 3); -lean_inc(x_400); -if (lean_is_exclusive(x_389)) { - lean_ctor_release(x_389, 0); - lean_ctor_release(x_389, 1); - lean_ctor_release(x_389, 2); - lean_ctor_release(x_389, 3); - x_401 = x_389; -} else { - lean_dec_ref(x_389); - x_401 = lean_box(0); -} -x_402 = 1; -lean_inc(x_349); -if (lean_is_scalar(x_401)) { - x_403 = lean_alloc_ctor(1, 4, 1); -} else { - x_403 = x_401; -} -lean_ctor_set(x_403, 0, x_349); -lean_ctor_set(x_403, 1, x_394); -lean_ctor_set(x_403, 2, x_395); -lean_ctor_set(x_403, 3, x_397); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_404 = x_349; -} else { - lean_dec_ref(x_349); - x_404 = lean_box(0); -} -lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); -if (lean_is_scalar(x_404)) { - x_405 = lean_alloc_ctor(1, 4, 1); -} else { - x_405 = x_404; -} -lean_ctor_set(x_405, 0, x_400); -lean_ctor_set(x_405, 1, x_343); -lean_ctor_set(x_405, 2, x_344); -lean_ctor_set(x_405, 3, x_345); -lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); -x_406 = 0; -if (lean_is_scalar(x_396)) { - x_407 = lean_alloc_ctor(1, 4, 1); -} else { - x_407 = x_396; -} -lean_ctor_set(x_407, 0, x_403); -lean_ctor_set(x_407, 1, x_398); -lean_ctor_set(x_407, 2, x_399); -lean_ctor_set(x_407, 3, x_405); -lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); -return x_407; -} -else -{ -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; -x_408 = lean_ctor_get(x_347, 1); -lean_inc(x_408); -x_409 = lean_ctor_get(x_347, 2); -lean_inc(x_409); -if (lean_is_exclusive(x_347)) { - lean_ctor_release(x_347, 0); - lean_ctor_release(x_347, 1); - lean_ctor_release(x_347, 2); - lean_ctor_release(x_347, 3); - x_410 = x_347; -} else { - lean_dec_ref(x_347); - x_410 = lean_box(0); -} -x_411 = lean_ctor_get(x_349, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_349, 1); -lean_inc(x_412); -x_413 = lean_ctor_get(x_349, 2); -lean_inc(x_413); -x_414 = lean_ctor_get(x_349, 3); -lean_inc(x_414); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - lean_ctor_release(x_349, 2); - lean_ctor_release(x_349, 3); - x_415 = x_349; -} else { - lean_dec_ref(x_349); - x_415 = lean_box(0); -} -if (lean_is_scalar(x_415)) { - x_416 = lean_alloc_ctor(1, 4, 1); -} else { - x_416 = x_415; -} -lean_ctor_set(x_416, 0, x_411); -lean_ctor_set(x_416, 1, x_412); -lean_ctor_set(x_416, 2, x_413); -lean_ctor_set(x_416, 3, x_414); -lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); -if (lean_is_scalar(x_410)) { - x_417 = lean_alloc_ctor(1, 4, 1); -} else { - x_417 = x_410; -} -lean_ctor_set(x_417, 0, x_416); -lean_ctor_set(x_417, 1, x_408); -lean_ctor_set(x_417, 2, x_409); -lean_ctor_set(x_417, 3, x_389); -lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); -x_418 = 1; -x_419 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_419, 0, x_417); -lean_ctor_set(x_419, 1, x_343); -lean_ctor_set(x_419, 2, x_344); -lean_ctor_set(x_419, 3, x_345); -lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); -return x_419; -} -} -} -} -} -else -{ -uint8_t x_420; lean_object* x_421; -x_420 = 1; -x_421 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_421, 0, x_347); -lean_ctor_set(x_421, 1, x_343); -lean_ctor_set(x_421, 2, x_344); -lean_ctor_set(x_421, 3, x_345); -lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); -return x_421; -} -} -case 1: -{ -uint8_t x_422; lean_object* x_423; -lean_dec(x_344); -lean_dec(x_343); -x_422 = 1; -x_423 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_423, 0, x_342); -lean_ctor_set(x_423, 1, x_2); -lean_ctor_set(x_423, 2, x_3); -lean_ctor_set(x_423, 3, x_345); -lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); -return x_423; -} -default: -{ -lean_object* x_424; uint8_t x_425; -x_424 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_345, x_2, x_3); -x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); -if (x_425 == 0) -{ -lean_object* x_426; -x_426 = lean_ctor_get(x_424, 0); -lean_inc(x_426); -if (lean_obj_tag(x_426) == 0) -{ -lean_object* x_427; -x_427 = lean_ctor_get(x_424, 3); -lean_inc(x_427); -if (lean_obj_tag(x_427) == 0) -{ -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; -x_428 = lean_ctor_get(x_424, 1); -lean_inc(x_428); -x_429 = lean_ctor_get(x_424, 2); -lean_inc(x_429); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_430 = x_424; -} else { - lean_dec_ref(x_424); - x_430 = lean_box(0); -} -if (lean_is_scalar(x_430)) { - x_431 = lean_alloc_ctor(1, 4, 1); -} else { - x_431 = x_430; -} -lean_ctor_set(x_431, 0, x_427); -lean_ctor_set(x_431, 1, x_428); -lean_ctor_set(x_431, 2, x_429); -lean_ctor_set(x_431, 3, x_427); -lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); -x_432 = 1; -x_433 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_433, 0, x_342); -lean_ctor_set(x_433, 1, x_343); -lean_ctor_set(x_433, 2, x_344); -lean_ctor_set(x_433, 3, x_431); -lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); -return x_433; -} -else -{ -uint8_t x_434; -x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); -if (x_434 == 0) -{ -lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; -x_435 = lean_ctor_get(x_424, 1); -lean_inc(x_435); -x_436 = lean_ctor_get(x_424, 2); -lean_inc(x_436); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_437 = x_424; -} else { - lean_dec_ref(x_424); - x_437 = lean_box(0); -} -x_438 = lean_ctor_get(x_427, 0); -lean_inc(x_438); -x_439 = lean_ctor_get(x_427, 1); -lean_inc(x_439); -x_440 = lean_ctor_get(x_427, 2); -lean_inc(x_440); -x_441 = lean_ctor_get(x_427, 3); -lean_inc(x_441); -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_442 = x_427; -} else { - lean_dec_ref(x_427); - x_442 = lean_box(0); -} -x_443 = 1; -if (lean_is_scalar(x_442)) { - x_444 = lean_alloc_ctor(1, 4, 1); -} else { - x_444 = x_442; -} -lean_ctor_set(x_444, 0, x_342); -lean_ctor_set(x_444, 1, x_343); -lean_ctor_set(x_444, 2, x_344); -lean_ctor_set(x_444, 3, x_426); -lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); -if (lean_is_scalar(x_437)) { - x_445 = lean_alloc_ctor(1, 4, 1); -} else { - x_445 = x_437; -} -lean_ctor_set(x_445, 0, x_438); -lean_ctor_set(x_445, 1, x_439); -lean_ctor_set(x_445, 2, x_440); -lean_ctor_set(x_445, 3, x_441); -lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); -x_446 = 0; -x_447 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_435); -lean_ctor_set(x_447, 2, x_436); -lean_ctor_set(x_447, 3, x_445); -lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); -return x_447; -} -else -{ -lean_object* x_448; uint8_t x_449; lean_object* x_450; -if (lean_is_exclusive(x_427)) { - lean_ctor_release(x_427, 0); - lean_ctor_release(x_427, 1); - lean_ctor_release(x_427, 2); - lean_ctor_release(x_427, 3); - x_448 = x_427; -} else { - lean_dec_ref(x_427); - x_448 = lean_box(0); -} -x_449 = 1; -if (lean_is_scalar(x_448)) { - x_450 = lean_alloc_ctor(1, 4, 1); -} else { - x_450 = x_448; -} -lean_ctor_set(x_450, 0, x_342); -lean_ctor_set(x_450, 1, x_343); -lean_ctor_set(x_450, 2, x_344); -lean_ctor_set(x_450, 3, x_424); -lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); -return x_450; -} -} -} -else -{ -uint8_t x_451; -x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); -if (x_451 == 0) -{ -lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; -x_452 = lean_ctor_get(x_424, 1); -lean_inc(x_452); -x_453 = lean_ctor_get(x_424, 2); -lean_inc(x_453); -x_454 = lean_ctor_get(x_424, 3); -lean_inc(x_454); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_455 = x_424; -} else { - lean_dec_ref(x_424); - x_455 = lean_box(0); -} -x_456 = lean_ctor_get(x_426, 0); -lean_inc(x_456); -x_457 = lean_ctor_get(x_426, 1); -lean_inc(x_457); -x_458 = lean_ctor_get(x_426, 2); -lean_inc(x_458); -x_459 = lean_ctor_get(x_426, 3); -lean_inc(x_459); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_460 = x_426; -} else { - lean_dec_ref(x_426); - x_460 = lean_box(0); -} -x_461 = 1; -if (lean_is_scalar(x_460)) { - x_462 = lean_alloc_ctor(1, 4, 1); -} else { - x_462 = x_460; -} -lean_ctor_set(x_462, 0, x_342); -lean_ctor_set(x_462, 1, x_343); -lean_ctor_set(x_462, 2, x_344); -lean_ctor_set(x_462, 3, x_456); -lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); -if (lean_is_scalar(x_455)) { - x_463 = lean_alloc_ctor(1, 4, 1); -} else { - x_463 = x_455; -} -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_452); -lean_ctor_set(x_463, 2, x_453); -lean_ctor_set(x_463, 3, x_454); -lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); -x_464 = 0; -x_465 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_465, 0, x_462); -lean_ctor_set(x_465, 1, x_457); -lean_ctor_set(x_465, 2, x_458); -lean_ctor_set(x_465, 3, x_463); -lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); -return x_465; -} -else -{ -lean_object* x_466; -x_466 = lean_ctor_get(x_424, 3); -lean_inc(x_466); -if (lean_obj_tag(x_466) == 0) -{ -lean_object* x_467; uint8_t x_468; lean_object* x_469; -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_467 = x_426; -} else { - lean_dec_ref(x_426); - x_467 = lean_box(0); -} -x_468 = 1; -if (lean_is_scalar(x_467)) { - x_469 = lean_alloc_ctor(1, 4, 1); -} else { - x_469 = x_467; -} -lean_ctor_set(x_469, 0, x_342); -lean_ctor_set(x_469, 1, x_343); -lean_ctor_set(x_469, 2, x_344); -lean_ctor_set(x_469, 3, x_424); -lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); -return x_469; -} -else -{ -uint8_t x_470; -x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); -if (x_470 == 0) -{ -lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; -x_471 = lean_ctor_get(x_424, 1); -lean_inc(x_471); -x_472 = lean_ctor_get(x_424, 2); -lean_inc(x_472); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_473 = x_424; -} else { - lean_dec_ref(x_424); - x_473 = lean_box(0); -} -x_474 = lean_ctor_get(x_466, 0); -lean_inc(x_474); -x_475 = lean_ctor_get(x_466, 1); -lean_inc(x_475); -x_476 = lean_ctor_get(x_466, 2); -lean_inc(x_476); -x_477 = lean_ctor_get(x_466, 3); -lean_inc(x_477); -if (lean_is_exclusive(x_466)) { - lean_ctor_release(x_466, 0); - lean_ctor_release(x_466, 1); - lean_ctor_release(x_466, 2); - lean_ctor_release(x_466, 3); - x_478 = x_466; -} else { - lean_dec_ref(x_466); - x_478 = lean_box(0); -} -x_479 = 1; -lean_inc(x_426); -if (lean_is_scalar(x_478)) { - x_480 = lean_alloc_ctor(1, 4, 1); -} else { - x_480 = x_478; -} -lean_ctor_set(x_480, 0, x_342); -lean_ctor_set(x_480, 1, x_343); -lean_ctor_set(x_480, 2, x_344); -lean_ctor_set(x_480, 3, x_426); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_481 = x_426; -} else { - lean_dec_ref(x_426); - x_481 = lean_box(0); -} -lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); -if (lean_is_scalar(x_481)) { - x_482 = lean_alloc_ctor(1, 4, 1); -} else { - x_482 = x_481; -} -lean_ctor_set(x_482, 0, x_474); -lean_ctor_set(x_482, 1, x_475); -lean_ctor_set(x_482, 2, x_476); -lean_ctor_set(x_482, 3, x_477); -lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); -x_483 = 0; -if (lean_is_scalar(x_473)) { - x_484 = lean_alloc_ctor(1, 4, 1); -} else { - x_484 = x_473; -} -lean_ctor_set(x_484, 0, x_480); -lean_ctor_set(x_484, 1, x_471); -lean_ctor_set(x_484, 2, x_472); -lean_ctor_set(x_484, 3, x_482); -lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); -return x_484; -} -else -{ -lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; -x_485 = lean_ctor_get(x_424, 1); -lean_inc(x_485); -x_486 = lean_ctor_get(x_424, 2); -lean_inc(x_486); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - lean_ctor_release(x_424, 2); - lean_ctor_release(x_424, 3); - x_487 = x_424; -} else { - lean_dec_ref(x_424); - x_487 = lean_box(0); -} -x_488 = lean_ctor_get(x_426, 0); -lean_inc(x_488); -x_489 = lean_ctor_get(x_426, 1); -lean_inc(x_489); -x_490 = lean_ctor_get(x_426, 2); -lean_inc(x_490); -x_491 = lean_ctor_get(x_426, 3); -lean_inc(x_491); -if (lean_is_exclusive(x_426)) { - lean_ctor_release(x_426, 0); - lean_ctor_release(x_426, 1); - lean_ctor_release(x_426, 2); - lean_ctor_release(x_426, 3); - x_492 = x_426; -} else { - lean_dec_ref(x_426); - x_492 = lean_box(0); -} -if (lean_is_scalar(x_492)) { - x_493 = lean_alloc_ctor(1, 4, 1); -} else { - x_493 = x_492; -} -lean_ctor_set(x_493, 0, x_488); -lean_ctor_set(x_493, 1, x_489); -lean_ctor_set(x_493, 2, x_490); -lean_ctor_set(x_493, 3, x_491); -lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); -if (lean_is_scalar(x_487)) { - x_494 = lean_alloc_ctor(1, 4, 1); -} else { - x_494 = x_487; -} -lean_ctor_set(x_494, 0, x_493); -lean_ctor_set(x_494, 1, x_485); -lean_ctor_set(x_494, 2, x_486); -lean_ctor_set(x_494, 3, x_466); -lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); -x_495 = 1; -x_496 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_496, 0, x_342); -lean_ctor_set(x_496, 1, x_343); -lean_ctor_set(x_496, 2, x_344); -lean_ctor_set(x_496, 3, x_494); -lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); -return x_496; -} -} -} -} -} -else -{ -uint8_t x_497; lean_object* x_498; -x_497 = 1; -x_498 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_498, 0, x_342); -lean_ctor_set(x_498, 1, x_343); -lean_ctor_set(x_498, 2, x_344); -lean_ctor_set(x_498, 3, x_424); -lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); -return x_498; -} -} -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = l_Lean_RBNode_isRed___rarg(x_1); -if (x_4 == 0) -{ -lean_object* x_5; -x_5 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_1, x_2, x_3); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_1, x_2, x_3); -x_7 = l_Lean_RBNode_setBlack___rarg(x_6); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; -x_10 = lean_usize_dec_lt(x_2, x_1); -if (x_10 == 0) -{ -lean_object* x_11; -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_3); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; -x_12 = lean_array_uget(x_3, x_2); -x_13 = lean_unsigned_to_nat(0u); -x_14 = lean_array_uset(x_3, x_2, x_13); -x_15 = 1; -x_16 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_2, x_20); -x_22 = lean_array_uset(x_14, x_2, x_19); -x_2 = x_21; -x_3 = x_22; -x_9 = x_18; -goto _start; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 2); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_nat_dec_lt(x_11, x_12); -if (x_13 == 0) -{ -size_t x_14; size_t x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_14 = 1; -x_15 = lean_usize_add(x_3, x_14); -x_3 = x_15; -goto _start; -} -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_get_size(x_10); -x_18 = lean_nat_dec_le(x_12, x_17); -if (x_18 == 0) -{ -uint8_t x_19; -lean_dec(x_12); -x_19 = lean_nat_dec_lt(x_11, x_17); -if (x_19 == 0) -{ -size_t x_20; size_t x_21; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -x_20 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_3 = x_21; -goto _start; -} -else -{ -size_t x_23; size_t x_24; uint8_t x_25; -x_23 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_24 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_10, x_23, x_24); -lean_dec(x_10); -if (x_25 == 0) -{ -size_t x_26; size_t x_27; -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -goto _start; -} -else -{ -uint8_t x_29; -x_29 = 1; -return x_29; -} -} -} -else -{ -size_t x_30; size_t x_31; uint8_t x_32; -lean_dec(x_17); -x_30 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_31 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_10, x_30, x_31); -lean_dec(x_10); -if (x_32 == 0) -{ -size_t x_33; size_t x_34; -x_33 = 1; -x_34 = lean_usize_add(x_3, x_33); -x_3 = x_34; -goto _start; -} -else -{ -uint8_t x_36; -x_36 = 1; -return x_36; -} -} -} -} -else -{ -uint8_t x_37; -x_37 = 0; -return x_37; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 2); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_nat_dec_lt(x_11, x_12); -if (x_13 == 0) -{ -size_t x_14; size_t x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_14 = 1; -x_15 = lean_usize_add(x_3, x_14); -x_3 = x_15; -goto _start; -} -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_get_size(x_10); -x_18 = lean_nat_dec_le(x_12, x_17); -if (x_18 == 0) -{ -uint8_t x_19; -lean_dec(x_12); -x_19 = lean_nat_dec_lt(x_11, x_17); -if (x_19 == 0) -{ -size_t x_20; size_t x_21; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -x_20 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_3 = x_21; -goto _start; -} -else -{ -size_t x_23; size_t x_24; uint8_t x_25; -x_23 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_24 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_10, x_23, x_24); -lean_dec(x_10); -if (x_25 == 0) -{ -size_t x_26; size_t x_27; -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -goto _start; -} -else -{ -uint8_t x_29; -x_29 = 1; -return x_29; -} -} -} -else -{ -size_t x_30; size_t x_31; uint8_t x_32; -lean_dec(x_17); -x_30 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_31 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_10, x_30, x_31); -lean_dec(x_10); -if (x_32 == 0) -{ -size_t x_33; size_t x_34; -x_33 = 1; -x_34 = lean_usize_add(x_3, x_33); -x_3 = x_34; -goto _start; -} -else -{ -uint8_t x_36; -x_36 = 1; -return x_36; -} -} -} -} -else -{ -uint8_t x_37; -x_37 = 0; -return x_37; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 2); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_nat_dec_lt(x_11, x_12); -if (x_13 == 0) -{ -size_t x_14; size_t x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_14 = 1; -x_15 = lean_usize_add(x_3, x_14); -x_3 = x_15; -goto _start; -} -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_get_size(x_10); -x_18 = lean_nat_dec_le(x_12, x_17); -if (x_18 == 0) -{ -uint8_t x_19; -lean_dec(x_12); -x_19 = lean_nat_dec_lt(x_11, x_17); -if (x_19 == 0) -{ -size_t x_20; size_t x_21; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -x_20 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_3 = x_21; -goto _start; -} -else -{ -size_t x_23; size_t x_24; uint8_t x_25; -x_23 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_24 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_10, x_23, x_24); -lean_dec(x_10); -if (x_25 == 0) -{ -size_t x_26; size_t x_27; -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -goto _start; -} -else -{ -uint8_t x_29; -x_29 = 1; -return x_29; -} -} -} -else -{ -size_t x_30; size_t x_31; uint8_t x_32; -lean_dec(x_17); -x_30 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_31 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_10, x_30, x_31); -lean_dec(x_10); -if (x_32 == 0) -{ -size_t x_33; size_t x_34; -x_33 = 1; -x_34 = lean_usize_add(x_3, x_33); -x_3 = x_34; -goto _start; -} -else -{ -uint8_t x_36; -x_36 = 1; -return x_36; -} -} -} -} -else -{ -uint8_t x_37; -x_37 = 0; -return x_37; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_array_get_size(x_6); -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 2); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_nat_dec_lt(x_11, x_12); -if (x_13 == 0) -{ -size_t x_14; size_t x_15; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_14 = 1; -x_15 = lean_usize_add(x_3, x_14); -x_3 = x_15; -goto _start; -} -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_get_size(x_10); -x_18 = lean_nat_dec_le(x_12, x_17); -if (x_18 == 0) -{ -uint8_t x_19; -lean_dec(x_12); -x_19 = lean_nat_dec_lt(x_11, x_17); -if (x_19 == 0) -{ -size_t x_20; size_t x_21; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); -x_20 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_3 = x_21; -goto _start; -} -else -{ -size_t x_23; size_t x_24; uint8_t x_25; -x_23 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_24 = lean_usize_of_nat(x_17); -lean_dec(x_17); -x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_10, x_23, x_24); -lean_dec(x_10); -if (x_25 == 0) -{ -size_t x_26; size_t x_27; -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -goto _start; -} -else -{ -uint8_t x_29; -x_29 = 1; -return x_29; -} -} -} -else -{ -size_t x_30; size_t x_31; uint8_t x_32; -lean_dec(x_17); -x_30 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_31 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_10, x_30, x_31); -lean_dec(x_10); -if (x_32 == 0) -{ -size_t x_33; size_t x_34; -x_33 = 1; -x_34 = lean_usize_add(x_3, x_33); -x_3 = x_34; -goto _start; -} -else -{ -uint8_t x_36; -x_36 = 1; -return x_36; -} -} -} -} -else -{ -uint8_t x_37; -x_37 = 0; -return x_37; -} -} -} -static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(0u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { -_start: -{ -uint8_t x_20; -x_20 = lean_nat_dec_lt(x_11, x_8); -if (x_20 == 0) -{ -lean_object* x_21; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_13); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -else -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_nat_dec_eq(x_10, x_22); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; -lean_dec(x_13); -x_24 = lean_unsigned_to_nat(1u); -x_25 = lean_nat_sub(x_10, x_24); -lean_dec(x_10); -x_26 = lean_nat_dec_lt(x_11, x_2); -lean_inc(x_11); -lean_inc(x_1); -x_27 = l_Array_toSubarray___rarg(x_1, x_22, x_11); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -x_30 = lean_ctor_get(x_27, 2); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_nat_dec_lt(x_29, x_30); -if (x_26 == 0) -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_85 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; -x_86 = l_outOfBounds___rarg(x_85); -x_87 = lean_array_get_size(x_86); -x_88 = lean_nat_dec_lt(x_22, x_87); -lean_dec(x_87); -if (x_88 == 0) -{ -lean_object* x_89; lean_object* x_90; -lean_dec(x_86); -x_89 = l_Lean_instInhabitedName; -x_90 = l_outOfBounds___rarg(x_89); -x_32 = x_90; -goto block_84; -} -else -{ -lean_object* x_91; -x_91 = lean_array_fget(x_86, x_22); -lean_dec(x_86); -x_32 = x_91; -goto block_84; -} -} -else -{ -lean_object* x_92; lean_object* x_93; uint8_t x_94; -x_92 = lean_array_fget(x_1, x_11); -x_93 = lean_array_get_size(x_92); -x_94 = lean_nat_dec_lt(x_22, x_93); -lean_dec(x_93); -if (x_94 == 0) -{ -lean_object* x_95; lean_object* x_96; -lean_dec(x_92); -x_95 = l_Lean_instInhabitedName; -x_96 = l_outOfBounds___rarg(x_95); -x_32 = x_96; -goto block_84; -} -else -{ -lean_object* x_97; -x_97 = lean_array_fget(x_92, x_22); -lean_dec(x_92); -x_32 = x_97; -goto block_84; -} -} -block_84: -{ -lean_object* x_33; lean_object* x_42; -if (x_31 == 0) -{ -lean_object* x_67; -lean_dec(x_30); -lean_dec(x_29); -lean_dec(x_28); -x_67 = lean_box(0); -x_42 = x_67; -goto block_66; -} -else -{ -lean_object* x_68; uint8_t x_69; -x_68 = lean_array_get_size(x_28); -x_69 = lean_nat_dec_le(x_30, x_68); -if (x_69 == 0) -{ -uint8_t x_70; -lean_dec(x_30); -x_70 = lean_nat_dec_lt(x_29, x_68); -if (x_70 == 0) -{ -lean_object* x_71; -lean_dec(x_68); -lean_dec(x_29); -lean_dec(x_28); -x_71 = lean_box(0); -x_42 = x_71; -goto block_66; -} -else -{ -size_t x_72; size_t x_73; uint8_t x_74; -x_72 = lean_usize_of_nat(x_29); -lean_dec(x_29); -x_73 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_74 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(x_32, x_28, x_72, x_73); -lean_dec(x_28); -if (x_74 == 0) -{ -lean_object* x_75; -x_75 = lean_box(0); -x_42 = x_75; -goto block_66; -} -else -{ -lean_object* x_76; -lean_dec(x_32); -x_76 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -lean_inc(x_3); -{ -lean_object* _tmp_9 = x_25; -lean_object* _tmp_10 = x_76; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_3; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -} -goto _start; -} -} -} -else -{ -size_t x_78; size_t x_79; uint8_t x_80; -lean_dec(x_68); -x_78 = lean_usize_of_nat(x_29); -lean_dec(x_29); -x_79 = lean_usize_of_nat(x_30); -lean_dec(x_30); -x_80 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(x_32, x_28, x_78, x_79); -lean_dec(x_28); -if (x_80 == 0) -{ -lean_object* x_81; -x_81 = lean_box(0); -x_42 = x_81; -goto block_66; -} -else -{ -lean_object* x_82; -lean_dec(x_32); -x_82 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -lean_inc(x_3); -{ -lean_object* _tmp_9 = x_25; -lean_object* _tmp_10 = x_82; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_3; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -} -goto _start; -} -} -} -block_41: -{ -uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_33); -x_34 = lean_nat_dec_eq(x_4, x_22); -x_35 = lean_box(x_34); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_32); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_box(0); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_19); -return x_40; -} -block_66: -{ -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -lean_dec(x_42); -x_43 = lean_nat_add(x_11, x_24); -lean_inc(x_5); -lean_inc(x_1); -x_44 = l_Array_toSubarray___rarg(x_1, x_43, x_5); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -x_47 = lean_ctor_get(x_44, 2); -lean_inc(x_47); -lean_dec(x_44); -x_48 = lean_nat_dec_lt(x_46, x_47); -if (x_48 == 0) -{ -lean_object* x_49; -lean_dec(x_47); -lean_dec(x_46); -lean_dec(x_45); -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_49 = lean_box(0); -x_33 = x_49; -goto block_41; -} -else -{ -lean_object* x_50; uint8_t x_51; -x_50 = lean_array_get_size(x_45); -x_51 = lean_nat_dec_le(x_47, x_50); -if (x_51 == 0) -{ -uint8_t x_52; -lean_dec(x_47); -x_52 = lean_nat_dec_lt(x_46, x_50); -if (x_52 == 0) -{ -lean_object* x_53; -lean_dec(x_50); -lean_dec(x_46); -lean_dec(x_45); -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_53 = lean_box(0); -x_33 = x_53; -goto block_41; -} -else -{ -size_t x_54; size_t x_55; uint8_t x_56; -x_54 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_55 = lean_usize_of_nat(x_50); -lean_dec(x_50); -x_56 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(x_32, x_45, x_54, x_55); -lean_dec(x_45); -if (x_56 == 0) -{ -lean_object* x_57; -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_57 = lean_box(0); -x_33 = x_57; -goto block_41; -} -else -{ -lean_object* x_58; -lean_dec(x_32); -x_58 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -lean_inc(x_3); -{ -lean_object* _tmp_9 = x_25; -lean_object* _tmp_10 = x_58; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_3; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -} -goto _start; -} -} -} -else -{ -size_t x_60; size_t x_61; uint8_t x_62; -lean_dec(x_50); -x_60 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_61 = lean_usize_of_nat(x_47); -lean_dec(x_47); -x_62 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(x_32, x_45, x_60, x_61); -lean_dec(x_45); -if (x_62 == 0) -{ -lean_object* x_63; -lean_dec(x_25); -lean_dec(x_11); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_63 = lean_box(0); -x_33 = x_63; -goto block_41; -} -else -{ -lean_object* x_64; -lean_dec(x_32); -x_64 = lean_nat_add(x_11, x_9); -lean_dec(x_11); -lean_inc(x_3); -{ -lean_object* _tmp_9 = x_25; -lean_object* _tmp_10 = x_64; -lean_object* _tmp_11 = lean_box(0); -lean_object* _tmp_12 = x_3; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_12 = _tmp_11; -x_13 = _tmp_12; -} -goto _start; -} -} -} -} -} -} -else -{ -lean_object* x_98; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_13); -lean_ctor_set(x_98, 1, x_19); -return x_98; -} -} -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: -{ -uint8_t x_18; -x_18 = lean_nat_dec_lt(x_9, x_6); -if (x_18 == 0) -{ -lean_object* x_19; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_1); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_11); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -else -{ -lean_object* x_20; uint8_t x_21; -x_20 = lean_unsigned_to_nat(0u); -x_21 = lean_nat_dec_eq(x_8, x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_11); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_sub(x_8, x_22); -lean_dec(x_8); -x_24 = lean_nat_sub(x_2, x_9); -lean_inc(x_24); -x_25 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_25, 0, x_20); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_25, 2, x_22); -lean_inc_n(x_24, 2); -lean_inc_n(x_4, 2); -lean_inc(x_1); -x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(x_1, x_2, x_4, x_9, x_24, x_25, x_20, x_24, x_22, x_24, x_20, lean_box(0), x_4, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_25); -lean_dec(x_24); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -lean_dec(x_27); -if (lean_obj_tag(x_28) == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_dec(x_26); -x_30 = lean_nat_add(x_9, x_7); -lean_dec(x_9); -lean_inc(x_4); -{ -lean_object* _tmp_7 = x_23; -lean_object* _tmp_8 = x_30; -lean_object* _tmp_9 = lean_box(0); -lean_object* _tmp_10 = x_4; -lean_object* _tmp_16 = x_29; -x_8 = _tmp_7; -x_9 = _tmp_8; -x_10 = _tmp_9; -x_11 = _tmp_10; -x_17 = _tmp_16; -} -goto _start; -} -else -{ -uint8_t x_32; -lean_dec(x_23); -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_1); -x_32 = !lean_is_exclusive(x_26); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_ctor_get(x_26, 0); -lean_dec(x_33); -x_34 = !lean_is_exclusive(x_28); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_28); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set(x_26, 0, x_36); -return x_26; -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_28, 0); -lean_inc(x_37); -lean_dec(x_28); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_box(0); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_26, 0, x_40); -return x_26; -} -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_41 = lean_ctor_get(x_26, 1); -lean_inc(x_41); -lean_dec(x_26); -x_42 = lean_ctor_get(x_28, 0); -lean_inc(x_42); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - x_43 = x_28; -} else { - lean_dec_ref(x_28); - x_43 = lean_box(0); -} -if (lean_is_scalar(x_43)) { - x_44 = lean_alloc_ctor(1, 1, 0); -} else { - x_44 = x_43; -} -lean_ctor_set(x_44, 0, x_42); -x_45 = lean_box(0); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_41); -return x_47; -} -} -} -else -{ -lean_object* x_48; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_4); -lean_dec(x_1); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_11); -lean_ctor_set(x_48, 1, x_17); -return x_48; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_nat_dec_lt(x_10, x_1); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; -x_13 = l_outOfBounds___rarg(x_12); -x_14 = lean_array_get_size(x_13); -x_15 = lean_nat_dec_lt(x_10, x_14); -lean_dec(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_13); -x_16 = l_Lean_instInhabitedName; -x_17 = l_outOfBounds___rarg(x_16); -x_18 = 0; -x_19 = lean_box(x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_9); -return x_21; -} -else -{ -lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_array_fget(x_13, x_10); -lean_dec(x_13); -x_23 = 0; -x_24 = lean_box(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_22); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_9); -return x_26; -} -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_array_fget(x_2, x_10); -x_28 = lean_array_get_size(x_27); -x_29 = lean_nat_dec_lt(x_10, x_28); -lean_dec(x_28); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_27); -x_30 = l_Lean_instInhabitedName; -x_31 = l_outOfBounds___rarg(x_30); -x_32 = 0; -x_33 = lean_box(x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_31); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_9); -return x_35; -} -else -{ -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_array_fget(x_27, x_10); -lean_dec(x_27); -x_37 = 0; -x_38 = lean_box(x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_9); -return x_40; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_8 = lean_array_get_size(x_1); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_unsigned_to_nat(1u); -lean_inc(x_8); -x_11 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_8); -lean_ctor_set(x_11, 2, x_10); -x_12 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -lean_inc(x_8); -lean_inc(x_1); -x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(x_1, x_8, x_11, x_12, x_9, x_8, x_10, x_8, x_9, lean_box(0), x_12, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_11); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_box(0); -x_18 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(x_8, x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_16); -lean_dec(x_1); -lean_dec(x_8); -return x_18; -} -else -{ -uint8_t x_19; -lean_dec(x_8); -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_13); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_13, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_15, 0); -lean_inc(x_21); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_21); -return x_13; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_dec(x_13); -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; uint8_t x_8; size_t x_9; size_t x_10; -x_7 = lean_array_uget(x_2, x_3); -x_8 = lean_name_eq(x_7, x_1); -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -if (x_8 == 0) -{ -lean_object* x_11; -x_11 = lean_array_push(x_5, x_7); -x_3 = x_10; -x_5 = x_11; -goto _start; -} -else -{ -lean_dec(x_7); -x_3 = x_10; -goto _start; -} -} -else -{ -return x_5; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(size_t x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_5, x_4); -if (x_7 == 0) -{ -lean_dec(x_2); -return x_6; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; size_t x_14; size_t x_15; -x_8 = lean_array_uget(x_6, x_5); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_array_uset(x_6, x_5, x_9); -x_11 = lean_array_get_size(x_8); -lean_inc(x_2); -x_12 = lean_array_mk(x_2); -x_13 = lean_nat_dec_lt(x_9, x_11); -x_14 = 1; -x_15 = lean_usize_add(x_5, x_14); -if (x_13 == 0) -{ -lean_object* x_16; -lean_dec(x_11); -lean_dec(x_8); -x_16 = lean_array_uset(x_10, x_5, x_12); -x_5 = x_15; -x_6 = x_16; -goto _start; -} -else -{ -uint8_t x_18; -x_18 = lean_nat_dec_le(x_11, x_11); -if (x_18 == 0) -{ -lean_object* x_19; -lean_dec(x_11); -lean_dec(x_8); -x_19 = lean_array_uset(x_10, x_5, x_12); -x_5 = x_15; -x_6 = x_19; -goto _start; -} -else -{ -size_t x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_22 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(x_3, x_8, x_1, x_21, x_12); -lean_dec(x_8); -x_23 = lean_array_uset(x_10, x_5, x_22); -x_5 = x_15; -x_6 = x_23; -goto _start; -} -} -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_uget(x_2, x_3); -x_7 = lean_name_eq(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -size_t x_8; size_t x_9; -x_8 = 1; -x_9 = lean_usize_add(x_3, x_8); -x_3 = x_9; -goto _start; -} -else -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) -{ -return x_4; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; -x_6 = lean_array_uget(x_4, x_3); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_array_uset(x_4, x_3, x_7); -x_9 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_1, x_6); -x_10 = lean_box(x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_6); -x_12 = 1; -x_13 = lean_usize_add(x_3, x_12); -x_14 = lean_array_uset(x_8, x_3, x_11); -x_3 = x_13; -x_4 = x_14; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; size_t x_15; size_t x_16; -x_7 = lean_array_uget(x_2, x_3); -x_8 = lean_array_get_size(x_7); -x_9 = lean_unsigned_to_nat(1u); -lean_inc(x_7); -x_10 = l_Array_toSubarray___rarg(x_7, x_9, x_8); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_10, 2); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_nat_dec_lt(x_12, x_13); -x_15 = 1; -x_16 = lean_usize_add(x_3, x_15); -if (x_14 == 0) -{ -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -x_3 = x_16; -goto _start; -} -else -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_array_get_size(x_11); -x_19 = lean_nat_dec_le(x_13, x_18); -if (x_19 == 0) -{ -uint8_t x_20; -lean_dec(x_13); -x_20 = lean_nat_dec_lt(x_12, x_18); -if (x_20 == 0) -{ -lean_dec(x_18); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -x_3 = x_16; -goto _start; -} -else -{ -size_t x_22; size_t x_23; uint8_t x_24; -x_22 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_23 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_24 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(x_1, x_11, x_22, x_23); -lean_dec(x_11); -if (x_24 == 0) -{ -lean_dec(x_7); -x_3 = x_16; -goto _start; -} -else -{ -lean_object* x_26; -x_26 = lean_array_push(x_5, x_7); -x_3 = x_16; -x_5 = x_26; -goto _start; -} -} -} -else -{ -size_t x_28; size_t x_29; uint8_t x_30; -lean_dec(x_18); -x_28 = lean_usize_of_nat(x_12); -lean_dec(x_12); -x_29 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_30 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(x_1, x_11, x_28, x_29); -lean_dec(x_11); -if (x_30 == 0) -{ -lean_dec(x_7); -x_3 = x_16; -goto _start; -} -else -{ -lean_object* x_32; -x_32 = lean_array_push(x_5, x_7); -x_3 = x_16; -x_5 = x_32; -goto _start; -} -} -} -} -else -{ -return x_5; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -lean_inc(x_1); -x_15 = lean_array_push(x_6, x_1); -x_16 = lean_array_size(x_7); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(x_2, x_3, x_1, x_16, x_2, x_7); -lean_dec(x_1); -x_18 = lean_array_get_size(x_17); -x_19 = lean_unsigned_to_nat(0u); -x_20 = lean_nat_dec_lt(x_19, x_18); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_18); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_4); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_5); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_14); -return x_24; -} -else -{ -uint8_t x_25; -x_25 = lean_nat_dec_le(x_18, x_18); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_18); -lean_dec(x_17); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_15); -lean_ctor_set(x_26, 1, x_4); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_5); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_14); -return x_29; -} -else -{ -size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_30 = lean_usize_of_nat(x_18); -lean_dec(x_18); -x_31 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_17, x_2, x_30, x_4); -lean_dec(x_17); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_15); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_5); -lean_ctor_set(x_33, 1, x_32); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_14); -return x_35; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_6); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_6, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_6, 0); -x_17 = lean_ctor_get(x_14, 0); -x_18 = lean_ctor_get(x_14, 1); -x_19 = l_Array_isEmpty___rarg(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_free_object(x_14); -lean_free_object(x_6); -lean_inc(x_18); -x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_array_get_size(x_18); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); -if (x_28 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -uint8_t x_49; -x_49 = lean_nat_dec_le(x_26, x_26); -if (x_49 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -size_t x_50; lean_object* x_51; -x_50 = lean_usize_of_nat(x_26); -lean_dec(x_26); -lean_inc(x_4); -x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); -x_30 = x_51; -goto block_48; -} -} -block_48: -{ -size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_31 = lean_array_size(x_30); -x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); -x_33 = lean_array_get_size(x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_33, x_34); -lean_dec(x_33); -x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); -lean_dec(x_35); -x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); -lean_dec(x_36); -x_38 = lean_array_size(x_37); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); -lean_inc(x_25); -x_40 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); -x_41 = lean_array_push(x_16, x_40); -x_42 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -lean_dec(x_44); -x_6 = x_46; -x_12 = x_45; -goto _start; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_52 = lean_ctor_get(x_20, 1); -lean_inc(x_52); -lean_dec(x_20); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_dec(x_21); -x_54 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -lean_dec(x_56); -x_6 = x_58; -x_12 = x_57; -goto _start; -} -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_60 = lean_ctor_get(x_20, 1); -lean_inc(x_60); -lean_dec(x_20); -x_61 = lean_ctor_get(x_21, 1); -lean_inc(x_61); -lean_dec(x_21); -x_62 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -lean_dec(x_64); -x_6 = x_66; -x_12 = x_65; -goto _start; -} -} -else -{ -lean_object* x_68; -lean_dec(x_5); -lean_dec(x_4); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_6); -lean_ctor_set(x_68, 1, x_12); -return x_68; -} -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_6, 0); -x_70 = lean_ctor_get(x_14, 0); -x_71 = lean_ctor_get(x_14, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_14); -x_72 = l_Array_isEmpty___rarg(x_71); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -lean_free_object(x_6); -lean_inc(x_71); -x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_unbox(x_75); -lean_dec(x_75); -if (x_76 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; -x_77 = lean_ctor_get(x_73, 1); -lean_inc(x_77); -lean_dec(x_73); -x_78 = lean_ctor_get(x_74, 1); -lean_inc(x_78); -lean_dec(x_74); -x_79 = lean_array_get_size(x_71); -x_80 = lean_unsigned_to_nat(0u); -x_81 = lean_nat_dec_lt(x_80, x_79); -x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); -if (x_81 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -uint8_t x_102; -x_102 = lean_nat_dec_le(x_79, x_79); -if (x_102 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -size_t x_103; lean_object* x_104; -x_103 = lean_usize_of_nat(x_79); -lean_dec(x_79); -lean_inc(x_4); -x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); -x_83 = x_104; -goto block_101; -} -} -block_101: -{ -size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_84 = lean_array_size(x_83); -x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); -x_86 = lean_array_get_size(x_85); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_sub(x_86, x_87); -lean_dec(x_86); -x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); -lean_dec(x_88); -x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); -lean_dec(x_89); -x_91 = lean_array_size(x_90); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); -lean_inc(x_78); -x_93 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_93, 0, x_78); -lean_ctor_set(x_93, 1, x_92); -lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); -x_94 = lean_array_push(x_69, x_93); -x_95 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = lean_ctor_get(x_97, 0); -lean_inc(x_99); -lean_dec(x_97); -x_6 = x_99; -x_12 = x_98; -goto _start; -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_105 = lean_ctor_get(x_73, 1); -lean_inc(x_105); -lean_dec(x_73); -x_106 = lean_ctor_get(x_74, 1); -lean_inc(x_106); -lean_dec(x_74); -x_107 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -lean_dec(x_109); -x_6 = x_111; -x_12 = x_110; -goto _start; -} -} -else -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_113 = lean_ctor_get(x_73, 1); -lean_inc(x_113); -lean_dec(x_73); -x_114 = lean_ctor_get(x_74, 1); -lean_inc(x_114); -lean_dec(x_74); -x_115 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_ctor_get(x_117, 0); -lean_inc(x_119); -lean_dec(x_117); -x_6 = x_119; -x_12 = x_118; -goto _start; -} -} -else -{ -lean_object* x_121; lean_object* x_122; -lean_dec(x_5); -lean_dec(x_4); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_70); -lean_ctor_set(x_121, 1, x_71); -lean_ctor_set(x_6, 1, x_121); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_6); -lean_ctor_set(x_122, 1, x_12); -return x_122; -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_123 = lean_ctor_get(x_6, 1); -x_124 = lean_ctor_get(x_6, 0); -lean_inc(x_123); -lean_inc(x_124); -lean_dec(x_6); -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_127 = x_123; -} else { - lean_dec_ref(x_123); - x_127 = lean_box(0); -} -x_128 = l_Array_isEmpty___rarg(x_126); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -lean_dec(x_127); -lean_inc(x_126); -x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_unbox(x_131); -lean_dec(x_131); -if (x_132 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; -x_133 = lean_ctor_get(x_129, 1); -lean_inc(x_133); -lean_dec(x_129); -x_134 = lean_ctor_get(x_130, 1); -lean_inc(x_134); -lean_dec(x_130); -x_135 = lean_array_get_size(x_126); -x_136 = lean_unsigned_to_nat(0u); -x_137 = lean_nat_dec_lt(x_136, x_135); -x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); -if (x_137 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -uint8_t x_158; -x_158 = lean_nat_dec_le(x_135, x_135); -if (x_158 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -size_t x_159; lean_object* x_160; -x_159 = lean_usize_of_nat(x_135); -lean_dec(x_135); -lean_inc(x_4); -x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); -x_139 = x_160; -goto block_157; -} -} -block_157: -{ -size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_140 = lean_array_size(x_139); -x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); -x_142 = lean_array_get_size(x_141); -x_143 = lean_unsigned_to_nat(1u); -x_144 = lean_nat_sub(x_142, x_143); -lean_dec(x_142); -x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); -lean_dec(x_144); -x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); -lean_dec(x_145); -x_147 = lean_array_size(x_146); -x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); -lean_inc(x_134); -x_149 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_149, 0, x_134); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); -x_150 = lean_array_push(x_124, x_149); -x_151 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_152, 1); -lean_inc(x_154); -lean_dec(x_152); -x_155 = lean_ctor_get(x_153, 0); -lean_inc(x_155); -lean_dec(x_153); -x_6 = x_155; -x_12 = x_154; -goto _start; -} -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_161 = lean_ctor_get(x_129, 1); -lean_inc(x_161); -lean_dec(x_129); -x_162 = lean_ctor_get(x_130, 1); -lean_inc(x_162); -lean_dec(x_130); -x_163 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_ctor_get(x_165, 0); -lean_inc(x_167); -lean_dec(x_165); -x_6 = x_167; -x_12 = x_166; -goto _start; -} -} -else -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_169 = lean_ctor_get(x_129, 1); -lean_inc(x_169); -lean_dec(x_129); -x_170 = lean_ctor_get(x_130, 1); -lean_inc(x_170); -lean_dec(x_130); -x_171 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = lean_ctor_get(x_173, 0); -lean_inc(x_175); -lean_dec(x_173); -x_6 = x_175; -x_12 = x_174; -goto _start; -} -} -else -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_5); -lean_dec(x_4); -if (lean_is_scalar(x_127)) { - x_177 = lean_alloc_ctor(0, 2, 0); -} else { - x_177 = x_127; -} -lean_ctor_set(x_177, 0, x_125); -lean_ctor_set(x_177, 1, x_126); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_124); -lean_ctor_set(x_178, 1, x_177); -x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_12); -return x_179; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_PersistentHashMap_insert___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__1(x_3, x_1, x_2); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_structureResolutionExt; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; -x_2 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -lean_ctor_set(x_2, 2, x_1); -lean_ctor_set(x_2, 3, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5; -x_2 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; -x_3 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6; -x_4 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_2); -lean_ctor_set(x_4, 3, x_2); -lean_ctor_set(x_4, 4, x_2); -lean_ctor_set(x_4, 5, x_3); -lean_ctor_set(x_4, 6, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_st_ref_take(x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_13 = lean_ctor_get(x_10, 0); -x_14 = lean_ctor_get(x_10, 4); -lean_dec(x_14); -x_15 = lean_alloc_closure((void*)(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1), 3, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -x_16 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; -x_17 = l_Lean_EnvExtension_modifyState___rarg(x_16, x_13, x_15); -x_18 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; -lean_ctor_set(x_10, 4, x_18); -lean_ctor_set(x_10, 0, x_17); -x_19 = lean_st_ref_set(x_7, x_10, x_11); -x_20 = lean_ctor_get(x_19, 1); -lean_inc(x_20); -lean_dec(x_19); -x_21 = lean_st_ref_take(x_5, x_20); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_22, 1); -lean_dec(x_25); -x_26 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; -lean_ctor_set(x_22, 1, x_26); -x_27 = lean_st_ref_set(x_5, x_22, x_23); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -x_30 = lean_box(0); -lean_ctor_set(x_27, 0, x_30); -return x_27; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 1); -lean_inc(x_31); -lean_dec(x_27); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; -} -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_34 = lean_ctor_get(x_22, 0); -x_35 = lean_ctor_get(x_22, 2); -x_36 = lean_ctor_get(x_22, 3); -x_37 = lean_ctor_get(x_22, 4); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_22); -x_38 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; -x_39 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_39, 0, x_34); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_39, 2, x_35); -lean_ctor_set(x_39, 3, x_36); -lean_ctor_set(x_39, 4, x_37); -x_40 = lean_st_ref_set(x_5, x_39, x_23); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_42 = x_40; -} else { - lean_dec_ref(x_40); - x_42 = lean_box(0); -} -x_43 = lean_box(0); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_41); -return x_44; -} -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_45 = lean_ctor_get(x_10, 0); -x_46 = lean_ctor_get(x_10, 1); -x_47 = lean_ctor_get(x_10, 2); -x_48 = lean_ctor_get(x_10, 3); -x_49 = lean_ctor_get(x_10, 5); -x_50 = lean_ctor_get(x_10, 6); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_10); -x_51 = lean_alloc_closure((void*)(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1), 3, 2); -lean_closure_set(x_51, 0, x_1); -lean_closure_set(x_51, 1, x_2); -x_52 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; -x_53 = l_Lean_EnvExtension_modifyState___rarg(x_52, x_45, x_51); -x_54 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; -x_55 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_46); -lean_ctor_set(x_55, 2, x_47); -lean_ctor_set(x_55, 3, x_48); -lean_ctor_set(x_55, 4, x_54); -lean_ctor_set(x_55, 5, x_49); -lean_ctor_set(x_55, 6, x_50); -x_56 = lean_st_ref_set(x_7, x_55, x_11); -x_57 = lean_ctor_get(x_56, 1); -lean_inc(x_57); -lean_dec(x_56); -x_58 = lean_st_ref_take(x_5, x_57); -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = lean_ctor_get(x_59, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_59, 2); -lean_inc(x_62); -x_63 = lean_ctor_get(x_59, 3); -lean_inc(x_63); -x_64 = lean_ctor_get(x_59, 4); -lean_inc(x_64); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - lean_ctor_release(x_59, 2); - lean_ctor_release(x_59, 3); - lean_ctor_release(x_59, 4); - x_65 = x_59; -} else { - lean_dec_ref(x_59); - x_65 = lean_box(0); -} -x_66 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 5, 0); -} else { - x_67 = x_65; -} -lean_ctor_set(x_67, 0, x_61); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_67, 2, x_62); -lean_ctor_set(x_67, 3, x_63); -lean_ctor_set(x_67, 4, x_64); -x_68 = lean_st_ref_set(x_5, x_67, x_60); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_70 = x_68; -} else { - lean_dec_ref(x_68); - x_70 = lean_box(0); -} -x_71 = lean_box(0); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(0, 2, 0); -} else { - x_72 = x_70; -} -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -return x_72; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_6); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_6, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_6, 0); -x_17 = lean_ctor_get(x_14, 0); -x_18 = lean_ctor_get(x_14, 1); -x_19 = l_Array_isEmpty___rarg(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_free_object(x_14); -lean_free_object(x_6); -lean_inc(x_18); -x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_array_get_size(x_18); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); -if (x_28 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -uint8_t x_49; -x_49 = lean_nat_dec_le(x_26, x_26); -if (x_49 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -size_t x_50; lean_object* x_51; -x_50 = lean_usize_of_nat(x_26); -lean_dec(x_26); -lean_inc(x_4); -x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); -x_30 = x_51; -goto block_48; -} -} -block_48: -{ -size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_31 = lean_array_size(x_30); -x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); -x_33 = lean_array_get_size(x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_33, x_34); -lean_dec(x_33); -x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); -lean_dec(x_35); -x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); -lean_dec(x_36); -x_38 = lean_array_size(x_37); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); -lean_inc(x_25); -x_40 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); -x_41 = lean_array_push(x_16, x_40); -x_42 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -lean_dec(x_44); -x_6 = x_46; -x_12 = x_45; -goto _start; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_52 = lean_ctor_get(x_20, 1); -lean_inc(x_52); -lean_dec(x_20); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_dec(x_21); -x_54 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -lean_dec(x_56); -x_6 = x_58; -x_12 = x_57; -goto _start; -} -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_60 = lean_ctor_get(x_20, 1); -lean_inc(x_60); -lean_dec(x_20); -x_61 = lean_ctor_get(x_21, 1); -lean_inc(x_61); -lean_dec(x_21); -x_62 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -lean_dec(x_64); -x_6 = x_66; -x_12 = x_65; -goto _start; -} -} -else -{ -lean_object* x_68; -lean_dec(x_5); -lean_dec(x_4); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_6); -lean_ctor_set(x_68, 1, x_12); -return x_68; -} -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_6, 0); -x_70 = lean_ctor_get(x_14, 0); -x_71 = lean_ctor_get(x_14, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_14); -x_72 = l_Array_isEmpty___rarg(x_71); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -lean_free_object(x_6); -lean_inc(x_71); -x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_unbox(x_75); -lean_dec(x_75); -if (x_76 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; -x_77 = lean_ctor_get(x_73, 1); -lean_inc(x_77); -lean_dec(x_73); -x_78 = lean_ctor_get(x_74, 1); -lean_inc(x_78); -lean_dec(x_74); -x_79 = lean_array_get_size(x_71); -x_80 = lean_unsigned_to_nat(0u); -x_81 = lean_nat_dec_lt(x_80, x_79); -x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); -if (x_81 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -uint8_t x_102; -x_102 = lean_nat_dec_le(x_79, x_79); -if (x_102 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -size_t x_103; lean_object* x_104; -x_103 = lean_usize_of_nat(x_79); -lean_dec(x_79); -lean_inc(x_4); -x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); -x_83 = x_104; -goto block_101; -} -} -block_101: -{ -size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_84 = lean_array_size(x_83); -x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); -x_86 = lean_array_get_size(x_85); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_sub(x_86, x_87); -lean_dec(x_86); -x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); -lean_dec(x_88); -x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); -lean_dec(x_89); -x_91 = lean_array_size(x_90); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); -lean_inc(x_78); -x_93 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_93, 0, x_78); -lean_ctor_set(x_93, 1, x_92); -lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); -x_94 = lean_array_push(x_69, x_93); -x_95 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = lean_ctor_get(x_97, 0); -lean_inc(x_99); -lean_dec(x_97); -x_6 = x_99; -x_12 = x_98; -goto _start; -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_105 = lean_ctor_get(x_73, 1); -lean_inc(x_105); -lean_dec(x_73); -x_106 = lean_ctor_get(x_74, 1); -lean_inc(x_106); -lean_dec(x_74); -x_107 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -lean_dec(x_109); -x_6 = x_111; -x_12 = x_110; -goto _start; -} -} -else -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_113 = lean_ctor_get(x_73, 1); -lean_inc(x_113); -lean_dec(x_73); -x_114 = lean_ctor_get(x_74, 1); -lean_inc(x_114); -lean_dec(x_74); -x_115 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_ctor_get(x_117, 0); -lean_inc(x_119); -lean_dec(x_117); -x_6 = x_119; -x_12 = x_118; -goto _start; -} -} -else -{ -lean_object* x_121; lean_object* x_122; -lean_dec(x_5); -lean_dec(x_4); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_70); -lean_ctor_set(x_121, 1, x_71); -lean_ctor_set(x_6, 1, x_121); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_6); -lean_ctor_set(x_122, 1, x_12); -return x_122; -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_123 = lean_ctor_get(x_6, 1); -x_124 = lean_ctor_get(x_6, 0); -lean_inc(x_123); -lean_inc(x_124); -lean_dec(x_6); -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_127 = x_123; -} else { - lean_dec_ref(x_123); - x_127 = lean_box(0); -} -x_128 = l_Array_isEmpty___rarg(x_126); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -lean_dec(x_127); -lean_inc(x_126); -x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_unbox(x_131); -lean_dec(x_131); -if (x_132 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; -x_133 = lean_ctor_get(x_129, 1); -lean_inc(x_133); -lean_dec(x_129); -x_134 = lean_ctor_get(x_130, 1); -lean_inc(x_134); -lean_dec(x_130); -x_135 = lean_array_get_size(x_126); -x_136 = lean_unsigned_to_nat(0u); -x_137 = lean_nat_dec_lt(x_136, x_135); -x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); -if (x_137 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -uint8_t x_158; -x_158 = lean_nat_dec_le(x_135, x_135); -if (x_158 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -size_t x_159; lean_object* x_160; -x_159 = lean_usize_of_nat(x_135); -lean_dec(x_135); -lean_inc(x_4); -x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); -x_139 = x_160; -goto block_157; -} -} -block_157: -{ -size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_140 = lean_array_size(x_139); -x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); -x_142 = lean_array_get_size(x_141); -x_143 = lean_unsigned_to_nat(1u); -x_144 = lean_nat_sub(x_142, x_143); -lean_dec(x_142); -x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); -lean_dec(x_144); -x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); -lean_dec(x_145); -x_147 = lean_array_size(x_146); -x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); -lean_inc(x_134); -x_149 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_149, 0, x_134); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); -x_150 = lean_array_push(x_124, x_149); -x_151 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_152, 1); -lean_inc(x_154); -lean_dec(x_152); -x_155 = lean_ctor_get(x_153, 0); -lean_inc(x_155); -lean_dec(x_153); -x_6 = x_155; -x_12 = x_154; -goto _start; -} -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_161 = lean_ctor_get(x_129, 1); -lean_inc(x_161); -lean_dec(x_129); -x_162 = lean_ctor_get(x_130, 1); -lean_inc(x_162); -lean_dec(x_130); -x_163 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_ctor_get(x_165, 0); -lean_inc(x_167); -lean_dec(x_165); -x_6 = x_167; -x_12 = x_166; -goto _start; -} -} -else -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_169 = lean_ctor_get(x_129, 1); -lean_inc(x_169); -lean_dec(x_129); -x_170 = lean_ctor_get(x_130, 1); -lean_inc(x_170); -lean_dec(x_130); -x_171 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = lean_ctor_get(x_173, 0); -lean_inc(x_175); -lean_dec(x_173); -x_6 = x_175; -x_12 = x_174; -goto _start; -} -} -else -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_5); -lean_dec(x_4); -if (lean_is_scalar(x_127)) { - x_177 = lean_alloc_ctor(0, 2, 0); -} else { - x_177 = x_127; -} -lean_ctor_set(x_177, 0, x_125); -lean_ctor_set(x_177, 1, x_126); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_124); -lean_ctor_set(x_178, 1, x_177); -x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_12); -return x_179; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_6); -if (x_13 == 0) -{ -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_6, 1); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_6, 0); -x_17 = lean_ctor_get(x_14, 0); -x_18 = lean_ctor_get(x_14, 1); -x_19 = l_Array_isEmpty___rarg(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -lean_free_object(x_14); -lean_free_object(x_6); -lean_inc(x_18); -x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = lean_array_get_size(x_18); -x_27 = lean_unsigned_to_nat(0u); -x_28 = lean_nat_dec_lt(x_27, x_26); -x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); -if (x_28 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -uint8_t x_49; -x_49 = lean_nat_dec_le(x_26, x_26); -if (x_49 == 0) -{ -lean_dec(x_26); -lean_inc(x_4); -x_30 = x_4; -goto block_48; -} -else -{ -size_t x_50; lean_object* x_51; -x_50 = lean_usize_of_nat(x_26); -lean_dec(x_26); -lean_inc(x_4); -x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); -x_30 = x_51; -goto block_48; -} -} -block_48: -{ -size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_31 = lean_array_size(x_30); -x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); -x_33 = lean_array_get_size(x_32); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_sub(x_33, x_34); -lean_dec(x_33); -x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); -lean_dec(x_35); -x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); -lean_dec(x_36); -x_38 = lean_array_size(x_37); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); -lean_inc(x_25); -x_40 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_40, 0, x_25); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); -x_41 = lean_array_push(x_16, x_40); -x_42 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_ctor_get(x_44, 0); -lean_inc(x_46); -lean_dec(x_44); -x_6 = x_46; -x_12 = x_45; -goto _start; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_52 = lean_ctor_get(x_20, 1); -lean_inc(x_52); -lean_dec(x_20); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_dec(x_21); -x_54 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -lean_dec(x_56); -x_6 = x_58; -x_12 = x_57; -goto _start; -} -} -else -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_60 = lean_ctor_get(x_20, 1); -lean_inc(x_60); -lean_dec(x_20); -x_61 = lean_ctor_get(x_21, 1); -lean_inc(x_61); -lean_dec(x_21); -x_62 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -lean_dec(x_64); -x_6 = x_66; -x_12 = x_65; -goto _start; -} -} -else -{ -lean_object* x_68; -lean_dec(x_5); -lean_dec(x_4); -x_68 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_68, 0, x_6); -lean_ctor_set(x_68, 1, x_12); -return x_68; -} -} -else -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_69 = lean_ctor_get(x_6, 0); -x_70 = lean_ctor_get(x_14, 0); -x_71 = lean_ctor_get(x_14, 1); -lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_14); -x_72 = l_Array_isEmpty___rarg(x_71); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -lean_free_object(x_6); -lean_inc(x_71); -x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_unbox(x_75); -lean_dec(x_75); -if (x_76 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; -x_77 = lean_ctor_get(x_73, 1); -lean_inc(x_77); -lean_dec(x_73); -x_78 = lean_ctor_get(x_74, 1); -lean_inc(x_78); -lean_dec(x_74); -x_79 = lean_array_get_size(x_71); -x_80 = lean_unsigned_to_nat(0u); -x_81 = lean_nat_dec_lt(x_80, x_79); -x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); -if (x_81 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -uint8_t x_102; -x_102 = lean_nat_dec_le(x_79, x_79); -if (x_102 == 0) -{ -lean_dec(x_79); -lean_inc(x_4); -x_83 = x_4; -goto block_101; -} -else -{ -size_t x_103; lean_object* x_104; -x_103 = lean_usize_of_nat(x_79); -lean_dec(x_79); -lean_inc(x_4); -x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); -x_83 = x_104; -goto block_101; -} -} -block_101: -{ -size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_84 = lean_array_size(x_83); -x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); -x_86 = lean_array_get_size(x_85); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_sub(x_86, x_87); -lean_dec(x_86); -x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); -lean_dec(x_88); -x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); -lean_dec(x_89); -x_91 = lean_array_size(x_90); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); -lean_inc(x_78); -x_93 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_93, 0, x_78); -lean_ctor_set(x_93, 1, x_92); -lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); -x_94 = lean_array_push(x_69, x_93); -x_95 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); -x_97 = lean_ctor_get(x_96, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_96, 1); -lean_inc(x_98); -lean_dec(x_96); -x_99 = lean_ctor_get(x_97, 0); -lean_inc(x_99); -lean_dec(x_97); -x_6 = x_99; -x_12 = x_98; -goto _start; -} -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_105 = lean_ctor_get(x_73, 1); -lean_inc(x_105); -lean_dec(x_73); -x_106 = lean_ctor_get(x_74, 1); -lean_inc(x_106); -lean_dec(x_74); -x_107 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -lean_dec(x_109); -x_6 = x_111; -x_12 = x_110; -goto _start; -} -} -else -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_113 = lean_ctor_get(x_73, 1); -lean_inc(x_113); -lean_dec(x_73); -x_114 = lean_ctor_get(x_74, 1); -lean_inc(x_114); -lean_dec(x_74); -x_115 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); -x_117 = lean_ctor_get(x_116, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_116, 1); -lean_inc(x_118); -lean_dec(x_116); -x_119 = lean_ctor_get(x_117, 0); -lean_inc(x_119); -lean_dec(x_117); -x_6 = x_119; -x_12 = x_118; -goto _start; -} -} -else -{ -lean_object* x_121; lean_object* x_122; -lean_dec(x_5); -lean_dec(x_4); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_70); -lean_ctor_set(x_121, 1, x_71); -lean_ctor_set(x_6, 1, x_121); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_6); -lean_ctor_set(x_122, 1, x_12); -return x_122; -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_123 = lean_ctor_get(x_6, 1); -x_124 = lean_ctor_get(x_6, 0); -lean_inc(x_123); -lean_inc(x_124); -lean_dec(x_6); -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_127 = x_123; -} else { - lean_dec_ref(x_123); - x_127 = lean_box(0); -} -x_128 = l_Array_isEmpty___rarg(x_126); -if (x_128 == 0) -{ -lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; -lean_dec(x_127); -lean_inc(x_126); -x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_unbox(x_131); -lean_dec(x_131); -if (x_132 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; -x_133 = lean_ctor_get(x_129, 1); -lean_inc(x_133); -lean_dec(x_129); -x_134 = lean_ctor_get(x_130, 1); -lean_inc(x_134); -lean_dec(x_130); -x_135 = lean_array_get_size(x_126); -x_136 = lean_unsigned_to_nat(0u); -x_137 = lean_nat_dec_lt(x_136, x_135); -x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); -if (x_137 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -uint8_t x_158; -x_158 = lean_nat_dec_le(x_135, x_135); -if (x_158 == 0) -{ -lean_dec(x_135); -lean_inc(x_4); -x_139 = x_4; -goto block_157; -} -else -{ -size_t x_159; lean_object* x_160; -x_159 = lean_usize_of_nat(x_135); -lean_dec(x_135); -lean_inc(x_4); -x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); -x_139 = x_160; -goto block_157; -} -} -block_157: -{ -size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; -x_140 = lean_array_size(x_139); -x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); -x_142 = lean_array_get_size(x_141); -x_143 = lean_unsigned_to_nat(1u); -x_144 = lean_nat_sub(x_142, x_143); -lean_dec(x_142); -x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); -lean_dec(x_144); -x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); -lean_dec(x_145); -x_147 = lean_array_size(x_146); -x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); -lean_inc(x_134); -x_149 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_149, 0, x_134); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); -x_150 = lean_array_push(x_124, x_149); -x_151 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_152, 1); -lean_inc(x_154); -lean_dec(x_152); -x_155 = lean_ctor_get(x_153, 0); -lean_inc(x_155); -lean_dec(x_153); -x_6 = x_155; -x_12 = x_154; -goto _start; -} -} -else -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_161 = lean_ctor_get(x_129, 1); -lean_inc(x_161); -lean_dec(x_129); -x_162 = lean_ctor_get(x_130, 1); -lean_inc(x_162); -lean_dec(x_130); -x_163 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = lean_ctor_get(x_165, 0); -lean_inc(x_167); -lean_dec(x_165); -x_6 = x_167; -x_12 = x_166; -goto _start; -} -} -else -{ -lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_169 = lean_ctor_get(x_129, 1); -lean_inc(x_169); -lean_dec(x_129); -x_170 = lean_ctor_get(x_130, 1); -lean_inc(x_170); -lean_dec(x_130); -x_171 = lean_box(0); -lean_inc(x_4); -lean_inc(x_5); -x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = lean_ctor_get(x_173, 0); -lean_inc(x_175); -lean_dec(x_173); -x_6 = x_175; -x_12 = x_174; -goto _start; -} -} -else -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_5); -lean_dec(x_4); -if (lean_is_scalar(x_127)) { - x_177 = lean_alloc_ctor(0, 2, 0); -} else { - x_177 = x_127; -} -lean_ctor_set(x_177, 0, x_125); -lean_ctor_set(x_177, 1, x_126); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_124); -lean_ctor_set(x_178, 1, x_177); -x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_12); -return x_179; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; size_t x_15; lean_object* x_16; uint8_t x_17; -lean_inc(x_2); -x_11 = l_Lean_getStructureParentInfo(x_1, x_2); -x_12 = lean_array_size(x_11); -x_13 = 0; -x_14 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__1(x_12, x_13, x_11); -x_15 = lean_array_size(x_14); -lean_inc(x_14); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(x_15, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -x_20 = lean_array_get_size(x_18); -x_21 = lean_unsigned_to_nat(1u); -x_22 = lean_nat_add(x_20, x_21); -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_nat_mod(x_23, x_22); -lean_dec(x_22); -lean_inc(x_14); -lean_inc(x_18); -x_25 = lean_array_push(x_18, x_14); -x_26 = l_Array_insertAt_loop___rarg(x_18, x_24, x_25, x_20); -lean_dec(x_24); -lean_dec(x_18); -x_27 = lean_array_get_size(x_26); -x_28 = lean_box(0); -x_29 = lean_nat_dec_lt(x_23, x_27); -lean_inc(x_2); -lean_ctor_set_tag(x_16, 1); -lean_ctor_set(x_16, 1, x_28); -lean_ctor_set(x_16, 0, x_2); -x_30 = lean_array_mk(x_16); -if (x_29 == 0) -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -lean_dec(x_27); -lean_dec(x_26); -x_31 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_3, x_13, x_14, x_31, x_28, x_33, x_5, x_6, x_7, x_8, x_9, x_19); -lean_dec(x_14); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_35, 1); -lean_inc(x_36); -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_ctor_get(x_35, 0); -lean_inc(x_38); -lean_dec(x_35); -x_39 = !lean_is_exclusive(x_36); -if (x_39 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_40 = lean_ctor_get(x_36, 0); -x_41 = lean_ctor_get(x_36, 1); -lean_dec(x_41); -lean_inc(x_40); -x_42 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_40, x_5, x_6, x_7, x_8, x_9, x_37); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) -{ -lean_object* x_44; -x_44 = lean_ctor_get(x_42, 0); -lean_dec(x_44); -lean_ctor_set(x_36, 1, x_38); -lean_ctor_set(x_42, 0, x_36); -return x_42; -} -else -{ -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -lean_ctor_set(x_36, 1, x_38); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_36); -lean_ctor_set(x_46, 1, x_45); -return x_46; -} -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_47 = lean_ctor_get(x_36, 0); -lean_inc(x_47); -lean_dec(x_36); -lean_inc(x_47); -x_48 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_47, x_5, x_6, x_7, x_8, x_9, x_37); -x_49 = lean_ctor_get(x_48, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - lean_ctor_release(x_48, 1); - x_50 = x_48; -} else { - lean_dec_ref(x_48); - x_50 = lean_box(0); -} -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_47); -lean_ctor_set(x_51, 1, x_38); -if (lean_is_scalar(x_50)) { - x_52 = lean_alloc_ctor(0, 2, 0); -} else { - x_52 = x_50; -} -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_49); -return x_52; -} -} -else -{ -uint8_t x_53; -x_53 = lean_nat_dec_le(x_27, x_27); -if (x_53 == 0) -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; -lean_dec(x_27); -lean_dec(x_26); -x_54 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_30); -lean_ctor_set(x_55, 1, x_54); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_3, x_13, x_14, x_54, x_28, x_56, x_5, x_6, x_7, x_8, x_9, x_19); -lean_dec(x_14); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -lean_dec(x_57); -x_61 = lean_ctor_get(x_58, 0); -lean_inc(x_61); -lean_dec(x_58); -x_62 = !lean_is_exclusive(x_59); -if (x_62 == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; -x_63 = lean_ctor_get(x_59, 0); -x_64 = lean_ctor_get(x_59, 1); -lean_dec(x_64); -lean_inc(x_63); -x_65 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_63, x_5, x_6, x_7, x_8, x_9, x_60); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) -{ -lean_object* x_67; -x_67 = lean_ctor_get(x_65, 0); -lean_dec(x_67); -lean_ctor_set(x_59, 1, x_61); -lean_ctor_set(x_65, 0, x_59); -return x_65; -} -else -{ -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_65, 1); -lean_inc(x_68); -lean_dec(x_65); -lean_ctor_set(x_59, 1, x_61); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_59); -lean_ctor_set(x_69, 1, x_68); -return x_69; -} -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_70 = lean_ctor_get(x_59, 0); -lean_inc(x_70); -lean_dec(x_59); -lean_inc(x_70); -x_71 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_70, x_5, x_6, x_7, x_8, x_9, x_60); -x_72 = lean_ctor_get(x_71, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_73 = x_71; -} else { - lean_dec_ref(x_71); - x_73 = lean_box(0); -} -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_70); -lean_ctor_set(x_74, 1, x_61); -if (lean_is_scalar(x_73)) { - x_75 = lean_alloc_ctor(0, 2, 0); -} else { - x_75 = x_73; -} -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_72); -return x_75; -} -} -else -{ -size_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; -x_76 = lean_usize_of_nat(x_27); -lean_dec(x_27); -x_77 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_78 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_26, x_13, x_76, x_77); -lean_dec(x_26); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_30); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_3, x_13, x_14, x_77, x_28, x_80, x_5, x_6, x_7, x_8, x_9, x_19); -lean_dec(x_14); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_82, 1); -lean_inc(x_83); -x_84 = lean_ctor_get(x_81, 1); -lean_inc(x_84); -lean_dec(x_81); -x_85 = lean_ctor_get(x_82, 0); -lean_inc(x_85); -lean_dec(x_82); -x_86 = !lean_is_exclusive(x_83); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; -x_87 = lean_ctor_get(x_83, 0); -x_88 = lean_ctor_get(x_83, 1); -lean_dec(x_88); -lean_inc(x_87); -x_89 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_87, x_5, x_6, x_7, x_8, x_9, x_84); -x_90 = !lean_is_exclusive(x_89); -if (x_90 == 0) -{ -lean_object* x_91; -x_91 = lean_ctor_get(x_89, 0); -lean_dec(x_91); -lean_ctor_set(x_83, 1, x_85); -lean_ctor_set(x_89, 0, x_83); -return x_89; -} -else -{ -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_89, 1); -lean_inc(x_92); -lean_dec(x_89); -lean_ctor_set(x_83, 1, x_85); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_83); -lean_ctor_set(x_93, 1, x_92); -return x_93; -} -} -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_94 = lean_ctor_get(x_83, 0); -lean_inc(x_94); -lean_dec(x_83); -lean_inc(x_94); -x_95 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_94, x_5, x_6, x_7, x_8, x_9, x_84); -x_96 = lean_ctor_get(x_95, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_97 = x_95; -} else { - lean_dec_ref(x_95); - x_97 = lean_box(0); -} -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_94); -lean_ctor_set(x_98, 1, x_85); -if (lean_is_scalar(x_97)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_97; -} -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_96); -return x_99; -} -} -} -} -else -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; -x_100 = lean_ctor_get(x_16, 0); -x_101 = lean_ctor_get(x_16, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_16); -x_102 = lean_array_get_size(x_100); -x_103 = lean_unsigned_to_nat(1u); -x_104 = lean_nat_add(x_102, x_103); -x_105 = lean_unsigned_to_nat(0u); -x_106 = lean_nat_mod(x_105, x_104); -lean_dec(x_104); -lean_inc(x_14); -lean_inc(x_100); -x_107 = lean_array_push(x_100, x_14); -x_108 = l_Array_insertAt_loop___rarg(x_100, x_106, x_107, x_102); -lean_dec(x_106); -lean_dec(x_100); -x_109 = lean_array_get_size(x_108); -x_110 = lean_box(0); -x_111 = lean_nat_dec_lt(x_105, x_109); -lean_inc(x_2); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_2); -lean_ctor_set(x_112, 1, x_110); -x_113 = lean_array_mk(x_112); -if (x_111 == 0) -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_109); -lean_dec(x_108); -x_114 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_3, x_13, x_14, x_114, x_110, x_116, x_5, x_6, x_7, x_8, x_9, x_101); -lean_dec(x_14); -x_118 = lean_ctor_get(x_117, 0); -lean_inc(x_118); -x_119 = lean_ctor_get(x_118, 1); -lean_inc(x_119); -x_120 = lean_ctor_get(x_117, 1); -lean_inc(x_120); -lean_dec(x_117); -x_121 = lean_ctor_get(x_118, 0); -lean_inc(x_121); -lean_dec(x_118); -x_122 = lean_ctor_get(x_119, 0); -lean_inc(x_122); -if (lean_is_exclusive(x_119)) { - lean_ctor_release(x_119, 0); - lean_ctor_release(x_119, 1); - x_123 = x_119; -} else { - lean_dec_ref(x_119); - x_123 = lean_box(0); -} -lean_inc(x_122); -x_124 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_122, x_5, x_6, x_7, x_8, x_9, x_120); -x_125 = lean_ctor_get(x_124, 1); -lean_inc(x_125); -if (lean_is_exclusive(x_124)) { - lean_ctor_release(x_124, 0); - lean_ctor_release(x_124, 1); - x_126 = x_124; -} else { - lean_dec_ref(x_124); - x_126 = lean_box(0); -} -if (lean_is_scalar(x_123)) { - x_127 = lean_alloc_ctor(0, 2, 0); -} else { - x_127 = x_123; -} -lean_ctor_set(x_127, 0, x_122); -lean_ctor_set(x_127, 1, x_121); -if (lean_is_scalar(x_126)) { - x_128 = lean_alloc_ctor(0, 2, 0); -} else { - x_128 = x_126; -} -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_125); -return x_128; -} -else -{ -uint8_t x_129; -x_129 = lean_nat_dec_le(x_109, x_109); -if (x_129 == 0) -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -lean_dec(x_109); -lean_dec(x_108); -x_130 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_113); -lean_ctor_set(x_131, 1, x_130); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -x_133 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_3, x_13, x_14, x_130, x_110, x_132, x_5, x_6, x_7, x_8, x_9, x_101); -lean_dec(x_14); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -x_136 = lean_ctor_get(x_133, 1); -lean_inc(x_136); -lean_dec(x_133); -x_137 = lean_ctor_get(x_134, 0); -lean_inc(x_137); -lean_dec(x_134); -x_138 = lean_ctor_get(x_135, 0); -lean_inc(x_138); -if (lean_is_exclusive(x_135)) { - lean_ctor_release(x_135, 0); - lean_ctor_release(x_135, 1); - x_139 = x_135; -} else { - lean_dec_ref(x_135); - x_139 = lean_box(0); -} -lean_inc(x_138); -x_140 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_138, x_5, x_6, x_7, x_8, x_9, x_136); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_142 = x_140; -} else { - lean_dec_ref(x_140); - x_142 = lean_box(0); -} -if (lean_is_scalar(x_139)) { - x_143 = lean_alloc_ctor(0, 2, 0); -} else { - x_143 = x_139; -} -lean_ctor_set(x_143, 0, x_138); -lean_ctor_set(x_143, 1, x_137); -if (lean_is_scalar(x_142)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_142; -} -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_141); -return x_144; -} -else -{ -size_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_145 = lean_usize_of_nat(x_109); -lean_dec(x_109); -x_146 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_147 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_108, x_13, x_145, x_146); -lean_dec(x_108); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_113); -lean_ctor_set(x_148, 1, x_147); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_146); -lean_ctor_set(x_149, 1, x_148); -x_150 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_3, x_13, x_14, x_146, x_110, x_149, x_5, x_6, x_7, x_8, x_9, x_101); -lean_dec(x_14); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_151, 1); -lean_inc(x_152); -x_153 = lean_ctor_get(x_150, 1); -lean_inc(x_153); -lean_dec(x_150); -x_154 = lean_ctor_get(x_151, 0); -lean_inc(x_154); -lean_dec(x_151); -x_155 = lean_ctor_get(x_152, 0); -lean_inc(x_155); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - x_156 = x_152; -} else { - lean_dec_ref(x_152); - x_156 = lean_box(0); -} -lean_inc(x_155); -x_157 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_155, x_5, x_6, x_7, x_8, x_9, x_153); -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_159 = x_157; -} else { - lean_dec_ref(x_157); - x_159 = lean_box(0); -} -if (lean_is_scalar(x_156)) { - x_160 = lean_alloc_ctor(0, 2, 0); -} else { - x_160 = x_156; -} -lean_ctor_set(x_160, 0, x_155); -lean_ctor_set(x_160, 1, x_154); -if (lean_is_scalar(x_159)) { - x_161 = lean_alloc_ctor(0, 2, 0); -} else { - x_161 = x_159; -} -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_158); -return x_161; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(x_13, x_1); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_free_object(x_9); -x_15 = lean_box(0); -x_16 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_13, x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_12); -lean_dec(x_13); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_13); -lean_dec(x_1); -x_17 = lean_ctor_get(x_14, 0); -lean_inc(x_17); -lean_dec(x_14); -x_18 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_9, 0, x_19); -return x_9; -} -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_9, 0); -x_21 = lean_ctor_get(x_9, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_9); -x_22 = lean_ctor_get(x_20, 0); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(x_22, x_1); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); -x_25 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_22, x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_7, x_21); -lean_dec(x_22); -return x_25; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_22); -lean_dec(x_1); -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_27 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_21); -return x_29; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; lean_object* x_9; uint8_t x_10; -x_8 = 1; -x_9 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -lean_ctor_set(x_9, 0, x_12); -return x_9; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_9); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; -} -} -} -LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; uint8_t x_9; -lean_inc(x_1); -x_8 = l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Array_erase___at_Lean_getAllParentStructures___spec__1(x_10, x_1); -lean_dec(x_1); -lean_ctor_set(x_8, 0, x_11); -return x_8; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_8, 0); -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_8); -x_14 = l_Array_erase___at_Lean_getAllParentStructures___spec__1(x_12, x_1); -lean_dec(x_1); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_lt(x_5, x_4); -if (x_13 == 0) -{ -lean_object* x_14; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_6); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; -lean_dec(x_6); -x_15 = lean_array_uget(x_3, x_5); -x_16 = lean_st_ref_take(x_7, x_12); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_box(0); -x_20 = l_Lean_RBNode_insert___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(x_17, x_15, x_19); -x_21 = lean_st_ref_set(x_7, x_20, x_18); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = 1; -x_24 = lean_usize_add(x_5, x_23); -x_5 = x_24; -x_6 = x_19; -x_12 = x_22; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_11 = !lean_is_exclusive(x_9); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 0); -lean_dec(x_12); -x_13 = lean_box(0); -lean_ctor_set(x_9, 0, x_13); -return x_9; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_dec(x_9); -x_15 = lean_box(0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; -} -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_9, 1); -lean_inc(x_17); -lean_dec(x_9); -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_18, x_3, x_4, x_5, x_6, x_7, x_17); -return x_19; -} -} -else -{ -uint8_t x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_20 = !lean_is_exclusive(x_9); -if (x_20 == 0) -{ -return x_9; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_9, 0); -x_22 = lean_ctor_get(x_9, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_9); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_Expr_getAppFn(x_1); -if (lean_obj_tag(x_8) == 4) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -lean_dec(x_8); -x_10 = lean_st_ref_take(x_2, x_7); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_box(0); -lean_inc(x_9); -x_14 = l_Lean_RBNode_insert___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(x_11, x_9, x_13); -x_15 = lean_st_ref_set(x_2, x_14, x_12); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_st_ref_get(x_6, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -lean_inc(x_9); -x_21 = l_Lean_isStructure(x_20, x_9); -lean_dec(x_20); -if (x_21 == 0) -{ -lean_object* x_22; -lean_dec(x_9); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_19); -return x_22; -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_23 = l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(x_9, x_2, x_3, x_4, x_5, x_6, x_19); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_box(0); -x_27 = lean_array_size(x_24); -x_28 = 0; -x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(x_24, x_26, x_24, x_27, x_28, x_13, x_2, x_3, x_4, x_5, x_6, x_25); -lean_dec(x_24); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -lean_dec(x_29); -x_31 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_30); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_32 = lean_box(0); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_7); -return x_33; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_11 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; -lean_object* x_19 = _args[18]; -_start: -{ -lean_object* x_20; -x_20 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -return x_20; -} -} -LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -_start: -{ -lean_object* x_18; -x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -size_t x_7; size_t x_8; size_t x_9; lean_object* x_10; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_9 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(x_7, x_2, x_3, x_8, x_9, x_6); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(x_1, x_2, x_5, x_6); -lean_dec(x_2); -lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; lean_object* x_16; -x_15 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_16 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -return x_16; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox(x_1); -lean_dec(x_1); -x_14 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -return x_15; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox(x_1); -lean_dec(x_1); -x_14 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox(x_1); -lean_dec(x_1); -x_14 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_3); -lean_dec(x_3); -x_12 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_2); -lean_dec(x_2); -x_10 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -size_t x_13; size_t x_14; lean_object* x_15; -x_13 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_14 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_15 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(x_1, x_2, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_15; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_box(0); -x_8 = lean_st_mk_ref(x_7, x_6); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_1, x_9, x_2, x_3, x_4, x_5, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_9, x_12); -lean_dec(x_9); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -return x_13; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -uint8_t x_18; -lean_dec(x_9); -x_18 = !lean_is_exclusive(x_11); -if (x_18 == 0) -{ -return x_11; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -lean_dec(x_2); -x_3 = lean_box(0); -return x_3; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 2); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 3); -lean_inc(x_7); -lean_dec(x_1); -lean_inc(x_5); -lean_inc(x_2); -x_8 = l_Lean_Server_Completion_cmpModPrivate(x_2, x_5); -switch (x_8) { -case 0: -{ -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_1 = x_4; -goto _start; -} -case 1: -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_2); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_5); -lean_ctor_set(x_10, 1, x_6); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -default: -{ -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_1 = x_7; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; -x_9 = l_Lean_Expr_hasMVar(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_1); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_12 = lean_st_ref_get(x_5, x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_instantiateMVarsCore(x_15, x_1); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_st_ref_take(x_5, x_14); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = !lean_is_exclusive(x_20); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_20, 0); -lean_dec(x_23); -lean_ctor_set(x_20, 0, x_18); -x_24 = lean_st_ref_set(x_5, x_20, x_21); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_27, 0, x_17); -lean_ctor_set(x_24, 0, x_27); -return x_24; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_17); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_31 = lean_ctor_get(x_20, 1); -x_32 = lean_ctor_get(x_20, 2); -x_33 = lean_ctor_get(x_20, 3); -x_34 = lean_ctor_get(x_20, 4); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_20); -x_35 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_35, 0, x_18); -lean_ctor_set(x_35, 1, x_31); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_35, 3, x_33); -lean_ctor_set(x_35, 4, x_34); -x_36 = lean_st_ref_set(x_5, x_35, x_21); -x_37 = lean_ctor_get(x_36, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_38 = x_36; -} else { - lean_dec_ref(x_36); - x_38 = lean_box(0); -} -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_17); -if (lean_is_scalar(x_38)) { - x_40 = lean_alloc_ctor(0, 2, 0); -} else { - x_40 = x_38; -} -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; -} -} -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_10 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_apply_8(x_2, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -return x_13; -} -else -{ -uint8_t x_14; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) -{ -return x_10; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_10, 0); -x_16 = lean_ctor_get(x_10, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_10); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -} -} -LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_10, 0, x_14); -return x_10; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -} -else -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -lean_dec(x_10); -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; double x_25; uint8_t x_26; lean_object* x_27; -x_20 = lean_ctor_get(x_11, 0); -x_21 = l_Lean_ConstantInfo_name(x_1); -x_22 = l_Lean_Name_getString_x21(x_21); -x_23 = lean_box(0); -x_24 = l_Lean_Name_str___override(x_23, x_22); -lean_ctor_set_tag(x_11, 0); -lean_ctor_set(x_11, 0, x_21); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_26 = lean_unbox(x_20); -lean_dec(x_20); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_24, x_11, x_26, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; uint8_t x_35; lean_object* x_36; -x_28 = lean_ctor_get(x_11, 0); -lean_inc(x_28); -lean_dec(x_11); -x_29 = l_Lean_ConstantInfo_name(x_1); -x_30 = l_Lean_Name_getString_x21(x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Name_str___override(x_31, x_30); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_29); -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_35 = lean_unbox(x_28); -lean_dec(x_28); -x_36 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_32, x_33, x_35, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_36; -} -} -} -else -{ -uint8_t x_37; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_37 = !lean_is_exclusive(x_10); -if (x_37 == 0) -{ -return x_10; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_10, 0); -x_39 = lean_ctor_get(x_10, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_10); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_6, x_7, x_8, x_9, x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -uint8_t x_13; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_13 = !lean_is_exclusive(x_11); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 0); -lean_dec(x_14); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_11, 0, x_15); -return x_11; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_11, 1); -lean_inc(x_19); -lean_dec(x_11); -x_20 = lean_ctor_get(x_12, 0); -lean_inc(x_20); -lean_dec(x_12); -x_21 = l_Lean_Name_getPrefix(x_20); -lean_dec(x_20); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod(x_21, x_2, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_unbox(x_23); -lean_dec(x_23); -if (x_24 == 0) -{ -uint8_t x_25; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_25 = !lean_is_exclusive(x_22); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_22, 0); -lean_dec(x_26); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_22, 0, x_27); -return x_22; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -lean_dec(x_22); -x_29 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_28); -return x_30; -} -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_22, 1); -lean_inc(x_31); -lean_dec(x_22); -x_32 = lean_box(0); -x_33 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); -return x_33; -} -} -else -{ -uint8_t x_34; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_34 = !lean_is_exclusive(x_22); -if (x_34 == 0) -{ -return x_22; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_22, 0); -x_36 = lean_ctor_get(x_22, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_22); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_Name_getPrefix(x_2); -x_12 = l_Lean_RBNode_findCore___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__1(x_1, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -x_15 = lean_box(0); -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3(x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_16; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4___boxed), 10, 1); -lean_closure_set(x_10, 0, x_1); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1___boxed), 8, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Parser", 6, 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Term", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("completion", 10, 10); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1; -x_15 = l_Lean_Syntax_isIdent(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -lean_dec(x_3); -lean_inc(x_13); -x_16 = l_Lean_Syntax_getKind(x_13); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5; -x_18 = lean_name_eq(x_16, x_17); -lean_dec(x_16); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_11); -return x_20; -} -else -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Lean_Syntax_getArg(x_13, x_21); -x_23 = l_Lean_Syntax_isIdent(x_22); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -lean_dec(x_22); -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_2); -x_24 = lean_box(0); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_11); -return x_25; -} -else -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; -x_26 = l_Lean_Syntax_getId(x_22); -lean_dec(x_22); -x_27 = lean_box(0); -x_28 = 1; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_29 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore(x_2, x_13, x_26, x_27, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_31 = !lean_is_exclusive(x_29); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_29, 0); -lean_dec(x_32); -x_33 = lean_box(0); -lean_ctor_set(x_29, 0, x_33); -return x_29; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_29, 1); -lean_inc(x_34); -lean_dec(x_29); -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; -} -} -else -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_29, 1); -lean_inc(x_37); -lean_dec(x_29); -x_38 = lean_ctor_get(x_30, 0); -lean_inc(x_38); -lean_dec(x_30); -x_39 = lean_apply_8(x_14, x_38, x_5, x_6, x_7, x_8, x_9, x_10, x_37); -return x_39; -} -} -else -{ -uint8_t x_40; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_40 = !lean_is_exclusive(x_29); -if (x_40 == 0) -{ -return x_29; -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_29, 0); -x_42 = lean_ctor_get(x_29, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_29); -x_43 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -return x_43; -} -} -} -} -} -else -{ -lean_object* x_44; uint8_t x_45; lean_object* x_46; -x_44 = l_Lean_Syntax_getId(x_13); -x_45 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_46 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore(x_2, x_13, x_44, x_3, x_45, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_13); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -if (lean_obj_tag(x_47) == 0) -{ -uint8_t x_48; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_48 = !lean_is_exclusive(x_46); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 0); -lean_dec(x_49); -x_50 = lean_box(0); -lean_ctor_set(x_46, 0, x_50); -return x_46; -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_box(0); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_51); -return x_53; -} -} -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_46, 1); -lean_inc(x_54); -lean_dec(x_46); -x_55 = lean_ctor_get(x_47, 0); -lean_inc(x_55); -lean_dec(x_47); -x_56 = lean_apply_8(x_14, x_55, x_5, x_6, x_7, x_8, x_9, x_10, x_54); -return x_56; -} -} -else -{ -uint8_t x_57; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_57 = !lean_is_exclusive(x_46); -if (x_57 == 0) -{ -return x_46; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_46, 0); -x_59 = lean_ctor_get(x_46, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_46); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} -} -} -} -else -{ -lean_object* x_61; lean_object* x_62; -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_61 = lean_box(0); -x_62 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5(x_4, x_61, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_62; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_18; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_18 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_20); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = !lean_is_exclusive(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 0); -x_26 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames(x_25, x_4, x_5, x_6, x_7, x_23); -if (lean_obj_tag(x_26) == 0) -{ -uint8_t x_27; -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_26, 0); -lean_ctor_set(x_22, 0, x_28); -lean_ctor_set(x_26, 0, x_22); -return x_26; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_26, 0); -x_30 = lean_ctor_get(x_26, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_26); -lean_ctor_set(x_22, 0, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_22); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; -lean_free_object(x_22); -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_26, 1); -lean_inc(x_33); -lean_dec(x_26); -x_9 = x_32; -x_10 = x_33; -goto block_17; -} -} -else -{ -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_22, 0); -lean_inc(x_34); -lean_dec(x_22); -x_35 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames(x_34, x_4, x_5, x_6, x_7, x_23); -if (lean_obj_tag(x_35) == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); -lean_inc(x_37); -if (lean_is_exclusive(x_35)) { - lean_ctor_release(x_35, 0); - lean_ctor_release(x_35, 1); - x_38 = x_35; -} else { - lean_dec_ref(x_35); - x_38 = lean_box(0); -} -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_36); -if (lean_is_scalar(x_38)) { - x_40 = lean_alloc_ctor(0, 2, 0); -} else { - x_40 = x_38; -} -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_37); -return x_40; -} -else -{ -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_35, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_35, 1); -lean_inc(x_42); -lean_dec(x_35); -x_9 = x_41; -x_10 = x_42; -goto block_17; -} -} -} -else -{ -lean_object* x_43; lean_object* x_44; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_43 = lean_ctor_get(x_18, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_18, 1); -lean_inc(x_44); -lean_dec(x_18); -x_9 = x_43; -x_10 = x_44; -goto block_17; -} -block_17: -{ -uint8_t x_11; -x_11 = l_Lean_Exception_isInterrupt(x_9); -if (x_11 == 0) -{ -uint8_t x_12; -x_12 = l_Lean_Exception_isRuntime(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_9); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1; -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_10); -return x_14; -} -else -{ -lean_object* x_15; -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_9); -lean_ctor_set(x_15, 1, x_10); -return x_15; -} -} -else -{ -lean_object* x_16; -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_9); -lean_ctor_set(x_16, 1, x_10); -return x_16; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -lean_dec(x_2); -x_13 = lean_apply_8(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_13; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_6 = lean_ctor_get(x_3, 1); -lean_inc(x_6); -lean_inc(x_2); -lean_inc(x_3); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6), 11, 3); -lean_closure_set(x_7, 0, x_3); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_4); -x_8 = lean_ctor_get(x_3, 3); -lean_inc(x_8); -lean_dec(x_3); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___boxed), 8, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__8), 9, 1); -lean_closure_set(x_10, 0, x_7); -x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); -lean_closure_set(x_11, 0, x_9); -lean_closure_set(x_11, 1, x_10); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(x_1, x_2, x_6, x_11, x_5); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_3); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_8); -return x_9; -} -} -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed), 8, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; -x_14 = 0; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_14, x_3, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -if (lean_obj_tag(x_16) == 0) -{ -uint8_t x_17; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_15, 0); -lean_dec(x_18); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_15, 0, x_19); -return x_15; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_dec(x_15); -x_21 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_16); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; double x_29; lean_object* x_30; -x_24 = lean_ctor_get(x_16, 0); -x_25 = lean_ctor_get(x_15, 1); -lean_inc(x_25); -lean_dec(x_15); -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = l_Lean_ConstantInfo_name(x_4); -lean_ctor_set_tag(x_16, 0); -lean_ctor_set(x_16, 0, x_28); -x_29 = lean_unbox_float(x_27); -lean_dec(x_27); -x_30 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_26, x_16, x_5, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_25); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -return x_30; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; double x_37; lean_object* x_38; -x_31 = lean_ctor_get(x_16, 0); -lean_inc(x_31); -lean_dec(x_16); -x_32 = lean_ctor_get(x_15, 1); -lean_inc(x_32); -lean_dec(x_15); -x_33 = lean_ctor_get(x_31, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 1); -lean_inc(x_34); -lean_dec(x_31); -x_35 = l_Lean_ConstantInfo_name(x_4); -x_36 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_unbox_float(x_34); -lean_dec(x_34); -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_33, x_36, x_5, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_32); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -return x_38; -} -} -} -else -{ -uint8_t x_39; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -x_39 = !lean_is_exclusive(x_15); -if (x_39 == 0) -{ -return x_15; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_15, 0); -x_41 = lean_ctor_get(x_15, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_15); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -x_15 = !lean_is_exclusive(x_13); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 0); -lean_dec(x_16); -x_17 = lean_box(0); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -return x_20; -} -} -else -{ -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_dec(x_13); -x_22 = !lean_is_exclusive(x_14); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_14, 0); -x_24 = l_Lean_Name_isAnonymous(x_3); -if (x_24 == 0) -{ -lean_object* x_25; uint8_t x_26; lean_object* x_27; -lean_free_object(x_14); -x_25 = lean_box(0); -x_26 = lean_unbox(x_23); -lean_dec(x_23); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1(x_2, x_3, x_4, x_1, x_26, x_25, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -return x_27; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; double x_32; uint8_t x_33; lean_object* x_34; uint8_t x_35; -lean_dec(x_4); -lean_dec(x_3); -x_28 = l_Lean_ConstantInfo_name(x_1); -x_29 = l_Lean_Name_getString_x21(x_28); -x_30 = lean_box(0); -x_31 = l_Lean_Name_str___override(x_30, x_29); -lean_ctor_set_tag(x_14, 0); -lean_ctor_set(x_14, 0, x_28); -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_33 = lean_unbox(x_23); -lean_dec(x_23); -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_31, x_14, x_33, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -x_37 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_34, 0, x_37); -return x_34; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_34, 1); -lean_inc(x_38); -lean_dec(x_34); -x_39 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -return x_40; -} -} -} -else -{ -lean_object* x_41; uint8_t x_42; -x_41 = lean_ctor_get(x_14, 0); -lean_inc(x_41); -lean_dec(x_14); -x_42 = l_Lean_Name_isAnonymous(x_3); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; lean_object* x_45; -x_43 = lean_box(0); -x_44 = lean_unbox(x_41); -lean_dec(x_41); -x_45 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1(x_2, x_3, x_4, x_1, x_44, x_43, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -return x_45; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; double x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_4); -lean_dec(x_3); -x_46 = l_Lean_ConstantInfo_name(x_1); -x_47 = l_Lean_Name_getString_x21(x_46); -x_48 = lean_box(0); -x_49 = l_Lean_Name_str___override(x_48, x_47); -x_50 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_50, 0, x_46); -x_51 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_52 = lean_unbox(x_41); -lean_dec(x_41); -x_53 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_49, x_50, x_52, x_51, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -if (lean_is_exclusive(x_53)) { - lean_ctor_release(x_53, 0); - lean_ctor_release(x_53, 1); - x_55 = x_53; -} else { - lean_dec_ref(x_53); - x_55 = lean_box(0); -} -x_56 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); -} else { - x_57 = x_55; -} -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; -} -} -} -} -else -{ -uint8_t x_58; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -x_58 = !lean_is_exclusive(x_13); -if (x_58 == 0) -{ -return x_13; -} -else -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_13, 0); -x_60 = lean_ctor_get(x_13, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_13); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_12, 0); -lean_dec(x_15); -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_12, 0, x_16); -return x_12; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); -lean_dec(x_12); -x_21 = lean_ctor_get(x_13, 0); -lean_inc(x_21); -lean_dec(x_13); -x_22 = l_Lean_Name_getPrefix(x_21); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_22); -x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotIdCompletionMethod(x_22, x_2, x_7, x_8, x_9, x_10, x_20); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; uint8_t x_25; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_unbox(x_24); -lean_dec(x_24); -if (x_25 == 0) -{ -uint8_t x_26; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_23, 0); -lean_dec(x_27); -x_28 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_23, 0, x_28); -return x_23; -} -else -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); -lean_dec(x_23); -x_30 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; -} -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_23, 1); -lean_inc(x_32); -lean_dec(x_23); -x_33 = lean_box(0); -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2(x_2, x_22, x_3, x_21, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_32); -lean_dec(x_22); -return x_34; -} -} -else -{ -uint8_t x_35; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -x_35 = !lean_is_exclusive(x_23); -if (x_35 == 0) -{ -return x_23; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_23, 0); -x_37 = lean_ctor_get(x_23, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_23); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_Name_getPrefix(x_3); -x_13 = l_Lean_RBNode_findCore___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__1(x_2, x_12); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_1); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_11); -return x_15; -} -else -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_13); -x_16 = lean_box(0); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3(x_3, x_4, x_1, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_17; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4___boxed), 11, 2); -lean_closure_set(x_10, 0, x_1); -lean_closure_set(x_10, 1, x_2); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_10; lean_object* x_11; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -lean_dec(x_2); -x_13 = l_Lean_Expr_cleanupAnnotations(x_12); -x_14 = l_Lean_Expr_getAppFn(x_13); -lean_dec(x_13); -x_15 = l_Lean_Expr_cleanupAnnotations(x_14); -if (lean_obj_tag(x_15) == 4) -{ -lean_object* x_16; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames(x_15, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -return x_19; -} -else -{ -uint8_t x_20; -x_20 = !lean_is_exclusive(x_16); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_16, 0); -x_22 = lean_ctor_get(x_16, 1); -x_23 = l_Lean_Exception_isInterrupt(x_21); -if (x_23 == 0) -{ -uint8_t x_24; -x_24 = l_Lean_Exception_isRuntime(x_21); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; -lean_free_object(x_16); -lean_dec(x_21); -x_25 = lean_box(0); -x_26 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_22); -return x_26; -} -else -{ -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_16; -} -} -else -{ -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_16; -} -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_16, 0); -x_28 = lean_ctor_get(x_16, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_16); -x_29 = l_Lean_Exception_isInterrupt(x_27); -if (x_29 == 0) -{ -uint8_t x_30; -x_30 = l_Lean_Exception_isRuntime(x_27); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; -lean_dec(x_27); -x_31 = lean_box(0); -x_32 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_31, x_3, x_4, x_5, x_6, x_7, x_8, x_28); -return x_32; -} -else -{ -lean_object* x_33; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_27); -lean_ctor_set(x_33, 1, x_28); -return x_33; -} -} -else -{ -lean_object* x_34; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_27); -lean_ctor_set(x_34, 1, x_28); -return x_34; -} -} -} -} -else -{ -lean_object* x_35; lean_object* x_36; -lean_dec(x_15); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_35 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_9); -return x_36; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_dec(x_4); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_8 = lean_alloc_closure((void*)(l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed), 8, 1); -lean_closure_set(x_8, 0, x_7); -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_8, x_6); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_ctor_get(x_5, 0); -lean_inc(x_10); -lean_dec(x_5); -x_11 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__2___boxed), 8, 1); -lean_closure_set(x_11, 0, x_10); -x_12 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__6), 9, 1); -lean_closure_set(x_12, 0, x_4); -x_13 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_12); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_13, x_6); -return x_14; -} -} -} -LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_ReaderT_pure___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_5); -lean_dec(x_5); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__1(x_1, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_1); -return x_15; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -lean_object* x_13; -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_2); -return x_12; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_6); -lean_dec(x_4); -return x_12; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("field", 5, 5); -return x_1; -} -} -static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; -x_15 = lean_usize_dec_lt(x_6, x_5); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -lean_dec(x_8); -lean_dec(x_1); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_7); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_14); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_7); -x_18 = lean_array_uget(x_4, x_6); -if (lean_obj_tag(x_18) == 1) -{ -lean_object* x_27; double x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_18, 1); -lean_inc(x_27); -lean_dec(x_18); -x_28 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -lean_inc(x_1); -x_29 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_27, x_28); -if (lean_obj_tag(x_29) == 0) -{ -lean_object* x_30; -lean_dec(x_27); -x_30 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_19 = x_30; -x_20 = x_14; -goto block_26; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_31 = lean_ctor_get(x_29, 0); -lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_box(0); -x_33 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2; -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3; -x_35 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_35, 0, x_27); -lean_ctor_set(x_35, 1, x_33); -lean_ctor_set(x_35, 2, x_32); -lean_ctor_set(x_35, 3, x_34); -lean_ctor_set(x_35, 4, x_32); -lean_ctor_set(x_35, 5, x_32); -lean_ctor_set(x_35, 6, x_32); -lean_ctor_set(x_35, 7, x_32); -x_36 = lean_unbox_float(x_31); -lean_dec(x_31); -lean_inc(x_8); -x_37 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem(x_35, x_36, x_32, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_19 = x_39; -x_20 = x_38; -goto block_26; -} -} -else -{ -lean_object* x_40; -lean_dec(x_18); -x_40 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1; -x_19 = x_40; -x_20 = x_14; -goto block_26; -} -block_26: -{ -lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec(x_21); -x_23 = 1; -x_24 = lean_usize_add(x_6, x_23); -x_6 = x_24; -x_7 = x_22; -x_14 = x_20; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_st_ref_get(x_6, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -lean_ctor_set(x_8, 0, x_11); -return x_8; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_8, 0); -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_8); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -else -{ -lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_13 = lean_ctor_get(x_3, 0); -x_14 = 0; -x_15 = l_Lean_getStructureFieldsFlattened(x_13, x_1, x_14); -x_16 = lean_box(0); -x_17 = lean_array_size(x_15); -x_18 = 0; -x_19 = lean_box(0); -x_20 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1(x_2, x_15, x_16, x_15, x_17, x_18, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_15); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -lean_ctor_set(x_20, 0, x_23); -return x_20; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 1); -lean_inc(x_24); -lean_dec(x_20); -x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1; -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed), 7, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed), 8, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2; -x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_7 = 1; -x_8 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -x_9 = l_Lean_Name_toString(x_4, x_7, x_8); -x_10 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3___boxed), 10, 2); -lean_closure_set(x_10, 0, x_5); -lean_closure_set(x_10, 1, x_9); -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3; -x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); -lean_closure_set(x_12, 0, x_11); -lean_closure_set(x_12, 1, x_10); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_12, x_6); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -size_t x_15; size_t x_16; lean_object* x_17; -x_15 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_16 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_17; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -return x_11; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 9; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("(", 1, 1); -return x_1; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("), ", 3, 3); -return x_1; -} -} -static lean_object* _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4() { -_start: -{ -uint32_t x_1; lean_object* x_2; -x_1 = 32; -x_2 = l_Char_utf8Size(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_9); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; -x_17 = lean_ctor_get(x_8, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_8, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_8, 2); -lean_inc(x_19); -x_20 = lean_ctor_get(x_8, 3); -lean_inc(x_20); -lean_dec(x_8); -lean_inc(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_21 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_17, x_9, x_10, x_11, x_12, x_13, x_14); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_24 = x_21; -} else { - lean_dec_ref(x_21); - x_24 = lean_box(0); -} -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -if (lean_is_exclusive(x_22)) { - lean_ctor_release(x_22, 0); - x_26 = x_22; -} else { - lean_dec_ref(x_22); - x_26 = lean_box(0); -} -x_33 = 1; -x_34 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1; -lean_inc(x_18); -x_35 = l_Lean_Name_toString(x_18, x_33, x_34); -x_36 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1; -lean_inc(x_5); -x_37 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_5, x_35, x_36); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; -lean_dec(x_35); -lean_dec(x_19); -lean_dec(x_18); -if (lean_is_scalar(x_26)) { - x_38 = lean_alloc_ctor(1, 1, 0); -} else { - x_38 = x_26; -} -lean_ctor_set(x_38, 0, x_25); -if (lean_is_scalar(x_24)) { - x_39 = lean_alloc_ctor(0, 2, 0); -} else { - x_39 = x_24; -} -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_23); -x_27 = x_39; -goto block_32; -} -else -{ -uint8_t x_40; -x_40 = !lean_is_exclusive(x_37); -if (x_40 == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_98; lean_object* x_115; lean_object* x_116; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_ctor_get(x_3, 0); -x_43 = lean_ctor_get(x_19, 1); -lean_inc(x_43); -x_44 = l_Lean_KVMap_findCore(x_7, x_18); -lean_dec(x_18); -x_45 = lean_ctor_get(x_19, 3); -lean_inc(x_45); -lean_dec(x_19); -x_46 = lean_box(0); -lean_inc(x_1); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_1); -lean_ctor_set(x_47, 1, x_46); -x_48 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(x_47); -lean_ctor_set(x_37, 0, x_48); -if (lean_obj_tag(x_42) == 0) -{ -lean_object* x_133; -x_133 = lean_box(0); -x_49 = x_133; -goto block_97; -} -else -{ -lean_object* x_134; lean_object* x_135; -x_134 = lean_ctor_get(x_42, 0); -x_135 = lean_ctor_get(x_134, 0); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; -x_136 = lean_box(0); -x_49 = x_136; -goto block_97; -} -else -{ -lean_object* x_137; -x_137 = lean_ctor_get(x_135, 0); -if (lean_obj_tag(x_137) == 0) -{ -lean_object* x_138; -x_138 = lean_box(0); -x_49 = x_138; -goto block_97; -} -else -{ -lean_object* x_139; -x_139 = lean_ctor_get(x_137, 0); -if (lean_obj_tag(x_139) == 0) -{ -lean_object* x_140; -x_140 = lean_box(0); -x_49 = x_140; -goto block_97; -} -else -{ -lean_object* x_141; uint8_t x_142; -x_141 = lean_ctor_get(x_139, 0); -x_142 = lean_unbox(x_141); -if (x_142 == 0) -{ -lean_object* x_143; -x_143 = lean_box(0); -x_49 = x_143; -goto block_97; -} -else -{ -uint8_t x_144; lean_object* x_145; -lean_dec(x_26); -lean_dec(x_24); -x_144 = 0; -x_145 = l_Lean_Syntax_getRange_x3f(x_4, x_144); -if (lean_obj_tag(x_145) == 0) -{ -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_46; -goto block_114; -} -else -{ -lean_object* x_146; -lean_dec(x_43); -x_146 = lean_ctor_get(x_44, 0); -lean_inc(x_146); -lean_dec(x_44); -x_115 = x_46; -x_116 = x_146; -goto block_132; -} -} -else -{ -uint8_t x_147; -x_147 = !lean_is_exclusive(x_145); -if (x_147 == 0) -{ -lean_object* x_148; uint8_t x_149; -x_148 = lean_ctor_get(x_145, 0); -x_149 = !lean_is_exclusive(x_148); -if (x_149 == 0) -{ -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_150 = lean_ctor_get(x_148, 0); -x_151 = lean_ctor_get(x_148, 1); -x_152 = lean_ctor_get(x_2, 0); -lean_inc(x_152); -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -lean_dec(x_152); -lean_inc(x_153); -x_154 = l_Lean_FileMap_utf8PosToLspPos(x_153, x_150); -lean_dec(x_150); -if (x_6 == 0) -{ -lean_object* x_155; lean_object* x_156; -x_155 = l_Lean_FileMap_utf8PosToLspPos(x_153, x_151); -lean_dec(x_151); -lean_ctor_set(x_148, 1, x_155); -lean_ctor_set(x_148, 0, x_154); -lean_inc(x_148); -lean_inc(x_35); -x_156 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_156, 0, x_35); -lean_ctor_set(x_156, 1, x_148); -lean_ctor_set(x_156, 2, x_148); -lean_ctor_set(x_145, 0, x_156); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_145; -goto block_114; -} -else -{ -lean_object* x_157; -lean_dec(x_43); -x_157 = lean_ctor_get(x_44, 0); -lean_inc(x_157); -lean_dec(x_44); -x_115 = x_145; -x_116 = x_157; -goto block_132; -} -} -else -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_158 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4; -x_159 = lean_nat_add(x_151, x_158); -lean_dec(x_151); -x_160 = l_Lean_FileMap_utf8PosToLspPos(x_153, x_159); -lean_dec(x_159); -lean_ctor_set(x_148, 1, x_160); -lean_ctor_set(x_148, 0, x_154); -lean_inc(x_148); -lean_inc(x_35); -x_161 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_161, 0, x_35); -lean_ctor_set(x_161, 1, x_148); -lean_ctor_set(x_161, 2, x_148); -lean_ctor_set(x_145, 0, x_161); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_145; -goto block_114; -} -else -{ -lean_object* x_162; -lean_dec(x_43); -x_162 = lean_ctor_get(x_44, 0); -lean_inc(x_162); -lean_dec(x_44); -x_115 = x_145; -x_116 = x_162; -goto block_132; -} -} -} -else -{ -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_163 = lean_ctor_get(x_148, 0); -x_164 = lean_ctor_get(x_148, 1); -lean_inc(x_164); -lean_inc(x_163); -lean_dec(x_148); -x_165 = lean_ctor_get(x_2, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_165, 1); -lean_inc(x_166); -lean_dec(x_165); -lean_inc(x_166); -x_167 = l_Lean_FileMap_utf8PosToLspPos(x_166, x_163); -lean_dec(x_163); -if (x_6 == 0) -{ -lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_168 = l_Lean_FileMap_utf8PosToLspPos(x_166, x_164); -lean_dec(x_164); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_167); -lean_ctor_set(x_169, 1, x_168); -lean_inc(x_169); -lean_inc(x_35); -x_170 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_170, 0, x_35); -lean_ctor_set(x_170, 1, x_169); -lean_ctor_set(x_170, 2, x_169); -lean_ctor_set(x_145, 0, x_170); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_145; -goto block_114; -} -else -{ -lean_object* x_171; -lean_dec(x_43); -x_171 = lean_ctor_get(x_44, 0); -lean_inc(x_171); -lean_dec(x_44); -x_115 = x_145; -x_116 = x_171; -goto block_132; -} -} -else -{ -lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_172 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4; -x_173 = lean_nat_add(x_164, x_172); -lean_dec(x_164); -x_174 = l_Lean_FileMap_utf8PosToLspPos(x_166, x_173); -lean_dec(x_173); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_167); -lean_ctor_set(x_175, 1, x_174); -lean_inc(x_175); -lean_inc(x_35); -x_176 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_176, 0, x_35); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_175); -lean_ctor_set(x_145, 0, x_176); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_145; -goto block_114; -} -else -{ -lean_object* x_177; -lean_dec(x_43); -x_177 = lean_ctor_get(x_44, 0); -lean_inc(x_177); -lean_dec(x_44); -x_115 = x_145; -x_116 = x_177; -goto block_132; -} -} -} -} -else -{ -lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_178 = lean_ctor_get(x_145, 0); -lean_inc(x_178); -lean_dec(x_145); -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_178, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_181 = x_178; -} else { - lean_dec_ref(x_178); - x_181 = lean_box(0); -} -x_182 = lean_ctor_get(x_2, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_182, 1); -lean_inc(x_183); -lean_dec(x_182); -lean_inc(x_183); -x_184 = l_Lean_FileMap_utf8PosToLspPos(x_183, x_179); -lean_dec(x_179); -if (x_6 == 0) -{ -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_185 = l_Lean_FileMap_utf8PosToLspPos(x_183, x_180); -lean_dec(x_180); -if (lean_is_scalar(x_181)) { - x_186 = lean_alloc_ctor(0, 2, 0); -} else { - x_186 = x_181; -} -lean_ctor_set(x_186, 0, x_184); -lean_ctor_set(x_186, 1, x_185); -lean_inc(x_186); -lean_inc(x_35); -x_187 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_187, 0, x_35); -lean_ctor_set(x_187, 1, x_186); -lean_ctor_set(x_187, 2, x_186); -x_188 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_188, 0, x_187); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_188; -goto block_114; -} -else -{ -lean_object* x_189; -lean_dec(x_43); -x_189 = lean_ctor_get(x_44, 0); -lean_inc(x_189); -lean_dec(x_44); -x_115 = x_188; -x_116 = x_189; -goto block_132; -} -} -else -{ -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_190 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4; -x_191 = lean_nat_add(x_180, x_190); -lean_dec(x_180); -x_192 = l_Lean_FileMap_utf8PosToLspPos(x_183, x_191); -lean_dec(x_191); -if (lean_is_scalar(x_181)) { - x_193 = lean_alloc_ctor(0, 2, 0); -} else { - x_193 = x_181; -} -lean_ctor_set(x_193, 0, x_184); -lean_ctor_set(x_193, 1, x_192); -lean_inc(x_193); -lean_inc(x_35); -x_194 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_194, 0, x_35); -lean_ctor_set(x_194, 1, x_193); -lean_ctor_set(x_194, 2, x_193); -x_195 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_195, 0, x_194); -if (lean_obj_tag(x_44) == 0) -{ -x_98 = x_195; -goto block_114; -} -else -{ -lean_object* x_196; -lean_dec(x_43); -x_196 = lean_ctor_get(x_44, 0); -lean_inc(x_196); -lean_dec(x_44); -x_115 = x_195; -x_116 = x_196; -goto block_132; -} -} -} -} -} -} -} -} -} -block_97: -{ -lean_dec(x_49); -if (lean_obj_tag(x_44) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_50 = lean_data_value_to_string(x_43); -x_51 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_52 = lean_string_append(x_51, x_50); -lean_dec(x_50); -x_53 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_54 = lean_string_append(x_52, x_53); -x_55 = lean_string_append(x_54, x_45); -lean_dec(x_45); -x_56 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_57 = lean_string_append(x_55, x_56); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_59 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_60 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_60, 0, x_35); -lean_ctor_set(x_60, 1, x_58); -lean_ctor_set(x_60, 2, x_46); -lean_ctor_set(x_60, 3, x_59); -lean_ctor_set(x_60, 4, x_46); -lean_ctor_set(x_60, 5, x_46); -lean_ctor_set(x_60, 6, x_37); -lean_ctor_set(x_60, 7, x_46); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_41); -x_62 = lean_array_push(x_25, x_61); -if (lean_is_scalar(x_26)) { - x_63 = lean_alloc_ctor(1, 1, 0); -} else { - x_63 = x_26; -} -lean_ctor_set(x_63, 0, x_62); -if (lean_is_scalar(x_24)) { - x_64 = lean_alloc_ctor(0, 2, 0); -} else { - x_64 = x_24; -} -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_23); -x_27 = x_64; -goto block_32; -} -else -{ -uint8_t x_65; -lean_dec(x_43); -x_65 = !lean_is_exclusive(x_44); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_66 = lean_ctor_get(x_44, 0); -x_67 = lean_data_value_to_string(x_66); -x_68 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_69 = lean_string_append(x_68, x_67); -lean_dec(x_67); -x_70 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_71 = lean_string_append(x_69, x_70); -x_72 = lean_string_append(x_71, x_45); -lean_dec(x_45); -x_73 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_74 = lean_string_append(x_72, x_73); -lean_ctor_set(x_44, 0, x_74); -x_75 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_76 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_76, 0, x_35); -lean_ctor_set(x_76, 1, x_44); -lean_ctor_set(x_76, 2, x_46); -lean_ctor_set(x_76, 3, x_75); -lean_ctor_set(x_76, 4, x_46); -lean_ctor_set(x_76, 5, x_46); -lean_ctor_set(x_76, 6, x_37); -lean_ctor_set(x_76, 7, x_46); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_41); -x_78 = lean_array_push(x_25, x_77); -if (lean_is_scalar(x_26)) { - x_79 = lean_alloc_ctor(1, 1, 0); -} else { - x_79 = x_26; -} -lean_ctor_set(x_79, 0, x_78); -if (lean_is_scalar(x_24)) { - x_80 = lean_alloc_ctor(0, 2, 0); -} else { - x_80 = x_24; -} -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_23); -x_27 = x_80; -goto block_32; -} -else -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_81 = lean_ctor_get(x_44, 0); -lean_inc(x_81); -lean_dec(x_44); -x_82 = lean_data_value_to_string(x_81); -x_83 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_84 = lean_string_append(x_83, x_82); -lean_dec(x_82); -x_85 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_86 = lean_string_append(x_84, x_85); -x_87 = lean_string_append(x_86, x_45); -lean_dec(x_45); -x_88 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_89 = lean_string_append(x_87, x_88); -x_90 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_90, 0, x_89); -x_91 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_92 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_92, 0, x_35); -lean_ctor_set(x_92, 1, x_90); -lean_ctor_set(x_92, 2, x_46); -lean_ctor_set(x_92, 3, x_91); -lean_ctor_set(x_92, 4, x_46); -lean_ctor_set(x_92, 5, x_46); -lean_ctor_set(x_92, 6, x_37); -lean_ctor_set(x_92, 7, x_46); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_41); -x_94 = lean_array_push(x_25, x_93); -if (lean_is_scalar(x_26)) { - x_95 = lean_alloc_ctor(1, 1, 0); -} else { - x_95 = x_26; -} -lean_ctor_set(x_95, 0, x_94); -if (lean_is_scalar(x_24)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_24; -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_23); -x_27 = x_96; -goto block_32; -} -} -} -block_114: -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_99 = lean_data_value_to_string(x_43); -x_100 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_101 = lean_string_append(x_100, x_99); -lean_dec(x_99); -x_102 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_103 = lean_string_append(x_101, x_102); -x_104 = lean_string_append(x_103, x_45); -lean_dec(x_45); -x_105 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_106 = lean_string_append(x_104, x_105); -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_106); -x_108 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_109 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_109, 0, x_35); -lean_ctor_set(x_109, 1, x_107); -lean_ctor_set(x_109, 2, x_46); -lean_ctor_set(x_109, 3, x_108); -lean_ctor_set(x_109, 4, x_98); -lean_ctor_set(x_109, 5, x_46); -lean_ctor_set(x_109, 6, x_37); -lean_ctor_set(x_109, 7, x_46); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_41); -x_111 = lean_array_push(x_25, x_110); -x_112 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_112, 0, x_111); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_23); -x_27 = x_113; -goto block_32; -} -block_132: -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_117 = lean_data_value_to_string(x_116); -x_118 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_119 = lean_string_append(x_118, x_117); -lean_dec(x_117); -x_120 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_121 = lean_string_append(x_119, x_120); -x_122 = lean_string_append(x_121, x_45); -lean_dec(x_45); -x_123 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_124 = lean_string_append(x_122, x_123); -x_125 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_125, 0, x_124); -x_126 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_127 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_127, 0, x_35); -lean_ctor_set(x_127, 1, x_125); -lean_ctor_set(x_127, 2, x_46); -lean_ctor_set(x_127, 3, x_126); -lean_ctor_set(x_127, 4, x_115); -lean_ctor_set(x_127, 5, x_46); -lean_ctor_set(x_127, 6, x_37); -lean_ctor_set(x_127, 7, x_46); -x_128 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_128, 0, x_127); -lean_ctor_set(x_128, 1, x_41); -x_129 = lean_array_push(x_25, x_128); -x_130 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_130, 0, x_129); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_23); -x_27 = x_131; -goto block_32; -} -} -else -{ -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_240; lean_object* x_257; lean_object* x_258; -x_197 = lean_ctor_get(x_37, 0); -lean_inc(x_197); -lean_dec(x_37); -x_198 = lean_ctor_get(x_3, 0); -x_199 = lean_ctor_get(x_19, 1); -lean_inc(x_199); -x_200 = l_Lean_KVMap_findCore(x_7, x_18); -lean_dec(x_18); -x_201 = lean_ctor_get(x_19, 3); -lean_inc(x_201); -lean_dec(x_19); -x_202 = lean_box(0); -lean_inc(x_1); -x_203 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_203, 0, x_1); -lean_ctor_set(x_203, 1, x_202); -x_204 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(x_203); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_204); -if (lean_obj_tag(x_198) == 0) -{ -lean_object* x_275; -x_275 = lean_box(0); -x_206 = x_275; -goto block_239; -} -else -{ -lean_object* x_276; lean_object* x_277; -x_276 = lean_ctor_get(x_198, 0); -x_277 = lean_ctor_get(x_276, 0); -if (lean_obj_tag(x_277) == 0) -{ -lean_object* x_278; -x_278 = lean_box(0); -x_206 = x_278; -goto block_239; -} -else -{ -lean_object* x_279; -x_279 = lean_ctor_get(x_277, 0); -if (lean_obj_tag(x_279) == 0) -{ -lean_object* x_280; -x_280 = lean_box(0); -x_206 = x_280; -goto block_239; -} -else -{ -lean_object* x_281; -x_281 = lean_ctor_get(x_279, 0); -if (lean_obj_tag(x_281) == 0) -{ -lean_object* x_282; -x_282 = lean_box(0); -x_206 = x_282; -goto block_239; -} -else -{ -lean_object* x_283; uint8_t x_284; -x_283 = lean_ctor_get(x_281, 0); -x_284 = lean_unbox(x_283); -if (x_284 == 0) -{ -lean_object* x_285; -x_285 = lean_box(0); -x_206 = x_285; -goto block_239; -} -else -{ -uint8_t x_286; lean_object* x_287; -lean_dec(x_26); -lean_dec(x_24); -x_286 = 0; -x_287 = l_Lean_Syntax_getRange_x3f(x_4, x_286); -if (lean_obj_tag(x_287) == 0) -{ -if (lean_obj_tag(x_200) == 0) -{ -x_240 = x_202; -goto block_256; -} -else -{ -lean_object* x_288; -lean_dec(x_199); -x_288 = lean_ctor_get(x_200, 0); -lean_inc(x_288); -lean_dec(x_200); -x_257 = x_202; -x_258 = x_288; -goto block_274; -} -} -else -{ -lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; -x_289 = lean_ctor_get(x_287, 0); -lean_inc(x_289); -if (lean_is_exclusive(x_287)) { - lean_ctor_release(x_287, 0); - x_290 = x_287; -} else { - lean_dec_ref(x_287); - x_290 = lean_box(0); -} -x_291 = lean_ctor_get(x_289, 0); -lean_inc(x_291); -x_292 = lean_ctor_get(x_289, 1); -lean_inc(x_292); -if (lean_is_exclusive(x_289)) { - lean_ctor_release(x_289, 0); - lean_ctor_release(x_289, 1); - x_293 = x_289; -} else { - lean_dec_ref(x_289); - x_293 = lean_box(0); -} -x_294 = lean_ctor_get(x_2, 0); -lean_inc(x_294); -x_295 = lean_ctor_get(x_294, 1); -lean_inc(x_295); -lean_dec(x_294); -lean_inc(x_295); -x_296 = l_Lean_FileMap_utf8PosToLspPos(x_295, x_291); -lean_dec(x_291); -if (x_6 == 0) -{ -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; -x_297 = l_Lean_FileMap_utf8PosToLspPos(x_295, x_292); -lean_dec(x_292); -if (lean_is_scalar(x_293)) { - x_298 = lean_alloc_ctor(0, 2, 0); -} else { - x_298 = x_293; -} -lean_ctor_set(x_298, 0, x_296); -lean_ctor_set(x_298, 1, x_297); -lean_inc(x_298); -lean_inc(x_35); -x_299 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_299, 0, x_35); -lean_ctor_set(x_299, 1, x_298); -lean_ctor_set(x_299, 2, x_298); -if (lean_is_scalar(x_290)) { - x_300 = lean_alloc_ctor(1, 1, 0); -} else { - x_300 = x_290; -} -lean_ctor_set(x_300, 0, x_299); -if (lean_obj_tag(x_200) == 0) -{ -x_240 = x_300; -goto block_256; -} -else -{ -lean_object* x_301; -lean_dec(x_199); -x_301 = lean_ctor_get(x_200, 0); -lean_inc(x_301); -lean_dec(x_200); -x_257 = x_300; -x_258 = x_301; -goto block_274; -} -} -else -{ -lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; -x_302 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4; -x_303 = lean_nat_add(x_292, x_302); -lean_dec(x_292); -x_304 = l_Lean_FileMap_utf8PosToLspPos(x_295, x_303); -lean_dec(x_303); -if (lean_is_scalar(x_293)) { - x_305 = lean_alloc_ctor(0, 2, 0); -} else { - x_305 = x_293; -} -lean_ctor_set(x_305, 0, x_296); -lean_ctor_set(x_305, 1, x_304); -lean_inc(x_305); -lean_inc(x_35); -x_306 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_306, 0, x_35); -lean_ctor_set(x_306, 1, x_305); -lean_ctor_set(x_306, 2, x_305); -if (lean_is_scalar(x_290)) { - x_307 = lean_alloc_ctor(1, 1, 0); -} else { - x_307 = x_290; -} -lean_ctor_set(x_307, 0, x_306); -if (lean_obj_tag(x_200) == 0) -{ -x_240 = x_307; -goto block_256; -} -else -{ -lean_object* x_308; -lean_dec(x_199); -x_308 = lean_ctor_get(x_200, 0); -lean_inc(x_308); -lean_dec(x_200); -x_257 = x_307; -x_258 = x_308; -goto block_274; -} -} -} -} -} -} -} -} -block_239: -{ -lean_dec(x_206); -if (lean_obj_tag(x_200) == 0) -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_207 = lean_data_value_to_string(x_199); -x_208 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_209 = lean_string_append(x_208, x_207); -lean_dec(x_207); -x_210 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_211 = lean_string_append(x_209, x_210); -x_212 = lean_string_append(x_211, x_201); -lean_dec(x_201); -x_213 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_214 = lean_string_append(x_212, x_213); -x_215 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_215, 0, x_214); -x_216 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_217 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_217, 0, x_35); -lean_ctor_set(x_217, 1, x_215); -lean_ctor_set(x_217, 2, x_202); -lean_ctor_set(x_217, 3, x_216); -lean_ctor_set(x_217, 4, x_202); -lean_ctor_set(x_217, 5, x_202); -lean_ctor_set(x_217, 6, x_205); -lean_ctor_set(x_217, 7, x_202); -x_218 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_218, 0, x_217); -lean_ctor_set(x_218, 1, x_197); -x_219 = lean_array_push(x_25, x_218); -if (lean_is_scalar(x_26)) { - x_220 = lean_alloc_ctor(1, 1, 0); -} else { - x_220 = x_26; -} -lean_ctor_set(x_220, 0, x_219); -if (lean_is_scalar(x_24)) { - x_221 = lean_alloc_ctor(0, 2, 0); -} else { - x_221 = x_24; -} -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_23); -x_27 = x_221; -goto block_32; -} -else -{ -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_199); -x_222 = lean_ctor_get(x_200, 0); -lean_inc(x_222); -if (lean_is_exclusive(x_200)) { - lean_ctor_release(x_200, 0); - x_223 = x_200; -} else { - lean_dec_ref(x_200); - x_223 = lean_box(0); -} -x_224 = lean_data_value_to_string(x_222); -x_225 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_226 = lean_string_append(x_225, x_224); -lean_dec(x_224); -x_227 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_228 = lean_string_append(x_226, x_227); -x_229 = lean_string_append(x_228, x_201); -lean_dec(x_201); -x_230 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_231 = lean_string_append(x_229, x_230); -if (lean_is_scalar(x_223)) { - x_232 = lean_alloc_ctor(1, 1, 0); -} else { - x_232 = x_223; -} -lean_ctor_set(x_232, 0, x_231); -x_233 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_234 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_234, 0, x_35); -lean_ctor_set(x_234, 1, x_232); -lean_ctor_set(x_234, 2, x_202); -lean_ctor_set(x_234, 3, x_233); -lean_ctor_set(x_234, 4, x_202); -lean_ctor_set(x_234, 5, x_202); -lean_ctor_set(x_234, 6, x_205); -lean_ctor_set(x_234, 7, x_202); -x_235 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_235, 0, x_234); -lean_ctor_set(x_235, 1, x_197); -x_236 = lean_array_push(x_25, x_235); -if (lean_is_scalar(x_26)) { - x_237 = lean_alloc_ctor(1, 1, 0); -} else { - x_237 = x_26; -} -lean_ctor_set(x_237, 0, x_236); -if (lean_is_scalar(x_24)) { - x_238 = lean_alloc_ctor(0, 2, 0); -} else { - x_238 = x_24; -} -lean_ctor_set(x_238, 0, x_237); -lean_ctor_set(x_238, 1, x_23); -x_27 = x_238; -goto block_32; -} -} -block_256: -{ -lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; -x_241 = lean_data_value_to_string(x_199); -x_242 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_243 = lean_string_append(x_242, x_241); -lean_dec(x_241); -x_244 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_245 = lean_string_append(x_243, x_244); -x_246 = lean_string_append(x_245, x_201); -lean_dec(x_201); -x_247 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_248 = lean_string_append(x_246, x_247); -x_249 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_249, 0, x_248); -x_250 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_251 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_251, 0, x_35); -lean_ctor_set(x_251, 1, x_249); -lean_ctor_set(x_251, 2, x_202); -lean_ctor_set(x_251, 3, x_250); -lean_ctor_set(x_251, 4, x_240); -lean_ctor_set(x_251, 5, x_202); -lean_ctor_set(x_251, 6, x_205); -lean_ctor_set(x_251, 7, x_202); -x_252 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_252, 0, x_251); -lean_ctor_set(x_252, 1, x_197); -x_253 = lean_array_push(x_25, x_252); -x_254 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_254, 0, x_253); -x_255 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_255, 0, x_254); -lean_ctor_set(x_255, 1, x_23); -x_27 = x_255; -goto block_32; -} -block_274: -{ -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; -x_259 = lean_data_value_to_string(x_258); -x_260 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2; -x_261 = lean_string_append(x_260, x_259); -lean_dec(x_259); -x_262 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3; -x_263 = lean_string_append(x_261, x_262); -x_264 = lean_string_append(x_263, x_201); -lean_dec(x_201); -x_265 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_266 = lean_string_append(x_264, x_265); -x_267 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_267, 0, x_266); -x_268 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1; -x_269 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_269, 0, x_35); -lean_ctor_set(x_269, 1, x_267); -lean_ctor_set(x_269, 2, x_202); -lean_ctor_set(x_269, 3, x_268); -lean_ctor_set(x_269, 4, x_257); -lean_ctor_set(x_269, 5, x_202); -lean_ctor_set(x_269, 6, x_205); -lean_ctor_set(x_269, 7, x_202); -x_270 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_197); -x_271 = lean_array_push(x_25, x_270); -x_272 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_272, 0, x_271); -x_273 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_273, 0, x_272); -lean_ctor_set(x_273, 1, x_23); -x_27 = x_273; -goto block_32; -} -} -} -block_32: -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -lean_dec(x_28); -x_8 = x_20; -x_9 = x_30; -x_14 = x_29; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_12 = l_Lean_getOptionDecls(x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_9, 2); -x_16 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1; -x_17 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_15, x_13, x_16, x_7, x_8, x_9, x_10, x_14); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_17, 0); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_19, 0); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_21); -x_23 = 1; -x_24 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_23); -lean_ctor_set(x_19, 0, x_24); -return x_17; -} -else -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_19, 0); -lean_inc(x_25); -lean_dec(x_19); -x_26 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_25); -x_27 = 1; -x_28 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_17, 0, x_29); -return x_17; -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_30 = lean_ctor_get(x_17, 0); -x_31 = lean_ctor_get(x_17, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_17); -x_32 = lean_ctor_get(x_30, 0); -lean_inc(x_32); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - x_33 = x_30; -} else { - lean_dec_ref(x_30); - x_33 = lean_box(0); -} -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_32); -x_35 = 1; -x_36 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_35); -if (lean_is_scalar(x_33)) { - x_37 = lean_alloc_ctor(1, 1, 0); -} else { - x_37 = x_33; -} -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_31); -return x_38; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(32u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3() { -_start: -{ -size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = 5; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1; -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); -lean_ctor_set(x_5, 0, x_2); -lean_ctor_set(x_5, 1, x_3); -lean_ctor_set(x_5, 2, x_4); -lean_ctor_set(x_5, 3, x_4); -lean_ctor_set_usize(x_5, 4, x_1); -return x_5; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; -x_6 = lean_unsigned_to_nat(1u); -x_7 = l_Lean_Syntax_getArg(x_3, x_6); -x_8 = 0; -x_9 = l_Lean_Syntax_getSubstring_x3f(x_7, x_8, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2; -x_11 = lean_box(x_8); -lean_inc(x_2); -x_12 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed), 11, 6); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_4); -lean_closure_set(x_12, 3, x_7); -lean_closure_set(x_12, 4, x_10); -lean_closure_set(x_12, 5, x_11); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4; -x_14 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_13, x_12, x_5); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_9, 0); -lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 2); -lean_inc(x_17); -x_18 = lean_string_utf8_at_end(x_16, x_17); -if (x_18 == 0) -{ -uint32_t x_19; uint32_t x_20; uint8_t x_21; -x_19 = lean_string_utf8_get(x_16, x_17); -lean_dec(x_17); -lean_dec(x_16); -x_20 = 46; -x_21 = lean_uint32_dec_eq(x_19, x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = lean_ctor_get(x_15, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -x_24 = lean_ctor_get(x_15, 2); -lean_inc(x_24); -lean_dec(x_15); -x_25 = lean_string_utf8_extract(x_22, x_23, x_24); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -x_26 = lean_box(x_8); -lean_inc(x_2); -x_27 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed), 11, 6); -lean_closure_set(x_27, 0, x_1); -lean_closure_set(x_27, 1, x_2); -lean_closure_set(x_27, 2, x_4); -lean_closure_set(x_27, 3, x_7); -lean_closure_set(x_27, 4, x_25); -lean_closure_set(x_27, 5, x_26); -x_28 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4; -x_29 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_28, x_27, x_5); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_30 = lean_ctor_get(x_15, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_15, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_15, 2); -lean_inc(x_32); -lean_dec(x_15); -x_33 = lean_string_utf8_extract(x_30, x_31, x_32); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -x_34 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7; -x_35 = lean_string_append(x_33, x_34); -x_36 = 1; -x_37 = lean_box(x_36); -lean_inc(x_2); -x_38 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed), 11, 6); -lean_closure_set(x_38, 0, x_1); -lean_closure_set(x_38, 1, x_2); -lean_closure_set(x_38, 2, x_4); -lean_closure_set(x_38, 3, x_7); -lean_closure_set(x_38, 4, x_35); -lean_closure_set(x_38, 5, x_37); -x_39 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4; -x_40 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_39, x_38, x_5); -return x_40; -} -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_17); -lean_dec(x_16); -x_41 = lean_ctor_get(x_15, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); -x_43 = lean_ctor_get(x_15, 2); -lean_inc(x_43); -lean_dec(x_15); -x_44 = lean_string_utf8_extract(x_41, x_42, x_43); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); -x_45 = lean_box(x_8); -lean_inc(x_2); -x_46 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed), 11, 6); -lean_closure_set(x_46, 0, x_1); -lean_closure_set(x_46, 1, x_2); -lean_closure_set(x_46, 2, x_4); -lean_closure_set(x_46, 3, x_7); -lean_closure_set(x_46, 4, x_44); -lean_closure_set(x_46, 5, x_45); -x_47 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4; -x_48 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_47, x_46, x_5); -return x_48; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_6); -lean_dec(x_6); -x_16 = l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -return x_16; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_6); -lean_dec(x_6); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_12, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -return x_6; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1() { -_start: -{ -double x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; -x_2 = lean_box_float(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) -{ -lean_dec(x_1); -return x_4; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; -x_6 = lean_array_uget(x_4, x_3); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_array_uset(x_4, x_3, x_7); -x_9 = lean_ctor_get(x_6, 1); -lean_inc(x_9); -x_10 = lean_box(0); -x_11 = lean_ctor_get(x_6, 3); -lean_inc(x_11); -lean_dec(x_6); -lean_inc(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_10); -x_13 = l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472_(x_12); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -x_15 = 1; -x_16 = lean_usize_add(x_3, x_15); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; -x_18 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_18, 0, x_9); -lean_ctor_set(x_18, 1, x_10); -lean_ctor_set(x_18, 2, x_10); -lean_ctor_set(x_18, 3, x_17); -lean_ctor_set(x_18, 4, x_10); -lean_ctor_set(x_18, 5, x_10); -lean_ctor_set(x_18, 6, x_14); -lean_ctor_set(x_18, 7, x_10); -x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1; -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_array_uset(x_8, x_3, x_20); -x_3 = x_16; -x_4 = x_21; -goto _start; -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_11); -if (x_23 == 0) -{ -lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_24 = lean_ctor_get(x_11, 0); -x_25 = 1; -x_26 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set_uint8(x_26, sizeof(void*)*1, x_25); -lean_ctor_set(x_11, 0, x_26); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; -x_28 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_28, 0, x_9); -lean_ctor_set(x_28, 1, x_10); -lean_ctor_set(x_28, 2, x_11); -lean_ctor_set(x_28, 3, x_27); -lean_ctor_set(x_28, 4, x_10); -lean_ctor_set(x_28, 5, x_10); -lean_ctor_set(x_28, 6, x_14); -lean_ctor_set(x_28, 7, x_10); -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1; -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = lean_array_uset(x_8, x_3, x_30); -x_3 = x_16; -x_4 = x_31; -goto _start; -} -else -{ -lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_33 = lean_ctor_get(x_11, 0); -lean_inc(x_33); -lean_dec(x_11); -x_34 = 1; -x_35 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_34); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; -x_38 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_38, 0, x_9); -lean_ctor_set(x_38, 1, x_10); -lean_ctor_set(x_38, 2, x_36); -lean_ctor_set(x_38, 3, x_37); -lean_ctor_set(x_38, 4, x_10); -lean_ctor_set(x_38, 5, x_10); -lean_ctor_set(x_38, 6, x_14); -lean_ctor_set(x_38, 7, x_10); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1; -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = lean_array_uset(x_8, x_3, x_40); -x_3 = x_16; -x_4 = x_41; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Elab_Tactic_Doc_allTacticDocs(x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_array_size(x_9); -x_11 = 0; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_10, x_11, x_9); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_12); -x_14 = 1; -x_15 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set_uint8(x_15, sizeof(void*)*1, x_14); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_7, 0, x_16); -return x_7; -} -else -{ -lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_7, 0); -x_18 = lean_ctor_get(x_7, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_7); -x_19 = lean_array_size(x_17); -x_20 = 0; -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_19, x_20, x_17); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_21); -x_23 = 1; -x_24 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_23); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_18); -return x_26; -} -} -else -{ -uint8_t x_27; -lean_dec(x_1); -x_27 = !lean_is_exclusive(x_7); -if (x_27 == 0) -{ -return x_7; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_7, 0); -x_29 = lean_ctor_get(x_7, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_7); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___lambda__1), 6, 1); -lean_closure_set(x_4, 0, x_1); -x_5 = l_Lean_LocalContext_empty; -x_6 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_2, x_5, x_4, x_3); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_5, x_6, x_4); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_4; lean_object* x_5; -lean_dec(x_3); -lean_dec(x_1); -x_4 = 0; -x_5 = lean_box(x_4); -return x_5; -} -else -{ -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_6; lean_object* x_7; -lean_dec(x_2); -lean_dec(x_1); -x_6 = 1; -x_7 = lean_box(x_6); -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_2, 0); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_ctor_get(x_3, 0); -lean_inc(x_9); -lean_dec(x_3); -x_10 = lean_apply_2(x_1, x_8, x_9); -return x_10; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; -} -else -{ -lean_object* x_4; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_1, 1); -x_7 = lean_ctor_get(x_1, 0); -lean_dec(x_7); -x_8 = lean_box(0); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_8); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; -} -} -else -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_1, 1); -x_16 = lean_ctor_get(x_1, 0); -lean_dec(x_16); -x_17 = lean_ctor_get(x_4, 0); -lean_inc(x_17); -lean_dec(x_4); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_17); -{ -lean_object* _tmp_0 = x_15; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; -} -goto _start; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_1, 1); -lean_inc(x_19); -lean_dec(x_1); -x_20 = lean_ctor_get(x_4, 0); -lean_inc(x_20); -lean_dec(x_4); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_2); -x_1 = x_19; -x_2 = x_21; -goto _start; -} -} -} -} -} -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg), 2, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_dec(x_1); -return x_2; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec(x_3); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_2, x_4); -x_7 = lean_unbox(x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -lean_dec(x_2); -x_2 = x_4; -x_3 = x_5; -goto _start; -} -else -{ -lean_dec(x_4); -x_3 = x_5; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_box(0); -x_8 = lean_box(0); -x_9 = l_List_mapTR_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(x_6, x_8); -lean_inc(x_1); -x_10 = l_List_foldl___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(x_1, x_7, x_9); -x_11 = lean_apply_3(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_11) == 0) -{ -lean_dec(x_1); -return x_10; -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -lean_inc(x_10); -lean_inc(x_11); -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_11, x_10); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_11); -return x_10; -} -else -{ -lean_dec(x_10); -return x_11; -} -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_11, 0); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_15); -lean_inc(x_10); -lean_inc(x_16); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_16, x_10); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_dec(x_16); -return x_10; -} -else -{ -lean_dec(x_10); -return x_16; -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___rarg), 6, 0); -return x_2; -} -} -static lean_object* _init_l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Id_instMonad; -x_3 = l_instInhabitedOfMonad___rarg(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg), 1, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_6; -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_6 = l_List_reverse___rarg(x_5); -return x_6; -} -else -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_4, 0); -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_10 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_1, x_2, x_3, x_8); -lean_ctor_set(x_4, 1, x_5); -lean_ctor_set(x_4, 0, x_10); -{ -lean_object* _tmp_3 = x_9; -lean_object* _tmp_4 = x_4; -x_4 = _tmp_3; -x_5 = _tmp_4; -} -goto _start; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_4, 0); -x_13 = lean_ctor_get(x_4, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_14 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_1, x_2, x_3, x_12); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_5); -x_4 = x_13; -x_5 = x_15; -goto _start; -} -} -} -} -LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg), 5, 0); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Server.InfoUtils", 21, 21); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Elab.InfoTree.visitM.go", 28, 28); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("unexpected context-free info tree node", 38, 38); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1; -x_2 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2; -x_3 = lean_unsigned_to_nat(62u); -x_4 = lean_unsigned_to_nat(21u); -x_5 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -switch (lean_obj_tag(x_4)) { -case 0: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_4, 1); -lean_inc(x_6); -lean_dec(x_4); -x_7 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_5, x_3); -x_3 = x_7; -x_4 = x_6; -goto _start; -} -case 1: -{ -lean_object* x_9; lean_object* x_10; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_9 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4; -x_10 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(x_9); -return x_10; -} -default: -{ -lean_object* x_11; -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_11 = lean_box(0); -return x_11; -} -} -} -else -{ -switch (lean_obj_tag(x_4)) { -case 0: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_4, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_4, 1); -lean_inc(x_13); -lean_dec(x_4); -x_14 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_12, x_3); -x_3 = x_14; -x_4 = x_13; -goto _start; -} -case 1: -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_3); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_3, 0); -x_18 = lean_ctor_get(x_4, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_4, 1); -lean_inc(x_19); -lean_dec(x_4); -lean_inc(x_1); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -x_20 = lean_apply_3(x_1, x_17, x_18, x_19); -x_21 = lean_unbox(x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -lean_dec(x_1); -x_22 = lean_box(0); -x_23 = lean_apply_4(x_2, x_17, x_18, x_19, x_22); -lean_ctor_set(x_3, 0, x_23); -return x_3; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_inc(x_17); -x_24 = l_Lean_Elab_Info_updateContext_x3f(x_3, x_18); -x_25 = l_Lean_PersistentArray_toList___rarg(x_19); -x_26 = lean_box(0); -lean_inc(x_2); -x_27 = l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(x_1, x_2, x_24, x_25, x_26); -x_28 = lean_apply_4(x_2, x_17, x_18, x_19, x_27); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -return x_29; -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_30 = lean_ctor_get(x_3, 0); -lean_inc(x_30); -lean_dec(x_3); -x_31 = lean_ctor_get(x_4, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_4, 1); -lean_inc(x_32); -lean_dec(x_4); -lean_inc(x_1); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -x_33 = lean_apply_3(x_1, x_30, x_31, x_32); -x_34 = lean_unbox(x_33); -lean_dec(x_33); -if (x_34 == 0) -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_1); -x_35 = lean_box(0); -x_36 = lean_apply_4(x_2, x_30, x_31, x_32, x_35); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_inc(x_30); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_30); -x_39 = l_Lean_Elab_Info_updateContext_x3f(x_38, x_31); -x_40 = l_Lean_PersistentArray_toList___rarg(x_32); -x_41 = lean_box(0); -lean_inc(x_2); -x_42 = l_List_mapM_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(x_1, x_2, x_39, x_40, x_41); -x_43 = lean_apply_4(x_2, x_30, x_31, x_32, x_42); -x_44 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_44, 0, x_43); -return x_44; -} -} -} -default: -{ -lean_object* x_45; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_45 = lean_box(0); -return x_45; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg), 4, 0); -return x_2; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = 1; -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed), 3, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f_choose___rarg), 6, 2); -lean_closure_set(x_4, 0, x_2); -lean_closure_set(x_4, 1, x_3); -x_5 = lean_box(0); -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1; -x_7 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_6, x_4, x_5, x_1); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; -x_8 = lean_box(0); -return x_8; -} -else -{ -lean_object* x_9; -x_9 = lean_ctor_get(x_7, 0); -lean_inc(x_9); -lean_dec(x_7); -return x_9; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg), 3, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_3 = lean_ctor_get(x_1, 1); -x_4 = lean_ctor_get(x_2, 1); -x_5 = l_Lean_Elab_Info_lctx(x_3); -x_6 = lean_local_ctx_is_empty(x_5); -if (x_6 == 0) -{ -lean_object* x_7; uint8_t x_8; -x_7 = l_Lean_Elab_Info_lctx(x_4); -x_8 = lean_local_ctx_is_empty(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -x_9 = l_Lean_Elab_Info_isSmaller(x_3, x_4); -return x_9; -} -else -{ -uint8_t x_10; -x_10 = 1; -return x_10; -} -} -else -{ -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Elab_Info_lctx(x_4); -x_12 = lean_local_ctx_is_empty(x_11); -if (x_12 == 0) -{ -uint8_t x_13; -x_13 = 0; -return x_13; -} -else -{ -uint8_t x_14; -x_14 = l_Lean_Elab_Info_isSmaller(x_3, x_4); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; -x_5 = l_Lean_Elab_Info_occursInOrOnBoundary(x_3, x_1); -if (x_5 == 0) -{ -lean_object* x_6; -lean_dec(x_3); -lean_dec(x_2); -x_6 = lean_box(0); -return x_6; -} -else -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_2); -lean_ctor_set(x_7, 1, x_3); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed), 4, 1); -lean_closure_set(x_3, 0, x_1); -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg(x_2, x_4, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_4); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = 0; -x_4 = l_Lean_Syntax_getRange_x3f(x_2, x_3); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; uint8_t x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -lean_dec(x_4); -x_7 = 1; -x_8 = l_String_Range_contains(x_6, x_1, x_7); -lean_dec(x_6); -return x_8; -} -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = l_Lean_Syntax_hasArgs(x_1); -if (x_2 == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("ident", 5, 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -lean_dec(x_1); -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; -lean_inc(x_2); -x_4 = l_Lean_Syntax_isOfKind(x_2, x_3); -if (x_4 == 0) -{ -lean_object* x_5; uint8_t x_6; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5; -lean_inc(x_2); -x_6 = l_Lean_Syntax_isOfKind(x_2, x_5); -if (x_6 == 0) -{ -uint8_t x_7; -lean_dec(x_2); -x_7 = 1; -return x_7; -} -else -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_Syntax_getArg(x_2, x_8); -lean_dec(x_2); -x_10 = l_Lean_Syntax_isOfKind(x_9, x_3); -if (x_10 == 0) -{ -uint8_t x_11; -x_11 = 1; -return x_11; -} -else -{ -uint8_t x_12; -x_12 = 0; -return x_12; -} -} -} -else -{ -uint8_t x_13; -lean_dec(x_2); -x_13 = 0; -return x_13; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("dotIdent", 8, 8); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -lean_dec(x_1); -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2; -lean_inc(x_2); -x_4 = l_Lean_Syntax_isOfKind(x_2, x_3); -if (x_4 == 0) -{ -uint8_t x_5; -lean_dec(x_2); -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_unsigned_to_nat(1u); -x_7 = l_Lean_Syntax_getArg(x_2, x_6); -lean_dec(x_2); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; -x_9 = l_Lean_Syntax_isOfKind(x_7, x_8); -return x_9; -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("value is none", 13, 13); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_73; uint8_t x_74; -x_73 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; -lean_inc(x_1); -x_74 = l_Lean_Syntax_isOfKind(x_1, x_73); -if (x_74 == 0) -{ -lean_object* x_75; uint8_t x_76; -x_75 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5; -lean_inc(x_1); -x_76 = l_Lean_Syntax_isOfKind(x_1, x_75); -if (x_76 == 0) -{ -lean_object* x_77; -lean_dec(x_3); -lean_dec(x_1); -x_77 = lean_box(0); -return x_77; -} -else -{ -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = lean_unsigned_to_nat(0u); -x_79 = l_Lean_Syntax_getArg(x_1, x_78); -lean_inc(x_79); -x_80 = l_Lean_Syntax_isOfKind(x_79, x_73); -if (x_80 == 0) -{ -lean_object* x_81; -lean_dec(x_79); -lean_dec(x_3); -lean_dec(x_1); -x_81 = lean_box(0); -return x_81; -} -else -{ -lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; -x_82 = l_Lean_Syntax_getId(x_79); -lean_dec(x_79); -x_83 = 1; -x_84 = lean_box(x_83); -x_85 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_84); -x_6 = x_85; -goto block_72; -} -} -} -else -{ -lean_object* x_86; uint8_t x_87; lean_object* x_88; lean_object* x_89; -x_86 = l_Lean_Syntax_getId(x_1); -x_87 = 0; -x_88 = lean_box(x_87); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_88); -x_6 = x_89; -goto block_72; -} -block_72: -{ -uint8_t x_7; -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_6, 1); -x_10 = 0; -x_11 = l_Lean_Syntax_getTailPos_x3f(x_1, x_10); -x_12 = l_Lean_Elab_Info_lctx(x_2); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_8); -lean_ctor_set(x_14, 2, x_12); -lean_ctor_set(x_14, 3, x_13); -x_15 = lean_unbox(x_9); -lean_dec(x_9); -lean_ctor_set_uint8(x_14, sizeof(void*)*4, x_15); -lean_ctor_set(x_6, 1, x_14); -lean_ctor_set(x_6, 0, x_3); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_17 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_16); -x_18 = lean_nat_dec_lt(x_4, x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_17); -x_19 = lean_box(0); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_6); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_20); -return x_21; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_nat_sub(x_17, x_4); -lean_dec(x_17); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_6); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -return x_25; -} -} -else -{ -uint8_t x_26; -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) -{ -lean_object* x_27; uint8_t x_28; -x_27 = lean_ctor_get(x_11, 0); -x_28 = lean_nat_dec_lt(x_4, x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_27); -x_29 = lean_box(0); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_6); -lean_ctor_set(x_11, 0, x_30); -return x_11; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_nat_sub(x_27, x_4); -lean_dec(x_27); -x_32 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_6); -lean_ctor_set(x_11, 0, x_33); -return x_11; -} -} -else -{ -lean_object* x_34; uint8_t x_35; -x_34 = lean_ctor_get(x_11, 0); -lean_inc(x_34); -lean_dec(x_11); -x_35 = lean_nat_dec_lt(x_4, x_34); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec(x_34); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_6); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -return x_38; -} -else -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_39 = lean_nat_sub(x_34, x_4); -lean_dec(x_34); -x_40 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_40, 0, x_39); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_6); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -return x_42; -} -} -} -} -else -{ -lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; -x_43 = lean_ctor_get(x_6, 0); -x_44 = lean_ctor_get(x_6, 1); -lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_6); -x_45 = 0; -x_46 = l_Lean_Syntax_getTailPos_x3f(x_1, x_45); -x_47 = l_Lean_Elab_Info_lctx(x_2); -x_48 = lean_box(0); -x_49 = lean_alloc_ctor(1, 4, 1); -lean_ctor_set(x_49, 0, x_1); -lean_ctor_set(x_49, 1, x_43); -lean_ctor_set(x_49, 2, x_47); -lean_ctor_set(x_49, 3, x_48); -x_50 = lean_unbox(x_44); -lean_dec(x_44); -lean_ctor_set_uint8(x_49, sizeof(void*)*4, x_50); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_3); -lean_ctor_set(x_51, 1, x_49); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_52 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_53 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_52); -x_54 = lean_nat_dec_lt(x_4, x_53); -if (x_54 == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_53); -x_55 = lean_box(0); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_51); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_56); -return x_57; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_58 = lean_nat_sub(x_53, x_4); -lean_dec(x_53); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_51); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_60); -return x_61; -} -} -else -{ -lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_62 = lean_ctor_get(x_46, 0); -lean_inc(x_62); -if (lean_is_exclusive(x_46)) { - lean_ctor_release(x_46, 0); - x_63 = x_46; -} else { - lean_dec_ref(x_46); - x_63 = lean_box(0); -} -x_64 = lean_nat_dec_lt(x_4, x_62); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_65 = lean_box(0); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_51); -if (lean_is_scalar(x_63)) { - x_67 = lean_alloc_ctor(1, 1, 0); -} else { - x_67 = x_63; -} -lean_ctor_set(x_67, 0, x_66); -return x_67; -} -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_nat_sub(x_62, x_4); -lean_dec(x_62); -x_69 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_69, 0, x_68); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_51); -if (lean_is_scalar(x_63)) { - x_71 = lean_alloc_ctor(1, 1, 0); -} else { - x_71 = x_63; -} -lean_ctor_set(x_71, 0, x_70); -return x_71; -} -} -} -} -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed), 1, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -lean_inc(x_1); -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -lean_dec(x_1); -x_4 = lean_box(0); -return x_4; -} -else -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); -lean_dec(x_3); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = l_Lean_Elab_Info_stx(x_7); -lean_inc(x_1); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed), 2, 1); -lean_closure_set(x_9, 0, x_1); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1; -x_11 = l_Lean_Syntax_findStack_x3f(x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_12 = lean_box(0); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2; -x_15 = l_List_dropWhile___rarg(x_14, x_13); -x_16 = l_List_head_x3f___rarg(x_15); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; -lean_dec(x_15); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_17 = lean_box(0); -return x_17; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3; -x_21 = l_List_any___rarg(x_15, x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(x_19, x_7, x_6, x_1, x_22); -lean_dec(x_1); -lean_dec(x_7); -return x_23; -} -else -{ -lean_object* x_24; -lean_dec(x_19); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_24 = lean_box(0); -return x_24; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -x_7 = lean_string_utf8_next(x_5, x_6); -lean_dec(x_6); -lean_ctor_set(x_2, 1, x_7); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_2); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_10 = lean_ctor_get(x_2, 0); -x_11 = lean_ctor_get(x_2, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_2); -x_12 = lean_string_utf8_next(x_10, x_11); -lean_dec(x_11); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_10); -lean_ctor_set(x_13, 1, x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_1); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} -} -} -static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed), 3, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -x_7 = lean_nat_dec_lt(x_6, x_1); -lean_dec(x_6); -if (x_7 == 0) -{ -return x_2; -} -else -{ -lean_object* x_8; uint32_t x_9; lean_object* x_10; lean_object* x_11; uint32_t x_12; uint8_t x_13; -x_8 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; -x_9 = 32; -x_10 = lean_ctor_get(x_5, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_5, 1); -lean_inc(x_11); -x_12 = lean_string_utf8_get(x_10, x_11); -lean_dec(x_11); -lean_dec(x_10); -x_13 = lean_uint32_dec_eq(x_12, x_9); -if (x_13 == 0) -{ -uint32_t x_14; uint8_t x_15; -x_14 = 9; -x_15 = lean_uint32_dec_eq(x_12, x_14); -if (x_15 == 0) -{ -return x_2; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_free_object(x_2); -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_add(x_4, x_16); -lean_dec(x_4); -x_18 = lean_box(0); -x_19 = lean_apply_3(x_8, x_17, x_5, x_18); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -return x_20; -} -else -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_2 = x_21; -goto _start; -} -} -} -else -{ -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_free_object(x_2); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_4, x_23); -lean_dec(x_4); -x_25 = lean_box(0); -x_26 = lean_apply_3(x_8, x_24, x_5, x_25); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_27; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec(x_26); -return x_27; -} -else -{ -lean_object* x_28; -x_28 = lean_ctor_get(x_26, 0); -lean_inc(x_28); -lean_dec(x_26); -x_2 = x_28; -goto _start; -} -} -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_30 = lean_ctor_get(x_2, 0); -x_31 = lean_ctor_get(x_2, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_2); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -x_33 = lean_nat_dec_lt(x_32, x_1); -lean_dec(x_32); -if (x_33 == 0) -{ -lean_object* x_34; -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_30); -lean_ctor_set(x_34, 1, x_31); -return x_34; -} -else -{ -lean_object* x_35; uint32_t x_36; lean_object* x_37; lean_object* x_38; uint32_t x_39; uint8_t x_40; -x_35 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; -x_36 = 32; -x_37 = lean_ctor_get(x_31, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_31, 1); -lean_inc(x_38); -x_39 = lean_string_utf8_get(x_37, x_38); -lean_dec(x_38); -lean_dec(x_37); -x_40 = lean_uint32_dec_eq(x_39, x_36); -if (x_40 == 0) -{ -uint32_t x_41; uint8_t x_42; -x_41 = 9; -x_42 = lean_uint32_dec_eq(x_39, x_41); -if (x_42 == 0) -{ -lean_object* x_43; -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_30); -lean_ctor_set(x_43, 1, x_31); -return x_43; -} -else -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_unsigned_to_nat(1u); -x_45 = lean_nat_add(x_30, x_44); -lean_dec(x_30); -x_46 = lean_box(0); -x_47 = lean_apply_3(x_35, x_45, x_31, x_46); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -lean_dec(x_47); -return x_48; -} -else -{ -lean_object* x_49; -x_49 = lean_ctor_get(x_47, 0); -lean_inc(x_49); -lean_dec(x_47); -x_2 = x_49; -goto _start; -} -} -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_unsigned_to_nat(1u); -x_52 = lean_nat_add(x_30, x_51); -lean_dec(x_30); -x_53 = lean_box(0); -x_54 = lean_apply_3(x_35, x_52, x_31, x_53); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -lean_dec(x_54); -return x_55; -} -else -{ -lean_object* x_56; -x_56 = lean_ctor_get(x_54, 0); -lean_inc(x_56); -lean_dec(x_54); -x_2 = x_56; -goto _start; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = l_Lean_FileMap_lineStart(x_1, x_2); -x_4 = lean_unsigned_to_nat(1u); -x_5 = lean_nat_add(x_2, x_4); -x_6 = l_Lean_FileMap_lineStart(x_1, x_5); -lean_dec(x_5); -x_7 = !lean_is_exclusive(x_1); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = lean_ctor_get(x_1, 1); -lean_dec(x_8); -lean_ctor_set(x_1, 1, x_3); -x_9 = lean_unsigned_to_nat(0u); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_1); -x_11 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_6, x_10); -lean_dec(x_6); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_3); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -x_17 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_6, x_16); -lean_dec(x_6); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -lean_dec(x_17); -return x_18; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_20; -x_3 = lean_ctor_get(x_1, 0); -x_20 = lean_string_utf8_at_end(x_3, x_2); -if (x_20 == 0) -{ -uint32_t x_21; uint32_t x_22; uint8_t x_23; -x_21 = lean_string_utf8_get(x_3, x_2); -x_22 = 32; -x_23 = lean_uint32_dec_eq(x_21, x_22); -if (x_23 == 0) -{ -uint32_t x_24; uint8_t x_25; -x_24 = 9; -x_25 = lean_uint32_dec_eq(x_21, x_24); -if (x_25 == 0) -{ -uint32_t x_26; uint8_t x_27; -x_26 = 13; -x_27 = lean_uint32_dec_eq(x_21, x_26); -if (x_27 == 0) -{ -uint32_t x_28; uint8_t x_29; -x_28 = 10; -x_29 = lean_uint32_dec_eq(x_21, x_28); -if (x_29 == 0) -{ -uint8_t x_30; -x_30 = 0; -return x_30; -} -else -{ -lean_object* x_31; -x_31 = lean_box(0); -x_4 = x_31; -goto block_19; -} -} -else -{ -lean_object* x_32; -x_32 = lean_box(0); -x_4 = x_32; -goto block_19; -} -} -else -{ -lean_object* x_33; -x_33 = lean_box(0); -x_4 = x_33; -goto block_19; -} -} -else -{ -lean_object* x_34; -x_34 = lean_box(0); -x_4 = x_34; -goto block_19; -} -} -else -{ -lean_object* x_35; -x_35 = lean_box(0); -x_4 = x_35; -goto block_19; -} -block_19: -{ -lean_object* x_5; lean_object* x_6; uint32_t x_7; uint32_t x_8; uint8_t x_9; -lean_dec(x_4); -x_5 = lean_unsigned_to_nat(1u); -x_6 = lean_nat_sub(x_2, x_5); -x_7 = lean_string_utf8_get(x_3, x_6); -lean_dec(x_6); -x_8 = 32; -x_9 = lean_uint32_dec_eq(x_7, x_8); -if (x_9 == 0) -{ -uint32_t x_10; uint8_t x_11; -x_10 = 9; -x_11 = lean_uint32_dec_eq(x_7, x_10); -if (x_11 == 0) -{ -uint32_t x_12; uint8_t x_13; -x_12 = 13; -x_13 = lean_uint32_dec_eq(x_7, x_12); -if (x_13 == 0) -{ -uint32_t x_14; uint8_t x_15; -x_14 = 10; -x_15 = lean_uint32_dec_eq(x_7, x_14); -return x_15; -} -else -{ -uint8_t x_16; -x_16 = 1; -return x_16; -} -} -else -{ -uint8_t x_17; -x_17 = 1; -return x_17; -} -} -else -{ -uint8_t x_18; -x_18 = 1; -return x_18; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Tactic", 6, 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeqBracketed", 18, 18); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -lean_inc(x_1); -x_2 = l_Lean_Syntax_getKind(x_1); -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; -x_4 = lean_name_eq(x_2, x_3); -if (x_4 == 0) -{ -lean_object* x_5; uint8_t x_6; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; -x_6 = lean_name_eq(x_2, x_5); -lean_dec(x_2); -if (x_6 == 0) -{ -lean_object* x_7; -lean_dec(x_1); -x_7 = lean_box(0); -return x_7; -} -else +if (lean_obj_tag(x_2) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_unsigned_to_nat(1u); -x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_object* x_3; lean_dec(x_1); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -return x_10; -} +x_3 = lean_box(0); +return x_3; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_2); -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -lean_dec(x_1); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; -} -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(x_1); -if (lean_obj_tag(x_6) == 0) -{ -uint8_t x_7; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_2, 2); +lean_inc(x_6); lean_dec(x_2); -x_7 = 0; -return x_7; -} -else -{ -lean_object* x_8; uint8_t x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_6, 0); +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_4, 1); lean_inc(x_8); -lean_dec(x_6); -x_9 = 0; -x_10 = l_Lean_Syntax_getPos_x3f(x_8, x_9); -lean_dec(x_8); -if (lean_obj_tag(x_10) == 0) +lean_dec(x_4); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +x_11 = lean_string_dec_eq(x_7, x_9); +lean_dec(x_9); +lean_dec(x_7); +if (x_11 == 0) { -uint8_t x_11; -lean_dec(x_2); -x_11 = 0; -return x_11; +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_5); +x_2 = x_6; +goto _start; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -lean_inc(x_2); -x_13 = l_Lean_FileMap_toPosition(x_2, x_12); -lean_dec(x_12); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_8, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_8, 1); lean_inc(x_14); -lean_dec(x_13); -lean_inc(x_2); -x_15 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount(x_2, x_14); -lean_dec(x_14); -x_16 = lean_nat_dec_eq(x_3, x_15); +lean_dec(x_8); +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_dec(x_10); +x_17 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_13, x_15); lean_dec(x_15); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace(x_2, x_4); -lean_dec(x_2); +lean_dec(x_13); if (x_17 == 0) { -uint8_t x_18; -x_18 = 0; -return x_18; -} -else -{ -return x_16; -} -} -} -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_ctor_get(x_3, 1); -x_7 = lean_nat_dec_le(x_6, x_4); -if (x_7 == 0) -{ -uint8_t x_8; +lean_dec(x_16); +lean_dec(x_14); lean_dec(x_5); -lean_dec(x_1); -x_8 = 0; -return x_8; +x_2 = x_6; +goto _start; } else { -lean_object* x_9; uint8_t x_10; -x_9 = lean_box(0); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(x_5, x_1, x_4, x_2, x_9); -return x_10; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_14, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +x_21 = lean_ctor_get(x_16, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_19, x_21); +lean_dec(x_21); +lean_dec(x_19); +if (x_23 == 0) { -uint8_t x_6; lean_object* x_7; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_22); +lean_dec(x_20); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_7 = lean_box(x_6); -return x_7; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -uint8_t x_6; lean_object* x_7; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_7 = lean_box(x_6); -return x_7; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; uint8_t x_4; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_string_utf8_at_end(x_3, x_2); -if (x_4 == 0) -{ -uint32_t x_5; uint32_t x_6; uint8_t x_7; -x_5 = lean_string_utf8_get(x_3, x_2); -x_6 = 32; -x_7 = lean_uint32_dec_eq(x_5, x_6); -if (x_7 == 0) -{ -uint32_t x_8; uint8_t x_9; -x_8 = 9; -x_9 = lean_uint32_dec_eq(x_5, x_8); -if (x_9 == 0) -{ -uint32_t x_10; uint8_t x_11; -x_10 = 13; -x_11 = lean_uint32_dec_eq(x_5, x_10); -if (x_11 == 0) -{ -uint32_t x_12; uint8_t x_13; -x_12 = 10; -x_13 = lean_uint32_dec_eq(x_5, x_12); -return x_13; -} -else -{ -uint8_t x_14; -x_14 = 1; -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = 1; -return x_15; -} -} -else -{ -uint8_t x_16; -x_16 = 1; -return x_16; -} +x_2 = x_6; +goto _start; } else { -uint8_t x_17; -x_17 = 1; -return x_17; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; -} -} -static lean_object* _init_l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(";", 1, 1); -return x_1; -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: -{ -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) -{ -lean_object* x_6; uint8_t x_7; lean_object* x_8; -x_6 = lean_array_uget(x_2, x_3); -x_7 = 0; -x_8 = l_Lean_Syntax_getTailPos_x3f(x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = lean_ctor_get(x_22, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec(x_22); +x_29 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_25, x_27); +if (x_29 == 0) { -size_t x_9; size_t x_10; -lean_dec(x_6); -x_9 = 1; -x_10 = lean_usize_add(x_3, x_9); -x_3 = x_10; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_5); +x_2 = x_6; goto _start; } else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_8, 0); -lean_inc(x_12); -lean_dec(x_8); -x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1; -x_14 = l_Lean_Syntax_isToken(x_13, x_6); -if (x_14 == 0) +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +x_35 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_31, x_33); +lean_dec(x_33); +lean_dec(x_31); +if (x_35 == 0) { -size_t x_15; size_t x_16; -lean_dec(x_12); -lean_dec(x_6); -x_15 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; +lean_dec(x_34); +lean_dec(x_32); +lean_dec(x_5); +x_2 = x_6; goto _start; } else { -uint8_t x_18; -x_18 = lean_nat_dec_le(x_12, x_1); -if (x_18 == 0) -{ -size_t x_19; size_t x_20; -lean_dec(x_12); -lean_dec(x_6); -x_19 = 1; -x_20 = lean_usize_add(x_3, x_19); -x_3 = x_20; +uint8_t x_37; +x_37 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_32, x_34); +lean_dec(x_34); +lean_dec(x_32); +if (x_37 == 0) +{ +lean_dec(x_5); +x_2 = x_6; goto _start; } else { -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = l_Lean_Syntax_getTrailingSize(x_6); +lean_object* x_39; lean_dec(x_6); -x_23 = lean_nat_add(x_12, x_22); -lean_dec(x_22); -lean_dec(x_12); -x_24 = lean_nat_dec_le(x_1, x_23); -lean_dec(x_23); -if (x_24 == 0) -{ -size_t x_25; size_t x_26; -x_25 = 1; -x_26 = lean_usize_add(x_3, x_25); -x_3 = x_26; -goto _start; +lean_dec(x_1); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_5); +return x_39; } -else -{ -uint8_t x_28; -x_28 = 1; -return x_28; } } } } } -else -{ -uint8_t x_29; -x_29 = 0; -return x_29; } } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(x_3); -if (lean_obj_tag(x_4) == 0) +if (lean_obj_tag(x_2) == 0) { -uint8_t x_5; -x_5 = 0; -return x_5; +uint8_t x_3; +lean_dec(x_1); +x_3 = 0; +return x_3; } else { -lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +lean_dec(x_2); x_6 = lean_ctor_get(x_4, 0); lean_inc(x_6); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); lean_dec(x_4); -x_7 = l_Lean_Syntax_getArgs(x_6); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +x_10 = lean_string_dec_eq(x_6, x_8); +lean_dec(x_8); lean_dec(x_6); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorOnWhitspace(x_1, x_2); -if (x_8 == 0) +if (x_10 == 0) { -uint8_t x_9; +lean_dec(x_9); lean_dec(x_7); -x_9 = 0; -return x_9; +x_2 = x_5; +goto _start; } else { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_array_get_size(x_7); -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_nat_dec_lt(x_11, x_10); -if (x_12 == 0) -{ -uint8_t x_13; -lean_dec(x_10); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_7, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); lean_dec(x_7); -x_13 = 0; -return x_13; +x_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +if (x_16 == 0) +{ +lean_dec(x_15); +lean_dec(x_13); +x_2 = x_5; +goto _start; } else { -size_t x_14; size_t x_15; uint8_t x_16; -x_14 = 0; -x_15 = lean_usize_of_nat(x_10); -lean_dec(x_10); -x_16 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(x_2, x_7, x_14, x_15); -lean_dec(x_7); -return x_16; -} +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_13, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +if (x_22 == 0) +{ +lean_dec(x_21); +lean_dec(x_19); +x_2 = x_5; +goto _start; } +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_19, 1); +lean_inc(x_25); +lean_dec(x_19); +x_26 = lean_ctor_get(x_21, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_24, x_26); +if (x_28 == 0) +{ +lean_dec(x_27); +lean_dec(x_25); +x_2 = x_5; +goto _start; } +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_30 = lean_ctor_get(x_25, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_ctor_get(x_27, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_27, 1); +lean_inc(x_33); +lean_dec(x_27); +x_34 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_30, x_32); +lean_dec(x_32); +lean_dec(x_30); +if (x_34 == 0) +{ +lean_dec(x_33); +lean_dec(x_31); +x_2 = x_5; +goto _start; } +else +{ +uint8_t x_36; +x_36 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_31, x_33); +lean_dec(x_33); +lean_dec(x_31); +if (x_36 == 0) +{ +x_2 = x_5; +goto _start; } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(x_1, x_2, x_5, x_6); -lean_dec(x_2); +uint8_t x_38; +lean_dec(x_5); lean_dec(x_1); -x_8 = lean_box(x_7); -return x_8; +x_38 = 1; +return x_38; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(lean_object* x_1, size_t x_2, size_t x_3) { +} +} +} +} +} +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_eq(x_2, x_3); -if (x_4 == 0) -{ -lean_object* x_5; uint8_t x_6; -x_5 = lean_array_uget(x_1, x_2); -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_5); -lean_dec(x_5); -if (x_6 == 0) -{ -uint8_t x_7; -x_7 = 1; -return x_7; -} -else +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -size_t x_8; size_t x_9; +uint64_t x_6; uint64_t x_7; size_t x_8; size_t x_9; +x_6 = 0; +x_7 = lean_uint64_mix_hash(x_4, x_6); x_8 = 1; x_9 = lean_usize_add(x_2, x_8); x_2 = x_9; +x_4 = x_7; goto _start; } -} else { -uint8_t x_11; -x_11 = 0; -return x_11; +return x_4; } } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -switch (lean_obj_tag(x_1)) { -case 0: +if (lean_obj_tag(x_3) == 0) { -uint8_t x_2; -x_2 = 1; +lean_dec(x_1); return x_2; } -case 1: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_3 = lean_ctor_get(x_1, 2); -x_4 = lean_array_get_size(x_3); -x_5 = lean_unsigned_to_nat(0u); -x_6 = lean_nat_dec_lt(x_5, x_4); -if (x_6 == 0) -{ -uint8_t x_7; -lean_dec(x_4); -x_7 = 1; -return x_7; -} else { -size_t x_8; size_t x_9; uint8_t x_10; -x_8 = 0; -x_9 = lean_usize_of_nat(x_4); -lean_dec(x_4); -x_10 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(x_3, x_8, x_9); -if (x_10 == 0) +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) { -uint8_t x_11; -x_11 = 1; -return x_11; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; } else { -uint8_t x_12; -x_12 = 0; -return x_12; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; } } -default: -{ -uint8_t x_13; -x_13 = 0; -return x_13; } } +static uint64_t _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1() { +_start: +{ +uint64_t x_1; uint64_t x_2; +x_1 = 11; +x_2 = lean_uint64_mix_hash(x_1, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static uint64_t _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2() { _start: { -size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_6 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(x_1, x_4, x_5); -lean_dec(x_1); -x_7 = lean_box(x_6); -return x_7; +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 7; +x_2 = 13; +x_3 = lean_uint64_mix_hash(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___boxed(lean_object* x_1) { +static uint64_t _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3() { _start: { -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 11; +x_2 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +x_3 = lean_uint64_mix_hash(x_2, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1() { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +if (lean_obj_tag(x_2) == 0) +{ return x_1; } +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; size_t x_10; size_t x_11; size_t x_12; uint64_t x_13; lean_object* x_25; lean_object* x_26; uint64_t x_27; uint64_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_124; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + x_6 = x_2; +} else { + lean_dec_ref(x_2); + x_6 = lean_box(0); +} +x_7 = lean_array_get_size(x_1); +x_8 = 32; +x_9 = 16; +x_10 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_11 = 1; +x_12 = lean_usize_sub(x_10, x_11); +x_25 = lean_ctor_get(x_3, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_3, 1); +lean_inc(x_26); +x_27 = lean_string_hash(x_25); +lean_dec(x_25); +x_124 = lean_ctor_get(x_26, 0); +lean_inc(x_124); +if (lean_obj_tag(x_124) == 0) +{ +lean_object* x_125; uint64_t x_126; lean_object* x_127; lean_object* x_128; +x_125 = lean_ctor_get(x_26, 1); +lean_inc(x_125); +lean_dec(x_26); +x_126 = 11; +x_127 = lean_ctor_get(x_125, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +lean_dec(x_125); +x_28 = x_126; +x_29 = x_127; +x_30 = x_128; +goto block_123; +} +else +{ +lean_object* x_129; lean_object* x_130; uint64_t x_131; uint64_t x_132; uint64_t x_133; lean_object* x_134; lean_object* x_135; +x_129 = lean_ctor_get(x_26, 1); +lean_inc(x_129); +lean_dec(x_26); +x_130 = lean_ctor_get(x_124, 0); +lean_inc(x_130); +lean_dec(x_124); +x_131 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(x_130); +lean_dec(x_130); +x_132 = 13; +x_133 = lean_uint64_mix_hash(x_131, x_132); +x_134 = lean_ctor_get(x_129, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_129, 1); +lean_inc(x_135); +lean_dec(x_129); +x_28 = x_133; +x_29 = x_134; +x_30 = x_135; +goto block_123; +} +block_24: +{ +uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = lean_uint64_shift_right(x_13, x_8); +x_15 = lean_uint64_xor(x_13, x_14); +x_16 = lean_uint64_shift_right(x_15, x_9); +x_17 = lean_uint64_xor(x_15, x_16); +x_18 = lean_uint64_to_usize(x_17); +x_19 = lean_usize_land(x_18, x_12); +x_20 = lean_array_uget(x_1, x_19); +if (lean_is_scalar(x_6)) { + x_21 = lean_alloc_ctor(1, 3, 0); +} else { + x_21 = x_6; +} +lean_ctor_set(x_21, 0, x_3); +lean_ctor_set(x_21, 1, x_4); +lean_ctor_set(x_21, 2, x_20); +x_22 = lean_array_uset(x_1, x_19, x_21); +x_1 = x_22; +x_2 = x_5; +goto _start; +} +block_123: +{ +uint64_t x_31; +if (lean_obj_tag(x_29) == 0) +{ +uint64_t x_118; +x_118 = 11; +x_31 = x_118; +goto block_117; } -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_119; uint64_t x_120; uint64_t x_121; uint64_t x_122; +x_119 = lean_ctor_get(x_29, 0); +lean_inc(x_119); +lean_dec(x_29); +x_120 = lean_string_hash(x_119); +lean_dec(x_119); +x_121 = 13; +x_122 = lean_uint64_mix_hash(x_120, x_121); +x_31 = x_122; +goto block_117; } +block_117: +{ +uint64_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_104; +x_104 = lean_ctor_get(x_30, 0); +lean_inc(x_104); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_105; uint64_t x_106; lean_object* x_107; lean_object* x_108; +x_105 = lean_ctor_get(x_30, 1); +lean_inc(x_105); +lean_dec(x_30); +x_106 = 11; +x_107 = lean_ctor_get(x_105, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_105, 1); +lean_inc(x_108); +lean_dec(x_105); +x_32 = x_106; +x_33 = x_107; +x_34 = x_108; +goto block_103; } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; uint8_t x_4; -lean_inc(x_1); -x_2 = l_Lean_Syntax_getKind(x_1); -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2; -x_4 = lean_name_eq(x_2, x_3); -if (x_4 == 0) +lean_object* x_109; lean_object* x_110; uint8_t x_111; uint64_t x_112; uint64_t x_113; uint64_t x_114; lean_object* x_115; lean_object* x_116; +x_109 = lean_ctor_get(x_30, 1); +lean_inc(x_109); +lean_dec(x_30); +x_110 = lean_ctor_get(x_104, 0); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_unbox(x_110); +lean_dec(x_110); +x_112 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(x_111); +x_113 = 13; +x_114 = lean_uint64_mix_hash(x_112, x_113); +x_115 = lean_ctor_get(x_109, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_109, 1); +lean_inc(x_116); +lean_dec(x_109); +x_32 = x_114; +x_33 = x_115; +x_34 = x_116; +goto block_103; +} +block_103: { -lean_object* x_5; uint8_t x_6; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; -x_6 = lean_name_eq(x_2, x_5); -if (x_6 == 0) +if (lean_obj_tag(x_33) == 0) { -lean_object* x_7; uint8_t x_8; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; -x_8 = lean_name_eq(x_2, x_7); -lean_dec(x_2); -if (x_8 == 0) +if (lean_obj_tag(x_34) == 0) { -uint8_t x_9; -lean_dec(x_1); -x_9 = 0; -return x_9; +uint64_t x_35; uint64_t x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; +x_35 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1; +x_36 = lean_uint64_mix_hash(x_32, x_35); +x_37 = lean_uint64_mix_hash(x_31, x_36); +x_38 = lean_uint64_mix_hash(x_28, x_37); +x_39 = lean_uint64_mix_hash(x_27, x_38); +x_13 = x_39; +goto block_24; } else { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = lean_unsigned_to_nat(1u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); -lean_dec(x_1); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_11); -lean_dec(x_11); -return x_12; +lean_object* x_40; uint64_t x_41; uint64_t x_42; uint64_t x_43; uint64_t x_44; uint64_t x_45; uint64_t x_46; uint64_t x_47; uint64_t x_48; uint64_t x_49; +x_40 = lean_ctor_get(x_34, 0); +lean_inc(x_40); +lean_dec(x_34); +x_41 = 11; +x_42 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_40); +lean_dec(x_40); +x_43 = 13; +x_44 = lean_uint64_mix_hash(x_42, x_43); +x_45 = lean_uint64_mix_hash(x_41, x_44); +x_46 = lean_uint64_mix_hash(x_32, x_45); +x_47 = lean_uint64_mix_hash(x_31, x_46); +x_48 = lean_uint64_mix_hash(x_28, x_47); +x_49 = lean_uint64_mix_hash(x_27, x_48); +x_13 = x_49; +goto block_24; } } else { -uint8_t x_13; -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); -if (x_13 == 0) +lean_object* x_50; uint64_t x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; uint64_t x_55; +x_50 = lean_ctor_get(x_33, 0); +lean_inc(x_50); +lean_dec(x_33); +x_51 = 7; +x_52 = lean_array_get_size(x_50); +x_53 = lean_unsigned_to_nat(0u); +x_54 = lean_nat_dec_lt(x_53, x_52); +x_55 = 13; +if (x_54 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; -x_15 = lean_name_eq(x_2, x_14); -lean_dec(x_2); -if (x_15 == 0) +lean_dec(x_52); +lean_dec(x_50); +if (lean_obj_tag(x_34) == 0) { -uint8_t x_16; -lean_dec(x_1); -x_16 = 0; -return x_16; +uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; +x_56 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3; +x_57 = lean_uint64_mix_hash(x_32, x_56); +x_58 = lean_uint64_mix_hash(x_31, x_57); +x_59 = lean_uint64_mix_hash(x_28, x_58); +x_60 = lean_uint64_mix_hash(x_27, x_59); +x_13 = x_60; +goto block_24; } else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_unsigned_to_nat(1u); -x_18 = l_Lean_Syntax_getArg(x_1, x_17); -lean_dec(x_1); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_18); -lean_dec(x_18); -return x_19; +lean_object* x_61; uint64_t x_62; uint64_t x_63; uint64_t x_64; uint64_t x_65; uint64_t x_66; uint64_t x_67; uint64_t x_68; uint64_t x_69; +x_61 = lean_ctor_get(x_34, 0); +lean_inc(x_61); +lean_dec(x_34); +x_62 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_61); +lean_dec(x_61); +x_63 = lean_uint64_mix_hash(x_62, x_55); +x_64 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +x_65 = lean_uint64_mix_hash(x_64, x_63); +x_66 = lean_uint64_mix_hash(x_32, x_65); +x_67 = lean_uint64_mix_hash(x_31, x_66); +x_68 = lean_uint64_mix_hash(x_28, x_67); +x_69 = lean_uint64_mix_hash(x_27, x_68); +x_13 = x_69; +goto block_24; } } else { -uint8_t x_20; -lean_dec(x_2); -lean_dec(x_1); -x_20 = 1; -return x_20; +uint8_t x_70; +x_70 = lean_nat_dec_le(x_52, x_52); +if (x_70 == 0) +{ +lean_dec(x_52); +lean_dec(x_50); +if (lean_obj_tag(x_34) == 0) +{ +uint64_t x_71; uint64_t x_72; uint64_t x_73; uint64_t x_74; uint64_t x_75; +x_71 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3; +x_72 = lean_uint64_mix_hash(x_32, x_71); +x_73 = lean_uint64_mix_hash(x_31, x_72); +x_74 = lean_uint64_mix_hash(x_28, x_73); +x_75 = lean_uint64_mix_hash(x_27, x_74); +x_13 = x_75; +goto block_24; } +else +{ +lean_object* x_76; uint64_t x_77; uint64_t x_78; uint64_t x_79; uint64_t x_80; uint64_t x_81; uint64_t x_82; uint64_t x_83; uint64_t x_84; +x_76 = lean_ctor_get(x_34, 0); +lean_inc(x_76); +lean_dec(x_34); +x_77 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_76); +lean_dec(x_76); +x_78 = lean_uint64_mix_hash(x_77, x_55); +x_79 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +x_80 = lean_uint64_mix_hash(x_79, x_78); +x_81 = lean_uint64_mix_hash(x_32, x_80); +x_82 = lean_uint64_mix_hash(x_31, x_81); +x_83 = lean_uint64_mix_hash(x_28, x_82); +x_84 = lean_uint64_mix_hash(x_27, x_83); +x_13 = x_84; +goto block_24; } } else { -uint8_t x_21; -x_21 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); -if (x_21 == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; -x_23 = lean_name_eq(x_2, x_22); -lean_dec(x_2); -if (x_23 == 0) +size_t x_85; size_t x_86; uint64_t x_87; uint64_t x_88; +x_85 = 0; +x_86 = lean_usize_of_nat(x_52); +lean_dec(x_52); +x_87 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6(x_50, x_85, x_86, x_51); +lean_dec(x_50); +x_88 = lean_uint64_mix_hash(x_87, x_55); +if (lean_obj_tag(x_34) == 0) { -uint8_t x_24; -lean_dec(x_1); -x_24 = 0; -return x_24; +uint64_t x_89; uint64_t x_90; uint64_t x_91; uint64_t x_92; uint64_t x_93; uint64_t x_94; +x_89 = 11; +x_90 = lean_uint64_mix_hash(x_88, x_89); +x_91 = lean_uint64_mix_hash(x_32, x_90); +x_92 = lean_uint64_mix_hash(x_31, x_91); +x_93 = lean_uint64_mix_hash(x_28, x_92); +x_94 = lean_uint64_mix_hash(x_27, x_93); +x_13 = x_94; +goto block_24; } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_unsigned_to_nat(1u); -x_26 = l_Lean_Syntax_getArg(x_1, x_25); -lean_dec(x_1); -x_27 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_26); -lean_dec(x_26); -return x_27; +lean_object* x_95; uint64_t x_96; uint64_t x_97; uint64_t x_98; uint64_t x_99; uint64_t x_100; uint64_t x_101; uint64_t x_102; +x_95 = lean_ctor_get(x_34, 0); +lean_inc(x_95); +lean_dec(x_34); +x_96 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_95); +lean_dec(x_95); +x_97 = lean_uint64_mix_hash(x_96, x_55); +x_98 = lean_uint64_mix_hash(x_88, x_97); +x_99 = lean_uint64_mix_hash(x_32, x_98); +x_100 = lean_uint64_mix_hash(x_31, x_99); +x_101 = lean_uint64_mix_hash(x_28, x_100); +x_102 = lean_uint64_mix_hash(x_27, x_101); +x_13 = x_102; +goto block_24; } } -else -{ -uint8_t x_28; -lean_dec(x_2); -lean_dec(x_1); -x_28 = 1; -return x_28; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(x_1); -x_3 = lean_box(x_2); -return x_3; } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_4; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCursorInProperWhitespace(x_1, x_2); -if (x_4 == 0) +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) { -uint8_t x_5; -lean_dec(x_3); -x_5 = 0; -return x_5; +lean_dec(x_2); +lean_dec(x_1); +return x_3; } else { -uint8_t x_6; -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(x_3); -return x_6; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__4(lean_object* x_1) { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_3); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); lean_dec(x_2); -lean_dec(x_1); -x_5 = lean_box(x_4); -return x_5; +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__5(x_7, x_1, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_Syntax_getTrailingSize(x_1); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_5); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, size_t x_9, size_t x_10, lean_object* x_11) { -_start: +else { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_10, x_9); -if (x_12 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) { -lean_dec(x_7); -lean_dec(x_1); -return x_11; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_6, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +x_13 = lean_string_dec_eq(x_9, x_11); +lean_dec(x_11); +lean_dec(x_9); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_14); +return x_3; } else { -lean_object* x_13; uint8_t x_14; -x_13 = lean_array_uget(x_8, x_10); -x_14 = !lean_is_exclusive(x_11); -if (x_14 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_10, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_dec(x_10); +x_17 = lean_ctor_get(x_12, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +if (x_19 == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_11, 1); -x_16 = lean_ctor_get(x_11, 0); +lean_object* x_20; +lean_dec(x_18); lean_dec(x_16); -lean_inc(x_15); -lean_inc(x_13); -lean_inc(x_1); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_13, x_15); -if (x_17 == 0) +x_20 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_20); +return x_3; +} +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; -lean_free_object(x_11); -x_18 = lean_box(0); -lean_inc(x_7); -x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_13, x_7, x_15, x_18); -lean_dec(x_15); -lean_dec(x_13); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -lean_dec(x_19); -x_21 = 1; -x_22 = lean_usize_add(x_10, x_21); -x_10 = x_22; -x_11 = x_20; -goto _start; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_ctor_get(x_16, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = lean_ctor_get(x_18, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_dec(x_18); +x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_21, x_23); +lean_dec(x_23); +lean_dec(x_21); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_24); +lean_dec(x_22); +x_26 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_26); +return x_3; } else { -lean_object* x_24; -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_1); -x_24 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1; -lean_ctor_set(x_11, 0, x_24); -return x_11; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_27 = lean_ctor_get(x_22, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec(x_22); +x_29 = lean_ctor_get(x_24, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +x_31 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_27, x_29); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec(x_30); +lean_dec(x_28); +x_32 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_32); +return x_3; } +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +x_35 = lean_ctor_get(x_30, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_30, 1); +lean_inc(x_36); +lean_dec(x_30); +x_37 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_33, x_35); +lean_dec(x_35); +lean_dec(x_33); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_36); +lean_dec(x_34); +x_38 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_38); +return x_3; } else { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_11, 1); -lean_inc(x_25); -lean_dec(x_11); -lean_inc(x_25); -lean_inc(x_13); -lean_inc(x_1); -x_26 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_13, x_25); -if (x_26 == 0) +uint8_t x_39; +x_39 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_34, x_36); +lean_dec(x_36); +lean_dec(x_34); +if (x_39 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; -x_27 = lean_box(0); -lean_inc(x_7); -x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_13, x_7, x_25, x_27); -lean_dec(x_25); -lean_dec(x_13); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -lean_dec(x_28); -x_30 = 1; -x_31 = lean_usize_add(x_10, x_30); -x_10 = x_31; -x_11 = x_29; -goto _start; +lean_object* x_40; +x_40 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_40); +return x_3; } else { -lean_object* x_33; lean_object* x_34; -lean_dec(x_13); lean_dec(x_7); -lean_dec(x_1); -x_33 = l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1; -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_25); -return x_34; +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} } } } } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; -lean_inc(x_3); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_3); -if (x_7 == 0) -{ -uint8_t x_8; -lean_inc(x_3); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(x_1, x_2, x_3); -if (x_8 == 0) -{ -uint8_t x_9; -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(x_1, x_2, x_4, x_5, x_3); -return x_9; } else { -uint8_t x_10; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_41 = lean_ctor_get(x_3, 0); +x_42 = lean_ctor_get(x_3, 1); +x_43 = lean_ctor_get(x_3, 2); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); lean_dec(x_3); -lean_dec(x_1); -x_10 = 1; -return x_10; +x_44 = lean_ctor_get(x_41, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_41, 1); +lean_inc(x_45); +x_46 = lean_ctor_get(x_1, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_1, 1); +lean_inc(x_47); +x_48 = lean_string_dec_eq(x_44, x_46); +lean_dec(x_46); +lean_dec(x_44); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_dec(x_47); +lean_dec(x_45); +x_49 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_50 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 2, x_49); +return x_50; } +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_51 = lean_ctor_get(x_45, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_45, 1); +lean_inc(x_52); +lean_dec(x_45); +x_53 = lean_ctor_get(x_47, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_47, 1); +lean_inc(x_54); +lean_dec(x_47); +x_55 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_51, x_53); +lean_dec(x_53); +lean_dec(x_51); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_54); +lean_dec(x_52); +x_56 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_57 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_57, 0, x_41); +lean_ctor_set(x_57, 1, x_42); +lean_ctor_set(x_57, 2, x_56); +return x_57; } else { -uint8_t x_11; -lean_dec(x_3); -lean_dec(x_1); -x_11 = 1; -return x_11; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_58 = lean_ctor_get(x_52, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_52, 1); +lean_inc(x_59); +lean_dec(x_52); +x_60 = lean_ctor_get(x_54, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_54, 1); +lean_inc(x_61); +lean_dec(x_54); +x_62 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_58, x_60); +lean_dec(x_60); +lean_dec(x_58); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_61); +lean_dec(x_59); +x_63 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_41); +lean_ctor_set(x_64, 1, x_42); +lean_ctor_set(x_64, 2, x_63); +return x_64; } +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_65 = lean_ctor_get(x_59, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_59, 1); +lean_inc(x_66); +lean_dec(x_59); +x_67 = lean_ctor_get(x_61, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_61, 1); +lean_inc(x_68); +lean_dec(x_61); +x_69 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_65, x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +lean_dec(x_68); +lean_dec(x_66); +x_70 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_71 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_71, 0, x_41); +lean_ctor_set(x_71, 1, x_42); +lean_ctor_set(x_71, 2, x_70); +return x_71; } +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_72 = lean_ctor_get(x_66, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_66, 1); +lean_inc(x_73); +lean_dec(x_66); +x_74 = lean_ctor_get(x_68, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_68, 1); +lean_inc(x_75); +lean_dec(x_68); +x_76 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_72, x_74); +lean_dec(x_74); +lean_dec(x_72); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_75); +lean_dec(x_73); +x_77 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_41); +lean_ctor_set(x_78, 1, x_42); +lean_ctor_set(x_78, 2, x_77); +return x_78; } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -x_8 = l_Lean_Syntax_getArgs(x_1); -x_9 = lean_box(0); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_2); -x_12 = lean_array_size(x_8); -x_13 = 0; -lean_inc(x_3); -x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(x_3, x_4, x_5, x_6, x_8, x_9, x_10, x_8, x_12, x_13, x_11); -lean_dec(x_8); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) +uint8_t x_79; +x_79 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_73, x_75); +lean_dec(x_75); +lean_dec(x_73); +if (x_79 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_box(0); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(x_3, x_4, x_1, x_5, x_6, x_16); -return x_17; +lean_object* x_80; lean_object* x_81; +x_80 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_1, x_2, x_43); +x_81 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_81, 0, x_41); +lean_ctor_set(x_81, 1, x_42); +lean_ctor_set(x_81, 2, x_80); +return x_81; } else { -lean_object* x_18; uint8_t x_19; -lean_dec(x_3); -lean_dec(x_1); -x_18 = lean_ctor_get(x_15, 0); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -return x_19; +lean_object* x_82; +lean_dec(x_42); +lean_dec(x_41); +x_82 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_82, 0, x_1); +lean_ctor_set(x_82, 1, x_2); +lean_ctor_set(x_82, 2, x_43); +return x_82; } } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_7; lean_object* x_8; -x_7 = 0; -x_8 = l_Lean_Syntax_getPos_x3f(x_5, x_7); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_2) == 0) { -uint8_t x_9; -lean_dec(x_6); -x_9 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_5); +lean_object* x_3; lean_dec(x_1); -return x_9; +x_3 = lean_box(0); +return x_3; } else { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_8, 0); +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 0); lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +x_12 = lean_string_dec_eq(x_8, x_10); +lean_dec(x_10); lean_dec(x_8); -x_11 = l_Lean_Syntax_getTailPos_x3f(x_5, x_7); -if (lean_obj_tag(x_11) == 0) +if (x_12 == 0) { -uint8_t x_12; -lean_dec(x_10); -lean_dec(x_6); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_5); -lean_dec(x_1); -return x_12; +lean_object* x_13; +lean_dec(x_11); +lean_dec(x_9); +x_13 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_13); +return x_2; } else { -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_9, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = lean_ctor_get(x_11, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); lean_dec(x_11); -x_14 = lean_nat_sub(x_10, x_6); -lean_dec(x_10); -x_15 = lean_nat_dec_le(x_14, x_2); +x_18 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_14, x_16); +lean_dec(x_16); lean_dec(x_14); -if (x_15 == 0) +if (x_18 == 0) { -uint8_t x_16; -lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_16 = 0; -return x_16; +lean_object* x_19; +lean_dec(x_17); +lean_dec(x_15); +x_19 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_19); +return x_2; } else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_17 = l_Lean_Syntax_getTrailingSize(x_5); -x_18 = lean_nat_add(x_13, x_17); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_dec(x_15); +x_22 = lean_ctor_get(x_17, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); lean_dec(x_17); -lean_dec(x_13); -x_19 = lean_nat_dec_le(x_2, x_18); -lean_dec(x_18); -if (x_19 == 0) +x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_20, x_22); +lean_dec(x_22); +lean_dec(x_20); +if (x_24 == 0) { -uint8_t x_20; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_20 = 0; -return x_20; +lean_object* x_25; +lean_dec(x_23); +lean_dec(x_21); +x_25 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_25); +return x_2; } else { -lean_object* x_21; uint8_t x_22; -x_21 = lean_box(0); -x_22 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(x_5, x_6, x_1, x_2, x_3, x_4, x_21); -return x_22; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_26 = lean_ctor_get(x_21, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_dec(x_21); +x_28 = lean_ctor_get(x_23, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_dec(x_23); +x_30 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_26, x_28); +if (x_30 == 0) { -lean_object* x_5; -x_5 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_5; -} +lean_object* x_31; +lean_dec(x_29); +lean_dec(x_27); +x_31 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_31); +return x_2; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: +else { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_9); -lean_dec(x_9); -x_13 = lean_unbox_usize(x_10); -lean_dec(x_10); -x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_13, x_11); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_27, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_27, 1); +lean_inc(x_33); +lean_dec(x_27); +x_34 = lean_ctor_get(x_29, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_29, 1); +lean_inc(x_35); +lean_dec(x_29); +x_36 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_32, x_34); +lean_dec(x_34); +lean_dec(x_32); +if (x_36 == 0) { -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_8 = lean_box(x_7); -return x_8; +lean_object* x_37; +lean_dec(x_35); +lean_dec(x_33); +x_37 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_37); +return x_2; } +else +{ +uint8_t x_38; +x_38 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_33, x_35); +lean_dec(x_35); +lean_dec(x_33); +if (x_38 == 0) +{ +lean_object* x_39; +x_39 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_7); +lean_ctor_set(x_2, 2, x_39); +return x_2; } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: +else { -uint8_t x_8; lean_object* x_9; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_7); +lean_free_object(x_2); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_9 = lean_box(x_8); -return x_9; +lean_dec(x_1); +return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_8 = lean_box(x_7); -return x_8; } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_5, x_4, x_7); -return x_8; } } -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Id_instMonad; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -x_3 = lean_ctor_get(x_2, 1); -lean_inc(x_3); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_40 = lean_ctor_get(x_2, 0); +x_41 = lean_ctor_get(x_2, 1); +x_42 = lean_ctor_get(x_2, 2); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); lean_dec(x_2); -x_4 = lean_box(0); -x_5 = lean_apply_2(x_3, lean_box(0), x_4); -return x_5; -} +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +x_45 = lean_ctor_get(x_1, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_1, 1); +lean_inc(x_46); +x_47 = lean_string_dec_eq(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_46); +lean_dec(x_44); +x_48 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_40); +lean_ctor_set(x_49, 1, x_41); +lean_ctor_set(x_49, 2, x_48); +return x_49; } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -lean_inc(x_1); -x_4 = l_Lean_FileMap_toPosition(x_1, x_2); -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -lean_inc(x_1); -x_6 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount(x_1, x_5); -lean_dec(x_5); -x_7 = lean_ctor_get(x_4, 1); -lean_inc(x_7); -x_8 = lean_nat_dec_lt(x_7, x_6); -if (x_8 == 0) +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_50 = lean_ctor_get(x_44, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_44, 1); +lean_inc(x_51); +lean_dec(x_44); +x_52 = lean_ctor_get(x_46, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_46, 1); +lean_inc(x_53); +lean_dec(x_46); +x_54 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__4(x_50, x_52); +lean_dec(x_52); +lean_dec(x_50); +if (x_54 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -lean_dec(x_7); -x_9 = l_Id_instMonad; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -x_11 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed), 6, 5); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_2); -lean_closure_set(x_11, 2, x_4); -lean_closure_set(x_11, 3, x_3); -lean_closure_set(x_11, 4, x_6); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; -x_13 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_12, x_11); -return x_13; +lean_object* x_55; lean_object* x_56; +lean_dec(x_53); +lean_dec(x_51); +x_55 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_40); +lean_ctor_set(x_56, 1, x_41); +lean_ctor_set(x_56, 2, x_55); +return x_56; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -lean_dec(x_6); -x_14 = l_Id_instMonad; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed), 6, 5); -lean_closure_set(x_16, 0, x_1); -lean_closure_set(x_16, 1, x_2); -lean_closure_set(x_16, 2, x_4); -lean_closure_set(x_16, 3, x_3); -lean_closure_set(x_16, 4, x_7); -x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; -x_18 = lean_apply_4(x_15, lean_box(0), lean_box(0), x_17, x_16); -return x_18; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_57 = lean_ctor_get(x_51, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_51, 1); +lean_inc(x_58); +lean_dec(x_51); +x_59 = lean_ctor_get(x_53, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_53, 1); +lean_inc(x_60); +lean_dec(x_53); +x_61 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_57, x_59); +lean_dec(x_59); +lean_dec(x_57); +if (x_61 == 0) { -uint8_t x_7; lean_object* x_8; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_8 = lean_box(x_7); -return x_8; -} +lean_object* x_62; lean_object* x_63; +lean_dec(x_60); +lean_dec(x_58); +x_62 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_63 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_63, 0, x_40); +lean_ctor_set(x_63, 1, x_41); +lean_ctor_set(x_63, 2, x_62); +return x_63; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: +else { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_64 = lean_ctor_get(x_58, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_58, 1); +lean_inc(x_65); +lean_dec(x_58); +x_66 = lean_ctor_get(x_60, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_60, 1); +lean_inc(x_67); +lean_dec(x_60); +x_68 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__3(x_64, x_66); +if (x_68 == 0) { -lean_inc(x_7); -return x_7; +lean_object* x_69; lean_object* x_70; +lean_dec(x_67); +lean_dec(x_65); +x_69 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_70 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_70, 0, x_40); +lean_ctor_set(x_70, 1, x_41); +lean_ctor_set(x_70, 2, x_69); +return x_70; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_9); -lean_dec(x_9); -if (lean_obj_tag(x_10) == 0) -{ -size_t x_11; size_t x_12; -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_71 = lean_ctor_get(x_65, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_65, 1); +lean_inc(x_72); +lean_dec(x_65); +x_73 = lean_ctor_get(x_67, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_67, 1); +lean_inc(x_74); +lean_dec(x_67); +x_75 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__6(x_71, x_73); +lean_dec(x_73); +lean_dec(x_71); +if (x_75 == 0) { -size_t _tmp_5 = x_12; -lean_object* _tmp_6 = x_3; -x_6 = _tmp_5; -x_7 = _tmp_6; -} -goto _start; +lean_object* x_76; lean_object* x_77; +lean_dec(x_74); +lean_dec(x_72); +x_76 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_40); +lean_ctor_set(x_77, 1, x_41); +lean_ctor_set(x_77, 2, x_76); +return x_77; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +uint8_t x_78; +x_78 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__2(x_72, x_74); +lean_dec(x_74); +lean_dec(x_72); +if (x_78 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_10); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_79; lean_object* x_80; +x_79 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_1, x_42); +x_80 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_80, 0, x_40); +lean_ctor_set(x_80, 1, x_41); +lean_ctor_set(x_80, 2, x_79); +return x_80; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_1); +return x_42; +} +} +} } } } } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1() { _start: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_inc(x_7); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, size_t x_9, lean_object* x_10) { +_start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_9); -if (lean_obj_tag(x_10) == 0) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_9, x_8); +if (x_11 == 0) { -size_t x_11; size_t x_12; -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); +lean_dec(x_5); +lean_dec(x_1); +return x_10; +} +else { -size_t _tmp_5 = x_12; -lean_object* _tmp_6 = x_3; -x_6 = _tmp_5; -x_7 = _tmp_6; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; size_t x_20; size_t x_21; size_t x_22; uint64_t x_23; lean_object* x_131; lean_object* x_132; uint64_t x_133; uint64_t x_134; lean_object* x_135; lean_object* x_136; lean_object* x_230; +x_12 = lean_array_uget(x_7, x_9); +lean_inc(x_1); +lean_inc(x_12); +x_13 = lean_apply_1(x_1, x_12); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_15); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_16 = x_10; +} else { + lean_dec_ref(x_10); + x_16 = lean_box(0); } -goto _start; +x_17 = lean_array_get_size(x_15); +x_18 = 32; +x_19 = 16; +x_20 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_21 = 1; +x_22 = lean_usize_sub(x_20, x_21); +x_131 = lean_ctor_get(x_13, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_13, 1); +lean_inc(x_132); +x_133 = lean_string_hash(x_131); +lean_dec(x_131); +x_230 = lean_ctor_get(x_132, 0); +lean_inc(x_230); +if (lean_obj_tag(x_230) == 0) +{ +lean_object* x_231; uint64_t x_232; lean_object* x_233; lean_object* x_234; +x_231 = lean_ctor_get(x_132, 1); +lean_inc(x_231); +lean_dec(x_132); +x_232 = 11; +x_233 = lean_ctor_get(x_231, 0); +lean_inc(x_233); +x_234 = lean_ctor_get(x_231, 1); +lean_inc(x_234); +lean_dec(x_231); +x_134 = x_232; +x_135 = x_233; +x_136 = x_234; +goto block_229; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +lean_object* x_235; lean_object* x_236; uint64_t x_237; uint64_t x_238; uint64_t x_239; lean_object* x_240; lean_object* x_241; +x_235 = lean_ctor_get(x_132, 1); +lean_inc(x_235); +lean_dec(x_132); +x_236 = lean_ctor_get(x_230, 0); +lean_inc(x_236); +lean_dec(x_230); +x_237 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1699_(x_236); +lean_dec(x_236); +x_238 = 13; +x_239 = lean_uint64_mix_hash(x_237, x_238); +x_240 = lean_ctor_get(x_235, 0); +lean_inc(x_240); +x_241 = lean_ctor_get(x_235, 1); +lean_inc(x_241); +lean_dec(x_235); +x_134 = x_239; +x_135 = x_240; +x_136 = x_241; +goto block_229; +} +block_130: +{ +uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_uint64_shift_right(x_23, x_18); +x_25 = lean_uint64_xor(x_23, x_24); +x_26 = lean_uint64_shift_right(x_25, x_19); +x_27 = lean_uint64_xor(x_25, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_land(x_28, x_22); +x_30 = lean_array_uget(x_15, x_29); +lean_inc(x_30); +lean_inc(x_13); +x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__2(x_13, x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_10); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -else +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +x_33 = lean_array_push(x_32, x_12); +lean_inc(x_30); +lean_inc(x_13); +x_34 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(x_13, x_30); +if (x_34 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_35 = lean_unsigned_to_nat(1u); +x_36 = lean_nat_add(x_14, x_35); +lean_dec(x_14); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_13); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_30); +x_38 = lean_array_uset(x_15, x_29, x_37); +x_39 = lean_unsigned_to_nat(4u); +x_40 = lean_nat_mul(x_36, x_39); +x_41 = lean_unsigned_to_nat(3u); +x_42 = lean_nat_div(x_40, x_41); +lean_dec(x_40); +x_43 = lean_array_get_size(x_38); +x_44 = lean_nat_dec_le(x_42, x_43); +lean_dec(x_43); +lean_dec(x_42); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; size_t x_47; +x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__4(x_38); +if (lean_is_scalar(x_16)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_16; } +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_usize_add(x_9, x_21); +x_9 = x_47; +x_10 = x_46; +goto _start; } +else +{ +lean_object* x_49; size_t x_50; +if (lean_is_scalar(x_16)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_16; } +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_38); +x_50 = lean_usize_add(x_9, x_21); +x_9 = x_50; +x_10 = x_49; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_box(0); -return x_2; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; size_t x_56; +lean_inc(x_5); +x_52 = lean_array_uset(x_15, x_29, x_5); +x_53 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_13, x_33, x_30); +x_54 = lean_array_uset(x_52, x_29, x_53); +if (lean_is_scalar(x_16)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_16; +} +lean_ctor_set(x_55, 0, x_14); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_usize_add(x_9, x_21); +x_9 = x_56; +x_10 = x_55; +goto _start; } } -static lean_object* _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed), 1, 0); -return x_1; +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_122; +x_58 = lean_ctor_get(x_31, 0); +lean_inc(x_58); +lean_dec(x_31); +x_59 = lean_array_push(x_58, x_12); +lean_inc(x_30); +lean_inc(x_13); +x_122 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(x_13, x_30); +if (x_122 == 0) +{ +lean_object* x_123; +lean_dec(x_30); +if (lean_is_scalar(x_16)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_16; } +lean_ctor_set(x_123, 0, x_14); +lean_ctor_set(x_123, 1, x_15); +x_60 = x_123; +goto block_121; } -static lean_object* _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; -x_2 = lean_box(0); -x_3 = lean_apply_1(x_1, x_2); -return x_3; +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_inc(x_5); +x_124 = lean_array_uset(x_15, x_29, x_5); +x_125 = lean_unsigned_to_nat(1u); +x_126 = lean_nat_sub(x_14, x_125); +lean_dec(x_14); +lean_inc(x_13); +x_127 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__10(x_13, x_30); +x_128 = lean_array_uset(x_124, x_29, x_127); +if (lean_is_scalar(x_16)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_16; } +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_128); +x_60 = x_129; +goto block_121; } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(lean_object* x_1) { -_start: +block_121: { -if (lean_obj_tag(x_1) == 0) +uint8_t x_61; +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) { -lean_object* x_2; lean_object* x_3; size_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = lean_ctor_get(x_1, 0); -x_3 = lean_box(0); -x_4 = lean_array_size(x_2); -x_5 = 0; -x_6 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -x_7 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(x_2, x_3, x_6, x_2, x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -lean_dec(x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; size_t x_67; lean_object* x_68; uint8_t x_69; +x_62 = lean_ctor_get(x_60, 0); +x_63 = lean_ctor_get(x_60, 1); +x_64 = lean_array_get_size(x_63); +x_65 = lean_usize_of_nat(x_64); +lean_dec(x_64); +x_66 = lean_usize_sub(x_65, x_21); +x_67 = lean_usize_land(x_28, x_66); +x_68 = lean_array_uget(x_63, x_67); +lean_inc(x_68); +lean_inc(x_13); +x_69 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(x_13, x_68); +if (x_69 == 0) { -lean_object* x_9; -x_9 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; -return x_9; -} -else +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_70 = lean_unsigned_to_nat(1u); +x_71 = lean_nat_add(x_62, x_70); +lean_dec(x_62); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_13); +lean_ctor_set(x_72, 1, x_59); +lean_ctor_set(x_72, 2, x_68); +x_73 = lean_array_uset(x_63, x_67, x_72); +x_74 = lean_unsigned_to_nat(4u); +x_75 = lean_nat_mul(x_71, x_74); +x_76 = lean_unsigned_to_nat(3u); +x_77 = lean_nat_div(x_75, x_76); +lean_dec(x_75); +x_78 = lean_array_get_size(x_73); +x_79 = lean_nat_dec_le(x_77, x_78); +lean_dec(x_78); +lean_dec(x_77); +if (x_79 == 0) { -lean_object* x_10; -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -lean_dec(x_8); -return x_10; -} +lean_object* x_80; size_t x_81; +x_80 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__4(x_73); +lean_ctor_set(x_60, 1, x_80); +lean_ctor_set(x_60, 0, x_71); +x_81 = lean_usize_add(x_9, x_21); +x_9 = x_81; +x_10 = x_60; +goto _start; } else { -lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_11 = lean_ctor_get(x_1, 0); -x_12 = lean_box(0); -x_13 = lean_array_size(x_11); -x_14 = 0; -x_15 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(x_11, x_12, x_15, x_11, x_13, x_14, x_15); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; -x_18 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; -return x_18; +size_t x_83; +lean_ctor_set(x_60, 1, x_73); +lean_ctor_set(x_60, 0, x_71); +x_83 = lean_usize_add(x_9, x_21); +x_9 = x_83; +x_10 = x_60; +goto _start; +} } else { -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -return x_19; -} -} -} +lean_object* x_85; lean_object* x_86; lean_object* x_87; size_t x_88; +lean_inc(x_5); +x_85 = lean_array_uset(x_63, x_67, x_5); +x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_13, x_59, x_68); +x_87 = lean_array_uset(x_85, x_67, x_86); +lean_ctor_set(x_60, 1, x_87); +x_88 = lean_usize_add(x_9, x_21); +x_9 = x_88; +x_10 = x_60; +goto _start; } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: -{ -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) -{ -lean_inc(x_7); -return x_7; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_array_uget(x_4, x_6); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; size_t x_94; size_t x_95; lean_object* x_96; uint8_t x_97; +x_90 = lean_ctor_get(x_60, 0); +x_91 = lean_ctor_get(x_60, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_60); +x_92 = lean_array_get_size(x_91); +x_93 = lean_usize_of_nat(x_92); +lean_dec(x_92); +x_94 = lean_usize_sub(x_93, x_21); +x_95 = lean_usize_land(x_28, x_94); +x_96 = lean_array_uget(x_91, x_95); +lean_inc(x_96); +lean_inc(x_13); +x_97 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(x_13, x_96); +if (x_97 == 0) { -size_t x_11; size_t x_12; -x_11 = 1; -x_12 = lean_usize_add(x_6, x_11); +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_98 = lean_unsigned_to_nat(1u); +x_99 = lean_nat_add(x_90, x_98); +lean_dec(x_90); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_13); +lean_ctor_set(x_100, 1, x_59); +lean_ctor_set(x_100, 2, x_96); +x_101 = lean_array_uset(x_91, x_95, x_100); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_mul(x_99, x_102); +x_104 = lean_unsigned_to_nat(3u); +x_105 = lean_nat_div(x_103, x_104); +lean_dec(x_103); +x_106 = lean_array_get_size(x_101); +x_107 = lean_nat_dec_le(x_105, x_106); +lean_dec(x_106); +lean_dec(x_105); +if (x_107 == 0) { -size_t _tmp_5 = x_12; -lean_object* _tmp_6 = x_3; -x_6 = _tmp_5; -x_7 = _tmp_6; -} +lean_object* x_108; lean_object* x_109; size_t x_110; +x_108 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__4(x_101); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_99); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_usize_add(x_9, x_21); +x_9 = x_110; +x_10 = x_109; goto _start; } else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_10); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_112; size_t x_113; +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_99); +lean_ctor_set(x_112, 1, x_101); +x_113 = lean_usize_add(x_9, x_21); +x_9 = x_113; +x_10 = x_112; +goto _start; +} } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_10, 0); -lean_inc(x_18); -lean_dec(x_10); -x_19 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_19); -x_21 = lean_box(0); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -return x_22; +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; +lean_inc(x_5); +x_115 = lean_array_uset(x_91, x_95, x_5); +x_116 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__9(x_13, x_59, x_96); +x_117 = lean_array_uset(x_115, x_95, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_90); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_usize_add(x_9, x_21); +x_9 = x_119; +x_10 = x_118; +goto _start; } } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = lean_ctor_get(x_1, 0); -x_3 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_2); -if (lean_obj_tag(x_3) == 0) +block_229: { -lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 1); -x_5 = lean_box(0); -x_6 = lean_array_size(x_4); -x_7 = 0; -x_8 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; -x_9 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(x_4, x_5, x_8, x_4, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec(x_9); -if (lean_obj_tag(x_10) == 0) +uint64_t x_137; +if (lean_obj_tag(x_135) == 0) { -lean_object* x_11; -x_11 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; -return x_11; +uint64_t x_224; +x_224 = 11; +x_137 = x_224; +goto block_223; } else { -lean_object* x_12; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -return x_12; +lean_object* x_225; uint64_t x_226; uint64_t x_227; uint64_t x_228; +x_225 = lean_ctor_get(x_135, 0); +lean_inc(x_225); +lean_dec(x_135); +x_226 = lean_string_hash(x_225); +lean_dec(x_225); +x_227 = 13; +x_228 = lean_uint64_mix_hash(x_226, x_227); +x_137 = x_228; +goto block_223; } +block_223: +{ +uint64_t x_138; lean_object* x_139; lean_object* x_140; lean_object* x_210; +x_210 = lean_ctor_get(x_136, 0); +lean_inc(x_210); +if (lean_obj_tag(x_210) == 0) +{ +lean_object* x_211; uint64_t x_212; lean_object* x_213; lean_object* x_214; +x_211 = lean_ctor_get(x_136, 1); +lean_inc(x_211); +lean_dec(x_136); +x_212 = 11; +x_213 = lean_ctor_get(x_211, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_211, 1); +lean_inc(x_214); +lean_dec(x_211); +x_138 = x_212; +x_139 = x_213; +x_140 = x_214; +goto block_209; +} +else +{ +lean_object* x_215; lean_object* x_216; uint8_t x_217; uint64_t x_218; uint64_t x_219; uint64_t x_220; lean_object* x_221; lean_object* x_222; +x_215 = lean_ctor_get(x_136, 1); +lean_inc(x_215); +lean_dec(x_136); +x_216 = lean_ctor_get(x_210, 0); +lean_inc(x_216); +lean_dec(x_210); +x_217 = lean_unbox(x_216); +lean_dec(x_216); +x_218 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1168_(x_217); +x_219 = 13; +x_220 = lean_uint64_mix_hash(x_218, x_219); +x_221 = lean_ctor_get(x_215, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_215, 1); +lean_inc(x_222); +lean_dec(x_215); +x_138 = x_220; +x_139 = x_221; +x_140 = x_222; +goto block_209; } -else +block_209: { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_3); -if (x_13 == 0) +if (lean_obj_tag(x_139) == 0) { -return x_3; +if (lean_obj_tag(x_140) == 0) +{ +uint64_t x_141; uint64_t x_142; uint64_t x_143; uint64_t x_144; uint64_t x_145; +x_141 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1; +x_142 = lean_uint64_mix_hash(x_138, x_141); +x_143 = lean_uint64_mix_hash(x_137, x_142); +x_144 = lean_uint64_mix_hash(x_134, x_143); +x_145 = lean_uint64_mix_hash(x_133, x_144); +x_23 = x_145; +goto block_130; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_3, 0); -lean_inc(x_14); -lean_dec(x_3); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; +lean_object* x_146; uint64_t x_147; uint64_t x_148; uint64_t x_149; uint64_t x_150; uint64_t x_151; uint64_t x_152; uint64_t x_153; uint64_t x_154; uint64_t x_155; +x_146 = lean_ctor_get(x_140, 0); +lean_inc(x_146); +lean_dec(x_140); +x_147 = 11; +x_148 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_146); +lean_dec(x_146); +x_149 = 13; +x_150 = lean_uint64_mix_hash(x_148, x_149); +x_151 = lean_uint64_mix_hash(x_147, x_150); +x_152 = lean_uint64_mix_hash(x_138, x_151); +x_153 = lean_uint64_mix_hash(x_137, x_152); +x_154 = lean_uint64_mix_hash(x_134, x_153); +x_155 = lean_uint64_mix_hash(x_133, x_154); +x_23 = x_155; +goto block_130; } } +else +{ +lean_object* x_156; uint64_t x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; uint64_t x_161; +x_156 = lean_ctor_get(x_139, 0); +lean_inc(x_156); +lean_dec(x_139); +x_157 = 7; +x_158 = lean_array_get_size(x_156); +x_159 = lean_unsigned_to_nat(0u); +x_160 = lean_nat_dec_lt(x_159, x_158); +x_161 = 13; +if (x_160 == 0) +{ +lean_dec(x_158); +lean_dec(x_156); +if (lean_obj_tag(x_140) == 0) +{ +uint64_t x_162; uint64_t x_163; uint64_t x_164; uint64_t x_165; uint64_t x_166; +x_162 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3; +x_163 = lean_uint64_mix_hash(x_138, x_162); +x_164 = lean_uint64_mix_hash(x_137, x_163); +x_165 = lean_uint64_mix_hash(x_134, x_164); +x_166 = lean_uint64_mix_hash(x_133, x_165); +x_23 = x_166; +goto block_130; +} +else +{ +lean_object* x_167; uint64_t x_168; uint64_t x_169; uint64_t x_170; uint64_t x_171; uint64_t x_172; uint64_t x_173; uint64_t x_174; uint64_t x_175; +x_167 = lean_ctor_get(x_140, 0); +lean_inc(x_167); +lean_dec(x_140); +x_168 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_167); +lean_dec(x_167); +x_169 = lean_uint64_mix_hash(x_168, x_161); +x_170 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +x_171 = lean_uint64_mix_hash(x_170, x_169); +x_172 = lean_uint64_mix_hash(x_138, x_171); +x_173 = lean_uint64_mix_hash(x_137, x_172); +x_174 = lean_uint64_mix_hash(x_134, x_173); +x_175 = lean_uint64_mix_hash(x_133, x_174); +x_23 = x_175; +goto block_130; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(lean_object* x_1) { -_start: -{ -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_2; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -if (lean_obj_tag(x_2) == 0) +else { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_1); -if (x_3 == 0) +uint8_t x_176; +x_176 = lean_nat_dec_le(x_158, x_158); +if (x_176 == 0) { -lean_object* x_4; lean_object* x_5; uint8_t x_6; -x_4 = lean_ctor_get(x_1, 1); -lean_dec(x_4); -x_5 = lean_ctor_get(x_1, 0); -lean_dec(x_5); -x_6 = !lean_is_exclusive(x_2); -if (x_6 == 0) +lean_dec(x_158); +lean_dec(x_156); +if (lean_obj_tag(x_140) == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_box(0); -lean_ctor_set(x_1, 1, x_8); -lean_ctor_set(x_1, 0, x_7); -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_1); -return x_2; +uint64_t x_177; uint64_t x_178; uint64_t x_179; uint64_t x_180; uint64_t x_181; +x_177 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3; +x_178 = lean_uint64_mix_hash(x_138, x_177); +x_179 = lean_uint64_mix_hash(x_137, x_178); +x_180 = lean_uint64_mix_hash(x_134, x_179); +x_181 = lean_uint64_mix_hash(x_133, x_180); +x_23 = x_181; +goto block_130; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_2, 0); -lean_inc(x_9); -lean_dec(x_2); -x_10 = lean_box(0); -lean_ctor_set(x_1, 1, x_10); -lean_ctor_set(x_1, 0, x_9); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_1); -return x_11; +lean_object* x_182; uint64_t x_183; uint64_t x_184; uint64_t x_185; uint64_t x_186; uint64_t x_187; uint64_t x_188; uint64_t x_189; uint64_t x_190; +x_182 = lean_ctor_get(x_140, 0); +lean_inc(x_182); +lean_dec(x_140); +x_183 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_182); +lean_dec(x_182); +x_184 = lean_uint64_mix_hash(x_183, x_161); +x_185 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2; +x_186 = lean_uint64_mix_hash(x_185, x_184); +x_187 = lean_uint64_mix_hash(x_138, x_186); +x_188 = lean_uint64_mix_hash(x_137, x_187); +x_189 = lean_uint64_mix_hash(x_134, x_188); +x_190 = lean_uint64_mix_hash(x_133, x_189); +x_23 = x_190; +goto block_130; +} +} +else +{ +size_t x_191; size_t x_192; uint64_t x_193; uint64_t x_194; +x_191 = 0; +x_192 = lean_usize_of_nat(x_158); +lean_dec(x_158); +x_193 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6(x_156, x_191, x_192, x_157); +lean_dec(x_156); +x_194 = lean_uint64_mix_hash(x_193, x_161); +if (lean_obj_tag(x_140) == 0) +{ +uint64_t x_195; uint64_t x_196; uint64_t x_197; uint64_t x_198; uint64_t x_199; uint64_t x_200; +x_195 = 11; +x_196 = lean_uint64_mix_hash(x_194, x_195); +x_197 = lean_uint64_mix_hash(x_138, x_196); +x_198 = lean_uint64_mix_hash(x_137, x_197); +x_199 = lean_uint64_mix_hash(x_134, x_198); +x_200 = lean_uint64_mix_hash(x_133, x_199); +x_23 = x_200; +goto block_130; +} +else +{ +lean_object* x_201; uint64_t x_202; uint64_t x_203; uint64_t x_204; uint64_t x_205; uint64_t x_206; uint64_t x_207; uint64_t x_208; +x_201 = lean_ctor_get(x_140, 0); +lean_inc(x_201); +lean_dec(x_140); +x_202 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_6452_(x_201); +lean_dec(x_201); +x_203 = lean_uint64_mix_hash(x_202, x_161); +x_204 = lean_uint64_mix_hash(x_194, x_203); +x_205 = lean_uint64_mix_hash(x_138, x_204); +x_206 = lean_uint64_mix_hash(x_137, x_205); +x_207 = lean_uint64_mix_hash(x_134, x_206); +x_208 = lean_uint64_mix_hash(x_133, x_207); +x_23 = x_208; +goto block_130; } } -else +} +} +} +} +} +} +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1() { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -if (lean_is_exclusive(x_2)) { - lean_ctor_release(x_2, 0); - x_13 = x_2; -} else { - lean_dec_ref(x_2); - x_13 = lean_box(0); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instDecidableEqString___boxed), 2, 0); +return x_1; } -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_12); -lean_ctor_set(x_15, 1, x_14); -if (lean_is_scalar(x_13)) { - x_16 = lean_alloc_ctor(1, 1, 0); -} else { - x_16 = x_13; - lean_ctor_set_tag(x_16, 1); } -lean_ctor_set(x_16, 0, x_15); -return x_16; +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3() { +_start: { -lean_object* x_17; -lean_dec(x_2); -x_17 = lean_ctor_get(x_1, 1); -lean_inc(x_17); -lean_dec(x_1); -x_1 = x_17; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instBEqInsertReplaceEdit; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -case 1: +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4() { +_start: { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_1, 1); -lean_inc(x_19); -lean_dec(x_1); -x_20 = l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(x_19); -lean_dec(x_19); -return x_20; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -default: +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5() { +_start: { -lean_object* x_21; -lean_dec(x_1); -x_21 = lean_box(0); -return x_21; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_instDecidableEqCompletionItemKind___boxed), 2, 0); +return x_1; } } +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7() { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8() { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_instDecidableEqCompletionItemTag___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9() { _start: { -lean_object* x_2; -x_2 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___boxed(lean_object* x_1) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10() { _start: { -lean_object* x_2; -x_2 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9; +x_2 = lean_alloc_closure((void*)(l_Array_instBEq___rarg___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11() { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(x_1, x_2, x_3, x_4, x_8, x_9, x_7); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12() { _start: { -lean_object* x_2; -x_2 = l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_instDecidableEqMarkupContent___boxed), 2, 0); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f(lean_object* x_1) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13() { _start: { -lean_object* x_2; -x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1() { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(7, 1, 0); -lean_ctor_set(x_2, 0, x_1); +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_3); -x_5 = lean_box(0); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_4); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16() { _start: { -lean_object* x_5; -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_4); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_6 = lean_box(0); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } -else -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_5, 0); -lean_inc(x_7); -lean_dec(x_5); -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion(x_1, x_2, x_3); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -lean_object* x_10; -lean_dec(x_7); -x_10 = lean_box(0); -return x_10; } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17() { +_start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_box(0); -x_12 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(x_7, x_11); -return x_12; -} -} +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18() { _start: { -lean_object* x_3; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(x_1, x_2); -lean_dec(x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19() { _start: { -if (lean_obj_tag(x_1) == 6) -{ -if (lean_obj_tag(x_2) == 6) -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; uint8_t x_15; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_ctor_get(x_2, 0); -x_15 = l_Lean_Elab_Info_isSmaller(x_1, x_2); -if (x_15 == 0) -{ -uint8_t x_16; -x_16 = l_Lean_Elab_Info_isSmaller(x_2, x_1); -if (x_16 == 0) -{ -if (lean_obj_tag(x_3) == 1) -{ -uint8_t x_17; -x_17 = 1; -x_5 = x_17; -goto block_14; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20() { +_start: { -uint8_t x_18; -x_18 = 0; -x_5 = x_18; -goto block_14; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instHashableInsertReplaceEdit; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21() { +_start: { -uint8_t x_19; -x_19 = 0; -return x_19; +lean_object* x_1; lean_object* x_2; +x_1 = l_instHashableString; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22() { +_start: { -uint8_t x_20; -x_20 = 1; -return x_20; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instHashableCompletionItemKind; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -block_14: -{ -lean_object* x_6; -if (x_5 == 0) -{ -if (lean_obj_tag(x_4) == 1) -{ -uint8_t x_11; -x_11 = 1; -return x_11; } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23() { +_start: { -lean_object* x_12; -x_12 = lean_box(0); -x_6 = x_12; -goto block_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instHashableCompletionItemTag; +x_2 = lean_alloc_closure((void*)(l_instHashableArray___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24() { +_start: { -lean_object* x_13; -x_13 = lean_box(0); -x_6 = x_13; -goto block_10; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -block_10: -{ -lean_dec(x_6); -if (x_5 == 0) -{ -uint8_t x_7; -x_7 = 1; -return x_7; } -else -{ -if (lean_obj_tag(x_4) == 1) +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25() { +_start: { -uint8_t x_8; -x_8 = 1; -return x_8; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Lsp_instHashableMarkupContent; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -else -{ -uint8_t x_9; -x_9 = 0; -return x_9; } +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } -else +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28() { +_start: { -uint8_t x_21; -x_21 = 1; -return x_21; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } -else -{ -if (lean_obj_tag(x_2) == 6) +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29() { +_start: { -uint8_t x_22; -x_22 = 0; -return x_22; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30() { +_start: { -uint8_t x_23; -x_23 = 1; -return x_23; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_instHashableString; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; } } +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32() { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31; +x_3 = lean_mk_array(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33() { _start: { -uint8_t x_7; -x_7 = l_Lean_Elab_Info_isCompletion(x_5); -if (x_7 == 0) -{ -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_6; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } -else -{ -uint8_t x_8; -x_8 = l_Lean_Elab_Info_occursInOrOnBoundary(x_5, x_2); -if (x_8 == 0) -{ -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_6; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_Elab_Info_pos_x3f(x_5); -x_10 = l_Lean_Elab_Info_tailPos_x3f(x_5); -if (lean_obj_tag(x_9) == 0) +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_93; lean_object* x_94; -x_93 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_94 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_93); -x_11 = x_94; -goto block_92; +lean_object* x_3; lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_array_size(x_2); +x_6 = 0; +x_7 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19; +x_8 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30; +x_9 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33; +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11(x_1, x_2, x_7, x_8, x_3, x_4, x_2, x_5, x_6, x_9); +return x_10; } -else -{ -lean_object* x_95; -x_95 = lean_ctor_get(x_9, 0); -lean_inc(x_95); -lean_dec(x_9); -x_11 = x_95; -goto block_92; } -block_92: +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_12; lean_object* x_13; -lean_inc(x_1); -x_12 = l_Lean_FileMap_toPosition(x_1, x_11); -if (lean_obj_tag(x_10) == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; double x_7; double x_8; uint8_t x_9; size_t x_10; size_t x_11; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get_float(x_4, sizeof(void*)*1); +x_8 = lean_ctor_get_float(x_6, sizeof(void*)*1); +x_9 = lean_float_decLt(x_7, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +if (x_9 == 0) { -lean_object* x_89; lean_object* x_90; -x_89 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_90 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_89); -x_13 = x_90; -goto block_88; +lean_dec(x_6); +x_2 = x_11; +goto _start; } else { -lean_object* x_91; -x_91 = lean_ctor_get(x_10, 0); -lean_inc(x_91); -lean_dec(x_10); -x_13 = x_91; -goto block_88; +lean_dec(x_4); +x_2 = x_11; +x_4 = x_6; +goto _start; } -block_88: -{ -uint8_t x_14; lean_object* x_15; -x_14 = lean_nat_dec_lt(x_2, x_13); -if (x_14 == 0) -{ -lean_object* x_85; -lean_dec(x_11); -x_85 = lean_box(0); -x_15 = x_85; -goto block_84; } else { -lean_object* x_86; lean_object* x_87; -x_86 = lean_nat_sub(x_2, x_11); -lean_dec(x_11); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_86); -x_15 = x_87; -goto block_84; +return x_4; } -block_84: -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_12, 0); -x_18 = lean_ctor_get(x_12, 1); -lean_dec(x_18); -x_19 = l_Lean_FileMap_toPosition(x_1, x_13); -lean_dec(x_13); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_dec(x_22); -x_23 = lean_nat_dec_eq(x_17, x_3); -if (x_23 == 0) -{ -lean_free_object(x_19); -lean_dec(x_21); -lean_free_object(x_12); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; } -else -{ -uint8_t x_24; -x_24 = lean_nat_dec_eq(x_17, x_21); -lean_dec(x_21); -lean_dec(x_17); -if (x_24 == 0) -{ -lean_free_object(x_19); -lean_free_object(x_12); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; } -else +LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12(lean_object* x_1) { +_start: { -if (lean_obj_tag(x_6) == 0) +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_lt(x_3, x_2); +if (x_4 == 0) { -lean_object* x_25; -lean_ctor_set(x_19, 1, x_5); -lean_ctor_set(x_19, 0, x_4); -lean_ctor_set(x_12, 1, x_19); -lean_ctor_set(x_12, 0, x_15); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_12); -return x_25; +lean_object* x_5; +lean_dec(x_2); +x_5 = lean_box(0); +return x_5; } else { -lean_object* x_26; uint8_t x_27; -lean_free_object(x_19); -lean_free_object(x_12); -x_26 = lean_ctor_get(x_6, 0); -lean_inc(x_26); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -lean_object* x_28; lean_object* x_29; uint8_t x_30; -x_28 = lean_ctor_get(x_26, 1); -x_29 = lean_ctor_get(x_26, 0); -lean_dec(x_29); -x_30 = !lean_is_exclusive(x_28); -if (x_30 == 0) -{ -lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = lean_ctor_get(x_28, 1); -x_32 = lean_ctor_get(x_28, 0); -lean_dec(x_32); -x_33 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_5, x_31); -lean_dec(x_31); -if (x_33 == 0) +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_array_fget(x_1, x_3); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_dec_lt(x_7, x_2); +if (x_8 == 0) { -lean_free_object(x_28); -lean_free_object(x_26); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +lean_object* x_9; +lean_dec(x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_6); +return x_9; } else { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_6); -if (x_34 == 0) +uint8_t x_10; +x_10 = lean_nat_dec_le(x_2, x_2); +if (x_10 == 0) { -lean_object* x_35; -x_35 = lean_ctor_get(x_6, 0); -lean_dec(x_35); -lean_ctor_set(x_28, 1, x_5); -lean_ctor_set(x_28, 0, x_4); -lean_ctor_set(x_26, 0, x_15); -return x_6; +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_6); +return x_11; } else { -lean_object* x_36; -lean_dec(x_6); -lean_ctor_set(x_28, 1, x_5); -lean_ctor_set(x_28, 0, x_4); -lean_ctor_set(x_26, 0, x_15); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_26); -return x_36; +size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_12 = 1; +x_13 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_14 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13(x_1, x_12, x_13, x_6); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; } } } -else +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__14(lean_object* x_1) { +_start: { -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_28, 1); -lean_inc(x_37); -lean_dec(x_28); -x_38 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_5, x_37); -lean_dec(x_37); -if (x_38 == 0) +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__15(lean_object* x_1, lean_object* x_2) { +_start: { -lean_free_object(x_26); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +if (lean_obj_tag(x_2) == 0) +{ +return x_1; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - x_39 = x_6; -} else { - lean_dec_ref(x_6); - x_39 = lean_box(0); -} -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_4); -lean_ctor_set(x_40, 1, x_5); -lean_ctor_set(x_26, 1, x_40); -lean_ctor_set(x_26, 0, x_15); -if (lean_is_scalar(x_39)) { - x_41 = lean_alloc_ctor(1, 1, 0); -} else { - x_41 = x_39; -} -lean_ctor_set(x_41, 0, x_26); -return x_41; +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 2); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_array_push(x_1, x_3); +x_1 = x_5; +x_2 = x_4; +goto _start; } } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_42 = lean_ctor_get(x_26, 1); -lean_inc(x_42); -lean_dec(x_26); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); -} -x_45 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_5, x_43); -lean_dec(x_43); -if (x_45 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) { -lean_dec(x_44); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +lean_dec(x_1); +return x_4; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - x_46 = x_6; -} else { - lean_dec_ref(x_6); - x_46 = lean_box(0); -} -if (lean_is_scalar(x_44)) { - x_47 = lean_alloc_ctor(0, 2, 0); -} else { - x_47 = x_44; -} -lean_ctor_set(x_47, 0, x_4); -lean_ctor_set(x_47, 1, x_5); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_15); -lean_ctor_set(x_48, 1, x_47); -if (lean_is_scalar(x_46)) { - x_49 = lean_alloc_ctor(1, 1, 0); -} else { - x_49 = x_46; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = lean_box(0); +lean_inc(x_1); +x_10 = l_Std_DHashMap_Internal_AssocList_map_go___rarg(x_1, x_9, x_6); +x_11 = 1; +x_12 = lean_usize_add(x_3, x_11); +x_13 = lean_array_uset(x_8, x_3, x_10); +x_3 = x_12; +x_4 = x_13; +goto _start; } -lean_ctor_set(x_49, 0, x_48); -return x_49; } } +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; } } +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; } } -else +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3() { +_start: { -lean_object* x_50; uint8_t x_51; -x_50 = lean_ctor_get(x_19, 0); -lean_inc(x_50); -lean_dec(x_19); -x_51 = lean_nat_dec_eq(x_17, x_3); -if (x_51 == 0) +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4() { +_start: { -lean_dec(x_50); -lean_free_object(x_12); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1; +x_2 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } -else +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_52; -x_52 = lean_nat_dec_eq(x_17, x_50); -lean_dec(x_50); -lean_dec(x_17); -if (x_52 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_free_object(x_12); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +return x_1; } else { +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12(x_4); +lean_dec(x_4); if (lean_obj_tag(x_6) == 0) { -lean_object* x_53; lean_object* x_54; -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_4); -lean_ctor_set(x_53, 1, x_5); -lean_ctor_set(x_12, 1, x_53); -lean_ctor_set(x_12, 0, x_15); -x_54 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_54, 0, x_12); -return x_54; +lean_object* x_7; lean_object* x_8; +x_7 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; +x_8 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__14(x_7); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 1, x_8); +x_1 = x_2; +x_2 = x_5; +goto _start; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -lean_free_object(x_12); -x_55 = lean_ctor_get(x_6, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_57 = x_55; -} else { - lean_dec_ref(x_55); - x_57 = lean_box(0); +lean_object* x_10; +x_10 = lean_ctor_get(x_6, 0); +lean_inc(x_10); +lean_dec(x_6); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 1, x_10); +x_1 = x_2; +x_2 = x_5; +goto _start; } -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_59 = x_56; -} else { - lean_dec_ref(x_56); - x_59 = lean_box(0); } -x_60 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_5, x_58); -lean_dec(x_58); -if (x_60 == 0) +else { -lean_dec(x_59); -lean_dec(x_57); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_ctor_get(x_2, 2); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_2); +x_15 = l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12(x_13); +lean_dec(x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; +x_17 = l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__14(x_16); +x_18 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_18, 2, x_1); +x_1 = x_18; +x_2 = x_14; +goto _start; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - x_61 = x_6; -} else { - lean_dec_ref(x_6); - x_61 = lean_box(0); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_1); +x_1 = x_21; +x_2 = x_14; +goto _start; } -if (lean_is_scalar(x_59)) { - x_62 = lean_alloc_ctor(0, 2, 0); -} else { - x_62 = x_59; } -lean_ctor_set(x_62, 0, x_4); -lean_ctor_set(x_62, 1, x_5); -if (lean_is_scalar(x_57)) { - x_63 = lean_alloc_ctor(0, 2, 0); -} else { - x_63 = x_57; } -lean_ctor_set(x_63, 0, x_15); -lean_ctor_set(x_63, 1, x_62); -if (lean_is_scalar(x_61)) { - x_64 = lean_alloc_ctor(1, 1, 0); -} else { - x_64 = x_61; } -lean_ctor_set(x_64, 0, x_63); -return x_64; } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; } +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_box(0); +x_9 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18(x_8, x_5); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_12 = lean_array_uset(x_7, x_2, x_9); +x_2 = x_11; +x_3 = x_12; +goto _start; } } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__15(x_4, x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +x_4 = x_7; +goto _start; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_65 = lean_ctor_get(x_12, 0); -lean_inc(x_65); -lean_dec(x_12); -x_66 = l_Lean_FileMap_toPosition(x_1, x_13); -lean_dec(x_13); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; -} else { - lean_dec_ref(x_66); - x_68 = lean_box(0); +return x_4; } -x_69 = lean_nat_dec_eq(x_65, x_3); -if (x_69 == 0) +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1(lean_object* x_1) { +_start: { -lean_dec(x_68); -lean_dec(x_67); -lean_dec(x_65); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get(x_2, 3); +x_7 = lean_ctor_get(x_2, 4); +x_8 = lean_ctor_get(x_2, 7); +lean_inc(x_5); +lean_inc(x_8); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +lean_inc(x_6); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_6); +lean_ctor_set(x_10, 1, x_9); +lean_inc(x_4); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_10); +lean_inc(x_7); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_7); +lean_ctor_set(x_12, 1, x_11); +lean_inc(x_3); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1___boxed), 1, 0); +return x_1; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems(lean_object* x_1) { +_start: { -uint8_t x_70; -x_70 = lean_nat_dec_eq(x_65, x_67); -lean_dec(x_67); -lean_dec(x_65); -if (x_70 == 0) +lean_object* x_2; lean_object* x_3; lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1; +x_3 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1(x_2, x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_array_size(x_4); +x_6 = 0; +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(x_5, x_6, x_4); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_nat_dec_lt(x_9, x_8); +if (x_10 == 0) { -lean_dec(x_68); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_7); +x_11 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +return x_11; } else { -if (lean_obj_tag(x_6) == 0) +uint8_t x_12; +x_12 = lean_nat_dec_le(x_8, x_8); +if (x_12 == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -if (lean_is_scalar(x_68)) { - x_71 = lean_alloc_ctor(0, 2, 0); -} else { - x_71 = x_68; -} -lean_ctor_set(x_71, 0, x_4); -lean_ctor_set(x_71, 1, x_5); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_15); -lean_ctor_set(x_72, 1, x_71); -x_73 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_73, 0, x_72); -return x_73; +lean_object* x_13; +lean_dec(x_8); +lean_dec(x_7); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +return x_13; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; -lean_dec(x_68); -x_74 = lean_ctor_get(x_6, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_74, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_74)) { - lean_ctor_release(x_74, 0); - lean_ctor_release(x_74, 1); - x_76 = x_74; -} else { - lean_dec_ref(x_74); - x_76 = lean_box(0); +size_t x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_usize_of_nat(x_8); +lean_dec(x_8); +x_15 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +x_16 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19(x_7, x_6, x_14, x_15); +lean_dec(x_7); +return x_16; } -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_78 = x_75; -} else { - lean_dec_ref(x_75); - x_78 = lean_box(0); } -x_79 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_isBetter(x_5, x_77); -lean_dec(x_77); -if (x_79 == 0) -{ -lean_dec(x_78); -lean_dec(x_76); -lean_dec(x_15); -lean_dec(x_5); -lean_dec(x_4); -return x_6; } -else +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -if (lean_is_exclusive(x_6)) { - lean_ctor_release(x_6, 0); - x_80 = x_6; -} else { - lean_dec_ref(x_6); - x_80 = lean_box(0); +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__3(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; } -if (lean_is_scalar(x_78)) { - x_81 = lean_alloc_ctor(0, 2, 0); -} else { - x_81 = x_78; } -lean_ctor_set(x_81, 0, x_4); -lean_ctor_set(x_81, 1, x_5); -if (lean_is_scalar(x_76)) { - x_82 = lean_alloc_ctor(0, 2, 0); -} else { - x_82 = x_76; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint64_t x_7; uint64_t x_8; lean_object* x_9; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_uint64(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__6(x_1, x_5, x_6, x_7); +lean_dec(x_1); +x_9 = lean_box_uint64(x_8); +return x_9; } -lean_ctor_set(x_82, 0, x_15); -lean_ctor_set(x_82, 1, x_81); -if (lean_is_scalar(x_80)) { - x_83 = lean_alloc_ctor(1, 1, 0); -} else { - x_83 = x_80; } -lean_ctor_set(x_83, 0, x_82); -return x_83; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; } } +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__13(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; } } +LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Array_getMax_x3f___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__12(x_1); +lean_dec(x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16(x_1, x_5, x_6, x_4); +return x_7; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__16___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__17(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_7; -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__19(x_1, x_5, x_6, x_4); +lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1___boxed(lean_object* x_1) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -lean_inc(x_1); -x_5 = l_Lean_FileMap_toPosition(x_1, x_2); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -lean_inc(x_2); -lean_inc(x_1); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f_choose___boxed), 6, 3); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -lean_closure_set(x_7, 2, x_6); -x_8 = lean_box(0); -lean_inc(x_4); -x_9 = l_Lean_Elab_InfoTree_foldInfo___rarg(x_7, x_8, x_4); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; -lean_inc(x_4); -lean_inc(x_2); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_10) == 0) +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___boxed(lean_object* x_1) { +_start: { -lean_object* x_11; -x_11 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(x_2, x_4); -return x_11; +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems(x_1); +lean_dec(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_12; -lean_dec(x_4); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_10); -if (x_12 == 0) +uint8_t x_3; +x_3 = lean_string_utf8_at_end(x_2, x_1); +if (x_3 == 0) { -return x_10; +uint32_t x_4; lean_object* x_5; uint32_t x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_string_utf8_get(x_2, x_1); +x_5 = l_Char_toLower(x_4); +x_6 = lean_unbox_uint32(x_5); +lean_dec(x_5); +x_7 = lean_string_utf8_set(x_2, x_1, x_6); +x_8 = lean_string_utf8_next(x_7, x_1); +lean_dec(x_1); +x_1 = x_8; +x_2 = x_7; +goto _start; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; +lean_dec(x_1); +return x_2; } } } -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 6) +LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_19; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +lean_object* x_3; double x_4; lean_object* x_5; double x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get_float(x_1, sizeof(void*)*1); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) -{ -lean_object* x_20; uint8_t x_21; -x_20 = lean_ctor_get(x_16, 1); -lean_dec(x_20); -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +x_6 = lean_ctor_get_float(x_2, sizeof(void*)*1); +lean_dec(x_2); +x_7 = lean_float_beq(x_4, x_6); +if (x_7 == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_17, 1); -lean_dec(x_22); -x_23 = lean_ctor_get(x_18, 0); -lean_inc(x_23); -lean_dec(x_18); -lean_ctor_set(x_17, 1, x_23); -return x_9; +uint8_t x_8; +lean_dec(x_5); +lean_dec(x_3); +x_8 = lean_float_decLt(x_6, x_4); +return x_8; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_17, 0); -lean_inc(x_24); -lean_dec(x_17); -x_25 = lean_ctor_get(x_18, 0); -lean_inc(x_25); -lean_dec(x_18); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_16, 1, x_26); -return x_9; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_unsigned_to_nat(0u); +x_11 = l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(x_10, x_9); +x_12 = lean_ctor_get(x_5, 0); +lean_inc(x_12); +lean_dec(x_5); +x_13 = l_String_mapAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__1(x_10, x_12); +x_14 = lean_string_dec_lt(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +return x_14; } } -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_27 = lean_ctor_get(x_16, 0); -lean_inc(x_27); -lean_dec(x_16); -x_28 = lean_ctor_get(x_17, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - lean_ctor_release(x_17, 1); - x_29 = x_17; -} else { - lean_dec_ref(x_17); - x_29 = lean_box(0); } -x_30 = lean_ctor_get(x_18, 0); -lean_inc(x_30); -lean_dec(x_18); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(0, 2, 0); -} else { - x_31 = x_29; +static lean_object* _init_l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1___boxed), 2, 0); +return x_1; } -lean_ctor_set(x_31, 0, x_28); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_9, 0, x_32); -return x_9; } +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) +{ +lean_dec(x_2); +return x_1; } else { -lean_object* x_33; -lean_dec(x_18); -lean_dec(x_17); -lean_free_object(x_9); -lean_dec(x_16); -lean_inc(x_4); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1; lean_inc(x_2); -x_33 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_33) == 0) +x_6 = l_Array_qpartition___rarg(x_1, x_5, x_2, x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_nat_dec_le(x_3, x_7); +if (x_9 == 0) { -lean_object* x_34; -x_34 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(x_2, x_4); -return x_34; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_8, x_2, x_7); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_7, x_11); +lean_dec(x_7); +x_1 = x_10; +x_2 = x_12; +goto _start; } else { -uint8_t x_35; -lean_dec(x_4); +lean_dec(x_7); lean_dec(x_2); -x_35 = !lean_is_exclusive(x_33); -if (x_35 == 0) -{ -return x_33; -} -else -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_33, 0); -lean_inc(x_36); -lean_dec(x_33); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +return x_8; } } } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(size_t x_1, size_t x_2, lean_object* x_3) { +_start: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_9, 0); -lean_inc(x_38); -lean_dec(x_9); -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 6) +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_41 = lean_ctor_get(x_38, 0); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_ctor_get(x_39, 0); -lean_inc(x_43); -if (lean_is_exclusive(x_39)) { - lean_ctor_release(x_39, 0); - lean_ctor_release(x_39, 1); - x_44 = x_39; -} else { - lean_dec_ref(x_39); - x_44 = lean_box(0); -} -x_45 = lean_ctor_get(x_40, 0); -lean_inc(x_45); -lean_dec(x_40); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; -} -lean_ctor_set(x_46, 0, x_43); -lean_ctor_set(x_46, 1, x_45); -if (lean_is_scalar(x_42)) { - x_47 = lean_alloc_ctor(0, 2, 0); -} else { - x_47 = x_42; -} -lean_ctor_set(x_47, 0, x_41); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -return x_48; +return x_3; } else { -lean_object* x_49; -lean_dec(x_40); -lean_dec(x_39); -lean_dec(x_38); -lean_inc(x_4); -lean_inc(x_2); -x_49 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; -x_50 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(x_2, x_4); -return x_50; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(lean_object* x_1) { +_start: { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_dec(x_4); +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; size_t x_7; size_t x_8; lean_object* x_9; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(1u); +x_4 = lean_nat_sub(x_2, x_3); lean_dec(x_2); -x_51 = lean_ctor_get(x_49, 0); -lean_inc(x_51); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - x_52 = x_49; -} else { - lean_dec_ref(x_49); - x_52 = lean_box(0); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_1, x_5, x_4); +lean_dec(x_4); +x_7 = lean_array_size(x_6); +x_8 = 0; +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(x_7, x_8, x_6); +return x_9; } -if (lean_is_scalar(x_52)) { - x_53 = lean_alloc_ctor(1, 1, 0); -} else { - x_53 = x_52; } -lean_ctor_set(x_53, 0, x_51); -return x_53; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___lambda__1(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; } } +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__3(x_4, x_5, x_3); +return x_6; } } LEAN_EXPORT lean_object* l_Array_mapFinIdxM_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -34689,7 +3538,7 @@ x_12 = lean_usize_add(x_3, x_11); if (lean_obj_tag(x_10) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_13 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_13 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; x_14 = l_panic___at_String_fromUTF8_x21___spec__1(x_13); x_15 = lean_string_length(x_14); x_16 = lean_nat_sub(x_1, x_15); @@ -34778,7 +3627,7 @@ x_53 = lean_usize_add(x_3, x_52); if (lean_obj_tag(x_49) == 0) { lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_54 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_54 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; x_55 = l_panic___at_String_fromUTF8_x21___spec__1(x_54); x_56 = lean_string_length(x_55); x_57 = lean_nat_sub(x_1, x_56); @@ -34851,118 +3700,102 @@ goto _start; LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; size_t x_14; size_t x_15; lean_object* x_16; -x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - x_5 = x_1; -} else { - lean_dec_ref(x_1); - x_5 = lean_box(0); -} -x_6 = lean_array_get_size(x_4); -x_7 = lean_mk_empty_array_with_capacity(x_6); -x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Array_mapFinIdxM_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__1(x_4, x_4, x_6, x_8, lean_box(0), x_7); -lean_dec(x_4); -x_10 = lean_array_get_size(x_9); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_sub(x_10, x_11); -x_13 = lean_nat_dec_lt(x_12, x_10); -lean_dec(x_10); -x_14 = lean_array_size(x_9); -x_15 = 0; -if (x_13 == 0) +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; size_t x_11; size_t x_12; +x_3 = lean_array_get_size(x_1); +x_4 = lean_mk_empty_array_with_capacity(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_mapFinIdxM_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__1(x_1, x_1, x_3, x_5, lean_box(0), x_4); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_sub(x_7, x_8); +x_10 = lean_nat_dec_lt(x_9, x_7); +lean_dec(x_7); +x_11 = lean_array_size(x_6); +x_12 = 0; +if (x_10 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_12); -x_21 = l_Lean_Lsp_instInhabitedCompletionItem; -x_22 = l_outOfBounds___rarg(x_21); -x_23 = lean_ctor_get(x_22, 5); -lean_inc(x_23); -lean_dec(x_22); -if (lean_obj_tag(x_23) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_9); +x_13 = l_Lean_Lsp_instInhabitedCompletionItem; +x_14 = l_outOfBounds___rarg(x_13); +x_15 = lean_ctor_get(x_14, 5); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_25 = l_panic___at_String_fromUTF8_x21___spec__1(x_24); -x_16 = x_25; -goto block_20; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; +x_17 = l_panic___at_String_fromUTF8_x21___spec__1(x_16); +x_18 = lean_string_length(x_17); +lean_dec(x_17); +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(x_18, x_11, x_12, x_6); +lean_dec(x_18); +return x_19; } else { -lean_object* x_26; -x_26 = lean_ctor_get(x_23, 0); -lean_inc(x_26); -lean_dec(x_23); -x_16 = x_26; -goto block_20; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_string_length(x_20); +lean_dec(x_20); +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(x_21, x_11, x_12, x_6); +lean_dec(x_21); +return x_22; } } else { -lean_object* x_27; lean_object* x_28; -x_27 = lean_array_fget(x_9, x_12); -lean_dec(x_12); -x_28 = lean_ctor_get(x_27, 5); -lean_inc(x_28); -lean_dec(x_27); -if (lean_obj_tag(x_28) == 0) +lean_object* x_23; lean_object* x_24; +x_23 = lean_array_fget(x_6, x_9); +lean_dec(x_9); +x_24 = lean_ctor_get(x_23, 5); +lean_inc(x_24); +lean_dec(x_23); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; -x_30 = l_panic___at_String_fromUTF8_x21___spec__1(x_29); -x_16 = x_30; -goto block_20; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4; +x_26 = l_panic___at_String_fromUTF8_x21___spec__1(x_25); +x_27 = lean_string_length(x_26); +lean_dec(x_26); +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(x_27, x_11, x_12, x_6); +lean_dec(x_27); +return x_28; } else { -lean_object* x_31; -x_31 = lean_ctor_get(x_28, 0); -lean_inc(x_31); -lean_dec(x_28); -x_16 = x_31; -goto block_20; -} -} -block_20: -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_string_length(x_16); -lean_dec(x_16); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(x_17, x_14, x_15, x_9); -lean_dec(x_17); -if (lean_is_scalar(x_5)) { - x_19 = lean_alloc_ctor(0, 1, 1); -} else { - x_19 = x_5; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_24, 0); +lean_inc(x_29); +lean_dec(x_24); +x_30 = lean_string_length(x_29); +lean_dec(x_29); +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2(x_30, x_11, x_12, x_6); +lean_dec(x_30); +return x_31; } -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_3); -return x_19; } } } LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(lean_object* x_1) { _start: { -lean_object* x_2; uint8_t x_3; -x_2 = lean_ctor_get(x_1, 0); -lean_inc(x_2); -x_3 = l_Array_isEmpty___rarg(x_2); -lean_dec(x_2); -if (x_3 == 0) +uint8_t x_2; +x_2 = l_Array_isEmpty___rarg(x_1); +if (x_2 == 0) { -lean_object* x_4; lean_object* x_5; -x_4 = lean_box(0); -x_5 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1(x_1, x_4); -return x_5; +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1(x_1, x_3); +return x_4; } else { -return x_1; +lean_object* x_5; +x_5 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +return x_5; } } } @@ -34995,1193 +3828,806 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Compl lean_object* x_3; x_3 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___lambda__1(x_1, x_2); lean_dec(x_2); +lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_Completion_find_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___boxed(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; lean_object* x_4; -x_3 = lean_box(0); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; -} -else -{ -uint8_t x_5; -x_5 = !lean_is_exclusive(x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(x_6); -lean_ctor_set(x_1, 0, x_7); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_2); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(x_1); lean_dec(x_1); -x_10 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(x_9); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -return x_12; +return x_2; } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = l_Array_append___rarg(x_1, x_2); +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_3); +return x_6; } } -static lean_object* _init_l_Lean_Server_Completion_find_x3f___closed__1() { +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_Completion_find_x3f___lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1___boxed), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_Completion_find_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f(x_2, x_3, x_4, x_5); -if (lean_obj_tag(x_8) == 0) +lean_object* x_11; uint8_t x_29; +x_29 = lean_usize_dec_lt(x_8, x_7); +if (x_29 == 0) { -lean_object* x_9; lean_object* x_10; -lean_dec(x_6); +lean_object* x_30; +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_7); -return x_10; +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_9); +lean_ctor_set(x_30, 1, x_10); +return x_30; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_8, 0); -lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 0); -x_16 = lean_ctor_get(x_12, 1); -x_17 = l_Lean_Server_Completion_find_x3f___closed__1; -switch (lean_obj_tag(x_16)) { +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_array_uget(x_6, x_8); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1; +x_35 = lean_ctor_get(x_32, 2); +lean_inc(x_35); +switch (lean_obj_tag(x_35)) { case 0: { -lean_object* x_18; lean_object* x_19; -lean_free_object(x_12); -lean_dec(x_6); -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion(x_1, x_15, x_18, x_13, x_7); -if (lean_obj_tag(x_19) == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); +lean_dec(x_32); +lean_inc(x_1); +x_38 = l_Lean_Server_Completion_dotCompletion(x_1, x_33, x_37, x_36, x_10); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_apply_2(x_17, x_20, x_21); -return x_22; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_apply_3(x_34, x_9, x_39, x_40); +x_11 = x_41; +goto block_28; } else { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_19); -if (x_23 == 0) +uint8_t x_42; +lean_dec(x_9); +x_42 = !lean_is_exclusive(x_38); +if (x_42 == 0) { -return x_19; +x_11 = x_38; +goto block_28; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_19, 0); -x_25 = lean_ctor_get(x_19, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_19); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_38, 0); +x_44 = lean_ctor_get(x_38, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_38); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_11 = x_45; +goto block_28; } } } case 1: { -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; -lean_free_object(x_12); -lean_dec(x_6); -x_27 = lean_ctor_get(x_16, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_16, 1); -lean_inc(x_28); -x_29 = lean_ctor_get_uint8(x_16, sizeof(void*)*4); -x_30 = lean_ctor_get(x_16, 2); -lean_inc(x_30); -lean_dec(x_16); -x_31 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion(x_1, x_15, x_30, x_27, x_28, x_13, x_29, x_7); -if (lean_obj_tag(x_31) == 0) +lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_46 = lean_ctor_get(x_35, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_35, 1); +lean_inc(x_47); +x_48 = lean_ctor_get_uint8(x_35, sizeof(void*)*4); +x_49 = lean_ctor_get(x_35, 2); +lean_inc(x_49); +lean_dec(x_35); +x_50 = lean_ctor_get(x_32, 1); +lean_inc(x_50); +x_51 = lean_ctor_get(x_32, 0); +lean_inc(x_51); +lean_dec(x_32); +lean_inc(x_1); +x_52 = l_Lean_Server_Completion_idCompletion(x_1, x_33, x_50, x_49, x_46, x_47, x_51, x_48, x_10); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_apply_2(x_17, x_32, x_33); -return x_34; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_apply_3(x_34, x_9, x_53, x_54); +x_11 = x_55; +goto block_28; } else { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_31); -if (x_35 == 0) +uint8_t x_56; +lean_dec(x_9); +x_56 = !lean_is_exclusive(x_52); +if (x_56 == 0) { -return x_31; +x_11 = x_52; +goto block_28; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_31, 0); -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_31); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_52, 0); +x_58 = lean_ctor_get(x_52, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_52); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_11 = x_59; +goto block_28; } } } case 2: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -lean_free_object(x_12); -lean_dec(x_13); -lean_dec(x_6); -x_39 = lean_ctor_get(x_16, 1); -lean_inc(x_39); -x_40 = lean_ctor_get(x_16, 2); -lean_inc(x_40); -x_41 = lean_ctor_get(x_16, 3); -lean_inc(x_41); -lean_dec(x_16); -x_42 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion(x_1, x_15, x_40, x_39, x_41, x_7); -if (lean_obj_tag(x_42) == 0) +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_60 = lean_ctor_get(x_35, 1); +lean_inc(x_60); +x_61 = lean_ctor_get(x_35, 2); +lean_inc(x_61); +x_62 = lean_ctor_get(x_35, 3); +lean_inc(x_62); +lean_dec(x_35); +x_63 = lean_ctor_get(x_32, 1); +lean_inc(x_63); +lean_dec(x_32); +lean_inc(x_1); +x_64 = l_Lean_Server_Completion_dotIdCompletion(x_1, x_33, x_63, x_61, x_60, x_62, x_10); +if (lean_obj_tag(x_64) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -lean_dec(x_42); -x_45 = lean_apply_2(x_17, x_43, x_44); -return x_45; +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_apply_3(x_34, x_9, x_65, x_66); +x_11 = x_67; +goto block_28; } else { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_42); -if (x_46 == 0) +uint8_t x_68; +lean_dec(x_9); +x_68 = !lean_is_exclusive(x_64); +if (x_68 == 0) { -return x_42; +x_11 = x_64; +goto block_28; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_42, 0); -x_48 = lean_ctor_get(x_42, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_42); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_64, 0); +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_64); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_11 = x_71; +goto block_28; } } } case 3: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -lean_free_object(x_12); -lean_dec(x_13); -lean_dec(x_6); -x_50 = lean_ctor_get(x_16, 1); -lean_inc(x_50); -x_51 = lean_ctor_get(x_16, 2); -lean_inc(x_51); -x_52 = lean_ctor_get(x_16, 3); -lean_inc(x_52); -lean_dec(x_16); -x_53 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion(x_1, x_15, x_51, x_50, x_52, x_7); -if (lean_obj_tag(x_53) == 0) +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_72 = lean_ctor_get(x_35, 1); +lean_inc(x_72); +x_73 = lean_ctor_get(x_35, 2); +lean_inc(x_73); +x_74 = lean_ctor_get(x_35, 3); +lean_inc(x_74); +lean_dec(x_35); +x_75 = lean_ctor_get(x_32, 1); +lean_inc(x_75); +lean_dec(x_32); +lean_inc(x_1); +x_76 = l_Lean_Server_Completion_fieldIdCompletion(x_1, x_33, x_75, x_73, x_72, x_74, x_10); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_apply_2(x_17, x_54, x_55); -return x_56; +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_apply_3(x_34, x_9, x_77, x_78); +x_11 = x_79; +goto block_28; } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_53); -if (x_57 == 0) +uint8_t x_80; +lean_dec(x_9); +x_80 = !lean_is_exclusive(x_76); +if (x_80 == 0) { -return x_53; +x_11 = x_76; +goto block_28; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_53, 0); -x_59 = lean_ctor_get(x_53, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_53); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; -} +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_76, 0); +x_82 = lean_ctor_get(x_76, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_76); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +x_11 = x_83; +goto block_28; } } -case 4: -{ -lean_object* x_61; -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_1); -x_61 = lean_box(0); -lean_ctor_set(x_12, 1, x_7); -lean_ctor_set(x_12, 0, x_61); -return x_12; } case 5: { -lean_object* x_62; lean_object* x_63; -lean_free_object(x_12); -lean_dec(x_13); -x_62 = lean_ctor_get(x_16, 0); -lean_inc(x_62); -lean_dec(x_16); -x_63 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(x_1, x_15, x_62, x_6, x_7); -lean_dec(x_62); -if (lean_obj_tag(x_63) == 0) +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_35, 0); +lean_inc(x_84); +lean_dec(x_35); +x_85 = lean_ctor_get(x_32, 1); +lean_inc(x_85); +lean_dec(x_32); +lean_inc(x_2); +lean_inc(x_1); +x_86 = l_Lean_Server_Completion_optionCompletion(x_1, x_33, x_85, x_84, x_2, x_10); +lean_dec(x_84); +if (lean_obj_tag(x_86) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_apply_2(x_17, x_64, x_65); -return x_66; +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_apply_3(x_34, x_9, x_87, x_88); +x_11 = x_89; +goto block_28; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_63); -if (x_67 == 0) +uint8_t x_90; +lean_dec(x_9); +x_90 = !lean_is_exclusive(x_86); +if (x_90 == 0) { -return x_63; +x_11 = x_86; +goto block_28; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_63, 0); -x_69 = lean_ctor_get(x_63, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_63); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_86, 0); +x_92 = lean_ctor_get(x_86, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_86); +x_93 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_93, 0, x_91); +lean_ctor_set(x_93, 1, x_92); +x_11 = x_93; +goto block_28; } } } -case 6: +case 7: { -uint8_t x_71; -lean_free_object(x_12); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_1); -x_71 = !lean_is_exclusive(x_16); -if (x_71 == 0) +lean_object* x_94; lean_object* x_95; +lean_dec(x_35); +x_94 = lean_ctor_get(x_32, 1); +lean_inc(x_94); +lean_dec(x_32); +lean_inc(x_1); +x_95 = l_Lean_Server_Completion_tacticCompletion(x_1, x_33, x_94, x_10); +if (lean_obj_tag(x_95) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_16, 1); -lean_dec(x_72); -x_73 = lean_ctor_get(x_16, 0); -lean_dec(x_73); -x_74 = lean_box(0); -lean_ctor_set_tag(x_16, 0); -lean_ctor_set(x_16, 1, x_7); -lean_ctor_set(x_16, 0, x_74); -return x_16; +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_apply_3(x_34, x_9, x_96, x_97); +x_11 = x_98; +goto block_28; } else { -lean_object* x_75; lean_object* x_76; -lean_dec(x_16); -x_75 = lean_box(0); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_7); -return x_76; +uint8_t x_99; +lean_dec(x_9); +x_99 = !lean_is_exclusive(x_95); +if (x_99 == 0) +{ +x_11 = x_95; +goto block_28; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_95, 0); +x_101 = lean_ctor_get(x_95, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_95); +x_102 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_102, 0, x_100); +lean_ctor_set(x_102, 1, x_101); +x_11 = x_102; +goto block_28; +} } } default: { -lean_object* x_77; -lean_free_object(x_12); -lean_dec(x_16); -lean_dec(x_13); -lean_dec(x_6); -x_77 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion(x_1, x_15, x_7); -if (lean_obj_tag(x_77) == 0) +lean_object* x_103; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_32); +lean_inc(x_3); +x_103 = lean_apply_3(x_34, x_9, x_3, x_10); +x_11 = x_103; +goto block_28; +} +} +} +block_28: { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_apply_2(x_17, x_78, x_79); -return x_80; +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_ctor_get(x_12, 0); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_ctor_get(x_12, 0); +lean_inc(x_20); +lean_dec(x_12); +x_21 = 1; +x_22 = lean_usize_add(x_8, x_21); +x_8 = x_22; +x_9 = x_20; +x_10 = x_19; +goto _start; +} } else { -uint8_t x_81; -x_81 = !lean_is_exclusive(x_77); -if (x_81 == 0) +uint8_t x_24; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_11); +if (x_24 == 0) { -return x_77; +return x_11; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_77, 0); -x_83 = lean_ctor_get(x_77, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_77); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_11, 0); +x_26 = lean_ctor_get(x_11, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_11); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } } } -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_12, 0); -x_86 = lean_ctor_get(x_12, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_12); -x_87 = l_Lean_Server_Completion_find_x3f___closed__1; -switch (lean_obj_tag(x_86)) { -case 0: +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_88; lean_object* x_89; -lean_dec(x_6); -x_88 = lean_ctor_get(x_86, 0); -lean_inc(x_88); -lean_dec(x_86); -x_89 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion(x_1, x_85, x_88, x_13, x_7); -if (lean_obj_tag(x_89) == 0) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_8, x_7); +if (x_11 == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_89, 1); -lean_inc(x_91); -lean_dec(x_89); -x_92 = lean_apply_2(x_87, x_90, x_91); -return x_92; +lean_object* x_12; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_93 = lean_ctor_get(x_89, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_89, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_95 = x_89; -} else { - lean_dec_ref(x_89); - x_95 = lean_box(0); -} -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); -} else { - x_96 = x_95; -} -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; -} -} -case 1: +lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; +x_13 = lean_array_uget(x_6, x_8); +x_14 = lean_box(0); +x_15 = lean_array_size(x_13); +x_16 = 0; +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1(x_1, x_2, x_4, x_13, x_14, x_13, x_15, x_16, x_9, x_10); +lean_dec(x_13); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_6); -x_97 = lean_ctor_get(x_86, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_86, 1); -lean_inc(x_98); -x_99 = lean_ctor_get_uint8(x_86, sizeof(void*)*4); -x_100 = lean_ctor_get(x_86, 2); -lean_inc(x_100); -lean_dec(x_86); -x_101 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletion(x_1, x_85, x_100, x_97, x_98, x_13, x_99, x_7); -if (lean_obj_tag(x_101) == 0) -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_101, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_101, 1); -lean_inc(x_103); -lean_dec(x_101); -x_104 = lean_apply_2(x_87, x_102, x_103); -return x_104; +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +x_21 = l_Array_isEmpty___rarg(x_19); +if (x_21 == 0) +{ +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_17; } else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_105 = lean_ctor_get(x_101, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_101, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_107 = x_101; -} else { - lean_dec_ref(x_101); - x_107 = lean_box(0); -} -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); -} else { - x_108 = x_107; -} -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_106); -return x_108; +size_t x_22; size_t x_23; +lean_free_object(x_17); +x_22 = 1; +x_23 = lean_usize_add(x_8, x_22); +x_8 = x_23; +x_9 = x_19; +x_10 = x_20; +goto _start; } } -case 2: +else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_13); -lean_dec(x_6); -x_109 = lean_ctor_get(x_86, 1); -lean_inc(x_109); -x_110 = lean_ctor_get(x_86, 2); -lean_inc(x_110); -x_111 = lean_ctor_get(x_86, 3); -lean_inc(x_111); -lean_dec(x_86); -x_112 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotIdCompletion(x_1, x_85, x_110, x_109, x_111, x_7); -if (lean_obj_tag(x_112) == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); -lean_inc(x_114); -lean_dec(x_112); -x_115 = lean_apply_2(x_87, x_113, x_114); -return x_115; +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_17); +x_27 = l_Array_isEmpty___rarg(x_25); +if (x_27 == 0) +{ +lean_object* x_28; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_26); +return x_28; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_116 = lean_ctor_get(x_112, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_112, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_118 = x_112; -} else { - lean_dec_ref(x_112); - x_118 = lean_box(0); -} -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(1, 2, 0); -} else { - x_119 = x_118; +size_t x_29; size_t x_30; +x_29 = 1; +x_30 = lean_usize_add(x_8, x_29); +x_8 = x_30; +x_9 = x_25; +x_10 = x_26; +goto _start; } -lean_ctor_set(x_119, 0, x_116); -lean_ctor_set(x_119, 1, x_117); -return x_119; } } -case 3: +else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -lean_dec(x_13); -lean_dec(x_6); -x_120 = lean_ctor_get(x_86, 1); -lean_inc(x_120); -x_121 = lean_ctor_get(x_86, 2); -lean_inc(x_121); -x_122 = lean_ctor_get(x_86, 3); -lean_inc(x_122); -lean_dec(x_86); -x_123 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion(x_1, x_85, x_121, x_120, x_122, x_7); -if (lean_obj_tag(x_123) == 0) +uint8_t x_32; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_17); +if (x_32 == 0) { -lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); -lean_inc(x_125); -lean_dec(x_123); -x_126 = lean_apply_2(x_87, x_124, x_125); -return x_126; +return x_17; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_127 = lean_ctor_get(x_123, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_123, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_129 = x_123; -} else { - lean_dec_ref(x_123); - x_129 = lean_box(0); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_17, 0); +x_34 = lean_ctor_get(x_17, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_17); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } -if (lean_is_scalar(x_129)) { - x_130 = lean_alloc_ctor(1, 2, 0); -} else { - x_130 = x_129; } -lean_ctor_set(x_130, 0, x_127); -lean_ctor_set(x_130, 1, x_128); -return x_130; } } -case 4: -{ -lean_object* x_131; lean_object* x_132; -lean_dec(x_86); -lean_dec(x_85); -lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_1); -x_131 = lean_box(0); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_131); -lean_ctor_set(x_132, 1, x_7); -return x_132; } -case 5: +LEAN_EXPORT lean_object* l_Lean_Server_Completion_find_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: { -lean_object* x_133; lean_object* x_134; -lean_dec(x_13); -x_133 = lean_ctor_get(x_86, 0); -lean_inc(x_133); -lean_dec(x_86); -x_134 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion(x_1, x_85, x_133, x_6, x_7); -lean_dec(x_133); -if (lean_obj_tag(x_134) == 0) +lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_8 = l_Lean_Server_Completion_findPrioritizedCompletionPartitionsAt(x_2, x_3, x_4, x_5); +x_9 = lean_box(0); +x_10 = lean_array_size(x_8); +x_11 = 0; +x_12 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1; +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2(x_1, x_6, x_8, x_12, x_9, x_8, x_10, x_11, x_12, x_7); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); -x_137 = lean_apply_2(x_87, x_135, x_136); -return x_137; -} -else +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_138 = lean_ctor_get(x_134, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_134, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_140 = x_134; -} else { - lean_dec_ref(x_134); - x_140 = lean_box(0); -} -if (lean_is_scalar(x_140)) { - x_141 = lean_alloc_ctor(1, 2, 0); -} else { - x_141 = x_140; -} -lean_ctor_set(x_141, 0, x_138); -lean_ctor_set(x_141, 1, x_139); -return x_141; -} +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems(x_15); +lean_dec(x_15); +x_17 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_16); +x_18 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(x_17); +lean_dec(x_17); +x_19 = 1; +x_20 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_19); +lean_ctor_set(x_13, 0, x_20); +return x_13; } -case 6: +else { -lean_object* x_142; lean_object* x_143; lean_object* x_144; -lean_dec(x_85); +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); lean_dec(x_13); -lean_dec(x_6); -lean_dec(x_1); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_142 = x_86; -} else { - lean_dec_ref(x_86); - x_142 = lean_box(0); -} -x_143 = lean_box(0); -if (lean_is_scalar(x_142)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_142; - lean_ctor_set_tag(x_144, 0); +x_23 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems(x_21); +lean_dec(x_21); +x_24 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems(x_23); +x_25 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts(x_24); +lean_dec(x_24); +x_26 = 1; +x_27 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_22); +return x_28; } -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_7); -return x_144; } -default: +else { -lean_object* x_145; -lean_dec(x_86); -lean_dec(x_13); -lean_dec(x_6); -x_145 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion(x_1, x_85, x_7); -if (lean_obj_tag(x_145) == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_13); +if (x_29 == 0) { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = lean_apply_2(x_87, x_146, x_147); -return x_148; +return x_13; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_149 = lean_ctor_get(x_145, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_145, 1); -lean_inc(x_150); -if (lean_is_exclusive(x_145)) { - lean_ctor_release(x_145, 0); - lean_ctor_release(x_145, 1); - x_151 = x_145; -} else { - lean_dec_ref(x_145); - x_151 = lean_box(0); -} -if (lean_is_scalar(x_151)) { - x_152 = lean_alloc_ctor(1, 2, 0); -} else { - x_152 = x_151; -} -lean_ctor_set(x_152, 0, x_149); -lean_ctor_set(x_152, 1, x_150); -return x_152; -} +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_13, 0); +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_13); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } } +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_8; -x_8 = l___private_Lean_Server_Completion_0__Lean_Server_Completion_findCompletionInfoAt_x3f(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_9, x_10); lean_dec(x_6); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_5); -lean_ctor_set(x_9, 1, x_7); -return x_9; +lean_dec(x_5); +lean_dec(x_4); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = lean_unbox_usize(x_8); lean_dec(x_8); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Elab_CompletionInfo_lctx(x_13); -lean_dec(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve), 7, 2); -lean_closure_set(x_15, 0, x_5); -lean_closure_set(x_15, 1, x_6); -x_16 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_12, x_14, x_15, x_7); -return x_16; -} +x_13 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_12, x_9, x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_13; } } -lean_object* initialize_Lean_Environment(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Parser_Term(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_FuzzyMatching(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Lsp_LanguageFeatures(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Lsp_Capabilities(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Lsp_Utf16(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_CompletionName(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Tactic_Apply(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Meta_Match_MatcherInfo(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Elab_Tactic_Doc(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Server_InfoUtils(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Parser_Extension(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Server_FileSource(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Server_CompletionItemData(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_CompletionCollectors(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Server_Completion(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Environment(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Parser_Term(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Data_FuzzyMatching(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Data_Lsp_LanguageFeatures(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Data_Lsp_Capabilities(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Data_Lsp_Utf16(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_CompletionName(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Tactic_Apply(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Meta_Match_MatcherInfo(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Elab_Tactic_Doc(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Server_InfoUtils(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Parser_Extension(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Server_FileSource(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Server_CompletionItemData(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__1___closed__2); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__1); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__2); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__3); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__4); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__5); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__6); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__7); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__8); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__9); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____lambda__3___closed__10); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__1); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__2); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__3); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__4); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__5); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__6); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__7); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_136____closed__8); -l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1); -l_Lean_Lsp_instFromJsonCompletionIdentifier = _init_l_Lean_Lsp_instFromJsonCompletionIdentifier(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionIdentifier); -l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion___hyg_270____closed__1); -l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1); -l_Lean_Lsp_instToJsonCompletionIdentifier = _init_l_Lean_Lsp_instToJsonCompletionIdentifier(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionIdentifier); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____spec__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__1); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__2); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__3); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__4); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__5); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__6); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__7); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__8); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__9); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__10); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__11); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__12); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__13); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__14); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__15); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__16); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__17); -l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366____closed__18); -l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemDataWithId___closed__1); -l_Lean_Lsp_instFromJsonCompletionItemDataWithId = _init_l_Lean_Lsp_instFromJsonCompletionItemDataWithId(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemDataWithId); -l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Lsp_toJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_472____closed__1); -l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemDataWithId___closed__1); -l_Lean_Lsp_instToJsonCompletionItemDataWithId = _init_l_Lean_Lsp_instToJsonCompletionItemDataWithId(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemDataWithId); -l_Lean_Lsp_CompletionItem_resolve___closed__1 = _init_l_Lean_Lsp_CompletionItem_resolve___closed__1(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_resolve___closed__1); -l_Lean_Lsp_CompletionItem_resolve___closed__2 = _init_l_Lean_Lsp_CompletionItem_resolve___closed__2(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_resolve___closed__2); -if (builtin) {res = l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion___hyg_743_(lean_io_mk_world()); +res = initialize_Lean_Server_Completion_CompletionCollectors(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; -l_Lean_Server_Completion_eligibleHeaderDeclsRef = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_Server_Completion_eligibleHeaderDeclsRef); lean_dec_ref(res); -}l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1(); -lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1); -l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2(); -lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2); -l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3(); -lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3); -l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4(); -lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4); -l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5(); -lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5); +l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__1(); +l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__2(); +l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__7___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__8___closed__3(); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__11___closed__1); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__1); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__2); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__3); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__4); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__5); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__6); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__7); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__8); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__9); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__10); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__11); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__12); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__13); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__14); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__15); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__16); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__17); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__18); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__19); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__20); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__21); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__22); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__23); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__24); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__25); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__26); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__27); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__28); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__29); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__30); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__31); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__32); +l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__1___closed__33); +l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__1); +l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__2); +l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3 = _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__3); +l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4 = _init_l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4(); +lean_mark_persistent(l_Std_DHashMap_Internal_AssocList_map_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___spec__18___closed__4); +l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_filterDuplicateCompletionItems___closed__1); l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1 = _init_l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1(); lean_mark_persistent(l_Array_qsort_sort___at___private_Lean_Server_Completion_0__Lean_Server_Completion_sortCompletionItems___spec__2___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addItem___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8); -l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1 = _init_l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1(); -lean_mark_persistent(l_Lean_isProjectionFn___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1); -l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1(); -lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1); -l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2(); -lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2); -l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3(); -lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3); -l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1 = _init_l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1(); -lean_mark_persistent(l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__1); -l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2 = _init_l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2(); -lean_mark_persistent(l_List_allM___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__5); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__6); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__7); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_getCompletionKindForDecl___closed__8); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchAtomic___closed__1(); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2(); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1); -l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1 = _init_l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1(); -lean_mark_persistent(l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___spec__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_truncate_go___closed__4); -l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1 = _init_l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1(); -lean_mark_persistent(l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1); -l_Lean_Server_Completion_matchNamespace___closed__1 = _init_l_Lean_Server_Completion_matchNamespace___closed__1(); -lean_mark_persistent(l_Lean_Server_Completion_matchNamespace___closed__1); -l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1); -l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2); -l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3 = _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3(); -lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__3); -l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1 = _init_l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1); -l_Lean_Server_Completion_completeNamespaces___closed__1 = _init_l_Lean_Server_Completion_completeNamespaces___closed__1(); -lean_mark_persistent(l_Lean_Server_Completion_completeNamespaces___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1); -l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1(); -lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_idCompletionCore___spec__17___closed__1); -l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1); -l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1(); -lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6); -l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7(); -lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__6___closed__5); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_dotCompletion___lambda__7___closed__1); -l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1); -l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2(); -lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_fieldIdCompletion___closed__3); -l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__1); -l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__2); -l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__3); -l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4 = _init_l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4(); -lean_mark_persistent(l_Lean_RBNode_forIn_visit___at___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___spec__1___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_optionCompletion___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_tacticCompletion___spec__1___boxed__const__1); -l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1 = _init_l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1(); -lean_mark_persistent(l_panic___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1); -l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1); -l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2); -l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3); -l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3); -l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1 = _init_l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5); -l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1 = _init_l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1); -l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1 = _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1(); -lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1); -l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2 = _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2(); -lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2); -l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1 = _init_l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Server_Completion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2___boxed__const__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2___boxed__const__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_0__Lean_Server_Completion_assignSortTexts___spec__2___boxed__const__1); -l_Lean_Server_Completion_find_x3f___closed__1 = _init_l_Lean_Server_Completion_find_x3f___closed__1(); -lean_mark_persistent(l_Lean_Server_Completion_find_x3f___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_find_x3f___spec__1___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c b/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c new file mode 100644 index 000000000000..09481773c710 --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/CompletionCollectors.c @@ -0,0 +1,23164 @@ +// Lean compiler output +// Module: Lean.Server.Completion.CompletionCollectors +// Imports: Lean.Data.FuzzyMatching Lean.Elab.Tactic.Doc Lean.Server.Completion.CompletionResolution Lean.Server.Completion.EligibleHeaderDecls +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_getNumParts(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(lean_object*, lean_object*, size_t, size_t); +lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getStructureParentInfo(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5; +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint32_t lean_string_utf8_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; +static lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1; +lean_object* lean_private_to_user_name(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_ConstantInfo_type(lean_object*); +lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Server_Completion_allowCompletion(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double lean_float_div(double, double); +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Exception_isInterrupt(lean_object*); +LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_matchNamespace___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_idCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_optionCompletion___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7(lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Name_isAnonymous(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_idCompletion___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_insertAt_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_ConstantInfo_name(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2; +LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_environment_find(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_LocalContext_empty; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6; +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; +lean_object* l_Lean_getAliasState(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(size_t, lean_object*, lean_object*, size_t, size_t, lean_object*); +uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_InductiveVal_numTypeFormers(lean_object*); +static lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_instInhabitedNat; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_getTokenTable(lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(lean_object*, lean_object*, size_t, size_t); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate(lean_object*, lean_object*); +extern lean_object* l_Lean_projectionFnInfoExt; +uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Linter_instInhabitedDeprecationEntry; +static double l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Server_Completion_dotCompletion___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Server_Completion_fieldIdCompletion___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_List_elem___at_Lean_addAliasEntry___spec__16(lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_isTheorem(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(lean_object*, lean_object*, size_t, size_t); +lean_object* l_Lean_Data_Trie_findPrefix_go___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_erase___at_Lean_getAllParentStructures___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___boxed(lean_object**); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +lean_object* l_Lean_Name_getPrefix(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_isCtor(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2; +lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1; +lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at_Lean_Server_Completion_dotCompletion___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_List_isEmpty___rarg(lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +extern lean_object* l_Lean_NameSSet_instInhabited; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_isPrivatePrefix(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getOptionDecls(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_tacticCompletion(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_optionCompletion___closed__4; +uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; +lean_object* l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(lean_object*, lean_object*, double); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_optionCompletion___closed__2; +LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1; +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_name_eq(lean_object*, lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(lean_object*); +static lean_object* l_Lean_Server_Completion_optionCompletion___closed__5; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2; +lean_object* l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3; +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_fvarId(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___boxed(lean_object**); +lean_object* l_Lean_Elab_Tactic_Doc_allTacticDocs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go(lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; +lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; +lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +extern lean_object* l_Lean_structureResolutionExt; +static lean_object* l_Lean_Server_Completion_optionCompletion___closed__1; +static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_userName(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Linter_deprecatedAttr; +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2; +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4; +lean_object* l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_tacticCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Name_hasMacroScopes(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_searchAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Name_isAtomic(lean_object*); +uint8_t l_Lean_Meta_allowCompletion(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3; +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(lean_object*, lean_object*, size_t, size_t); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_mod(lean_object*, lean_object*); +uint8_t l_Lean_ConstantInfo_isInductive(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(lean_object*, size_t, size_t, lean_object*); +lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_type(lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +uint8_t l_Lean_isStructure(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1(lean_object*); +uint8_t l_Lean_Expr_isProp(lean_object*); +lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix(lean_object*); +lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_getString_x21(lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_erase_macro_scopes(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_namespacesExt; +lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +extern lean_object* l_Lean_instInhabitedName; +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_completeNamespaces___closed__1; +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8(lean_object*, lean_object*, lean_object*, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(lean_object*, lean_object*, size_t, size_t); +uint8_t l_Lean_Name_isSuffixOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(lean_object*, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(lean_object*); +static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Char_utf8Size(uint32_t); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Exception_isRuntime(lean_object*); +uint8_t lean_string_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4; +uint8_t l_Lean_Expr_isForall(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(uint8_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_is_class(lean_object*, lean_object*); +lean_object* l_Lean_Expr_consumeMData(lean_object*); +uint8_t l_Lean_Name_isInternal(lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; +lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_data_value_to_string(lean_object*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +static lean_object* _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_3 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_1); +lean_ctor_set(x_3, 3, x_1); +lean_ctor_set(x_3, 4, x_1); +lean_ctor_set(x_3, 5, x_1); +lean_ctor_set(x_3, 6, x_1); +lean_ctor_set(x_3, 7, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; double x_4; +x_1 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2; +x_2 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3; +x_3 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_3, 0, x_1); +x_4 = lean_unbox_float(x_2); +lean_ctor_set_float(x_3, sizeof(void*)*1, x_4); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_4, 0); +x_12 = lean_ctor_get(x_4, 1); +lean_inc(x_12); +lean_inc(x_11); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +lean_ctor_set(x_13, 2, x_3); +x_14 = !lean_is_exclusive(x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_15 = lean_ctor_get(x_1, 6); +lean_dec(x_15); +x_16 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(x_13); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_1, 6, x_17); +x_18 = lean_st_ref_take(x_5, x_10); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set_float(x_21, sizeof(void*)*1, x_2); +x_22 = lean_array_push(x_19, x_21); +x_23 = lean_st_ref_set(x_5, x_22, x_20); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = lean_box(0); +lean_ctor_set(x_23, 0, x_26); +return x_23; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +x_34 = lean_ctor_get(x_1, 4); +x_35 = lean_ctor_get(x_1, 5); +x_36 = lean_ctor_get(x_1, 7); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_1); +x_37 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(x_13); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_39, 0, x_30); +lean_ctor_set(x_39, 1, x_31); +lean_ctor_set(x_39, 2, x_32); +lean_ctor_set(x_39, 3, x_33); +lean_ctor_set(x_39, 4, x_34); +lean_ctor_set(x_39, 5, x_35); +lean_ctor_set(x_39, 6, x_38); +lean_ctor_set(x_39, 7, x_36); +x_40 = lean_st_ref_take(x_5, x_10); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_43, 0, x_39); +lean_ctor_set_float(x_43, sizeof(void*)*1, x_2); +x_44 = lean_array_push(x_41, x_43); +x_45 = lean_st_ref_set(x_5, x_44, x_42); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_47 = x_45; +} else { + lean_dec_ref(x_45); + x_47 = lean_box(0); +} +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_47; +} +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +double x_11; lean_object* x_12; +x_11 = lean_unbox_float(x_2); +lean_dec(x_2); +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_2 = 1; +x_3 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +x_4 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_4, 0, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__1), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\n\n", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, double x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1; +x_16 = 1; +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_18 = l_Lean_Name_toString(x_1, x_16, x_17); +x_19 = lean_box(0); +x_20 = lean_box(x_2); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_3); +if (lean_obj_tag(x_4) == 0) +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_19); +lean_ctor_set(x_23, 2, x_19); +lean_ctor_set(x_23, 3, x_21); +lean_ctor_set(x_23, 4, x_19); +lean_ctor_set(x_23, 5, x_19); +lean_ctor_set(x_23, 6, x_19); +lean_ctor_set(x_23, 7, x_5); +x_24 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_23, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_7, 0); +lean_inc(x_25); +lean_dec(x_7); +x_26 = lean_apply_1(x_15, x_25); +x_27 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_26); +lean_ctor_set(x_27, 3, x_21); +lean_ctor_set(x_27, 4, x_19); +lean_ctor_set(x_27, 5, x_19); +lean_ctor_set(x_27, 6, x_19); +lean_ctor_set(x_27, 7, x_5); +x_28 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_27, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_28; +} +} +else +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_4, 0); +lean_inc(x_29); +lean_dec(x_4); +x_30 = lean_apply_1(x_15, x_29); +x_31 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_31, 0, x_18); +lean_ctor_set(x_31, 1, x_19); +lean_ctor_set(x_31, 2, x_30); +lean_ctor_set(x_31, 3, x_21); +lean_ctor_set(x_31, 4, x_19); +lean_ctor_set(x_31, 5, x_19); +lean_ctor_set(x_31, 6, x_19); +lean_ctor_set(x_31, 7, x_5); +x_32 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_31, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_ctor_get(x_4, 0); +lean_inc(x_33); +lean_dec(x_4); +x_34 = lean_ctor_get(x_7, 0); +lean_inc(x_34); +lean_dec(x_7); +x_35 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_36 = lean_string_append(x_35, x_33); +lean_dec(x_33); +x_37 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3; +x_38 = lean_string_append(x_36, x_37); +x_39 = lean_string_append(x_38, x_34); +lean_dec(x_34); +x_40 = lean_string_append(x_39, x_35); +x_41 = lean_apply_1(x_15, x_40); +x_42 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_42, 0, x_18); +lean_ctor_set(x_42, 1, x_19); +lean_ctor_set(x_42, 2, x_41); +lean_ctor_set(x_42, 3, x_21); +lean_ctor_set(x_42, 4, x_19); +lean_ctor_set(x_42, 5, x_19); +lean_ctor_set(x_42, 6, x_19); +lean_ctor_set(x_42, 7, x_5); +x_43 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_42, x_6, x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_43; +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Linter_deprecatedAttr; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("`", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` has been deprecated.", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` has been deprecated, use `", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("` instead.", 10, 10); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(lean_object* x_1, lean_object* x_2, uint8_t x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_st_ref_get(x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_2, 0); +lean_inc(x_27); +x_28 = l_Lean_Linter_instInhabitedDeprecationEntry; +x_29 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1; +lean_inc(x_27); +x_30 = l_Lean_ParametricAttribute_getParam_x3f___rarg(x_28, x_29, x_15, x_27); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +lean_dec(x_27); +x_31 = lean_box(0); +x_16 = x_31; +x_17 = x_31; +goto block_26; +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_30, 0); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +if (lean_obj_tag(x_35) == 0) +{ +uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_36 = 1; +x_37 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_38 = l_Lean_Name_toString(x_27, x_36, x_37); +x_39 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +x_40 = lean_string_append(x_39, x_38); +lean_dec(x_38); +x_41 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; +x_42 = lean_string_append(x_40, x_41); +lean_ctor_set(x_30, 0, x_42); +x_43 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_30; +x_17 = x_43; +goto block_26; +} +else +{ +uint8_t x_44; +lean_free_object(x_30); +x_44 = !lean_is_exclusive(x_35); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_45 = lean_ctor_get(x_35, 0); +x_46 = 1; +x_47 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_48 = l_Lean_Name_toString(x_27, x_46, x_47); +x_49 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +x_50 = lean_string_append(x_49, x_48); +lean_dec(x_48); +x_51 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; +x_52 = lean_string_append(x_50, x_51); +x_53 = l_Lean_Name_toString(x_45, x_46, x_47); +x_54 = lean_string_append(x_52, x_53); +lean_dec(x_53); +x_55 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; +x_56 = lean_string_append(x_54, x_55); +lean_ctor_set(x_35, 0, x_56); +x_57 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_35; +x_17 = x_57; +goto block_26; +} +else +{ +lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_58 = lean_ctor_get(x_35, 0); +lean_inc(x_58); +lean_dec(x_35); +x_59 = 1; +x_60 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_61 = l_Lean_Name_toString(x_27, x_59, x_60); +x_62 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +x_63 = lean_string_append(x_62, x_61); +lean_dec(x_61); +x_64 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; +x_65 = lean_string_append(x_63, x_64); +x_66 = l_Lean_Name_toString(x_58, x_59, x_60); +x_67 = lean_string_append(x_65, x_66); +lean_dec(x_66); +x_68 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; +x_69 = lean_string_append(x_67, x_68); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_69); +x_71 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_70; +x_17 = x_71; +goto block_26; +} +} +} +else +{ +uint8_t x_72; +lean_free_object(x_30); +lean_dec(x_33); +lean_dec(x_27); +x_72 = !lean_is_exclusive(x_34); +if (x_72 == 0) +{ +lean_object* x_73; +x_73 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_34; +x_17 = x_73; +goto block_26; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_34, 0); +lean_inc(x_74); +lean_dec(x_34); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_74); +x_76 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_75; +x_17 = x_76; +goto block_26; +} +} +} +else +{ +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_30, 0); +lean_inc(x_77); +lean_dec(x_30); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_77, 0); +lean_inc(x_79); +lean_dec(x_77); +if (lean_obj_tag(x_79) == 0) +{ +uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_80 = 1; +x_81 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_82 = l_Lean_Name_toString(x_27, x_80, x_81); +x_83 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +x_84 = lean_string_append(x_83, x_82); +lean_dec(x_82); +x_85 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6; +x_86 = lean_string_append(x_84, x_85); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_86); +x_88 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_87; +x_17 = x_88; +goto block_26; +} +else +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_89 = lean_ctor_get(x_79, 0); +lean_inc(x_89); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + x_90 = x_79; +} else { + lean_dec_ref(x_79); + x_90 = lean_box(0); +} +x_91 = 1; +x_92 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_93 = l_Lean_Name_toString(x_27, x_91, x_92); +x_94 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5; +x_95 = lean_string_append(x_94, x_93); +lean_dec(x_93); +x_96 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7; +x_97 = lean_string_append(x_95, x_96); +x_98 = l_Lean_Name_toString(x_89, x_91, x_92); +x_99 = lean_string_append(x_97, x_98); +lean_dec(x_98); +x_100 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8; +x_101 = lean_string_append(x_99, x_100); +if (lean_is_scalar(x_90)) { + x_102 = lean_alloc_ctor(1, 1, 0); +} else { + x_102 = x_90; +} +lean_ctor_set(x_102, 0, x_101); +x_103 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_102; +x_17 = x_103; +goto block_26; +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_77); +lean_dec(x_27); +x_104 = lean_ctor_get(x_78, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + x_105 = x_78; +} else { + lean_dec_ref(x_78); + x_105 = lean_box(0); +} +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(1, 1, 0); +} else { + x_106 = x_105; +} +lean_ctor_set(x_106, 0, x_104); +x_107 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4; +x_16 = x_106; +x_17 = x_107; +goto block_26; +} +} +} +} +else +{ +lean_object* x_108; +x_108 = lean_box(0); +x_16 = x_108; +x_17 = x_108; +goto block_26; +} +block_26: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = 1; +x_20 = l_Lean_findDocString_x3f(x_15, x_18, x_19, x_14); +lean_dec(x_15); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3(x_1, x_3, x_2, x_16, x_17, x_4, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_15); +x_24 = lean_box(0); +x_25 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3(x_1, x_3, x_2, x_16, x_17, x_4, x_24, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__2(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; double x_16; lean_object* x_17; +x_15 = lean_unbox(x_2); +lean_dec(x_2); +x_16 = lean_unbox_float(x_6); +lean_dec(x_6); +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3(x_1, x_15, x_3, x_4, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; double x_13; lean_object* x_14; +x_12 = lean_unbox(x_3); +lean_dec(x_3); +x_13 = lean_unbox_float(x_4); +lean_dec(x_4); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_2, x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_14; +} +} +static lean_object* _init_l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_projectionFnInfoExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_instInhabitedProjectionFunctionInfo; +x_14 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; +x_15 = l_Lean_MapDeclarationExtension_contains___rarg(x_13, x_14, x_12, x_1); +lean_dec(x_12); +x_16 = lean_box(x_15); +lean_ctor_set(x_9, 0, x_16); +return x_9; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_instInhabitedProjectionFunctionInfo; +x_21 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1; +x_22 = l_Lean_MapDeclarationExtension_contains___rarg(x_20, x_21, x_19, x_1); +lean_dec(x_19); +x_23 = lean_box(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_18); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_6, 5); +x_10 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_12); +lean_ctor_set_tag(x_10, 1); +lean_ctor_set(x_10, 0, x_13); +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_10, 0); +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_10); +lean_inc(x_9); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unknown constant '", 18, 18); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_1); +x_14 = lean_environment_find(x_13, x_1); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_free_object(x_9); +x_15 = 0; +x_16 = l_Lean_MessageData_ofConstName(x_1, x_15); +x_17 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +x_19 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4; +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_1); +x_22 = lean_ctor_get(x_14, 0); +lean_inc(x_22); +lean_dec(x_14); +lean_ctor_set(x_9, 0, x_22); +return x_9; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_9); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_1); +x_26 = lean_environment_find(x_25, x_1); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_27 = 0; +x_28 = l_Lean_MessageData_ofConstName(x_1, x_27); +x_29 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_24); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_1); +x_34 = lean_ctor_get(x_26, 0); +lean_inc(x_34); +lean_dec(x_26); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_24); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_9 = 1; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_22; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + x_14 = x_1; +} else { + lean_dec_ref(x_1); + x_14 = lean_box(0); +} +x_22 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +if (lean_obj_tag(x_23) == 6) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_25, 4); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_eq(x_26, x_27); +lean_dec(x_26); +x_15 = x_28; +x_16 = x_24; +goto block_21; +} +else +{ +lean_object* x_29; uint8_t x_30; +lean_dec(x_23); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_dec(x_22); +x_30 = 0; +x_15 = x_30; +x_16 = x_29; +goto block_21; +} +} +else +{ +uint8_t x_31; +lean_dec(x_14); +lean_dec(x_13); +x_31 = !lean_is_exclusive(x_22); +if (x_31 == 0) +{ +return x_22; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_22, 0); +x_33 = lean_ctor_get(x_22, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_22); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +block_21: +{ +if (x_15 == 0) +{ +uint8_t x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_13); +x_17 = 0; +x_18 = lean_box(x_17); +if (lean_is_scalar(x_14)) { + x_19 = lean_alloc_ctor(0, 2, 0); +} else { + x_19 = x_14; + lean_ctor_set_tag(x_19, 0); +} +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +return x_19; +} +else +{ +lean_dec(x_14); +x_1 = x_13; +x_8 = x_16; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 5) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_15, 2); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_Expr_isProp(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Lean_InductiveVal_numTypeFormers(x_14); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_dec_eq(x_18, x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; +lean_dec(x_14); +x_21 = 0; +x_22 = lean_box(x_21); +lean_ctor_set(x_9, 0, x_22); +return x_9; +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_14, 2); +lean_inc(x_23); +x_24 = lean_unsigned_to_nat(0u); +x_25 = lean_nat_dec_eq(x_23, x_24); +lean_dec(x_23); +if (x_25 == 0) +{ +uint8_t x_26; lean_object* x_27; +lean_dec(x_14); +x_26 = 0; +x_27 = lean_box(x_26); +lean_ctor_set(x_9, 0, x_27); +return x_9; +} +else +{ +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_14, 1); +lean_inc(x_28); +x_29 = lean_nat_dec_eq(x_28, x_24); +lean_dec(x_28); +if (x_29 == 0) +{ +uint8_t x_30; lean_object* x_31; +lean_dec(x_14); +x_30 = 0; +x_31 = lean_box(x_30); +lean_ctor_set(x_9, 0, x_31); +return x_9; +} +else +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_14, 4); +lean_inc(x_32); +x_33 = l_List_isEmpty___rarg(x_32); +if (x_33 == 0) +{ +uint8_t x_34; +x_34 = lean_ctor_get_uint8(x_14, sizeof(void*)*6); +if (x_34 == 0) +{ +uint8_t x_35; +x_35 = lean_ctor_get_uint8(x_14, sizeof(void*)*6 + 1); +lean_dec(x_14); +if (x_35 == 0) +{ +lean_object* x_36; +lean_free_object(x_9); +x_36 = l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +return x_36; +} +else +{ +uint8_t x_37; lean_object* x_38; +lean_dec(x_32); +x_37 = 0; +x_38 = lean_box(x_37); +lean_ctor_set(x_9, 0, x_38); +return x_9; +} +} +else +{ +uint8_t x_39; lean_object* x_40; +lean_dec(x_32); +lean_dec(x_14); +x_39 = 0; +x_40 = lean_box(x_39); +lean_ctor_set(x_9, 0, x_40); +return x_9; +} +} +else +{ +uint8_t x_41; lean_object* x_42; +lean_dec(x_32); +lean_dec(x_14); +x_41 = 0; +x_42 = lean_box(x_41); +lean_ctor_set(x_9, 0, x_42); +return x_9; +} +} +} +} +} +else +{ +uint8_t x_43; lean_object* x_44; +lean_dec(x_14); +x_43 = 0; +x_44 = lean_box(x_43); +lean_ctor_set(x_9, 0, x_44); +return x_9; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_45 = lean_ctor_get(x_9, 1); +lean_inc(x_45); +lean_dec(x_9); +x_46 = lean_ctor_get(x_10, 0); +lean_inc(x_46); +lean_dec(x_10); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_47, 2); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_Lean_Expr_isProp(x_48); +lean_dec(x_48); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = l_Lean_InductiveVal_numTypeFormers(x_46); +x_51 = lean_unsigned_to_nat(1u); +x_52 = lean_nat_dec_eq(x_50, x_51); +lean_dec(x_50); +if (x_52 == 0) +{ +uint8_t x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_46); +x_53 = 0; +x_54 = lean_box(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_45); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_56 = lean_ctor_get(x_46, 2); +lean_inc(x_56); +x_57 = lean_unsigned_to_nat(0u); +x_58 = lean_nat_dec_eq(x_56, x_57); +lean_dec(x_56); +if (x_58 == 0) +{ +uint8_t x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_46); +x_59 = 0; +x_60 = lean_box(x_59); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_45); +return x_61; +} +else +{ +lean_object* x_62; uint8_t x_63; +x_62 = lean_ctor_get(x_46, 1); +lean_inc(x_62); +x_63 = lean_nat_dec_eq(x_62, x_57); +lean_dec(x_62); +if (x_63 == 0) +{ +uint8_t x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_46); +x_64 = 0; +x_65 = lean_box(x_64); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_45); +return x_66; +} +else +{ +lean_object* x_67; uint8_t x_68; +x_67 = lean_ctor_get(x_46, 4); +lean_inc(x_67); +x_68 = l_List_isEmpty___rarg(x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = lean_ctor_get_uint8(x_46, sizeof(void*)*6); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = lean_ctor_get_uint8(x_46, sizeof(void*)*6 + 1); +lean_dec(x_46); +if (x_70 == 0) +{ +lean_object* x_71; +x_71 = l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_45); +return x_71; +} +else +{ +uint8_t x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_67); +x_72 = 0; +x_73 = lean_box(x_72); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_45); +return x_74; +} +} +else +{ +uint8_t x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_67); +lean_dec(x_46); +x_75 = 0; +x_76 = lean_box(x_75); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_45); +return x_77; +} +} +else +{ +uint8_t x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_67); +lean_dec(x_46); +x_78 = 0; +x_79 = lean_box(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_45); +return x_80; +} +} +} +} +} +else +{ +uint8_t x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_46); +x_81 = 0; +x_82 = lean_box(x_81); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_45); +return x_83; +} +} +} +else +{ +uint8_t x_84; +lean_dec(x_10); +x_84 = !lean_is_exclusive(x_9); +if (x_84 == 0) +{ +lean_object* x_85; uint8_t x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_9, 0); +lean_dec(x_85); +x_86 = 0; +x_87 = lean_box(x_86); +lean_ctor_set(x_9, 0, x_87); +return x_9; +} +else +{ +lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_9, 1); +lean_inc(x_88); +lean_dec(x_9); +x_89 = 0; +x_90 = lean_box(x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_88); +return x_91; +} +} +} +else +{ +uint8_t x_92; +x_92 = !lean_is_exclusive(x_9); +if (x_92 == 0) +{ +return x_9; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_9, 0); +x_94 = lean_ctor_get(x_9, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_9); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_ConstantInfo_isCtor(x_1); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_ConstantInfo_isInductive(x_1); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_13); +x_16 = l_Lean_ConstantInfo_isTheorem(x_1); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_free_object(x_9); +x_17 = l_Lean_ConstantInfo_name(x_1); +x_18 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = l_Lean_ConstantInfo_type(x_1); +x_23 = !lean_is_exclusive(x_6); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_6, 9); +lean_dec(x_24); +x_25 = lean_unsigned_to_nat(0u); +lean_ctor_set(x_6, 9, x_25); +x_26 = lean_whnf(x_22, x_4, x_5, x_6, x_7, x_21); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_26, 0); +x_29 = l_Lean_Expr_isForall(x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +uint8_t x_30; lean_object* x_31; +x_30 = 20; +x_31 = lean_box(x_30); +lean_ctor_set(x_26, 0, x_31); +return x_26; +} +else +{ +uint8_t x_32; lean_object* x_33; +x_32 = 2; +x_33 = lean_box(x_32); +lean_ctor_set(x_26, 0, x_33); +return x_26; +} +} +else +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_26, 0); +x_35 = lean_ctor_get(x_26, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_26); +x_36 = l_Lean_Expr_isForall(x_34); +lean_dec(x_34); +if (x_36 == 0) +{ +uint8_t x_37; lean_object* x_38; lean_object* x_39; +x_37 = 20; +x_38 = lean_box(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_35); +return x_39; +} +else +{ +uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_40 = 2; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_35); +return x_42; +} +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_26); +if (x_43 == 0) +{ +return x_26; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_26, 0); +x_45 = lean_ctor_get(x_26, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_26); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_47 = lean_ctor_get(x_6, 0); +x_48 = lean_ctor_get(x_6, 1); +x_49 = lean_ctor_get(x_6, 2); +x_50 = lean_ctor_get(x_6, 3); +x_51 = lean_ctor_get(x_6, 4); +x_52 = lean_ctor_get(x_6, 5); +x_53 = lean_ctor_get(x_6, 6); +x_54 = lean_ctor_get(x_6, 7); +x_55 = lean_ctor_get(x_6, 8); +x_56 = lean_ctor_get(x_6, 10); +x_57 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); +x_58 = lean_ctor_get(x_6, 11); +x_59 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +lean_inc(x_58); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_6); +x_60 = lean_unsigned_to_nat(0u); +x_61 = lean_alloc_ctor(0, 12, 2); +lean_ctor_set(x_61, 0, x_47); +lean_ctor_set(x_61, 1, x_48); +lean_ctor_set(x_61, 2, x_49); +lean_ctor_set(x_61, 3, x_50); +lean_ctor_set(x_61, 4, x_51); +lean_ctor_set(x_61, 5, x_52); +lean_ctor_set(x_61, 6, x_53); +lean_ctor_set(x_61, 7, x_54); +lean_ctor_set(x_61, 8, x_55); +lean_ctor_set(x_61, 9, x_60); +lean_ctor_set(x_61, 10, x_56); +lean_ctor_set(x_61, 11, x_58); +lean_ctor_set_uint8(x_61, sizeof(void*)*12, x_57); +lean_ctor_set_uint8(x_61, sizeof(void*)*12 + 1, x_59); +x_62 = lean_whnf(x_22, x_4, x_5, x_61, x_7, x_21); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_65 = x_62; +} else { + lean_dec_ref(x_62); + x_65 = lean_box(0); +} +x_66 = l_Lean_Expr_isForall(x_63); +lean_dec(x_63); +if (x_66 == 0) +{ +uint8_t x_67; lean_object* x_68; lean_object* x_69; +x_67 = 20; +x_68 = lean_box(x_67); +if (lean_is_scalar(x_65)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_65; +} +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_64); +return x_69; +} +else +{ +uint8_t x_70; lean_object* x_71; lean_object* x_72; +x_70 = 2; +x_71 = lean_box(x_70); +if (lean_is_scalar(x_65)) { + x_72 = lean_alloc_ctor(0, 2, 0); +} else { + x_72 = x_65; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_64); +return x_72; +} +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_62, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_62, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_75 = x_62; +} else { + lean_dec_ref(x_62); + x_75 = lean_box(0); +} +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(1, 2, 0); +} else { + x_76 = x_75; +} +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_74); +return x_76; +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_77 = !lean_is_exclusive(x_18); +if (x_77 == 0) +{ +lean_object* x_78; uint8_t x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_18, 0); +lean_dec(x_78); +x_79 = 4; +x_80 = lean_box(x_79); +lean_ctor_set(x_18, 0, x_80); +return x_18; +} +else +{ +lean_object* x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; +x_81 = lean_ctor_get(x_18, 1); +lean_inc(x_81); +lean_dec(x_18); +x_82 = 4; +x_83 = lean_box(x_82); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_81); +return x_84; +} +} +} +else +{ +uint8_t x_85; lean_object* x_86; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_85 = 22; +x_86 = lean_box(x_85); +lean_ctor_set(x_9, 0, x_86); +return x_9; +} +} +else +{ +lean_object* x_87; uint8_t x_88; +x_87 = l_Lean_ConstantInfo_name(x_1); +lean_inc(x_87); +x_88 = lean_is_class(x_13, x_87); +if (x_88 == 0) +{ +lean_object* x_89; +lean_free_object(x_9); +x_89 = l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_89) == 0) +{ +lean_object* x_90; uint8_t x_91; +x_90 = lean_ctor_get(x_89, 0); +lean_inc(x_90); +x_91 = lean_unbox(x_90); +lean_dec(x_90); +if (x_91 == 0) +{ +uint8_t x_92; +x_92 = !lean_is_exclusive(x_89); +if (x_92 == 0) +{ +lean_object* x_93; uint8_t x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_89, 0); +lean_dec(x_93); +x_94 = 21; +x_95 = lean_box(x_94); +lean_ctor_set(x_89, 0, x_95); +return x_89; +} +else +{ +lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_89, 1); +lean_inc(x_96); +lean_dec(x_89); +x_97 = 21; +x_98 = lean_box(x_97); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_96); +return x_99; +} +} +else +{ +uint8_t x_100; +x_100 = !lean_is_exclusive(x_89); +if (x_100 == 0) +{ +lean_object* x_101; uint8_t x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_89, 0); +lean_dec(x_101); +x_102 = 12; +x_103 = lean_box(x_102); +lean_ctor_set(x_89, 0, x_103); +return x_89; +} +else +{ +lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_89, 1); +lean_inc(x_104); +lean_dec(x_89); +x_105 = 12; +x_106 = lean_box(x_105); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; +} +} +} +else +{ +uint8_t x_108; +x_108 = !lean_is_exclusive(x_89); +if (x_108 == 0) +{ +return x_89; +} +else +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_89, 0); +x_110 = lean_ctor_get(x_89, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_89); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; +} +} +} +else +{ +uint8_t x_112; lean_object* x_113; +lean_dec(x_87); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_112 = 6; +x_113 = lean_box(x_112); +lean_ctor_set(x_9, 0, x_113); +return x_9; +} +} +} +else +{ +uint8_t x_114; lean_object* x_115; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_114 = 3; +x_115 = lean_box(x_114); +lean_ctor_set(x_9, 0, x_115); +return x_9; +} +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_116 = lean_ctor_get(x_9, 0); +x_117 = lean_ctor_get(x_9, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_9); +x_118 = lean_ctor_get(x_116, 0); +lean_inc(x_118); +lean_dec(x_116); +x_119 = l_Lean_ConstantInfo_isCtor(x_1); +if (x_119 == 0) +{ +uint8_t x_120; +x_120 = l_Lean_ConstantInfo_isInductive(x_1); +if (x_120 == 0) +{ +uint8_t x_121; +lean_dec(x_118); +x_121 = l_Lean_ConstantInfo_isTheorem(x_1); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +x_122 = l_Lean_ConstantInfo_name(x_1); +x_123 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_122, x_2, x_3, x_4, x_5, x_6, x_7, x_117); +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_unbox(x_124); +lean_dec(x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +lean_dec(x_123); +x_127 = l_Lean_ConstantInfo_type(x_1); +x_128 = lean_ctor_get(x_6, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_6, 1); +lean_inc(x_129); +x_130 = lean_ctor_get(x_6, 2); +lean_inc(x_130); +x_131 = lean_ctor_get(x_6, 3); +lean_inc(x_131); +x_132 = lean_ctor_get(x_6, 4); +lean_inc(x_132); +x_133 = lean_ctor_get(x_6, 5); +lean_inc(x_133); +x_134 = lean_ctor_get(x_6, 6); +lean_inc(x_134); +x_135 = lean_ctor_get(x_6, 7); +lean_inc(x_135); +x_136 = lean_ctor_get(x_6, 8); +lean_inc(x_136); +x_137 = lean_ctor_get(x_6, 10); +lean_inc(x_137); +x_138 = lean_ctor_get_uint8(x_6, sizeof(void*)*12); +x_139 = lean_ctor_get(x_6, 11); +lean_inc(x_139); +x_140 = lean_ctor_get_uint8(x_6, sizeof(void*)*12 + 1); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + lean_ctor_release(x_6, 2); + lean_ctor_release(x_6, 3); + lean_ctor_release(x_6, 4); + lean_ctor_release(x_6, 5); + lean_ctor_release(x_6, 6); + lean_ctor_release(x_6, 7); + lean_ctor_release(x_6, 8); + lean_ctor_release(x_6, 9); + lean_ctor_release(x_6, 10); + lean_ctor_release(x_6, 11); + x_141 = x_6; +} else { + lean_dec_ref(x_6); + x_141 = lean_box(0); +} +x_142 = lean_unsigned_to_nat(0u); +if (lean_is_scalar(x_141)) { + x_143 = lean_alloc_ctor(0, 12, 2); +} else { + x_143 = x_141; +} +lean_ctor_set(x_143, 0, x_128); +lean_ctor_set(x_143, 1, x_129); +lean_ctor_set(x_143, 2, x_130); +lean_ctor_set(x_143, 3, x_131); +lean_ctor_set(x_143, 4, x_132); +lean_ctor_set(x_143, 5, x_133); +lean_ctor_set(x_143, 6, x_134); +lean_ctor_set(x_143, 7, x_135); +lean_ctor_set(x_143, 8, x_136); +lean_ctor_set(x_143, 9, x_142); +lean_ctor_set(x_143, 10, x_137); +lean_ctor_set(x_143, 11, x_139); +lean_ctor_set_uint8(x_143, sizeof(void*)*12, x_138); +lean_ctor_set_uint8(x_143, sizeof(void*)*12 + 1, x_140); +x_144 = lean_whnf(x_127, x_4, x_5, x_143, x_7, x_126); +if (lean_obj_tag(x_144) == 0) +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_147 = x_144; +} else { + lean_dec_ref(x_144); + x_147 = lean_box(0); +} +x_148 = l_Lean_Expr_isForall(x_145); +lean_dec(x_145); +if (x_148 == 0) +{ +uint8_t x_149; lean_object* x_150; lean_object* x_151; +x_149 = 20; +x_150 = lean_box(x_149); +if (lean_is_scalar(x_147)) { + x_151 = lean_alloc_ctor(0, 2, 0); +} else { + x_151 = x_147; +} +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_146); +return x_151; +} +else +{ +uint8_t x_152; lean_object* x_153; lean_object* x_154; +x_152 = 2; +x_153 = lean_box(x_152); +if (lean_is_scalar(x_147)) { + x_154 = lean_alloc_ctor(0, 2, 0); +} else { + x_154 = x_147; +} +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_146); +return x_154; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_155 = lean_ctor_get(x_144, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_144, 1); +lean_inc(x_156); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_157 = x_144; +} else { + lean_dec_ref(x_144); + x_157 = lean_box(0); +} +if (lean_is_scalar(x_157)) { + x_158 = lean_alloc_ctor(1, 2, 0); +} else { + x_158 = x_157; +} +lean_ctor_set(x_158, 0, x_155); +lean_ctor_set(x_158, 1, x_156); +return x_158; +} +} +else +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_159 = lean_ctor_get(x_123, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_160 = x_123; +} else { + lean_dec_ref(x_123); + x_160 = lean_box(0); +} +x_161 = 4; +x_162 = lean_box(x_161); +if (lean_is_scalar(x_160)) { + x_163 = lean_alloc_ctor(0, 2, 0); +} else { + x_163 = x_160; +} +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_159); +return x_163; +} +} +else +{ +uint8_t x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_164 = 22; +x_165 = lean_box(x_164); +x_166 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_166, 0, x_165); +lean_ctor_set(x_166, 1, x_117); +return x_166; +} +} +else +{ +lean_object* x_167; uint8_t x_168; +x_167 = l_Lean_ConstantInfo_name(x_1); +lean_inc(x_167); +x_168 = lean_is_class(x_118, x_167); +if (x_168 == 0) +{ +lean_object* x_169; +x_169 = l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_167, x_2, x_3, x_4, x_5, x_6, x_7, x_117); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +if (lean_obj_tag(x_169) == 0) +{ +lean_object* x_170; uint8_t x_171; +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_unbox(x_170); +lean_dec(x_170); +if (x_171 == 0) +{ +lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; +x_172 = lean_ctor_get(x_169, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_173 = x_169; +} else { + lean_dec_ref(x_169); + x_173 = lean_box(0); +} +x_174 = 21; +x_175 = lean_box(x_174); +if (lean_is_scalar(x_173)) { + x_176 = lean_alloc_ctor(0, 2, 0); +} else { + x_176 = x_173; +} +lean_ctor_set(x_176, 0, x_175); +lean_ctor_set(x_176, 1, x_172); +return x_176; +} +else +{ +lean_object* x_177; lean_object* x_178; uint8_t x_179; lean_object* x_180; lean_object* x_181; +x_177 = lean_ctor_get(x_169, 1); +lean_inc(x_177); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_178 = x_169; +} else { + lean_dec_ref(x_169); + x_178 = lean_box(0); +} +x_179 = 12; +x_180 = lean_box(x_179); +if (lean_is_scalar(x_178)) { + x_181 = lean_alloc_ctor(0, 2, 0); +} else { + x_181 = x_178; +} +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_177); +return x_181; +} +} +else +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_169, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_169, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_184 = x_169; +} else { + lean_dec_ref(x_169); + x_184 = lean_box(0); +} +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(1, 2, 0); +} else { + x_185 = x_184; +} +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +return x_185; +} +} +else +{ +uint8_t x_186; lean_object* x_187; lean_object* x_188; +lean_dec(x_167); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_186 = 6; +x_187 = lean_box(x_186); +x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_117); +return x_188; +} +} +} +else +{ +uint8_t x_189; lean_object* x_190; lean_object* x_191; +lean_dec(x_118); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_189 = 3; +x_190 = lean_box(x_189); +x_191 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_117); +return x_191; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_List_allM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_isEnumType___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(lean_object* x_1, lean_object* x_2, double x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_2); +x_16 = lean_environment_find(x_15, x_2); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_box(0); +lean_ctor_set(x_11, 0, x_17); +return x_11; +} +else +{ +uint8_t x_18; +lean_free_object(x_11); +x_18 = !lean_is_exclusive(x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_20 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +lean_ctor_set_tag(x_16, 0); +lean_ctor_set(x_16, 0, x_2); +x_23 = lean_unbox(x_21); +lean_dec(x_21); +x_24 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_16, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_22); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_24; +} +else +{ +uint8_t x_25; +lean_free_object(x_16); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_20); +if (x_25 == 0) +{ +return x_20; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_20, 0); +x_27 = lean_ctor_get(x_20, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_20); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_16, 0); +lean_inc(x_29); +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_30 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_2); +x_34 = lean_unbox(x_31); +lean_dec(x_31); +x_35 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_33, x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_32); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_ctor_get(x_30, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_30, 1); +lean_inc(x_37); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_38 = x_30; +} else { + lean_dec_ref(x_30); + x_38 = lean_box(0); +} +if (lean_is_scalar(x_38)) { + x_39 = lean_alloc_ctor(1, 2, 0); +} else { + x_39 = x_38; +} +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_37); +return x_39; +} +} +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_11, 0); +x_41 = lean_ctor_get(x_11, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_11); +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +lean_inc(x_2); +x_43 = lean_environment_find(x_42, x_2); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_44 = lean_box(0); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_41); +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_43, 0); +lean_inc(x_46); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + x_47 = x_43; +} else { + lean_dec_ref(x_43); + x_47 = lean_box(0); +} +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_48 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_46, x_4, x_5, x_6, x_7, x_8, x_9, x_41); +lean_dec(x_46); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +if (lean_is_scalar(x_47)) { + x_51 = lean_alloc_ctor(0, 1, 0); +} else { + x_51 = x_47; + lean_ctor_set_tag(x_51, 0); +} +lean_ctor_set(x_51, 0, x_2); +x_52 = lean_unbox(x_49); +lean_dec(x_49); +x_53 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_1, x_51, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_50); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_47); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_54 = lean_ctor_get(x_48, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_48, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_56 = x_48; +} else { + lean_dec_ref(x_48); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +double x_11; lean_object* x_12; +x_11 = lean_unbox_float(x_3); +lean_dec(x_3); +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("keyword", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 13; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_box(0); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2; +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; +x_13 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +lean_ctor_set(x_13, 4, x_10); +lean_ctor_set(x_13, 5, x_10); +lean_ctor_set(x_13, 6, x_10); +lean_ctor_set(x_13, 7, x_10); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_13, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +double x_10; lean_object* x_11; +x_10 = lean_unbox_float(x_2); +lean_dec(x_2); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("namespace", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 8; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_10 = 1; +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_12 = l_Lean_Name_toString(x_1, x_10, x_11); +x_13 = lean_box(0); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2; +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3; +x_16 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_16, 0, x_12); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_15); +lean_ctor_set(x_16, 4, x_13); +lean_ctor_set(x_16, 5, x_13); +lean_ctor_set(x_16, 6, x_13); +lean_ctor_set(x_16, 7, x_13); +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_16, x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +double x_10; lean_object* x_11; +x_10 = lean_unbox_float(x_2); +lean_dec(x_2); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_st_mk_ref(x_1, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +lean_inc(x_10); +x_12 = lean_apply_7(x_2, x_3, x_10, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_get(x_10, x_13); +lean_dec(x_10); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +return x_14; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_10); +x_19 = !lean_is_exclusive(x_12); +if (x_19 == 0) +{ +return x_12; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 0); +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_12); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_2); +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___lambda__1), 8, 3); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_5); +lean_closure_set(x_9, 2, x_7); +x_10 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_4, x_9, x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +lean_inc(x_1); +x_7 = lean_private_to_user_name(x_1); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_st_ref_get(x_5, x_6); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_12, 0); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +lean_inc(x_11); +x_16 = l_Lean_mkPrivateName(x_15, x_11); +x_17 = lean_name_eq(x_16, x_1); +lean_dec(x_1); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_free_object(x_7); +lean_dec(x_11); +x_18 = lean_box(0); +lean_ctor_set(x_12, 0, x_18); +return x_12; +} +else +{ +lean_ctor_set(x_12, 0, x_7); +return x_12; +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_12, 0); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_12); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +lean_inc(x_11); +x_22 = l_Lean_mkPrivateName(x_21, x_11); +x_23 = lean_name_eq(x_22, x_1); +lean_dec(x_1); +lean_dec(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_free_object(x_7); +lean_dec(x_11); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_20); +return x_25; +} +else +{ +lean_object* x_26; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_7); +lean_ctor_set(x_26, 1, x_20); +return x_26; +} +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_27 = lean_ctor_get(x_7, 0); +lean_inc(x_27); +lean_dec(x_7); +x_28 = lean_st_ref_get(x_5, x_6); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_31 = x_28; +} else { + lean_dec_ref(x_28); + x_31 = lean_box(0); +} +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +lean_inc(x_27); +x_33 = l_Lean_mkPrivateName(x_32, x_27); +x_34 = lean_name_eq(x_33, x_1); +lean_dec(x_1); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_27); +x_35 = lean_box(0); +if (lean_is_scalar(x_31)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_31; +} +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_30); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_27); +if (lean_is_scalar(x_31)) { + x_38 = lean_alloc_ctor(0, 2, 0); +} else { + x_38 = x_31; +} +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_30); +return x_38; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed), 6, 0); +return x_1; +} +} +static double _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(1u); +x_2 = 1; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; +} +} +static double _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1() { +_start: +{ +double x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_2 = lean_box_float(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_box(0); +x_12 = l_Lean_Name_replacePrefix(x_1, x_2, x_11); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1; +if (x_3 == 0) +{ +if (lean_obj_tag(x_4) == 1) +{ +if (lean_obj_tag(x_12) == 1) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_4, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +lean_dec(x_4); +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +x_18 = lean_name_eq(x_14, x_16); +if (x_18 == 0) +{ +uint8_t x_19; +x_19 = l_Lean_Name_isAnonymous(x_14); +lean_dec(x_14); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_12); +x_20 = lean_box(0); +x_21 = lean_apply_6(x_13, x_20, x_6, x_7, x_8, x_9, x_10); +return x_21; +} +else +{ +double x_22; lean_object* x_23; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_22 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +x_23 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_15, x_17, x_22); +lean_dec(x_17); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_16); +lean_dec(x_12); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_10); +return x_25; +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_27 = lean_ctor_get(x_23, 0); +x_28 = l_Lean_Name_getNumParts(x_16); +lean_dec(x_16); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_28, x_29); +lean_dec(x_28); +x_31 = 0; +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_Float_ofScientific(x_30, x_31, x_32); +lean_dec(x_30); +x_34 = lean_unbox_float(x_27); +lean_dec(x_27); +x_35 = lean_float_div(x_34, x_33); +x_36 = lean_box_float(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_12); +lean_ctor_set(x_37, 1, x_36); +lean_ctor_set(x_23, 0, x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_23); +lean_ctor_set(x_38, 1, x_10); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; double x_45; double x_46; double x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_39 = lean_ctor_get(x_23, 0); +lean_inc(x_39); +lean_dec(x_23); +x_40 = l_Lean_Name_getNumParts(x_16); +lean_dec(x_16); +x_41 = lean_unsigned_to_nat(1u); +x_42 = lean_nat_add(x_40, x_41); +lean_dec(x_40); +x_43 = 0; +x_44 = lean_unsigned_to_nat(0u); +x_45 = l_Float_ofScientific(x_42, x_43, x_44); +lean_dec(x_42); +x_46 = lean_unbox_float(x_39); +lean_dec(x_39); +x_47 = lean_float_div(x_46, x_45); +x_48 = lean_box_float(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_12); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_10); +return x_51; +} +} +} +} +else +{ +double x_52; lean_object* x_53; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_52 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +x_53 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_15, x_17, x_52); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_17); +x_54 = lean_box(0); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_10); +return x_55; +} +else +{ +uint8_t x_56; +x_56 = !lean_is_exclusive(x_53); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_53, 0); +x_58 = l_Lean_Name_str___override(x_11, x_17); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +lean_ctor_set(x_53, 0, x_59); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_53); +lean_ctor_set(x_60, 1, x_10); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_53, 0); +lean_inc(x_61); +lean_dec(x_53); +x_62 = l_Lean_Name_str___override(x_11, x_17); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_10); +return x_65; +} +} +} +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_12); +lean_dec(x_4); +x_66 = lean_box(0); +x_67 = lean_apply_6(x_13, x_66, x_6, x_7, x_8, x_9, x_10); +return x_67; +} +} +else +{ +lean_object* x_68; lean_object* x_69; +lean_dec(x_12); +lean_dec(x_4); +x_68 = lean_box(0); +x_69 = lean_apply_6(x_13, x_68, x_6, x_7, x_8, x_9, x_10); +return x_69; +} +} +else +{ +uint8_t x_70; +x_70 = l_Lean_Name_isPrefixOf(x_4, x_12); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +lean_dec(x_12); +lean_dec(x_4); +x_71 = lean_box(0); +x_72 = lean_apply_6(x_13, x_71, x_6, x_7, x_8, x_9, x_10); +return x_72; +} +else +{ +lean_object* x_73; uint8_t x_74; +x_73 = l_Lean_Name_replacePrefix(x_12, x_4, x_11); +lean_dec(x_4); +x_74 = l_Lean_Name_isAtomic(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +lean_dec(x_73); +x_75 = lean_box(0); +x_76 = lean_apply_6(x_13, x_75, x_6, x_7, x_8, x_9, x_10); +return x_76; +} +else +{ +uint8_t x_77; +x_77 = l_Lean_Name_isAnonymous(x_73); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_78 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_73); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_10); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; +lean_dec(x_73); +x_82 = lean_box(0); +x_83 = lean_apply_6(x_13, x_82, x_6, x_7, x_8, x_9, x_10); +return x_83; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_dec(x_13); +x_14 = lean_box(0); +lean_ctor_set(x_10, 0, x_14); +return x_10; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_15); +lean_dec(x_10); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_10); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_10, 1); +x_20 = lean_ctor_get(x_10, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_11, 0); +lean_inc(x_21); +lean_dec(x_11); +x_22 = l_Lean_Name_isPrefixOf(x_1, x_21); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_21); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_23 = lean_box(0); +lean_ctor_set(x_10, 0, x_23); +return x_10; +} +else +{ +lean_object* x_24; lean_object* x_25; +lean_free_object(x_10); +x_24 = lean_box(0); +x_25 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_21, x_1, x_3, x_2, x_24, x_5, x_6, x_7, x_8, x_19); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_10, 1); +lean_inc(x_26); +lean_dec(x_10); +x_27 = lean_ctor_get(x_11, 0); +lean_inc(x_27); +lean_dec(x_11); +x_28 = l_Lean_Name_isPrefixOf(x_1, x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_27); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_26); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_box(0); +x_32 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_27, x_1, x_3, x_2, x_31, x_5, x_6, x_7, x_8, x_26); +return x_32; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_3); +lean_dec(x_3); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_10, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) +{ +if (lean_obj_tag(x_2) == 1) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 0); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; double x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_ctor_get(x_2, 1); +x_7 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +x_8 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_5, x_6, x_7); +return x_8; +} +else +{ +lean_object* x_9; +lean_dec(x_1); +x_9 = lean_box(0); +return x_9; +} +} +else +{ +lean_object* x_10; +lean_dec(x_1); +x_10 = lean_box(0); +return x_10; +} +} +else +{ +lean_object* x_11; +lean_dec(x_3); +lean_dec(x_1); +x_11 = lean_box(0); +return x_11; +} +} +else +{ +lean_object* x_12; +lean_dec(x_1); +x_12 = lean_box(0); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedName; +x_2 = l_instInhabitedNat; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Server.Completion.CompletionCollectors", 43, 43); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Lean.Server.Completion.CompletionCollectors.0.Lean.Server.Completion.truncate.go", 89, 89); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unreachable code has been reached", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1; +x_2 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2; +x_3 = lean_unsigned_to_nat(201u); +x_4 = lean_unsigned_to_nat(26u); +x_5 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_3; lean_object* x_4; +lean_dec(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +case 1: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_7 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go(x_1, x_5); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_10, x_11); +x_13 = lean_nat_dec_le(x_1, x_12); +if (x_13 == 0) +{ +uint8_t x_14; lean_object* x_15; +lean_dec(x_9); +x_14 = l_Lean_Name_isAnonymous(x_5); +x_15 = lean_string_length(x_6); +if (x_14 == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_nat_add(x_12, x_15); +lean_dec(x_15); +lean_dec(x_12); +x_17 = lean_nat_dec_le(x_16, x_1); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_16); +lean_dec(x_2); +x_18 = lean_nat_sub(x_1, x_11); +x_19 = lean_nat_sub(x_18, x_10); +lean_dec(x_10); +lean_dec(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_string_utf8_extract(x_6, x_20, x_19); +lean_dec(x_19); +lean_dec(x_6); +x_22 = l_Lean_Name_str___override(x_5, x_21); +lean_ctor_set(x_7, 1, x_1); +lean_ctor_set(x_7, 0, x_22); +return x_7; +} +else +{ +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +lean_ctor_set(x_7, 1, x_16); +lean_ctor_set(x_7, 0, x_2); +return x_7; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_12); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_add(x_10, x_23); +x_25 = lean_nat_add(x_24, x_15); +lean_dec(x_15); +lean_dec(x_24); +x_26 = lean_nat_dec_le(x_25, x_1); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_25); +lean_dec(x_2); +x_27 = lean_nat_sub(x_1, x_23); +x_28 = lean_nat_sub(x_27, x_10); +lean_dec(x_10); +lean_dec(x_27); +x_29 = lean_string_utf8_extract(x_6, x_23, x_28); +lean_dec(x_28); +lean_dec(x_6); +x_30 = l_Lean_Name_str___override(x_5, x_29); +lean_ctor_set(x_7, 1, x_1); +lean_ctor_set(x_7, 0, x_30); +return x_7; +} +else +{ +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +lean_ctor_set(x_7, 1, x_25); +lean_ctor_set(x_7, 0, x_2); +return x_7; +} +} +} +else +{ +lean_dec(x_12); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_32, x_33); +x_35 = lean_nat_dec_le(x_1, x_34); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; +lean_dec(x_31); +x_36 = l_Lean_Name_isAnonymous(x_5); +x_37 = lean_string_length(x_6); +if (x_36 == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_nat_add(x_34, x_37); +lean_dec(x_37); +lean_dec(x_34); +x_39 = lean_nat_dec_le(x_38, x_1); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_38); +lean_dec(x_2); +x_40 = lean_nat_sub(x_1, x_33); +x_41 = lean_nat_sub(x_40, x_32); +lean_dec(x_32); +lean_dec(x_40); +x_42 = lean_unsigned_to_nat(0u); +x_43 = lean_string_utf8_extract(x_6, x_42, x_41); +lean_dec(x_41); +lean_dec(x_6); +x_44 = l_Lean_Name_str___override(x_5, x_43); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_1); +return x_45; +} +else +{ +lean_object* x_46; +lean_dec(x_32); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_2); +lean_ctor_set(x_46, 1, x_38); +return x_46; +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +lean_dec(x_34); +x_47 = lean_unsigned_to_nat(0u); +x_48 = lean_nat_add(x_32, x_47); +x_49 = lean_nat_add(x_48, x_37); +lean_dec(x_37); +lean_dec(x_48); +x_50 = lean_nat_dec_le(x_49, x_1); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_49); +lean_dec(x_2); +x_51 = lean_nat_sub(x_1, x_47); +x_52 = lean_nat_sub(x_51, x_32); +lean_dec(x_32); +lean_dec(x_51); +x_53 = lean_string_utf8_extract(x_6, x_47, x_52); +lean_dec(x_52); +lean_dec(x_6); +x_54 = l_Lean_Name_str___override(x_5, x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_1); +return x_55; +} +else +{ +lean_object* x_56; +lean_dec(x_32); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_2); +lean_ctor_set(x_56, 1, x_49); +return x_56; +} +} +} +else +{ +lean_object* x_57; +lean_dec(x_34); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_31); +lean_ctor_set(x_57, 1, x_32); +return x_57; +} +} +} +default: +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_2); +lean_dec(x_1); +x_58 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4; +x_59 = l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1(x_58); +return x_59; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go(x_2, x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1() { +_start: +{ +double x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_2 = lean_box_float(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Server_Completion_matchNamespace___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace(lean_object* x_1, lean_object* x_2, uint8_t x_3) { +_start: +{ +if (x_3 == 0) +{ +if (lean_obj_tag(x_1) == 1) +{ +if (lean_obj_tag(x_2) == 1) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +lean_dec(x_2); +x_8 = lean_name_eq(x_4, x_6); +lean_dec(x_6); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_7); +x_9 = lean_box(0); +return x_9; +} +else +{ +double x_10; lean_object* x_11; +x_10 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +x_11 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_7, x_5, x_10); +return x_11; +} +} +else +{ +lean_object* x_12; +lean_dec(x_2); +x_12 = lean_box(0); +return x_12; +} +} +else +{ +lean_object* x_13; +lean_dec(x_2); +x_13 = lean_box(0); +return x_13; +} +} +else +{ +uint8_t x_14; +x_14 = lean_name_eq(x_2, x_1); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = l_Lean_Name_isPrefixOf(x_2, x_1); +lean_dec(x_2); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_box(0); +return x_16; +} +else +{ +lean_object* x_17; +x_17 = l_Lean_Server_Completion_matchNamespace___closed__1; +return x_17; +} +} +else +{ +lean_object* x_18; +lean_dec(x_2); +x_18 = lean_box(0); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_matchNamespace___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l_Lean_Server_Completion_matchNamespace(x_1, x_2, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; +lean_inc(x_1); +lean_inc(x_5); +x_13 = l_Lean_Name_append(x_5, x_1); +x_14 = l_Lean_Server_Completion_matchNamespace(x_4, x_13, x_2); +if (lean_obj_tag(x_14) == 0) +{ +if (lean_obj_tag(x_5) == 1) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +lean_dec(x_5); +x_5 = x_15; +goto _start; +} +else +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_12); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_1); +x_19 = lean_ctor_get(x_14, 0); +lean_inc(x_19); +lean_dec(x_14); +x_20 = lean_apply_10(x_3, x_4, x_5, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces_visitNamespaces___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_2); +lean_dec(x_2); +x_14 = l_Lean_Server_Completion_completeNamespaces_visitNamespaces(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_18; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +else +{ +lean_object* x_19; +lean_dec(x_9); +x_19 = lean_ctor_get(x_8, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_8, 1); +lean_inc(x_20); +lean_dec(x_8); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +lean_inc(x_1); +lean_inc(x_21); +x_22 = l_Lean_Name_append(x_21, x_1); +lean_inc(x_22); +x_23 = l_Lean_Server_Completion_matchNamespace(x_3, x_22, x_2); +if (lean_obj_tag(x_23) == 0) +{ +lean_dec(x_22); +lean_dec(x_21); +lean_inc(x_6); +{ +lean_object* _tmp_7 = x_20; +lean_object* _tmp_8 = x_6; +lean_object* _tmp_9 = lean_box(0); +x_8 = _tmp_7; +x_9 = _tmp_8; +x_10 = _tmp_9; +} +goto _start; +} +else +{ +lean_dec(x_20); +lean_dec(x_6); +lean_dec(x_1); +if (x_2 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; double x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_22); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = l_Lean_Name_replacePrefix(x_3, x_21, x_26); +lean_dec(x_21); +x_28 = lean_unbox_float(x_25); +lean_dec(x_25); +x_29 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(x_27, x_28, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +x_32 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; +lean_ctor_set(x_29, 0, x_32); +return x_29; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +x_34 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; double x_39; lean_object* x_40; uint8_t x_41; +lean_dec(x_21); +x_36 = lean_ctor_get(x_23, 0); +lean_inc(x_36); +lean_dec(x_23); +x_37 = lean_box(0); +x_38 = l_Lean_Name_replacePrefix(x_3, x_22, x_37); +lean_dec(x_22); +x_39 = lean_unbox_float(x_36); +lean_dec(x_36); +x_40 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(x_38, x_39, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 0); +lean_dec(x_42); +x_43 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; +lean_ctor_set(x_40, 0, x_43); +return x_40; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +lean_dec(x_40); +x_45 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2; +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +return x_46; +} +} +} +} +else +{ +lean_object* x_47; +lean_dec(x_19); +x_47 = lean_ctor_get(x_8, 1); +lean_inc(x_47); +lean_dec(x_8); +lean_inc(x_6); +{ +lean_object* _tmp_7 = x_47; +lean_object* _tmp_8 = x_6; +lean_object* _tmp_9 = lean_box(0); +x_8 = _tmp_7; +x_9 = _tmp_8; +x_10 = _tmp_9; +} +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_11; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 2); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = lean_apply_9(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_2 = x_16; +x_3 = x_14; +x_10 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_array_uget(x_2, x_3); +switch (lean_obj_tag(x_14)) { +case 0: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_17 = lean_apply_10(x_1, x_5, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +x_5 = x_18; +x_12 = x_19; +goto _start; +} +else +{ +uint8_t x_23; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +case 1: +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_14, 0); +lean_inc(x_27); +lean_dec(x_14); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_28 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_1, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = 1; +x_32 = lean_usize_add(x_3, x_31); +x_3 = x_32; +x_5 = x_29; +x_12 = x_30; +goto _start; +} +else +{ +uint8_t x_34; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +return x_28; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_28, 0); +x_36 = lean_ctor_get(x_28, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_28); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +default: +{ +size_t x_38; size_t x_39; +x_38 = 1; +x_39 = lean_usize_add(x_3, x_38); +x_3 = x_39; +goto _start; +} +} +} +else +{ +lean_object* x_41; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_5); +lean_ctor_set(x_41, 1, x_12); +return x_41; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed), 12, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_array_get_size(x_2); +x_15 = lean_nat_dec_lt(x_5, x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_fget(x_2, x_5); +x_18 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_19 = lean_apply_10(x_1, x_6, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_5, x_22); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_23; +x_6 = x_20; +x_13 = x_21; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_19); +if (x_25 == 0) +{ +return x_19; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_19, 0); +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_19); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed), 13, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = lean_nat_dec_le(x_12, x_12); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_3); +lean_ctor_set(x_17, 1, x_10); +return x_17; +} +else +{ +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = 0; +x_19 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(x_1, x_11, x_18, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_20; +} +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_2, 0); +x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(x_1, x_21, x_22, lean_box(0), x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg___boxed), 10, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = lean_apply_9(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = lean_box(0); +x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +x_14 = lean_array_uget(x_2, x_3); +x_15 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__3(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +x_5 = x_17; +x_12 = x_18; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_12); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_10, 1); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_15 = lean_ctor_get(x_1, 1); +x_16 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_16; +} +else +{ +uint8_t x_17; +x_17 = lean_nat_dec_le(x_12, x_12); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_12); +x_18 = lean_ctor_get(x_1, 1); +x_19 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_19; +} +else +{ +size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_20 = 0; +x_21 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_22 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_ctor_get(x_1, 1); +x_26 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_13 = lean_box(0); +x_14 = l_Lean_Name_replacePrefix(x_3, x_4, x_13); +lean_dec(x_4); +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = l_Lean_Name_append(x_4, x_2); +x_17 = lean_box(0); +x_18 = l_Lean_Name_replacePrefix(x_3, x_16, x_17); +lean_dec(x_16); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem(x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_1, 4); +lean_inc(x_14); +lean_dec(x_1); +x_15 = l_Lean_Server_Completion_completeNamespaces_visitNamespaces(x_2, x_3, x_4, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +static lean_object* _init_l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = l_Lean_Name_isInternal(x_6); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = l_Lean_Environment_contains(x_1, x_6); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_2, 0); +lean_inc(x_17); +lean_dec(x_2); +x_18 = lean_ctor_get(x_17, 5); +lean_inc(x_18); +x_19 = lean_box(0); +x_20 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; +lean_inc(x_18); +lean_inc(x_6); +lean_inc(x_3); +x_21 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(x_3, x_4, x_6, x_18, x_19, x_20, x_18, x_18, x_20, lean_box(0), x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_18); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = l_Lean_Server_Completion_completeNamespaces___lambda__2(x_17, x_3, x_4, x_5, x_6, x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_17); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_27 = !lean_is_exclusive(x_21); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_21, 0); +lean_dec(x_28); +x_29 = lean_ctor_get(x_23, 0); +lean_inc(x_29); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_29); +return x_21; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_dec(x_21); +x_31 = lean_ctor_get(x_23, 0); +lean_inc(x_31); +lean_dec(x_23); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +return x_32; +} +} +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_14); +return x_34; +} +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_14); +return x_36; +} +} +} +static lean_object* _init_l_Lean_Server_Completion_completeNamespaces___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_namespacesExt; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_11 = lean_st_ref_get(x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(x_3); +lean_inc(x_2); +x_16 = lean_alloc_closure((void*)(l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed), 12, 2); +lean_closure_set(x_16, 0, x_15); +lean_closure_set(x_16, 1, x_2); +x_17 = l_Lean_NameSSet_instInhabited; +x_18 = l_Lean_Server_Completion_completeNamespaces___closed__1; +x_19 = l_Lean_SimplePersistentEnvExtension_getState___rarg(x_17, x_18, x_14); +x_20 = lean_box(x_3); +x_21 = lean_alloc_closure((void*)(l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed), 14, 5); +lean_closure_set(x_21, 0, x_14); +lean_closure_set(x_21, 1, x_1); +lean_closure_set(x_21, 2, x_2); +lean_closure_set(x_21, 3, x_20); +lean_closure_set(x_21, 4, x_16); +x_22 = l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(x_19, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +lean_dec(x_19); +return x_22; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_2); +lean_dec(x_2); +x_19 = l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1(x_1, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__7___rarg(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Server_Completion_completeNamespaces___spec__8___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_completeNamespaces___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_completeNamespaces___spec__9(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_SMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; double x_14; lean_object* x_15; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_float(x_5); +lean_dec(x_5); +x_15 = l_Lean_Server_Completion_completeNamespaces___lambda__1(x_13, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_Lean_Server_Completion_completeNamespaces___lambda__2(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_6); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_Server_Completion_completeNamespaces___lambda__3(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_7); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_completeNamespaces___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_Server_Completion_completeNamespaces(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_unfoldDefinition_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) +{ +return x_7; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = l_Lean_Exception_isInterrupt(x_9); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = l_Lean_Exception_isRuntime(x_9); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_9); +x_12 = lean_box(0); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else +{ +return x_7; +} +} +else +{ +return x_7; +} +} +else +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_7, 0); +x_14 = lean_ctor_get(x_7, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_7); +x_15 = l_Lean_Exception_isInterrupt(x_13); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = l_Lean_Exception_isRuntime(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_13); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_14); +return x_18; +} +else +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +else +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_14); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = 0; +x_14 = lean_box(x_13); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = 0; +x_17 = lean_box(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_dec(x_9); +x_20 = lean_ctor_get(x_10, 0); +lean_inc(x_20); +lean_dec(x_10); +x_21 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(x_20, x_2, x_4, x_5, x_6, x_7, x_19); +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_22 = !lean_is_exclusive(x_9); +if (x_22 == 0) +{ +return x_9; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_9); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_8) == 4) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec(x_8); +lean_inc(x_9); +x_10 = lean_private_to_user_name(x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = lean_name_eq(x_9, x_2); +lean_dec(x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_12, x_3, x_4, x_5, x_6, x_7); +return x_13; +} +else +{ +uint8_t x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_14 = 1; +x_15 = lean_box(x_14); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_7); +return x_16; +} +} +else +{ +lean_object* x_17; uint8_t x_18; +lean_dec(x_9); +x_17 = lean_ctor_get(x_10, 0); +lean_inc(x_17); +lean_dec(x_10); +x_18 = lean_name_eq(x_17, x_2); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_19, x_3, x_4, x_5, x_6, x_7); +return x_20; +} +else +{ +uint8_t x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_21 = 1; +x_22 = lean_box(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_7); +return x_23; +} +} +} +else +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_8); +x_24 = lean_box(0); +x_25 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_7); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1; +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_lt(x_7, x_6); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_8); +x_16 = lean_array_uget(x_5, x_7); +x_17 = l_Lean_Expr_fvarId_x21(x_16); +lean_dec(x_16); +lean_inc(x_9); +x_18 = l_Lean_FVarId_getDecl(x_17, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_LocalDecl_type(x_19); +lean_dec(x_19); +x_22 = l_Lean_Expr_consumeMData(x_21); +lean_dec(x_21); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_23 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(x_22, x_1, x_9, x_10, x_11, x_12, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +lean_object* x_26; size_t x_27; size_t x_28; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = 1; +x_28 = lean_usize_add(x_7, x_27); +lean_inc(x_4); +{ +size_t _tmp_6 = x_28; +lean_object* _tmp_7 = x_4; +lean_object* _tmp_12 = x_26; +x_7 = _tmp_6; +x_8 = _tmp_7; +x_13 = _tmp_12; +} +goto _start; +} +else +{ +uint8_t x_30; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_23, 0); +lean_dec(x_31); +x_32 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2; +lean_ctor_set(x_23, 0, x_32); +return x_23; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_23, 1); +lean_inc(x_33); +lean_dec(x_23); +x_34 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +return x_35; +} +} +} +else +{ +uint8_t x_36; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +x_36 = !lean_is_exclusive(x_23); +if (x_36 == 0) +{ +return x_23; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_23, 0); +x_38 = lean_ctor_get(x_23, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_23); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_4); +x_40 = !lean_is_exclusive(x_18); +if (x_40 == 0) +{ +return x_18; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_18, 0); +x_42 = lean_ctor_get(x_18, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_18); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_7 = 0; +x_8 = lean_box(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_box(0); +x_10 = lean_array_size(x_2); +x_11 = 0; +x_12 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(x_1, x_2, x_9, x_12, x_2, x_10, x_11, x_12, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1; +x_18 = lean_box(0); +x_19 = lean_apply_6(x_17, x_18, x_4, x_5, x_6, x_7, x_16); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_13, 0); +lean_dec(x_21); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_22); +return x_13; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_13, 1); +lean_inc(x_23); +lean_dec(x_13); +x_24 = lean_ctor_get(x_15, 0); +lean_inc(x_24); +lean_dec(x_15); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_26 = !lean_is_exclusive(x_13); +if (x_26 == 0) +{ +return x_13; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_13, 0); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_13); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = l_Lean_ConstantInfo_type(x_2); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed), 8, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = 0; +x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +size_t x_14; size_t x_15; lean_object* x_16; +x_14 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_15 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1(x_1, x_2, x_3, x_4, x_5, x_14, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_Expr_consumeMData(x_3); +x_10 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDefEqToAppOf(x_9, x_1, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = l_Lean_ConstantInfo_type(x_2); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed), 8, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = 0; +x_11 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_8, x_9, x_10, x_3, x_4, x_5, x_6, x_7); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 2) +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_nat_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +uint8_t x_5; +x_5 = l_Lean_isPrivatePrefix(x_1); +if (x_5 == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_object* x_6; +x_6 = lean_box(0); +return x_6; +} +} +} +else +{ +lean_inc(x_1); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix(x_1); +lean_dec(x_1); +x_4 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_stripPrivatePrefix(x_2); +lean_dec(x_2); +switch (lean_obj_tag(x_3)) { +case 0: +{ +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +else +{ +uint8_t x_6; +lean_dec(x_4); +x_6 = 0; +return x_6; +} +} +case 1: +{ +if (lean_obj_tag(x_4) == 1) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_dec(x_3); +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_string_dec_lt(x_8, x_10); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = lean_string_dec_eq(x_8, x_10); +lean_dec(x_10); +lean_dec(x_8); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_7); +x_13 = 2; +return x_13; +} +else +{ +x_1 = x_7; +x_2 = x_9; +goto _start; +} +} +else +{ +uint8_t x_15; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_15 = 0; +return x_15; +} +} +else +{ +uint8_t x_16; +lean_dec(x_4); +lean_dec(x_3); +x_16 = 2; +return x_16; +} +} +default: +{ +switch (lean_obj_tag(x_4)) { +case 0: +{ +uint8_t x_17; +lean_dec(x_3); +x_17 = 2; +return x_17; +} +case 1: +{ +uint8_t x_18; +lean_dec(x_4); +lean_dec(x_3); +x_18 = 0; +return x_18; +} +default: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_3, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_3, 1); +lean_inc(x_20); +lean_dec(x_3); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_dec(x_4); +x_23 = lean_nat_dec_lt(x_20, x_22); +if (x_23 == 0) +{ +uint8_t x_24; +x_24 = lean_nat_dec_eq(x_20, x_22); +lean_dec(x_22); +lean_dec(x_20); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_21); +lean_dec(x_19); +x_25 = 2; +return x_25; +} +else +{ +x_1 = x_19; +x_2 = x_21; +goto _start; +} +} +else +{ +uint8_t x_27; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +x_27 = 0; +return x_27; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_box(0); +x_5 = 0; +x_6 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*4, x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +lean_inc(x_2); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_2, x_10); +switch (x_13) { +case 0: +{ +lean_object* x_14; uint8_t x_15; +x_14 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_9, x_2, x_3); +x_15 = 0; +lean_ctor_set(x_1, 0, x_14); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_15); +return x_1; +} +case 1: +{ +uint8_t x_16; +lean_dec(x_11); +lean_dec(x_10); +x_16 = 0; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_16); +return x_1; +} +default: +{ +lean_object* x_17; uint8_t x_18; +x_17 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_12, x_2, x_3); +x_18 = 0; +lean_ctor_set(x_1, 3, x_17); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_18); +return x_1; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_19 = lean_ctor_get(x_1, 0); +x_20 = lean_ctor_get(x_1, 1); +x_21 = lean_ctor_get(x_1, 2); +x_22 = lean_ctor_get(x_1, 3); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_1); +lean_inc(x_20); +lean_inc(x_2); +x_23 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_2, x_20); +switch (x_23) { +case 0: +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_19, x_2, x_3); +x_25 = 0; +x_26 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_21); +lean_ctor_set(x_26, 3, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*4, x_25); +return x_26; +} +case 1: +{ +uint8_t x_27; lean_object* x_28; +lean_dec(x_21); +lean_dec(x_20); +x_27 = 0; +x_28 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_28, 0, x_19); +lean_ctor_set(x_28, 1, x_2); +lean_ctor_set(x_28, 2, x_3); +lean_ctor_set(x_28, 3, x_22); +lean_ctor_set_uint8(x_28, sizeof(void*)*4, x_27); +return x_28; +} +default: +{ +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_22, x_2, x_3); +x_30 = 0; +x_31 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_31, 0, x_19); +lean_ctor_set(x_31, 1, x_20); +lean_ctor_set(x_31, 2, x_21); +lean_ctor_set(x_31, 3, x_29); +lean_ctor_set_uint8(x_31, sizeof(void*)*4, x_30); +return x_31; +} +} +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_1); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_33 = lean_ctor_get(x_1, 0); +x_34 = lean_ctor_get(x_1, 1); +x_35 = lean_ctor_get(x_1, 2); +x_36 = lean_ctor_get(x_1, 3); +lean_inc(x_34); +lean_inc(x_2); +x_37 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_2, x_34); +switch (x_37) { +case 0: +{ +lean_object* x_38; uint8_t x_39; +x_38 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_33, x_2, x_3); +x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*4); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 3); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_38); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_38, 3); +lean_dec(x_43); +x_44 = lean_ctor_get(x_38, 0); +lean_dec(x_44); +lean_ctor_set(x_38, 0, x_41); +x_45 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_45); +return x_1; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_38, 1); +x_47 = lean_ctor_get(x_38, 2); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_38); +x_48 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_48, 0, x_41); +lean_ctor_set(x_48, 1, x_46); +lean_ctor_set(x_48, 2, x_47); +lean_ctor_set(x_48, 3, x_41); +lean_ctor_set_uint8(x_48, sizeof(void*)*4, x_39); +x_49 = 1; +lean_ctor_set(x_1, 0, x_48); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_49); +return x_1; +} +} +else +{ +uint8_t x_50; +x_50 = lean_ctor_get_uint8(x_41, sizeof(void*)*4); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_38); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_38, 1); +x_53 = lean_ctor_get(x_38, 2); +x_54 = lean_ctor_get(x_38, 3); +lean_dec(x_54); +x_55 = lean_ctor_get(x_38, 0); +lean_dec(x_55); +x_56 = !lean_is_exclusive(x_41); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; uint8_t x_62; +x_57 = lean_ctor_get(x_41, 0); +x_58 = lean_ctor_get(x_41, 1); +x_59 = lean_ctor_get(x_41, 2); +x_60 = lean_ctor_get(x_41, 3); +x_61 = 1; +lean_ctor_set(x_41, 3, x_57); +lean_ctor_set(x_41, 2, x_53); +lean_ctor_set(x_41, 1, x_52); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_61); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_60); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_61); +x_62 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_59); +lean_ctor_set(x_1, 1, x_58); +lean_ctor_set(x_1, 0, x_41); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_62); +return x_1; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; uint8_t x_69; +x_63 = lean_ctor_get(x_41, 0); +x_64 = lean_ctor_get(x_41, 1); +x_65 = lean_ctor_get(x_41, 2); +x_66 = lean_ctor_get(x_41, 3); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_41); +x_67 = 1; +x_68 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_68, 0, x_40); +lean_ctor_set(x_68, 1, x_52); +lean_ctor_set(x_68, 2, x_53); +lean_ctor_set(x_68, 3, x_63); +lean_ctor_set_uint8(x_68, sizeof(void*)*4, x_67); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_66); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_67); +x_69 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_65); +lean_ctor_set(x_1, 1, x_64); +lean_ctor_set(x_1, 0, x_68); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_69); +return x_1; +} +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_70 = lean_ctor_get(x_38, 1); +x_71 = lean_ctor_get(x_38, 2); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_38); +x_72 = lean_ctor_get(x_41, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_41, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_41, 2); +lean_inc(x_74); +x_75 = lean_ctor_get(x_41, 3); +lean_inc(x_75); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + lean_ctor_release(x_41, 2); + lean_ctor_release(x_41, 3); + x_76 = x_41; +} else { + lean_dec_ref(x_41); + x_76 = lean_box(0); +} +x_77 = 1; +if (lean_is_scalar(x_76)) { + x_78 = lean_alloc_ctor(1, 4, 1); +} else { + x_78 = x_76; +} +lean_ctor_set(x_78, 0, x_40); +lean_ctor_set(x_78, 1, x_70); +lean_ctor_set(x_78, 2, x_71); +lean_ctor_set(x_78, 3, x_72); +lean_ctor_set_uint8(x_78, sizeof(void*)*4, x_77); +x_79 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_34); +lean_ctor_set(x_79, 2, x_35); +lean_ctor_set(x_79, 3, x_36); +lean_ctor_set_uint8(x_79, sizeof(void*)*4, x_77); +x_80 = 0; +lean_ctor_set(x_1, 3, x_79); +lean_ctor_set(x_1, 2, x_74); +lean_ctor_set(x_1, 1, x_73); +lean_ctor_set(x_1, 0, x_78); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_80); +return x_1; +} +} +else +{ +uint8_t x_81; +lean_free_object(x_1); +x_81 = !lean_is_exclusive(x_41); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_82 = lean_ctor_get(x_41, 3); +lean_dec(x_82); +x_83 = lean_ctor_get(x_41, 2); +lean_dec(x_83); +x_84 = lean_ctor_get(x_41, 1); +lean_dec(x_84); +x_85 = lean_ctor_get(x_41, 0); +lean_dec(x_85); +x_86 = 1; +lean_ctor_set(x_41, 3, x_36); +lean_ctor_set(x_41, 2, x_35); +lean_ctor_set(x_41, 1, x_34); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set_uint8(x_41, sizeof(void*)*4, x_86); +return x_41; +} +else +{ +uint8_t x_87; lean_object* x_88; +lean_dec(x_41); +x_87 = 1; +x_88 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_88, 0, x_38); +lean_ctor_set(x_88, 1, x_34); +lean_ctor_set(x_88, 2, x_35); +lean_ctor_set(x_88, 3, x_36); +lean_ctor_set_uint8(x_88, sizeof(void*)*4, x_87); +return x_88; +} +} +} +} +else +{ +uint8_t x_89; +x_89 = lean_ctor_get_uint8(x_40, sizeof(void*)*4); +if (x_89 == 0) +{ +uint8_t x_90; +x_90 = !lean_is_exclusive(x_38); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_91 = lean_ctor_get(x_38, 1); +x_92 = lean_ctor_get(x_38, 2); +x_93 = lean_ctor_get(x_38, 3); +x_94 = lean_ctor_get(x_38, 0); +lean_dec(x_94); +x_95 = !lean_is_exclusive(x_40); +if (x_95 == 0) +{ +uint8_t x_96; uint8_t x_97; +x_96 = 1; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_96); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_96); +x_97 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_40); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_97); +return x_1; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_98 = lean_ctor_get(x_40, 0); +x_99 = lean_ctor_get(x_40, 1); +x_100 = lean_ctor_get(x_40, 2); +x_101 = lean_ctor_get(x_40, 3); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_40); +x_102 = 1; +x_103 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_103, 0, x_98); +lean_ctor_set(x_103, 1, x_99); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_101); +lean_ctor_set_uint8(x_103, sizeof(void*)*4, x_102); +lean_ctor_set(x_38, 3, x_36); +lean_ctor_set(x_38, 2, x_35); +lean_ctor_set(x_38, 1, x_34); +lean_ctor_set(x_38, 0, x_93); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_102); +x_104 = 0; +lean_ctor_set(x_1, 3, x_38); +lean_ctor_set(x_1, 2, x_92); +lean_ctor_set(x_1, 1, x_91); +lean_ctor_set(x_1, 0, x_103); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_104); +return x_1; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_105 = lean_ctor_get(x_38, 1); +x_106 = lean_ctor_get(x_38, 2); +x_107 = lean_ctor_get(x_38, 3); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_38); +x_108 = lean_ctor_get(x_40, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_40, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_40, 2); +lean_inc(x_110); +x_111 = lean_ctor_get(x_40, 3); +lean_inc(x_111); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_112 = x_40; +} else { + lean_dec_ref(x_40); + x_112 = lean_box(0); +} +x_113 = 1; +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(1, 4, 1); +} else { + x_114 = x_112; +} +lean_ctor_set(x_114, 0, x_108); +lean_ctor_set(x_114, 1, x_109); +lean_ctor_set(x_114, 2, x_110); +lean_ctor_set(x_114, 3, x_111); +lean_ctor_set_uint8(x_114, sizeof(void*)*4, x_113); +x_115 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_115, 0, x_107); +lean_ctor_set(x_115, 1, x_34); +lean_ctor_set(x_115, 2, x_35); +lean_ctor_set(x_115, 3, x_36); +lean_ctor_set_uint8(x_115, sizeof(void*)*4, x_113); +x_116 = 0; +lean_ctor_set(x_1, 3, x_115); +lean_ctor_set(x_1, 2, x_106); +lean_ctor_set(x_1, 1, x_105); +lean_ctor_set(x_1, 0, x_114); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_116); +return x_1; +} +} +else +{ +lean_object* x_117; +x_117 = lean_ctor_get(x_38, 3); +lean_inc(x_117); +if (lean_obj_tag(x_117) == 0) +{ +uint8_t x_118; +lean_free_object(x_1); +x_118 = !lean_is_exclusive(x_40); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_119 = lean_ctor_get(x_40, 3); +lean_dec(x_119); +x_120 = lean_ctor_get(x_40, 2); +lean_dec(x_120); +x_121 = lean_ctor_get(x_40, 1); +lean_dec(x_121); +x_122 = lean_ctor_get(x_40, 0); +lean_dec(x_122); +x_123 = 1; +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_123); +return x_40; +} +else +{ +uint8_t x_124; lean_object* x_125; +lean_dec(x_40); +x_124 = 1; +x_125 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_125, 0, x_38); +lean_ctor_set(x_125, 1, x_34); +lean_ctor_set(x_125, 2, x_35); +lean_ctor_set(x_125, 3, x_36); +lean_ctor_set_uint8(x_125, sizeof(void*)*4, x_124); +return x_125; +} +} +else +{ +uint8_t x_126; +x_126 = lean_ctor_get_uint8(x_117, sizeof(void*)*4); +if (x_126 == 0) +{ +uint8_t x_127; +lean_free_object(x_1); +x_127 = !lean_is_exclusive(x_38); +if (x_127 == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_128 = lean_ctor_get(x_38, 1); +x_129 = lean_ctor_get(x_38, 2); +x_130 = lean_ctor_get(x_38, 3); +lean_dec(x_130); +x_131 = lean_ctor_get(x_38, 0); +lean_dec(x_131); +x_132 = !lean_is_exclusive(x_117); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; +x_133 = lean_ctor_get(x_117, 0); +x_134 = lean_ctor_get(x_117, 1); +x_135 = lean_ctor_get(x_117, 2); +x_136 = lean_ctor_get(x_117, 3); +x_137 = 1; +lean_inc(x_40); +lean_ctor_set(x_117, 3, x_133); +lean_ctor_set(x_117, 2, x_129); +lean_ctor_set(x_117, 1, x_128); +lean_ctor_set(x_117, 0, x_40); +x_138 = !lean_is_exclusive(x_40); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_139 = lean_ctor_get(x_40, 3); +lean_dec(x_139); +x_140 = lean_ctor_get(x_40, 2); +lean_dec(x_140); +x_141 = lean_ctor_get(x_40, 1); +lean_dec(x_141); +x_142 = lean_ctor_get(x_40, 0); +lean_dec(x_142); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 0, x_136); +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_137); +x_143 = 0; +lean_ctor_set(x_38, 3, x_40); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_143); +return x_38; +} +else +{ +lean_object* x_144; uint8_t x_145; +lean_dec(x_40); +lean_ctor_set_uint8(x_117, sizeof(void*)*4, x_137); +x_144 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_144, 0, x_136); +lean_ctor_set(x_144, 1, x_34); +lean_ctor_set(x_144, 2, x_35); +lean_ctor_set(x_144, 3, x_36); +lean_ctor_set_uint8(x_144, sizeof(void*)*4, x_137); +x_145 = 0; +lean_ctor_set(x_38, 3, x_144); +lean_ctor_set(x_38, 2, x_135); +lean_ctor_set(x_38, 1, x_134); +lean_ctor_set(x_38, 0, x_117); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_145); +return x_38; +} +} +else +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; uint8_t x_154; +x_146 = lean_ctor_get(x_117, 0); +x_147 = lean_ctor_get(x_117, 1); +x_148 = lean_ctor_get(x_117, 2); +x_149 = lean_ctor_get(x_117, 3); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_117); +x_150 = 1; +lean_inc(x_40); +x_151 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_151, 0, x_40); +lean_ctor_set(x_151, 1, x_128); +lean_ctor_set(x_151, 2, x_129); +lean_ctor_set(x_151, 3, x_146); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_152 = x_40; +} else { + lean_dec_ref(x_40); + x_152 = lean_box(0); +} +lean_ctor_set_uint8(x_151, sizeof(void*)*4, x_150); +if (lean_is_scalar(x_152)) { + x_153 = lean_alloc_ctor(1, 4, 1); +} else { + x_153 = x_152; +} +lean_ctor_set(x_153, 0, x_149); +lean_ctor_set(x_153, 1, x_34); +lean_ctor_set(x_153, 2, x_35); +lean_ctor_set(x_153, 3, x_36); +lean_ctor_set_uint8(x_153, sizeof(void*)*4, x_150); +x_154 = 0; +lean_ctor_set(x_38, 3, x_153); +lean_ctor_set(x_38, 2, x_148); +lean_ctor_set(x_38, 1, x_147); +lean_ctor_set(x_38, 0, x_151); +lean_ctor_set_uint8(x_38, sizeof(void*)*4, x_154); +return x_38; +} +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; +x_155 = lean_ctor_get(x_38, 1); +x_156 = lean_ctor_get(x_38, 2); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_38); +x_157 = lean_ctor_get(x_117, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_117, 1); +lean_inc(x_158); +x_159 = lean_ctor_get(x_117, 2); +lean_inc(x_159); +x_160 = lean_ctor_get(x_117, 3); +lean_inc(x_160); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + lean_ctor_release(x_117, 2); + lean_ctor_release(x_117, 3); + x_161 = x_117; +} else { + lean_dec_ref(x_117); + x_161 = lean_box(0); +} +x_162 = 1; +lean_inc(x_40); +if (lean_is_scalar(x_161)) { + x_163 = lean_alloc_ctor(1, 4, 1); +} else { + x_163 = x_161; +} +lean_ctor_set(x_163, 0, x_40); +lean_ctor_set(x_163, 1, x_155); +lean_ctor_set(x_163, 2, x_156); +lean_ctor_set(x_163, 3, x_157); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_164 = x_40; +} else { + lean_dec_ref(x_40); + x_164 = lean_box(0); +} +lean_ctor_set_uint8(x_163, sizeof(void*)*4, x_162); +if (lean_is_scalar(x_164)) { + x_165 = lean_alloc_ctor(1, 4, 1); +} else { + x_165 = x_164; +} +lean_ctor_set(x_165, 0, x_160); +lean_ctor_set(x_165, 1, x_34); +lean_ctor_set(x_165, 2, x_35); +lean_ctor_set(x_165, 3, x_36); +lean_ctor_set_uint8(x_165, sizeof(void*)*4, x_162); +x_166 = 0; +x_167 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_167, 0, x_163); +lean_ctor_set(x_167, 1, x_158); +lean_ctor_set(x_167, 2, x_159); +lean_ctor_set(x_167, 3, x_165); +lean_ctor_set_uint8(x_167, sizeof(void*)*4, x_166); +return x_167; +} +} +else +{ +uint8_t x_168; +x_168 = !lean_is_exclusive(x_38); +if (x_168 == 0) +{ +lean_object* x_169; lean_object* x_170; uint8_t x_171; +x_169 = lean_ctor_get(x_38, 3); +lean_dec(x_169); +x_170 = lean_ctor_get(x_38, 0); +lean_dec(x_170); +x_171 = !lean_is_exclusive(x_40); +if (x_171 == 0) +{ +uint8_t x_172; +lean_ctor_set_uint8(x_40, sizeof(void*)*4, x_126); +x_172 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_172); +return x_1; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; +x_173 = lean_ctor_get(x_40, 0); +x_174 = lean_ctor_get(x_40, 1); +x_175 = lean_ctor_get(x_40, 2); +x_176 = lean_ctor_get(x_40, 3); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_40); +x_177 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_174); +lean_ctor_set(x_177, 2, x_175); +lean_ctor_set(x_177, 3, x_176); +lean_ctor_set_uint8(x_177, sizeof(void*)*4, x_126); +lean_ctor_set(x_38, 0, x_177); +x_178 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_178); +return x_1; +} +} +else +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; +x_179 = lean_ctor_get(x_38, 1); +x_180 = lean_ctor_get(x_38, 2); +lean_inc(x_180); +lean_inc(x_179); +lean_dec(x_38); +x_181 = lean_ctor_get(x_40, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_40, 1); +lean_inc(x_182); +x_183 = lean_ctor_get(x_40, 2); +lean_inc(x_183); +x_184 = lean_ctor_get(x_40, 3); +lean_inc(x_184); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + lean_ctor_release(x_40, 2); + lean_ctor_release(x_40, 3); + x_185 = x_40; +} else { + lean_dec_ref(x_40); + x_185 = lean_box(0); +} +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 4, 1); +} else { + x_186 = x_185; +} +lean_ctor_set(x_186, 0, x_181); +lean_ctor_set(x_186, 1, x_182); +lean_ctor_set(x_186, 2, x_183); +lean_ctor_set(x_186, 3, x_184); +lean_ctor_set_uint8(x_186, sizeof(void*)*4, x_126); +x_187 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_179); +lean_ctor_set(x_187, 2, x_180); +lean_ctor_set(x_187, 3, x_117); +lean_ctor_set_uint8(x_187, sizeof(void*)*4, x_39); +x_188 = 1; +lean_ctor_set(x_1, 0, x_187); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_188); +return x_1; +} +} +} +} +} +} +else +{ +uint8_t x_189; +x_189 = 1; +lean_ctor_set(x_1, 0, x_38); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_189); +return x_1; +} +} +case 1: +{ +uint8_t x_190; +lean_dec(x_35); +lean_dec(x_34); +x_190 = 1; +lean_ctor_set(x_1, 2, x_3); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_190); +return x_1; +} +default: +{ +lean_object* x_191; uint8_t x_192; +x_191 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_36, x_2, x_3); +x_192 = lean_ctor_get_uint8(x_191, sizeof(void*)*4); +if (x_192 == 0) +{ +lean_object* x_193; +x_193 = lean_ctor_get(x_191, 0); +lean_inc(x_193); +if (lean_obj_tag(x_193) == 0) +{ +lean_object* x_194; +x_194 = lean_ctor_get(x_191, 3); +lean_inc(x_194); +if (lean_obj_tag(x_194) == 0) +{ +uint8_t x_195; +x_195 = !lean_is_exclusive(x_191); +if (x_195 == 0) +{ +lean_object* x_196; lean_object* x_197; uint8_t x_198; +x_196 = lean_ctor_get(x_191, 3); +lean_dec(x_196); +x_197 = lean_ctor_get(x_191, 0); +lean_dec(x_197); +lean_ctor_set(x_191, 0, x_194); +x_198 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_198); +return x_1; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; uint8_t x_202; +x_199 = lean_ctor_get(x_191, 1); +x_200 = lean_ctor_get(x_191, 2); +lean_inc(x_200); +lean_inc(x_199); +lean_dec(x_191); +x_201 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_201, 0, x_194); +lean_ctor_set(x_201, 1, x_199); +lean_ctor_set(x_201, 2, x_200); +lean_ctor_set(x_201, 3, x_194); +lean_ctor_set_uint8(x_201, sizeof(void*)*4, x_192); +x_202 = 1; +lean_ctor_set(x_1, 3, x_201); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_202); +return x_1; +} +} +else +{ +uint8_t x_203; +x_203 = lean_ctor_get_uint8(x_194, sizeof(void*)*4); +if (x_203 == 0) +{ +uint8_t x_204; +x_204 = !lean_is_exclusive(x_191); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +x_205 = lean_ctor_get(x_191, 1); +x_206 = lean_ctor_get(x_191, 2); +x_207 = lean_ctor_get(x_191, 3); +lean_dec(x_207); +x_208 = lean_ctor_get(x_191, 0); +lean_dec(x_208); +x_209 = !lean_is_exclusive(x_194); +if (x_209 == 0) +{ +lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; uint8_t x_215; +x_210 = lean_ctor_get(x_194, 0); +x_211 = lean_ctor_get(x_194, 1); +x_212 = lean_ctor_get(x_194, 2); +x_213 = lean_ctor_get(x_194, 3); +x_214 = 1; +lean_ctor_set(x_194, 3, x_193); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_214); +lean_ctor_set(x_191, 3, x_213); +lean_ctor_set(x_191, 2, x_212); +lean_ctor_set(x_191, 1, x_211); +lean_ctor_set(x_191, 0, x_210); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_214); +x_215 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_194); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_215); +return x_1; +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; uint8_t x_222; +x_216 = lean_ctor_get(x_194, 0); +x_217 = lean_ctor_get(x_194, 1); +x_218 = lean_ctor_get(x_194, 2); +x_219 = lean_ctor_get(x_194, 3); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_dec(x_194); +x_220 = 1; +x_221 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_221, 0, x_33); +lean_ctor_set(x_221, 1, x_34); +lean_ctor_set(x_221, 2, x_35); +lean_ctor_set(x_221, 3, x_193); +lean_ctor_set_uint8(x_221, sizeof(void*)*4, x_220); +lean_ctor_set(x_191, 3, x_219); +lean_ctor_set(x_191, 2, x_218); +lean_ctor_set(x_191, 1, x_217); +lean_ctor_set(x_191, 0, x_216); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_220); +x_222 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_206); +lean_ctor_set(x_1, 1, x_205); +lean_ctor_set(x_1, 0, x_221); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_222); +return x_1; +} +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; uint8_t x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_223 = lean_ctor_get(x_191, 1); +x_224 = lean_ctor_get(x_191, 2); +lean_inc(x_224); +lean_inc(x_223); +lean_dec(x_191); +x_225 = lean_ctor_get(x_194, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_194, 1); +lean_inc(x_226); +x_227 = lean_ctor_get(x_194, 2); +lean_inc(x_227); +x_228 = lean_ctor_get(x_194, 3); +lean_inc(x_228); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + lean_ctor_release(x_194, 2); + lean_ctor_release(x_194, 3); + x_229 = x_194; +} else { + lean_dec_ref(x_194); + x_229 = lean_box(0); +} +x_230 = 1; +if (lean_is_scalar(x_229)) { + x_231 = lean_alloc_ctor(1, 4, 1); +} else { + x_231 = x_229; +} +lean_ctor_set(x_231, 0, x_33); +lean_ctor_set(x_231, 1, x_34); +lean_ctor_set(x_231, 2, x_35); +lean_ctor_set(x_231, 3, x_193); +lean_ctor_set_uint8(x_231, sizeof(void*)*4, x_230); +x_232 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_232, 0, x_225); +lean_ctor_set(x_232, 1, x_226); +lean_ctor_set(x_232, 2, x_227); +lean_ctor_set(x_232, 3, x_228); +lean_ctor_set_uint8(x_232, sizeof(void*)*4, x_230); +x_233 = 0; +lean_ctor_set(x_1, 3, x_232); +lean_ctor_set(x_1, 2, x_224); +lean_ctor_set(x_1, 1, x_223); +lean_ctor_set(x_1, 0, x_231); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_233); +return x_1; +} +} +else +{ +uint8_t x_234; +lean_free_object(x_1); +x_234 = !lean_is_exclusive(x_194); +if (x_234 == 0) +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_235 = lean_ctor_get(x_194, 3); +lean_dec(x_235); +x_236 = lean_ctor_get(x_194, 2); +lean_dec(x_236); +x_237 = lean_ctor_get(x_194, 1); +lean_dec(x_237); +x_238 = lean_ctor_get(x_194, 0); +lean_dec(x_238); +x_239 = 1; +lean_ctor_set(x_194, 3, x_191); +lean_ctor_set(x_194, 2, x_35); +lean_ctor_set(x_194, 1, x_34); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set_uint8(x_194, sizeof(void*)*4, x_239); +return x_194; +} +else +{ +uint8_t x_240; lean_object* x_241; +lean_dec(x_194); +x_240 = 1; +x_241 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_241, 0, x_33); +lean_ctor_set(x_241, 1, x_34); +lean_ctor_set(x_241, 2, x_35); +lean_ctor_set(x_241, 3, x_191); +lean_ctor_set_uint8(x_241, sizeof(void*)*4, x_240); +return x_241; +} +} +} +} +else +{ +uint8_t x_242; +x_242 = lean_ctor_get_uint8(x_193, sizeof(void*)*4); +if (x_242 == 0) +{ +uint8_t x_243; +x_243 = !lean_is_exclusive(x_191); +if (x_243 == 0) +{ +lean_object* x_244; uint8_t x_245; +x_244 = lean_ctor_get(x_191, 0); +lean_dec(x_244); +x_245 = !lean_is_exclusive(x_193); +if (x_245 == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; uint8_t x_251; +x_246 = lean_ctor_get(x_193, 0); +x_247 = lean_ctor_get(x_193, 1); +x_248 = lean_ctor_get(x_193, 2); +x_249 = lean_ctor_get(x_193, 3); +x_250 = 1; +lean_ctor_set(x_193, 3, x_246); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_250); +lean_ctor_set(x_191, 0, x_249); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_250); +x_251 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_248); +lean_ctor_set(x_1, 1, x_247); +lean_ctor_set(x_1, 0, x_193); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_251); +return x_1; +} +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; uint8_t x_258; +x_252 = lean_ctor_get(x_193, 0); +x_253 = lean_ctor_get(x_193, 1); +x_254 = lean_ctor_get(x_193, 2); +x_255 = lean_ctor_get(x_193, 3); +lean_inc(x_255); +lean_inc(x_254); +lean_inc(x_253); +lean_inc(x_252); +lean_dec(x_193); +x_256 = 1; +x_257 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_257, 0, x_33); +lean_ctor_set(x_257, 1, x_34); +lean_ctor_set(x_257, 2, x_35); +lean_ctor_set(x_257, 3, x_252); +lean_ctor_set_uint8(x_257, sizeof(void*)*4, x_256); +lean_ctor_set(x_191, 0, x_255); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_256); +x_258 = 0; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set(x_1, 2, x_254); +lean_ctor_set(x_1, 1, x_253); +lean_ctor_set(x_1, 0, x_257); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_258); +return x_1; +} +} +else +{ +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; +x_259 = lean_ctor_get(x_191, 1); +x_260 = lean_ctor_get(x_191, 2); +x_261 = lean_ctor_get(x_191, 3); +lean_inc(x_261); +lean_inc(x_260); +lean_inc(x_259); +lean_dec(x_191); +x_262 = lean_ctor_get(x_193, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_193, 1); +lean_inc(x_263); +x_264 = lean_ctor_get(x_193, 2); +lean_inc(x_264); +x_265 = lean_ctor_get(x_193, 3); +lean_inc(x_265); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_266 = x_193; +} else { + lean_dec_ref(x_193); + x_266 = lean_box(0); +} +x_267 = 1; +if (lean_is_scalar(x_266)) { + x_268 = lean_alloc_ctor(1, 4, 1); +} else { + x_268 = x_266; +} +lean_ctor_set(x_268, 0, x_33); +lean_ctor_set(x_268, 1, x_34); +lean_ctor_set(x_268, 2, x_35); +lean_ctor_set(x_268, 3, x_262); +lean_ctor_set_uint8(x_268, sizeof(void*)*4, x_267); +x_269 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_269, 0, x_265); +lean_ctor_set(x_269, 1, x_259); +lean_ctor_set(x_269, 2, x_260); +lean_ctor_set(x_269, 3, x_261); +lean_ctor_set_uint8(x_269, sizeof(void*)*4, x_267); +x_270 = 0; +lean_ctor_set(x_1, 3, x_269); +lean_ctor_set(x_1, 2, x_264); +lean_ctor_set(x_1, 1, x_263); +lean_ctor_set(x_1, 0, x_268); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_270); +return x_1; +} +} +else +{ +lean_object* x_271; +x_271 = lean_ctor_get(x_191, 3); +lean_inc(x_271); +if (lean_obj_tag(x_271) == 0) +{ +uint8_t x_272; +lean_free_object(x_1); +x_272 = !lean_is_exclusive(x_193); +if (x_272 == 0) +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; +x_273 = lean_ctor_get(x_193, 3); +lean_dec(x_273); +x_274 = lean_ctor_get(x_193, 2); +lean_dec(x_274); +x_275 = lean_ctor_get(x_193, 1); +lean_dec(x_275); +x_276 = lean_ctor_get(x_193, 0); +lean_dec(x_276); +x_277 = 1; +lean_ctor_set(x_193, 3, x_191); +lean_ctor_set(x_193, 2, x_35); +lean_ctor_set(x_193, 1, x_34); +lean_ctor_set(x_193, 0, x_33); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_277); +return x_193; +} +else +{ +uint8_t x_278; lean_object* x_279; +lean_dec(x_193); +x_278 = 1; +x_279 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_279, 0, x_33); +lean_ctor_set(x_279, 1, x_34); +lean_ctor_set(x_279, 2, x_35); +lean_ctor_set(x_279, 3, x_191); +lean_ctor_set_uint8(x_279, sizeof(void*)*4, x_278); +return x_279; +} +} +else +{ +uint8_t x_280; +x_280 = lean_ctor_get_uint8(x_271, sizeof(void*)*4); +if (x_280 == 0) +{ +uint8_t x_281; +lean_free_object(x_1); +x_281 = !lean_is_exclusive(x_191); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; +x_282 = lean_ctor_get(x_191, 3); +lean_dec(x_282); +x_283 = lean_ctor_get(x_191, 0); +lean_dec(x_283); +x_284 = !lean_is_exclusive(x_271); +if (x_284 == 0) +{ +lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; uint8_t x_290; +x_285 = lean_ctor_get(x_271, 0); +x_286 = lean_ctor_get(x_271, 1); +x_287 = lean_ctor_get(x_271, 2); +x_288 = lean_ctor_get(x_271, 3); +x_289 = 1; +lean_inc(x_193); +lean_ctor_set(x_271, 3, x_193); +lean_ctor_set(x_271, 2, x_35); +lean_ctor_set(x_271, 1, x_34); +lean_ctor_set(x_271, 0, x_33); +x_290 = !lean_is_exclusive(x_193); +if (x_290 == 0) +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_291 = lean_ctor_get(x_193, 3); +lean_dec(x_291); +x_292 = lean_ctor_get(x_193, 2); +lean_dec(x_292); +x_293 = lean_ctor_get(x_193, 1); +lean_dec(x_293); +x_294 = lean_ctor_get(x_193, 0); +lean_dec(x_294); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +lean_ctor_set(x_193, 3, x_288); +lean_ctor_set(x_193, 2, x_287); +lean_ctor_set(x_193, 1, x_286); +lean_ctor_set(x_193, 0, x_285); +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_289); +x_295 = 0; +lean_ctor_set(x_191, 3, x_193); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_295); +return x_191; +} +else +{ +lean_object* x_296; uint8_t x_297; +lean_dec(x_193); +lean_ctor_set_uint8(x_271, sizeof(void*)*4, x_289); +x_296 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_296, 0, x_285); +lean_ctor_set(x_296, 1, x_286); +lean_ctor_set(x_296, 2, x_287); +lean_ctor_set(x_296, 3, x_288); +lean_ctor_set_uint8(x_296, sizeof(void*)*4, x_289); +x_297 = 0; +lean_ctor_set(x_191, 3, x_296); +lean_ctor_set(x_191, 0, x_271); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_297); +return x_191; +} +} +else +{ +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; uint8_t x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; uint8_t x_306; +x_298 = lean_ctor_get(x_271, 0); +x_299 = lean_ctor_get(x_271, 1); +x_300 = lean_ctor_get(x_271, 2); +x_301 = lean_ctor_get(x_271, 3); +lean_inc(x_301); +lean_inc(x_300); +lean_inc(x_299); +lean_inc(x_298); +lean_dec(x_271); +x_302 = 1; +lean_inc(x_193); +x_303 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_303, 0, x_33); +lean_ctor_set(x_303, 1, x_34); +lean_ctor_set(x_303, 2, x_35); +lean_ctor_set(x_303, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_304 = x_193; +} else { + lean_dec_ref(x_193); + x_304 = lean_box(0); +} +lean_ctor_set_uint8(x_303, sizeof(void*)*4, x_302); +if (lean_is_scalar(x_304)) { + x_305 = lean_alloc_ctor(1, 4, 1); +} else { + x_305 = x_304; +} +lean_ctor_set(x_305, 0, x_298); +lean_ctor_set(x_305, 1, x_299); +lean_ctor_set(x_305, 2, x_300); +lean_ctor_set(x_305, 3, x_301); +lean_ctor_set_uint8(x_305, sizeof(void*)*4, x_302); +x_306 = 0; +lean_ctor_set(x_191, 3, x_305); +lean_ctor_set(x_191, 0, x_303); +lean_ctor_set_uint8(x_191, sizeof(void*)*4, x_306); +return x_191; +} +} +else +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; uint8_t x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; uint8_t x_318; lean_object* x_319; +x_307 = lean_ctor_get(x_191, 1); +x_308 = lean_ctor_get(x_191, 2); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_191); +x_309 = lean_ctor_get(x_271, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_271, 1); +lean_inc(x_310); +x_311 = lean_ctor_get(x_271, 2); +lean_inc(x_311); +x_312 = lean_ctor_get(x_271, 3); +lean_inc(x_312); +if (lean_is_exclusive(x_271)) { + lean_ctor_release(x_271, 0); + lean_ctor_release(x_271, 1); + lean_ctor_release(x_271, 2); + lean_ctor_release(x_271, 3); + x_313 = x_271; +} else { + lean_dec_ref(x_271); + x_313 = lean_box(0); +} +x_314 = 1; +lean_inc(x_193); +if (lean_is_scalar(x_313)) { + x_315 = lean_alloc_ctor(1, 4, 1); +} else { + x_315 = x_313; +} +lean_ctor_set(x_315, 0, x_33); +lean_ctor_set(x_315, 1, x_34); +lean_ctor_set(x_315, 2, x_35); +lean_ctor_set(x_315, 3, x_193); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_316 = x_193; +} else { + lean_dec_ref(x_193); + x_316 = lean_box(0); +} +lean_ctor_set_uint8(x_315, sizeof(void*)*4, x_314); +if (lean_is_scalar(x_316)) { + x_317 = lean_alloc_ctor(1, 4, 1); +} else { + x_317 = x_316; +} +lean_ctor_set(x_317, 0, x_309); +lean_ctor_set(x_317, 1, x_310); +lean_ctor_set(x_317, 2, x_311); +lean_ctor_set(x_317, 3, x_312); +lean_ctor_set_uint8(x_317, sizeof(void*)*4, x_314); +x_318 = 0; +x_319 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_319, 0, x_315); +lean_ctor_set(x_319, 1, x_307); +lean_ctor_set(x_319, 2, x_308); +lean_ctor_set(x_319, 3, x_317); +lean_ctor_set_uint8(x_319, sizeof(void*)*4, x_318); +return x_319; +} +} +else +{ +uint8_t x_320; +x_320 = !lean_is_exclusive(x_191); +if (x_320 == 0) +{ +lean_object* x_321; lean_object* x_322; uint8_t x_323; +x_321 = lean_ctor_get(x_191, 3); +lean_dec(x_321); +x_322 = lean_ctor_get(x_191, 0); +lean_dec(x_322); +x_323 = !lean_is_exclusive(x_193); +if (x_323 == 0) +{ +uint8_t x_324; +lean_ctor_set_uint8(x_193, sizeof(void*)*4, x_280); +x_324 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_324); +return x_1; +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; +x_325 = lean_ctor_get(x_193, 0); +x_326 = lean_ctor_get(x_193, 1); +x_327 = lean_ctor_get(x_193, 2); +x_328 = lean_ctor_get(x_193, 3); +lean_inc(x_328); +lean_inc(x_327); +lean_inc(x_326); +lean_inc(x_325); +lean_dec(x_193); +x_329 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_329, 0, x_325); +lean_ctor_set(x_329, 1, x_326); +lean_ctor_set(x_329, 2, x_327); +lean_ctor_set(x_329, 3, x_328); +lean_ctor_set_uint8(x_329, sizeof(void*)*4, x_280); +lean_ctor_set(x_191, 0, x_329); +x_330 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_330); +return x_1; +} +} +else +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; +x_331 = lean_ctor_get(x_191, 1); +x_332 = lean_ctor_get(x_191, 2); +lean_inc(x_332); +lean_inc(x_331); +lean_dec(x_191); +x_333 = lean_ctor_get(x_193, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_193, 1); +lean_inc(x_334); +x_335 = lean_ctor_get(x_193, 2); +lean_inc(x_335); +x_336 = lean_ctor_get(x_193, 3); +lean_inc(x_336); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + lean_ctor_release(x_193, 2); + lean_ctor_release(x_193, 3); + x_337 = x_193; +} else { + lean_dec_ref(x_193); + x_337 = lean_box(0); +} +if (lean_is_scalar(x_337)) { + x_338 = lean_alloc_ctor(1, 4, 1); +} else { + x_338 = x_337; +} +lean_ctor_set(x_338, 0, x_333); +lean_ctor_set(x_338, 1, x_334); +lean_ctor_set(x_338, 2, x_335); +lean_ctor_set(x_338, 3, x_336); +lean_ctor_set_uint8(x_338, sizeof(void*)*4, x_280); +x_339 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_339, 0, x_338); +lean_ctor_set(x_339, 1, x_331); +lean_ctor_set(x_339, 2, x_332); +lean_ctor_set(x_339, 3, x_271); +lean_ctor_set_uint8(x_339, sizeof(void*)*4, x_192); +x_340 = 1; +lean_ctor_set(x_1, 3, x_339); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_340); +return x_1; +} +} +} +} +} +} +else +{ +uint8_t x_341; +x_341 = 1; +lean_ctor_set(x_1, 3, x_191); +lean_ctor_set_uint8(x_1, sizeof(void*)*4, x_341); +return x_1; +} +} +} +} +else +{ +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; +x_342 = lean_ctor_get(x_1, 0); +x_343 = lean_ctor_get(x_1, 1); +x_344 = lean_ctor_get(x_1, 2); +x_345 = lean_ctor_get(x_1, 3); +lean_inc(x_345); +lean_inc(x_344); +lean_inc(x_343); +lean_inc(x_342); +lean_dec(x_1); +lean_inc(x_343); +lean_inc(x_2); +x_346 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_2, x_343); +switch (x_346) { +case 0: +{ +lean_object* x_347; uint8_t x_348; +x_347 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_342, x_2, x_3); +x_348 = lean_ctor_get_uint8(x_347, sizeof(void*)*4); +if (x_348 == 0) +{ +lean_object* x_349; +x_349 = lean_ctor_get(x_347, 0); +lean_inc(x_349); +if (lean_obj_tag(x_349) == 0) +{ +lean_object* x_350; +x_350 = lean_ctor_get(x_347, 3); +lean_inc(x_350); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; +x_351 = lean_ctor_get(x_347, 1); +lean_inc(x_351); +x_352 = lean_ctor_get(x_347, 2); +lean_inc(x_352); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_353 = x_347; +} else { + lean_dec_ref(x_347); + x_353 = lean_box(0); +} +if (lean_is_scalar(x_353)) { + x_354 = lean_alloc_ctor(1, 4, 1); +} else { + x_354 = x_353; +} +lean_ctor_set(x_354, 0, x_350); +lean_ctor_set(x_354, 1, x_351); +lean_ctor_set(x_354, 2, x_352); +lean_ctor_set(x_354, 3, x_350); +lean_ctor_set_uint8(x_354, sizeof(void*)*4, x_348); +x_355 = 1; +x_356 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_356, 0, x_354); +lean_ctor_set(x_356, 1, x_343); +lean_ctor_set(x_356, 2, x_344); +lean_ctor_set(x_356, 3, x_345); +lean_ctor_set_uint8(x_356, sizeof(void*)*4, x_355); +return x_356; +} +else +{ +uint8_t x_357; +x_357 = lean_ctor_get_uint8(x_350, sizeof(void*)*4); +if (x_357 == 0) +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; uint8_t x_369; lean_object* x_370; +x_358 = lean_ctor_get(x_347, 1); +lean_inc(x_358); +x_359 = lean_ctor_get(x_347, 2); +lean_inc(x_359); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_360 = x_347; +} else { + lean_dec_ref(x_347); + x_360 = lean_box(0); +} +x_361 = lean_ctor_get(x_350, 0); +lean_inc(x_361); +x_362 = lean_ctor_get(x_350, 1); +lean_inc(x_362); +x_363 = lean_ctor_get(x_350, 2); +lean_inc(x_363); +x_364 = lean_ctor_get(x_350, 3); +lean_inc(x_364); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_365 = x_350; +} else { + lean_dec_ref(x_350); + x_365 = lean_box(0); +} +x_366 = 1; +if (lean_is_scalar(x_365)) { + x_367 = lean_alloc_ctor(1, 4, 1); +} else { + x_367 = x_365; +} +lean_ctor_set(x_367, 0, x_349); +lean_ctor_set(x_367, 1, x_358); +lean_ctor_set(x_367, 2, x_359); +lean_ctor_set(x_367, 3, x_361); +lean_ctor_set_uint8(x_367, sizeof(void*)*4, x_366); +if (lean_is_scalar(x_360)) { + x_368 = lean_alloc_ctor(1, 4, 1); +} else { + x_368 = x_360; +} +lean_ctor_set(x_368, 0, x_364); +lean_ctor_set(x_368, 1, x_343); +lean_ctor_set(x_368, 2, x_344); +lean_ctor_set(x_368, 3, x_345); +lean_ctor_set_uint8(x_368, sizeof(void*)*4, x_366); +x_369 = 0; +x_370 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_370, 0, x_367); +lean_ctor_set(x_370, 1, x_362); +lean_ctor_set(x_370, 2, x_363); +lean_ctor_set(x_370, 3, x_368); +lean_ctor_set_uint8(x_370, sizeof(void*)*4, x_369); +return x_370; +} +else +{ +lean_object* x_371; uint8_t x_372; lean_object* x_373; +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + lean_ctor_release(x_350, 2); + lean_ctor_release(x_350, 3); + x_371 = x_350; +} else { + lean_dec_ref(x_350); + x_371 = lean_box(0); +} +x_372 = 1; +if (lean_is_scalar(x_371)) { + x_373 = lean_alloc_ctor(1, 4, 1); +} else { + x_373 = x_371; +} +lean_ctor_set(x_373, 0, x_347); +lean_ctor_set(x_373, 1, x_343); +lean_ctor_set(x_373, 2, x_344); +lean_ctor_set(x_373, 3, x_345); +lean_ctor_set_uint8(x_373, sizeof(void*)*4, x_372); +return x_373; +} +} +} +else +{ +uint8_t x_374; +x_374 = lean_ctor_get_uint8(x_349, sizeof(void*)*4); +if (x_374 == 0) +{ +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; uint8_t x_384; lean_object* x_385; lean_object* x_386; uint8_t x_387; lean_object* x_388; +x_375 = lean_ctor_get(x_347, 1); +lean_inc(x_375); +x_376 = lean_ctor_get(x_347, 2); +lean_inc(x_376); +x_377 = lean_ctor_get(x_347, 3); +lean_inc(x_377); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_378 = x_347; +} else { + lean_dec_ref(x_347); + x_378 = lean_box(0); +} +x_379 = lean_ctor_get(x_349, 0); +lean_inc(x_379); +x_380 = lean_ctor_get(x_349, 1); +lean_inc(x_380); +x_381 = lean_ctor_get(x_349, 2); +lean_inc(x_381); +x_382 = lean_ctor_get(x_349, 3); +lean_inc(x_382); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_383 = x_349; +} else { + lean_dec_ref(x_349); + x_383 = lean_box(0); +} +x_384 = 1; +if (lean_is_scalar(x_383)) { + x_385 = lean_alloc_ctor(1, 4, 1); +} else { + x_385 = x_383; +} +lean_ctor_set(x_385, 0, x_379); +lean_ctor_set(x_385, 1, x_380); +lean_ctor_set(x_385, 2, x_381); +lean_ctor_set(x_385, 3, x_382); +lean_ctor_set_uint8(x_385, sizeof(void*)*4, x_384); +if (lean_is_scalar(x_378)) { + x_386 = lean_alloc_ctor(1, 4, 1); +} else { + x_386 = x_378; +} +lean_ctor_set(x_386, 0, x_377); +lean_ctor_set(x_386, 1, x_343); +lean_ctor_set(x_386, 2, x_344); +lean_ctor_set(x_386, 3, x_345); +lean_ctor_set_uint8(x_386, sizeof(void*)*4, x_384); +x_387 = 0; +x_388 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_388, 0, x_385); +lean_ctor_set(x_388, 1, x_375); +lean_ctor_set(x_388, 2, x_376); +lean_ctor_set(x_388, 3, x_386); +lean_ctor_set_uint8(x_388, sizeof(void*)*4, x_387); +return x_388; +} +else +{ +lean_object* x_389; +x_389 = lean_ctor_get(x_347, 3); +lean_inc(x_389); +if (lean_obj_tag(x_389) == 0) +{ +lean_object* x_390; uint8_t x_391; lean_object* x_392; +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_390 = x_349; +} else { + lean_dec_ref(x_349); + x_390 = lean_box(0); +} +x_391 = 1; +if (lean_is_scalar(x_390)) { + x_392 = lean_alloc_ctor(1, 4, 1); +} else { + x_392 = x_390; +} +lean_ctor_set(x_392, 0, x_347); +lean_ctor_set(x_392, 1, x_343); +lean_ctor_set(x_392, 2, x_344); +lean_ctor_set(x_392, 3, x_345); +lean_ctor_set_uint8(x_392, sizeof(void*)*4, x_391); +return x_392; +} +else +{ +uint8_t x_393; +x_393 = lean_ctor_get_uint8(x_389, sizeof(void*)*4); +if (x_393 == 0) +{ +lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; +x_394 = lean_ctor_get(x_347, 1); +lean_inc(x_394); +x_395 = lean_ctor_get(x_347, 2); +lean_inc(x_395); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_396 = x_347; +} else { + lean_dec_ref(x_347); + x_396 = lean_box(0); +} +x_397 = lean_ctor_get(x_389, 0); +lean_inc(x_397); +x_398 = lean_ctor_get(x_389, 1); +lean_inc(x_398); +x_399 = lean_ctor_get(x_389, 2); +lean_inc(x_399); +x_400 = lean_ctor_get(x_389, 3); +lean_inc(x_400); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + lean_ctor_release(x_389, 2); + lean_ctor_release(x_389, 3); + x_401 = x_389; +} else { + lean_dec_ref(x_389); + x_401 = lean_box(0); +} +x_402 = 1; +lean_inc(x_349); +if (lean_is_scalar(x_401)) { + x_403 = lean_alloc_ctor(1, 4, 1); +} else { + x_403 = x_401; +} +lean_ctor_set(x_403, 0, x_349); +lean_ctor_set(x_403, 1, x_394); +lean_ctor_set(x_403, 2, x_395); +lean_ctor_set(x_403, 3, x_397); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_404 = x_349; +} else { + lean_dec_ref(x_349); + x_404 = lean_box(0); +} +lean_ctor_set_uint8(x_403, sizeof(void*)*4, x_402); +if (lean_is_scalar(x_404)) { + x_405 = lean_alloc_ctor(1, 4, 1); +} else { + x_405 = x_404; +} +lean_ctor_set(x_405, 0, x_400); +lean_ctor_set(x_405, 1, x_343); +lean_ctor_set(x_405, 2, x_344); +lean_ctor_set(x_405, 3, x_345); +lean_ctor_set_uint8(x_405, sizeof(void*)*4, x_402); +x_406 = 0; +if (lean_is_scalar(x_396)) { + x_407 = lean_alloc_ctor(1, 4, 1); +} else { + x_407 = x_396; +} +lean_ctor_set(x_407, 0, x_403); +lean_ctor_set(x_407, 1, x_398); +lean_ctor_set(x_407, 2, x_399); +lean_ctor_set(x_407, 3, x_405); +lean_ctor_set_uint8(x_407, sizeof(void*)*4, x_406); +return x_407; +} +else +{ +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; uint8_t x_418; lean_object* x_419; +x_408 = lean_ctor_get(x_347, 1); +lean_inc(x_408); +x_409 = lean_ctor_get(x_347, 2); +lean_inc(x_409); +if (lean_is_exclusive(x_347)) { + lean_ctor_release(x_347, 0); + lean_ctor_release(x_347, 1); + lean_ctor_release(x_347, 2); + lean_ctor_release(x_347, 3); + x_410 = x_347; +} else { + lean_dec_ref(x_347); + x_410 = lean_box(0); +} +x_411 = lean_ctor_get(x_349, 0); +lean_inc(x_411); +x_412 = lean_ctor_get(x_349, 1); +lean_inc(x_412); +x_413 = lean_ctor_get(x_349, 2); +lean_inc(x_413); +x_414 = lean_ctor_get(x_349, 3); +lean_inc(x_414); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + lean_ctor_release(x_349, 2); + lean_ctor_release(x_349, 3); + x_415 = x_349; +} else { + lean_dec_ref(x_349); + x_415 = lean_box(0); +} +if (lean_is_scalar(x_415)) { + x_416 = lean_alloc_ctor(1, 4, 1); +} else { + x_416 = x_415; +} +lean_ctor_set(x_416, 0, x_411); +lean_ctor_set(x_416, 1, x_412); +lean_ctor_set(x_416, 2, x_413); +lean_ctor_set(x_416, 3, x_414); +lean_ctor_set_uint8(x_416, sizeof(void*)*4, x_393); +if (lean_is_scalar(x_410)) { + x_417 = lean_alloc_ctor(1, 4, 1); +} else { + x_417 = x_410; +} +lean_ctor_set(x_417, 0, x_416); +lean_ctor_set(x_417, 1, x_408); +lean_ctor_set(x_417, 2, x_409); +lean_ctor_set(x_417, 3, x_389); +lean_ctor_set_uint8(x_417, sizeof(void*)*4, x_348); +x_418 = 1; +x_419 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_419, 0, x_417); +lean_ctor_set(x_419, 1, x_343); +lean_ctor_set(x_419, 2, x_344); +lean_ctor_set(x_419, 3, x_345); +lean_ctor_set_uint8(x_419, sizeof(void*)*4, x_418); +return x_419; +} +} +} +} +} +else +{ +uint8_t x_420; lean_object* x_421; +x_420 = 1; +x_421 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_421, 0, x_347); +lean_ctor_set(x_421, 1, x_343); +lean_ctor_set(x_421, 2, x_344); +lean_ctor_set(x_421, 3, x_345); +lean_ctor_set_uint8(x_421, sizeof(void*)*4, x_420); +return x_421; +} +} +case 1: +{ +uint8_t x_422; lean_object* x_423; +lean_dec(x_344); +lean_dec(x_343); +x_422 = 1; +x_423 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_423, 0, x_342); +lean_ctor_set(x_423, 1, x_2); +lean_ctor_set(x_423, 2, x_3); +lean_ctor_set(x_423, 3, x_345); +lean_ctor_set_uint8(x_423, sizeof(void*)*4, x_422); +return x_423; +} +default: +{ +lean_object* x_424; uint8_t x_425; +x_424 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_345, x_2, x_3); +x_425 = lean_ctor_get_uint8(x_424, sizeof(void*)*4); +if (x_425 == 0) +{ +lean_object* x_426; +x_426 = lean_ctor_get(x_424, 0); +lean_inc(x_426); +if (lean_obj_tag(x_426) == 0) +{ +lean_object* x_427; +x_427 = lean_ctor_get(x_424, 3); +lean_inc(x_427); +if (lean_obj_tag(x_427) == 0) +{ +lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; uint8_t x_432; lean_object* x_433; +x_428 = lean_ctor_get(x_424, 1); +lean_inc(x_428); +x_429 = lean_ctor_get(x_424, 2); +lean_inc(x_429); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_430 = x_424; +} else { + lean_dec_ref(x_424); + x_430 = lean_box(0); +} +if (lean_is_scalar(x_430)) { + x_431 = lean_alloc_ctor(1, 4, 1); +} else { + x_431 = x_430; +} +lean_ctor_set(x_431, 0, x_427); +lean_ctor_set(x_431, 1, x_428); +lean_ctor_set(x_431, 2, x_429); +lean_ctor_set(x_431, 3, x_427); +lean_ctor_set_uint8(x_431, sizeof(void*)*4, x_425); +x_432 = 1; +x_433 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_433, 0, x_342); +lean_ctor_set(x_433, 1, x_343); +lean_ctor_set(x_433, 2, x_344); +lean_ctor_set(x_433, 3, x_431); +lean_ctor_set_uint8(x_433, sizeof(void*)*4, x_432); +return x_433; +} +else +{ +uint8_t x_434; +x_434 = lean_ctor_get_uint8(x_427, sizeof(void*)*4); +if (x_434 == 0) +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; lean_object* x_447; +x_435 = lean_ctor_get(x_424, 1); +lean_inc(x_435); +x_436 = lean_ctor_get(x_424, 2); +lean_inc(x_436); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_437 = x_424; +} else { + lean_dec_ref(x_424); + x_437 = lean_box(0); +} +x_438 = lean_ctor_get(x_427, 0); +lean_inc(x_438); +x_439 = lean_ctor_get(x_427, 1); +lean_inc(x_439); +x_440 = lean_ctor_get(x_427, 2); +lean_inc(x_440); +x_441 = lean_ctor_get(x_427, 3); +lean_inc(x_441); +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_442 = x_427; +} else { + lean_dec_ref(x_427); + x_442 = lean_box(0); +} +x_443 = 1; +if (lean_is_scalar(x_442)) { + x_444 = lean_alloc_ctor(1, 4, 1); +} else { + x_444 = x_442; +} +lean_ctor_set(x_444, 0, x_342); +lean_ctor_set(x_444, 1, x_343); +lean_ctor_set(x_444, 2, x_344); +lean_ctor_set(x_444, 3, x_426); +lean_ctor_set_uint8(x_444, sizeof(void*)*4, x_443); +if (lean_is_scalar(x_437)) { + x_445 = lean_alloc_ctor(1, 4, 1); +} else { + x_445 = x_437; +} +lean_ctor_set(x_445, 0, x_438); +lean_ctor_set(x_445, 1, x_439); +lean_ctor_set(x_445, 2, x_440); +lean_ctor_set(x_445, 3, x_441); +lean_ctor_set_uint8(x_445, sizeof(void*)*4, x_443); +x_446 = 0; +x_447 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_447, 0, x_444); +lean_ctor_set(x_447, 1, x_435); +lean_ctor_set(x_447, 2, x_436); +lean_ctor_set(x_447, 3, x_445); +lean_ctor_set_uint8(x_447, sizeof(void*)*4, x_446); +return x_447; +} +else +{ +lean_object* x_448; uint8_t x_449; lean_object* x_450; +if (lean_is_exclusive(x_427)) { + lean_ctor_release(x_427, 0); + lean_ctor_release(x_427, 1); + lean_ctor_release(x_427, 2); + lean_ctor_release(x_427, 3); + x_448 = x_427; +} else { + lean_dec_ref(x_427); + x_448 = lean_box(0); +} +x_449 = 1; +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(1, 4, 1); +} else { + x_450 = x_448; +} +lean_ctor_set(x_450, 0, x_342); +lean_ctor_set(x_450, 1, x_343); +lean_ctor_set(x_450, 2, x_344); +lean_ctor_set(x_450, 3, x_424); +lean_ctor_set_uint8(x_450, sizeof(void*)*4, x_449); +return x_450; +} +} +} +else +{ +uint8_t x_451; +x_451 = lean_ctor_get_uint8(x_426, sizeof(void*)*4); +if (x_451 == 0) +{ +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; +x_452 = lean_ctor_get(x_424, 1); +lean_inc(x_452); +x_453 = lean_ctor_get(x_424, 2); +lean_inc(x_453); +x_454 = lean_ctor_get(x_424, 3); +lean_inc(x_454); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_455 = x_424; +} else { + lean_dec_ref(x_424); + x_455 = lean_box(0); +} +x_456 = lean_ctor_get(x_426, 0); +lean_inc(x_456); +x_457 = lean_ctor_get(x_426, 1); +lean_inc(x_457); +x_458 = lean_ctor_get(x_426, 2); +lean_inc(x_458); +x_459 = lean_ctor_get(x_426, 3); +lean_inc(x_459); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_460 = x_426; +} else { + lean_dec_ref(x_426); + x_460 = lean_box(0); +} +x_461 = 1; +if (lean_is_scalar(x_460)) { + x_462 = lean_alloc_ctor(1, 4, 1); +} else { + x_462 = x_460; +} +lean_ctor_set(x_462, 0, x_342); +lean_ctor_set(x_462, 1, x_343); +lean_ctor_set(x_462, 2, x_344); +lean_ctor_set(x_462, 3, x_456); +lean_ctor_set_uint8(x_462, sizeof(void*)*4, x_461); +if (lean_is_scalar(x_455)) { + x_463 = lean_alloc_ctor(1, 4, 1); +} else { + x_463 = x_455; +} +lean_ctor_set(x_463, 0, x_459); +lean_ctor_set(x_463, 1, x_452); +lean_ctor_set(x_463, 2, x_453); +lean_ctor_set(x_463, 3, x_454); +lean_ctor_set_uint8(x_463, sizeof(void*)*4, x_461); +x_464 = 0; +x_465 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_457); +lean_ctor_set(x_465, 2, x_458); +lean_ctor_set(x_465, 3, x_463); +lean_ctor_set_uint8(x_465, sizeof(void*)*4, x_464); +return x_465; +} +else +{ +lean_object* x_466; +x_466 = lean_ctor_get(x_424, 3); +lean_inc(x_466); +if (lean_obj_tag(x_466) == 0) +{ +lean_object* x_467; uint8_t x_468; lean_object* x_469; +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_467 = x_426; +} else { + lean_dec_ref(x_426); + x_467 = lean_box(0); +} +x_468 = 1; +if (lean_is_scalar(x_467)) { + x_469 = lean_alloc_ctor(1, 4, 1); +} else { + x_469 = x_467; +} +lean_ctor_set(x_469, 0, x_342); +lean_ctor_set(x_469, 1, x_343); +lean_ctor_set(x_469, 2, x_344); +lean_ctor_set(x_469, 3, x_424); +lean_ctor_set_uint8(x_469, sizeof(void*)*4, x_468); +return x_469; +} +else +{ +uint8_t x_470; +x_470 = lean_ctor_get_uint8(x_466, sizeof(void*)*4); +if (x_470 == 0) +{ +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; +x_471 = lean_ctor_get(x_424, 1); +lean_inc(x_471); +x_472 = lean_ctor_get(x_424, 2); +lean_inc(x_472); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_473 = x_424; +} else { + lean_dec_ref(x_424); + x_473 = lean_box(0); +} +x_474 = lean_ctor_get(x_466, 0); +lean_inc(x_474); +x_475 = lean_ctor_get(x_466, 1); +lean_inc(x_475); +x_476 = lean_ctor_get(x_466, 2); +lean_inc(x_476); +x_477 = lean_ctor_get(x_466, 3); +lean_inc(x_477); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + lean_ctor_release(x_466, 2); + lean_ctor_release(x_466, 3); + x_478 = x_466; +} else { + lean_dec_ref(x_466); + x_478 = lean_box(0); +} +x_479 = 1; +lean_inc(x_426); +if (lean_is_scalar(x_478)) { + x_480 = lean_alloc_ctor(1, 4, 1); +} else { + x_480 = x_478; +} +lean_ctor_set(x_480, 0, x_342); +lean_ctor_set(x_480, 1, x_343); +lean_ctor_set(x_480, 2, x_344); +lean_ctor_set(x_480, 3, x_426); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_481 = x_426; +} else { + lean_dec_ref(x_426); + x_481 = lean_box(0); +} +lean_ctor_set_uint8(x_480, sizeof(void*)*4, x_479); +if (lean_is_scalar(x_481)) { + x_482 = lean_alloc_ctor(1, 4, 1); +} else { + x_482 = x_481; +} +lean_ctor_set(x_482, 0, x_474); +lean_ctor_set(x_482, 1, x_475); +lean_ctor_set(x_482, 2, x_476); +lean_ctor_set(x_482, 3, x_477); +lean_ctor_set_uint8(x_482, sizeof(void*)*4, x_479); +x_483 = 0; +if (lean_is_scalar(x_473)) { + x_484 = lean_alloc_ctor(1, 4, 1); +} else { + x_484 = x_473; +} +lean_ctor_set(x_484, 0, x_480); +lean_ctor_set(x_484, 1, x_471); +lean_ctor_set(x_484, 2, x_472); +lean_ctor_set(x_484, 3, x_482); +lean_ctor_set_uint8(x_484, sizeof(void*)*4, x_483); +return x_484; +} +else +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; +x_485 = lean_ctor_get(x_424, 1); +lean_inc(x_485); +x_486 = lean_ctor_get(x_424, 2); +lean_inc(x_486); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + lean_ctor_release(x_424, 2); + lean_ctor_release(x_424, 3); + x_487 = x_424; +} else { + lean_dec_ref(x_424); + x_487 = lean_box(0); +} +x_488 = lean_ctor_get(x_426, 0); +lean_inc(x_488); +x_489 = lean_ctor_get(x_426, 1); +lean_inc(x_489); +x_490 = lean_ctor_get(x_426, 2); +lean_inc(x_490); +x_491 = lean_ctor_get(x_426, 3); +lean_inc(x_491); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + lean_ctor_release(x_426, 2); + lean_ctor_release(x_426, 3); + x_492 = x_426; +} else { + lean_dec_ref(x_426); + x_492 = lean_box(0); +} +if (lean_is_scalar(x_492)) { + x_493 = lean_alloc_ctor(1, 4, 1); +} else { + x_493 = x_492; +} +lean_ctor_set(x_493, 0, x_488); +lean_ctor_set(x_493, 1, x_489); +lean_ctor_set(x_493, 2, x_490); +lean_ctor_set(x_493, 3, x_491); +lean_ctor_set_uint8(x_493, sizeof(void*)*4, x_470); +if (lean_is_scalar(x_487)) { + x_494 = lean_alloc_ctor(1, 4, 1); +} else { + x_494 = x_487; +} +lean_ctor_set(x_494, 0, x_493); +lean_ctor_set(x_494, 1, x_485); +lean_ctor_set(x_494, 2, x_486); +lean_ctor_set(x_494, 3, x_466); +lean_ctor_set_uint8(x_494, sizeof(void*)*4, x_425); +x_495 = 1; +x_496 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_496, 0, x_342); +lean_ctor_set(x_496, 1, x_343); +lean_ctor_set(x_496, 2, x_344); +lean_ctor_set(x_496, 3, x_494); +lean_ctor_set_uint8(x_496, sizeof(void*)*4, x_495); +return x_496; +} +} +} +} +} +else +{ +uint8_t x_497; lean_object* x_498; +x_497 = 1; +x_498 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_498, 0, x_342); +lean_ctor_set(x_498, 1, x_343); +lean_ctor_set(x_498, 2, x_344); +lean_ctor_set(x_498, 3, x_424); +lean_ctor_set_uint8(x_498, sizeof(void*)*4, x_497); +return x_498; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_RBNode_isRed___rarg(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_1, x_2, x_3); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_RBNode_ins___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__2(x_1, x_2, x_3); +x_7 = l_Lean_RBNode_setBlack___rarg(x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_2, x_1); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; +x_12 = lean_array_uget(x_3, x_2); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_array_uset(x_3, x_2, x_13); +x_15 = 1; +x_16 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_12, x_15, x_4, x_5, x_6, x_7, x_8, x_9); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_2, x_20); +x_22 = lean_array_uset(x_14, x_2, x_19); +x_2 = x_21; +x_3 = x_22; +x_9 = x_18; +goto _start; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_14 = 1; +x_15 = lean_usize_add(x_3, x_14); +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_get_size(x_10); +x_18 = lean_nat_dec_le(x_12, x_17); +if (x_18 == 0) +{ +uint8_t x_19; +lean_dec(x_12); +x_19 = lean_nat_dec_lt(x_11, x_17); +if (x_19 == 0) +{ +size_t x_20; size_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +goto _start; +} +else +{ +size_t x_23; size_t x_24; uint8_t x_25; +x_23 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_24 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_10, x_23, x_24); +lean_dec(x_10); +if (x_25 == 0) +{ +size_t x_26; size_t x_27; +x_26 = 1; +x_27 = lean_usize_add(x_3, x_26); +x_3 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +x_29 = 1; +return x_29; +} +} +} +else +{ +size_t x_30; size_t x_31; uint8_t x_32; +lean_dec(x_17); +x_30 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_31 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_10, x_30, x_31); +lean_dec(x_10); +if (x_32 == 0) +{ +size_t x_33; size_t x_34; +x_33 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +goto _start; +} +else +{ +uint8_t x_36; +x_36 = 1; +return x_36; +} +} +} +} +else +{ +uint8_t x_37; +x_37 = 0; +return x_37; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_14 = 1; +x_15 = lean_usize_add(x_3, x_14); +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_get_size(x_10); +x_18 = lean_nat_dec_le(x_12, x_17); +if (x_18 == 0) +{ +uint8_t x_19; +lean_dec(x_12); +x_19 = lean_nat_dec_lt(x_11, x_17); +if (x_19 == 0) +{ +size_t x_20; size_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +goto _start; +} +else +{ +size_t x_23; size_t x_24; uint8_t x_25; +x_23 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_24 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_10, x_23, x_24); +lean_dec(x_10); +if (x_25 == 0) +{ +size_t x_26; size_t x_27; +x_26 = 1; +x_27 = lean_usize_add(x_3, x_26); +x_3 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +x_29 = 1; +return x_29; +} +} +} +else +{ +size_t x_30; size_t x_31; uint8_t x_32; +lean_dec(x_17); +x_30 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_31 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_10, x_30, x_31); +lean_dec(x_10); +if (x_32 == 0) +{ +size_t x_33; size_t x_34; +x_33 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +goto _start; +} +else +{ +uint8_t x_36; +x_36 = 1; +return x_36; +} +} +} +} +else +{ +uint8_t x_37; +x_37 = 0; +return x_37; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_14 = 1; +x_15 = lean_usize_add(x_3, x_14); +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_get_size(x_10); +x_18 = lean_nat_dec_le(x_12, x_17); +if (x_18 == 0) +{ +uint8_t x_19; +lean_dec(x_12); +x_19 = lean_nat_dec_lt(x_11, x_17); +if (x_19 == 0) +{ +size_t x_20; size_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +goto _start; +} +else +{ +size_t x_23; size_t x_24; uint8_t x_25; +x_23 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_24 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_10, x_23, x_24); +lean_dec(x_10); +if (x_25 == 0) +{ +size_t x_26; size_t x_27; +x_26 = 1; +x_27 = lean_usize_add(x_3, x_26); +x_3 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +x_29 = 1; +return x_29; +} +} +} +else +{ +size_t x_30; size_t x_31; uint8_t x_32; +lean_dec(x_17); +x_30 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_31 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_10, x_30, x_31); +lean_dec(x_10); +if (x_32 == 0) +{ +size_t x_33; size_t x_34; +x_33 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +goto _start; +} +else +{ +uint8_t x_36; +x_36 = 1; +return x_36; +} +} +} +} +else +{ +uint8_t x_37; +x_37 = 0; +return x_37; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Array_toSubarray___rarg(x_6, x_8, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_nat_dec_lt(x_11, x_12); +if (x_13 == 0) +{ +size_t x_14; size_t x_15; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_14 = 1; +x_15 = lean_usize_add(x_3, x_14); +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_get_size(x_10); +x_18 = lean_nat_dec_le(x_12, x_17); +if (x_18 == 0) +{ +uint8_t x_19; +lean_dec(x_12); +x_19 = lean_nat_dec_lt(x_11, x_17); +if (x_19 == 0) +{ +size_t x_20; size_t x_21; +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_10); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +goto _start; +} +else +{ +size_t x_23; size_t x_24; uint8_t x_25; +x_23 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_24 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_25 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_10, x_23, x_24); +lean_dec(x_10); +if (x_25 == 0) +{ +size_t x_26; size_t x_27; +x_26 = 1; +x_27 = lean_usize_add(x_3, x_26); +x_3 = x_27; +goto _start; +} +else +{ +uint8_t x_29; +x_29 = 1; +return x_29; +} +} +} +else +{ +size_t x_30; size_t x_31; uint8_t x_32; +lean_dec(x_17); +x_30 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_31 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_32 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_10, x_30, x_31); +lean_dec(x_10); +if (x_32 == 0) +{ +size_t x_33; size_t x_34; +x_33 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +goto _start; +} +else +{ +uint8_t x_36; +x_36 = 1; +return x_36; +} +} +} +} +else +{ +uint8_t x_37; +x_37 = 0; +return x_37; +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +_start: +{ +uint8_t x_20; +x_20 = lean_nat_dec_lt(x_11, x_8); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +else +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_nat_dec_eq(x_10, x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +lean_dec(x_13); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_10, x_24); +lean_dec(x_10); +x_26 = lean_nat_dec_lt(x_11, x_2); +lean_inc(x_11); +lean_inc(x_1); +x_27 = l_Array_toSubarray___rarg(x_1, x_22, x_11); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +x_30 = lean_ctor_get(x_27, 2); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_nat_dec_lt(x_29, x_30); +if (x_26 == 0) +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_85 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; +x_86 = l_outOfBounds___rarg(x_85); +x_87 = lean_array_get_size(x_86); +x_88 = lean_nat_dec_lt(x_22, x_87); +lean_dec(x_87); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; +lean_dec(x_86); +x_89 = l_Lean_instInhabitedName; +x_90 = l_outOfBounds___rarg(x_89); +x_32 = x_90; +goto block_84; +} +else +{ +lean_object* x_91; +x_91 = lean_array_fget(x_86, x_22); +lean_dec(x_86); +x_32 = x_91; +goto block_84; +} +} +else +{ +lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_92 = lean_array_fget(x_1, x_11); +x_93 = lean_array_get_size(x_92); +x_94 = lean_nat_dec_lt(x_22, x_93); +lean_dec(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_92); +x_95 = l_Lean_instInhabitedName; +x_96 = l_outOfBounds___rarg(x_95); +x_32 = x_96; +goto block_84; +} +else +{ +lean_object* x_97; +x_97 = lean_array_fget(x_92, x_22); +lean_dec(x_92); +x_32 = x_97; +goto block_84; +} +} +block_84: +{ +lean_object* x_33; lean_object* x_42; +if (x_31 == 0) +{ +lean_object* x_67; +lean_dec(x_30); +lean_dec(x_29); +lean_dec(x_28); +x_67 = lean_box(0); +x_42 = x_67; +goto block_66; +} +else +{ +lean_object* x_68; uint8_t x_69; +x_68 = lean_array_get_size(x_28); +x_69 = lean_nat_dec_le(x_30, x_68); +if (x_69 == 0) +{ +uint8_t x_70; +lean_dec(x_30); +x_70 = lean_nat_dec_lt(x_29, x_68); +if (x_70 == 0) +{ +lean_object* x_71; +lean_dec(x_68); +lean_dec(x_29); +lean_dec(x_28); +x_71 = lean_box(0); +x_42 = x_71; +goto block_66; +} +else +{ +size_t x_72; size_t x_73; uint8_t x_74; +x_72 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_73 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_74 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(x_32, x_28, x_72, x_73); +lean_dec(x_28); +if (x_74 == 0) +{ +lean_object* x_75; +x_75 = lean_box(0); +x_42 = x_75; +goto block_66; +} +else +{ +lean_object* x_76; +lean_dec(x_32); +x_76 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +lean_inc(x_3); +{ +lean_object* _tmp_9 = x_25; +lean_object* _tmp_10 = x_76; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_3; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +} +goto _start; +} +} +} +else +{ +size_t x_78; size_t x_79; uint8_t x_80; +lean_dec(x_68); +x_78 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_79 = lean_usize_of_nat(x_30); +lean_dec(x_30); +x_80 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(x_32, x_28, x_78, x_79); +lean_dec(x_28); +if (x_80 == 0) +{ +lean_object* x_81; +x_81 = lean_box(0); +x_42 = x_81; +goto block_66; +} +else +{ +lean_object* x_82; +lean_dec(x_32); +x_82 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +lean_inc(x_3); +{ +lean_object* _tmp_9 = x_25; +lean_object* _tmp_10 = x_82; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_3; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +} +goto _start; +} +} +} +block_41: +{ +uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_33); +x_34 = lean_nat_dec_eq(x_4, x_22); +x_35 = lean_box(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_19); +return x_40; +} +block_66: +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +lean_dec(x_42); +x_43 = lean_nat_add(x_11, x_24); +lean_inc(x_5); +lean_inc(x_1); +x_44 = l_Array_toSubarray___rarg(x_1, x_43, x_5); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +x_47 = lean_ctor_get(x_44, 2); +lean_inc(x_47); +lean_dec(x_44); +x_48 = lean_nat_dec_lt(x_46, x_47); +if (x_48 == 0) +{ +lean_object* x_49; +lean_dec(x_47); +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_49 = lean_box(0); +x_33 = x_49; +goto block_41; +} +else +{ +lean_object* x_50; uint8_t x_51; +x_50 = lean_array_get_size(x_45); +x_51 = lean_nat_dec_le(x_47, x_50); +if (x_51 == 0) +{ +uint8_t x_52; +lean_dec(x_47); +x_52 = lean_nat_dec_lt(x_46, x_50); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_50); +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_53 = lean_box(0); +x_33 = x_53; +goto block_41; +} +else +{ +size_t x_54; size_t x_55; uint8_t x_56; +x_54 = lean_usize_of_nat(x_46); +lean_dec(x_46); +x_55 = lean_usize_of_nat(x_50); +lean_dec(x_50); +x_56 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(x_32, x_45, x_54, x_55); +lean_dec(x_45); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_57 = lean_box(0); +x_33 = x_57; +goto block_41; +} +else +{ +lean_object* x_58; +lean_dec(x_32); +x_58 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +lean_inc(x_3); +{ +lean_object* _tmp_9 = x_25; +lean_object* _tmp_10 = x_58; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_3; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +} +goto _start; +} +} +} +else +{ +size_t x_60; size_t x_61; uint8_t x_62; +lean_dec(x_50); +x_60 = lean_usize_of_nat(x_46); +lean_dec(x_46); +x_61 = lean_usize_of_nat(x_47); +lean_dec(x_47); +x_62 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(x_32, x_45, x_60, x_61); +lean_dec(x_45); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_25); +lean_dec(x_11); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_63 = lean_box(0); +x_33 = x_63; +goto block_41; +} +else +{ +lean_object* x_64; +lean_dec(x_32); +x_64 = lean_nat_add(x_11, x_9); +lean_dec(x_11); +lean_inc(x_3); +{ +lean_object* _tmp_9 = x_25; +lean_object* _tmp_10 = x_64; +lean_object* _tmp_11 = lean_box(0); +lean_object* _tmp_12 = x_3; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_12 = _tmp_11; +x_13 = _tmp_12; +} +goto _start; +} +} +} +} +} +} +else +{ +lean_object* x_98; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_13); +lean_ctor_set(x_98, 1, x_19); +return x_98; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +uint8_t x_18; +x_18 = lean_nat_dec_lt(x_9, x_6); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_1); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_11); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +else +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_nat_dec_eq(x_8, x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_11); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_sub(x_8, x_22); +lean_dec(x_8); +x_24 = lean_nat_sub(x_2, x_9); +lean_inc(x_24); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_20); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_22); +lean_inc_n(x_24, 2); +lean_inc_n(x_4, 2); +lean_inc(x_1); +x_26 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(x_1, x_2, x_4, x_9, x_24, x_25, x_20, x_24, x_22, x_24, x_20, lean_box(0), x_4, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_25); +lean_dec(x_24); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_dec(x_26); +x_30 = lean_nat_add(x_9, x_7); +lean_dec(x_9); +lean_inc(x_4); +{ +lean_object* _tmp_7 = x_23; +lean_object* _tmp_8 = x_30; +lean_object* _tmp_9 = lean_box(0); +lean_object* _tmp_10 = x_4; +lean_object* _tmp_16 = x_29; +x_8 = _tmp_7; +x_9 = _tmp_8; +x_10 = _tmp_9; +x_11 = _tmp_10; +x_17 = _tmp_16; +} +goto _start; +} +else +{ +uint8_t x_32; +lean_dec(x_23); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_1); +x_32 = !lean_is_exclusive(x_26); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_26, 0); +lean_dec(x_33); +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_28); +lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_26, 0, x_36); +return x_26; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_28, 0); +lean_inc(x_37); +lean_dec(x_28); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_26, 0, x_40); +return x_26; +} +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_41 = lean_ctor_get(x_26, 1); +lean_inc(x_41); +lean_dec(x_26); +x_42 = lean_ctor_get(x_28, 0); +lean_inc(x_42); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + x_43 = x_28; +} else { + lean_dec_ref(x_28); + x_43 = lean_box(0); +} +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(1, 1, 0); +} else { + x_44 = x_43; +} +lean_ctor_set(x_44, 0, x_42); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_41); +return x_47; +} +} +} +else +{ +lean_object* x_48; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_1); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_11); +lean_ctor_set(x_48, 1, x_17); +return x_48; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_nat_dec_lt(x_10, x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1; +x_13 = l_outOfBounds___rarg(x_12); +x_14 = lean_array_get_size(x_13); +x_15 = lean_nat_dec_lt(x_10, x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_13); +x_16 = l_Lean_instInhabitedName; +x_17 = l_outOfBounds___rarg(x_16); +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_9); +return x_21; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_array_fget(x_13, x_10); +lean_dec(x_13); +x_23 = 0; +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_9); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_array_fget(x_2, x_10); +x_28 = lean_array_get_size(x_27); +x_29 = lean_nat_dec_lt(x_10, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_27); +x_30 = l_Lean_instInhabitedName; +x_31 = l_outOfBounds___rarg(x_30); +x_32 = 0; +x_33 = lean_box(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_31); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_9); +return x_35; +} +else +{ +lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_36 = lean_array_fget(x_27, x_10); +lean_dec(x_27); +x_37 = 0; +x_38 = lean_box(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_9); +return x_40; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_8 = lean_array_get_size(x_1); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_unsigned_to_nat(1u); +lean_inc(x_8); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_8); +lean_ctor_set(x_11, 2, x_10); +x_12 = l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1; +lean_inc(x_8); +lean_inc(x_1); +x_13 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(x_1, x_8, x_11, x_12, x_9, x_8, x_10, x_8, x_9, lean_box(0), x_12, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_box(0); +x_18 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(x_8, x_1, x_17, x_2, x_3, x_4, x_5, x_6, x_16); +lean_dec(x_1); +lean_dec(x_8); +return x_18; +} +else +{ +uint8_t x_19; +lean_dec(x_8); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_13); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_13, 0); +lean_dec(x_20); +x_21 = lean_ctor_get(x_15, 0); +lean_inc(x_21); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_21); +return x_13; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec(x_15); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; size_t x_9; size_t x_10; +x_7 = lean_array_uget(x_2, x_3); +x_8 = lean_name_eq(x_7, x_1); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +if (x_8 == 0) +{ +lean_object* x_11; +x_11 = lean_array_push(x_5, x_7); +x_3 = x_10; +x_5 = x_11; +goto _start; +} +else +{ +lean_dec(x_7); +x_3 = x_10; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(size_t x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +lean_dec(x_2); +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; size_t x_14; size_t x_15; +x_8 = lean_array_uget(x_6, x_5); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_6, x_5, x_9); +x_11 = lean_array_get_size(x_8); +lean_inc(x_2); +x_12 = lean_array_mk(x_2); +x_13 = lean_nat_dec_lt(x_9, x_11); +x_14 = 1; +x_15 = lean_usize_add(x_5, x_14); +if (x_13 == 0) +{ +lean_object* x_16; +lean_dec(x_11); +lean_dec(x_8); +x_16 = lean_array_uset(x_10, x_5, x_12); +x_5 = x_15; +x_6 = x_16; +goto _start; +} +else +{ +uint8_t x_18; +x_18 = lean_nat_dec_le(x_11, x_11); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_11); +lean_dec(x_8); +x_19 = lean_array_uset(x_10, x_5, x_12); +x_5 = x_15; +x_6 = x_19; +goto _start; +} +else +{ +size_t x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_usize_of_nat(x_11); +lean_dec(x_11); +x_22 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(x_3, x_8, x_1, x_21, x_12); +lean_dec(x_8); +x_23 = lean_array_uset(x_10, x_5, x_22); +x_5 = x_15; +x_6 = x_23; +goto _start; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_name_eq(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_3, x_8); +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; +x_6 = lean_array_uget(x_4, x_3); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_array_uset(x_4, x_3, x_7); +x_9 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_1, x_6); +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_6); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_14 = lean_array_uset(x_8, x_3, x_11); +x_3 = x_13; +x_4 = x_14; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; size_t x_15; size_t x_16; +x_7 = lean_array_uget(x_2, x_3); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(1u); +lean_inc(x_7); +x_10 = l_Array_toSubarray___rarg(x_7, x_9, x_8); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_10, 2); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_nat_dec_lt(x_12, x_13); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +if (x_14 == 0) +{ +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_7); +x_3 = x_16; +goto _start; +} +else +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_array_get_size(x_11); +x_19 = lean_nat_dec_le(x_13, x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_13); +x_20 = lean_nat_dec_lt(x_12, x_18); +if (x_20 == 0) +{ +lean_dec(x_18); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_7); +x_3 = x_16; +goto _start; +} +else +{ +size_t x_22; size_t x_23; uint8_t x_24; +x_22 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_23 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_24 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(x_1, x_11, x_22, x_23); +lean_dec(x_11); +if (x_24 == 0) +{ +lean_dec(x_7); +x_3 = x_16; +goto _start; +} +else +{ +lean_object* x_26; +x_26 = lean_array_push(x_5, x_7); +x_3 = x_16; +x_5 = x_26; +goto _start; +} +} +} +else +{ +size_t x_28; size_t x_29; uint8_t x_30; +lean_dec(x_18); +x_28 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_29 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_30 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(x_1, x_11, x_28, x_29); +lean_dec(x_11); +if (x_30 == 0) +{ +lean_dec(x_7); +x_3 = x_16; +goto _start; +} +else +{ +lean_object* x_32; +x_32 = lean_array_push(x_5, x_7); +x_3 = x_16; +x_5 = x_32; +goto _start; +} +} +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_inc(x_1); +x_15 = lean_array_push(x_6, x_1); +x_16 = lean_array_size(x_7); +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(x_2, x_3, x_1, x_16, x_2, x_7); +lean_dec(x_1); +x_18 = lean_array_get_size(x_17); +x_19 = lean_unsigned_to_nat(0u); +x_20 = lean_nat_dec_lt(x_19, x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_18); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_15); +lean_ctor_set(x_21, 1, x_4); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_5); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_14); +return x_24; +} +else +{ +uint8_t x_25; +x_25 = lean_nat_dec_le(x_18, x_18); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_18); +lean_dec(x_17); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_15); +lean_ctor_set(x_26, 1, x_4); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_5); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_14); +return x_29; +} +else +{ +size_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_30 = lean_usize_of_nat(x_18); +lean_dec(x_18); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_17, x_2, x_30, x_4); +lean_dec(x_17); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_15); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_14); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_6); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_6, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +x_19 = l_Array_isEmpty___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_free_object(x_14); +lean_free_object(x_6); +lean_inc(x_18); +x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_array_get_size(x_18); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +uint8_t x_49; +x_49 = lean_nat_dec_le(x_26, x_26); +if (x_49 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +size_t x_50; lean_object* x_51; +x_50 = lean_usize_of_nat(x_26); +lean_dec(x_26); +lean_inc(x_4); +x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); +x_30 = x_51; +goto block_48; +} +} +block_48: +{ +size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_31 = lean_array_size(x_30); +x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +lean_dec(x_33); +x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); +lean_dec(x_35); +x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); +lean_dec(x_36); +x_38 = lean_array_size(x_37); +x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); +lean_inc(x_25); +x_40 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_40, 0, x_25); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); +x_41 = lean_array_push(x_16, x_40); +x_42 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +lean_dec(x_44); +x_6 = x_46; +x_12 = x_45; +goto _start; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_52 = lean_ctor_get(x_20, 1); +lean_inc(x_52); +lean_dec(x_20); +x_53 = lean_ctor_get(x_21, 1); +lean_inc(x_53); +lean_dec(x_21); +x_54 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_ctor_get(x_56, 0); +lean_inc(x_58); +lean_dec(x_56); +x_6 = x_58; +x_12 = x_57; +goto _start; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_60 = lean_ctor_get(x_20, 1); +lean_inc(x_60); +lean_dec(x_20); +x_61 = lean_ctor_get(x_21, 1); +lean_inc(x_61); +lean_dec(x_21); +x_62 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +lean_dec(x_64); +x_6 = x_66; +x_12 = x_65; +goto _start; +} +} +else +{ +lean_object* x_68; +lean_dec(x_5); +lean_dec(x_4); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_6); +lean_ctor_set(x_68, 1, x_12); +return x_68; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_6, 0); +x_70 = lean_ctor_get(x_14, 0); +x_71 = lean_ctor_get(x_14, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_14); +x_72 = l_Array_isEmpty___rarg(x_71); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +lean_free_object(x_6); +lean_inc(x_71); +x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_unbox(x_75); +lean_dec(x_75); +if (x_76 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_ctor_get(x_74, 1); +lean_inc(x_78); +lean_dec(x_74); +x_79 = lean_array_get_size(x_71); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_79); +x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); +if (x_81 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +uint8_t x_102; +x_102 = lean_nat_dec_le(x_79, x_79); +if (x_102 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +size_t x_103; lean_object* x_104; +x_103 = lean_usize_of_nat(x_79); +lean_dec(x_79); +lean_inc(x_4); +x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); +x_83 = x_104; +goto block_101; +} +} +block_101: +{ +size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_84 = lean_array_size(x_83); +x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); +x_86 = lean_array_get_size(x_85); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_sub(x_86, x_87); +lean_dec(x_86); +x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); +lean_dec(x_88); +x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); +lean_dec(x_89); +x_91 = lean_array_size(x_90); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); +lean_inc(x_78); +x_93 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_93, 0, x_78); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); +x_94 = lean_array_push(x_69, x_93); +x_95 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = lean_ctor_get(x_97, 0); +lean_inc(x_99); +lean_dec(x_97); +x_6 = x_99; +x_12 = x_98; +goto _start; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_105 = lean_ctor_get(x_73, 1); +lean_inc(x_105); +lean_dec(x_73); +x_106 = lean_ctor_get(x_74, 1); +lean_inc(x_106); +lean_dec(x_74); +x_107 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +lean_dec(x_109); +x_6 = x_111; +x_12 = x_110; +goto _start; +} +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_113 = lean_ctor_get(x_73, 1); +lean_inc(x_113); +lean_dec(x_73); +x_114 = lean_ctor_get(x_74, 1); +lean_inc(x_114); +lean_dec(x_74); +x_115 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_119 = lean_ctor_get(x_117, 0); +lean_inc(x_119); +lean_dec(x_117); +x_6 = x_119; +x_12 = x_118; +goto _start; +} +} +else +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_5); +lean_dec(x_4); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_70); +lean_ctor_set(x_121, 1, x_71); +lean_ctor_set(x_6, 1, x_121); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_6); +lean_ctor_set(x_122, 1, x_12); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_127 = x_123; +} else { + lean_dec_ref(x_123); + x_127 = lean_box(0); +} +x_128 = l_Array_isEmpty___rarg(x_126); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +lean_dec(x_127); +lean_inc(x_126); +x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_unbox(x_131); +lean_dec(x_131); +if (x_132 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; +x_133 = lean_ctor_get(x_129, 1); +lean_inc(x_133); +lean_dec(x_129); +x_134 = lean_ctor_get(x_130, 1); +lean_inc(x_134); +lean_dec(x_130); +x_135 = lean_array_get_size(x_126); +x_136 = lean_unsigned_to_nat(0u); +x_137 = lean_nat_dec_lt(x_136, x_135); +x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); +if (x_137 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +uint8_t x_158; +x_158 = lean_nat_dec_le(x_135, x_135); +if (x_158 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +size_t x_159; lean_object* x_160; +x_159 = lean_usize_of_nat(x_135); +lean_dec(x_135); +lean_inc(x_4); +x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); +x_139 = x_160; +goto block_157; +} +} +block_157: +{ +size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_140 = lean_array_size(x_139); +x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); +x_142 = lean_array_get_size(x_141); +x_143 = lean_unsigned_to_nat(1u); +x_144 = lean_nat_sub(x_142, x_143); +lean_dec(x_142); +x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); +lean_dec(x_144); +x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); +lean_dec(x_145); +x_147 = lean_array_size(x_146); +x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); +lean_inc(x_134); +x_149 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_149, 0, x_134); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); +x_150 = lean_array_push(x_124, x_149); +x_151 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); +x_155 = lean_ctor_get(x_153, 0); +lean_inc(x_155); +lean_dec(x_153); +x_6 = x_155; +x_12 = x_154; +goto _start; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_161 = lean_ctor_get(x_129, 1); +lean_inc(x_161); +lean_dec(x_129); +x_162 = lean_ctor_get(x_130, 1); +lean_inc(x_162); +lean_dec(x_130); +x_163 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_6 = x_167; +x_12 = x_166; +goto _start; +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_169 = lean_ctor_get(x_129, 1); +lean_inc(x_169); +lean_dec(x_129); +x_170 = lean_ctor_get(x_130, 1); +lean_inc(x_170); +lean_dec(x_130); +x_171 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); +lean_inc(x_174); +lean_dec(x_172); +x_175 = lean_ctor_get(x_173, 0); +lean_inc(x_175); +lean_dec(x_173); +x_6 = x_175; +x_12 = x_174; +goto _start; +} +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_5); +lean_dec(x_4); +if (lean_is_scalar(x_127)) { + x_177 = lean_alloc_ctor(0, 2, 0); +} else { + x_177 = x_127; +} +lean_ctor_set(x_177, 0, x_125); +lean_ctor_set(x_177, 1, x_126); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_124); +lean_ctor_set(x_178, 1, x_177); +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_179, 1, x_12); +return x_179; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_insert___at___private_Lean_Structure_0__Lean_setStructureResolutionOrder___spec__1(x_3, x_1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_structureResolutionExt; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +x_2 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5; +x_2 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +x_3 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6; +x_4 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set(x_4, 3, x_2); +lean_ctor_set(x_4, 4, x_2); +lean_ctor_set(x_4, 5, x_3); +lean_ctor_set(x_4, 6, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_13 = lean_ctor_get(x_10, 0); +x_14 = lean_ctor_get(x_10, 4); +lean_dec(x_14); +x_15 = lean_alloc_closure((void*)(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1), 3, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +x_16 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; +x_17 = l_Lean_EnvExtension_modifyState___rarg(x_16, x_13, x_15); +x_18 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; +lean_ctor_set(x_10, 4, x_18); +lean_ctor_set(x_10, 0, x_17); +x_19 = lean_st_ref_set(x_7, x_10, x_11); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_st_ref_take(x_5, x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = !lean_is_exclusive(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_22, 1); +lean_dec(x_25); +x_26 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; +lean_ctor_set(x_22, 1, x_26); +x_27 = lean_st_ref_set(x_5, x_22, x_23); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_34 = lean_ctor_get(x_22, 0); +x_35 = lean_ctor_get(x_22, 2); +x_36 = lean_ctor_get(x_22, 3); +x_37 = lean_ctor_get(x_22, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_22); +x_38 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; +x_39 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_39, 0, x_34); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_35); +lean_ctor_set(x_39, 3, x_36); +lean_ctor_set(x_39, 4, x_37); +x_40 = lean_st_ref_set(x_5, x_39, x_23); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_42 = x_40; +} else { + lean_dec_ref(x_40); + x_42 = lean_box(0); +} +x_43 = lean_box(0); +if (lean_is_scalar(x_42)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_42; +} +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_41); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_45 = lean_ctor_get(x_10, 0); +x_46 = lean_ctor_get(x_10, 1); +x_47 = lean_ctor_get(x_10, 2); +x_48 = lean_ctor_get(x_10, 3); +x_49 = lean_ctor_get(x_10, 5); +x_50 = lean_ctor_get(x_10, 6); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_10); +x_51 = lean_alloc_closure((void*)(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___lambda__1), 3, 2); +lean_closure_set(x_51, 0, x_1); +lean_closure_set(x_51, 1, x_2); +x_52 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1; +x_53 = l_Lean_EnvExtension_modifyState___rarg(x_52, x_45, x_51); +x_54 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4; +x_55 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_46); +lean_ctor_set(x_55, 2, x_47); +lean_ctor_set(x_55, 3, x_48); +lean_ctor_set(x_55, 4, x_54); +lean_ctor_set(x_55, 5, x_49); +lean_ctor_set(x_55, 6, x_50); +x_56 = lean_st_ref_set(x_7, x_55, x_11); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_st_ref_take(x_5, x_57); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_59, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 4); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + lean_ctor_release(x_59, 2); + lean_ctor_release(x_59, 3); + lean_ctor_release(x_59, 4); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); +} +x_66 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7; +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 5, 0); +} else { + x_67 = x_65; +} +lean_ctor_set(x_67, 0, x_61); +lean_ctor_set(x_67, 1, x_66); +lean_ctor_set(x_67, 2, x_62); +lean_ctor_set(x_67, 3, x_63); +lean_ctor_set(x_67, 4, x_64); +x_68 = lean_st_ref_set(x_5, x_67, x_60); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; +} else { + lean_dec_ref(x_68); + x_70 = lean_box(0); +} +x_71 = lean_box(0); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(0, 2, 0); +} else { + x_72 = x_70; +} +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_6); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_6, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +x_19 = l_Array_isEmpty___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_free_object(x_14); +lean_free_object(x_6); +lean_inc(x_18); +x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_array_get_size(x_18); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +uint8_t x_49; +x_49 = lean_nat_dec_le(x_26, x_26); +if (x_49 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +size_t x_50; lean_object* x_51; +x_50 = lean_usize_of_nat(x_26); +lean_dec(x_26); +lean_inc(x_4); +x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); +x_30 = x_51; +goto block_48; +} +} +block_48: +{ +size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_31 = lean_array_size(x_30); +x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +lean_dec(x_33); +x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); +lean_dec(x_35); +x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); +lean_dec(x_36); +x_38 = lean_array_size(x_37); +x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); +lean_inc(x_25); +x_40 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_40, 0, x_25); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); +x_41 = lean_array_push(x_16, x_40); +x_42 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +lean_dec(x_44); +x_6 = x_46; +x_12 = x_45; +goto _start; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_52 = lean_ctor_get(x_20, 1); +lean_inc(x_52); +lean_dec(x_20); +x_53 = lean_ctor_get(x_21, 1); +lean_inc(x_53); +lean_dec(x_21); +x_54 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_ctor_get(x_56, 0); +lean_inc(x_58); +lean_dec(x_56); +x_6 = x_58; +x_12 = x_57; +goto _start; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_60 = lean_ctor_get(x_20, 1); +lean_inc(x_60); +lean_dec(x_20); +x_61 = lean_ctor_get(x_21, 1); +lean_inc(x_61); +lean_dec(x_21); +x_62 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +lean_dec(x_64); +x_6 = x_66; +x_12 = x_65; +goto _start; +} +} +else +{ +lean_object* x_68; +lean_dec(x_5); +lean_dec(x_4); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_6); +lean_ctor_set(x_68, 1, x_12); +return x_68; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_6, 0); +x_70 = lean_ctor_get(x_14, 0); +x_71 = lean_ctor_get(x_14, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_14); +x_72 = l_Array_isEmpty___rarg(x_71); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +lean_free_object(x_6); +lean_inc(x_71); +x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_unbox(x_75); +lean_dec(x_75); +if (x_76 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_ctor_get(x_74, 1); +lean_inc(x_78); +lean_dec(x_74); +x_79 = lean_array_get_size(x_71); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_79); +x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); +if (x_81 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +uint8_t x_102; +x_102 = lean_nat_dec_le(x_79, x_79); +if (x_102 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +size_t x_103; lean_object* x_104; +x_103 = lean_usize_of_nat(x_79); +lean_dec(x_79); +lean_inc(x_4); +x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); +x_83 = x_104; +goto block_101; +} +} +block_101: +{ +size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_84 = lean_array_size(x_83); +x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); +x_86 = lean_array_get_size(x_85); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_sub(x_86, x_87); +lean_dec(x_86); +x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); +lean_dec(x_88); +x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); +lean_dec(x_89); +x_91 = lean_array_size(x_90); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); +lean_inc(x_78); +x_93 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_93, 0, x_78); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); +x_94 = lean_array_push(x_69, x_93); +x_95 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = lean_ctor_get(x_97, 0); +lean_inc(x_99); +lean_dec(x_97); +x_6 = x_99; +x_12 = x_98; +goto _start; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_105 = lean_ctor_get(x_73, 1); +lean_inc(x_105); +lean_dec(x_73); +x_106 = lean_ctor_get(x_74, 1); +lean_inc(x_106); +lean_dec(x_74); +x_107 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +lean_dec(x_109); +x_6 = x_111; +x_12 = x_110; +goto _start; +} +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_113 = lean_ctor_get(x_73, 1); +lean_inc(x_113); +lean_dec(x_73); +x_114 = lean_ctor_get(x_74, 1); +lean_inc(x_114); +lean_dec(x_74); +x_115 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_119 = lean_ctor_get(x_117, 0); +lean_inc(x_119); +lean_dec(x_117); +x_6 = x_119; +x_12 = x_118; +goto _start; +} +} +else +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_5); +lean_dec(x_4); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_70); +lean_ctor_set(x_121, 1, x_71); +lean_ctor_set(x_6, 1, x_121); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_6); +lean_ctor_set(x_122, 1, x_12); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_127 = x_123; +} else { + lean_dec_ref(x_123); + x_127 = lean_box(0); +} +x_128 = l_Array_isEmpty___rarg(x_126); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +lean_dec(x_127); +lean_inc(x_126); +x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_unbox(x_131); +lean_dec(x_131); +if (x_132 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; +x_133 = lean_ctor_get(x_129, 1); +lean_inc(x_133); +lean_dec(x_129); +x_134 = lean_ctor_get(x_130, 1); +lean_inc(x_134); +lean_dec(x_130); +x_135 = lean_array_get_size(x_126); +x_136 = lean_unsigned_to_nat(0u); +x_137 = lean_nat_dec_lt(x_136, x_135); +x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); +if (x_137 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +uint8_t x_158; +x_158 = lean_nat_dec_le(x_135, x_135); +if (x_158 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +size_t x_159; lean_object* x_160; +x_159 = lean_usize_of_nat(x_135); +lean_dec(x_135); +lean_inc(x_4); +x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); +x_139 = x_160; +goto block_157; +} +} +block_157: +{ +size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_140 = lean_array_size(x_139); +x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); +x_142 = lean_array_get_size(x_141); +x_143 = lean_unsigned_to_nat(1u); +x_144 = lean_nat_sub(x_142, x_143); +lean_dec(x_142); +x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); +lean_dec(x_144); +x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); +lean_dec(x_145); +x_147 = lean_array_size(x_146); +x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); +lean_inc(x_134); +x_149 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_149, 0, x_134); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); +x_150 = lean_array_push(x_124, x_149); +x_151 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); +x_155 = lean_ctor_get(x_153, 0); +lean_inc(x_155); +lean_dec(x_153); +x_6 = x_155; +x_12 = x_154; +goto _start; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_161 = lean_ctor_get(x_129, 1); +lean_inc(x_161); +lean_dec(x_129); +x_162 = lean_ctor_get(x_130, 1); +lean_inc(x_162); +lean_dec(x_130); +x_163 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_6 = x_167; +x_12 = x_166; +goto _start; +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_169 = lean_ctor_get(x_129, 1); +lean_inc(x_169); +lean_dec(x_129); +x_170 = lean_ctor_get(x_130, 1); +lean_inc(x_170); +lean_dec(x_130); +x_171 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); +lean_inc(x_174); +lean_dec(x_172); +x_175 = lean_ctor_get(x_173, 0); +lean_inc(x_175); +lean_dec(x_173); +x_6 = x_175; +x_12 = x_174; +goto _start; +} +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_5); +lean_dec(x_4); +if (lean_is_scalar(x_127)) { + x_177 = lean_alloc_ctor(0, 2, 0); +} else { + x_177 = x_127; +} +lean_ctor_set(x_177, 0, x_125); +lean_ctor_set(x_177, 1, x_126); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_124); +lean_ctor_set(x_178, 1, x_177); +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_179, 1, x_12); +return x_179; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(uint8_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_6); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_6, 1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +x_19 = l_Array_isEmpty___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_free_object(x_14); +lean_free_object(x_6); +lean_inc(x_18); +x_20 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_18, x_7, x_8, x_9, x_10, x_11, x_12); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_20, 1); +lean_inc(x_24); +lean_dec(x_20); +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_array_get_size(x_18); +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_nat_dec_lt(x_27, x_26); +x_29 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_25); +if (x_28 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +uint8_t x_49; +x_49 = lean_nat_dec_le(x_26, x_26); +if (x_49 == 0) +{ +lean_dec(x_26); +lean_inc(x_4); +x_30 = x_4; +goto block_48; +} +else +{ +size_t x_50; lean_object* x_51; +x_50 = lean_usize_of_nat(x_26); +lean_dec(x_26); +lean_inc(x_4); +x_51 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_25, x_18, x_2, x_50, x_4); +x_30 = x_51; +goto block_48; +} +} +block_48: +{ +size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; size_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_31 = lean_array_size(x_30); +x_32 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_31, x_2, x_30); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +lean_dec(x_33); +x_36 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_32, x_27, x_35); +lean_dec(x_35); +x_37 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_36); +lean_dec(x_36); +x_38 = lean_array_size(x_37); +x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_38, x_2, x_37); +lean_inc(x_25); +x_40 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_40, 0, x_25); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set_uint8(x_40, sizeof(void*)*2, x_29); +x_41 = lean_array_push(x_16, x_40); +x_42 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_43 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_25, x_2, x_5, x_4, x_41, x_17, x_18, x_42, x_7, x_8, x_9, x_10, x_11, x_24); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_ctor_get(x_44, 0); +lean_inc(x_46); +lean_dec(x_44); +x_6 = x_46; +x_12 = x_45; +goto _start; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_52 = lean_ctor_get(x_20, 1); +lean_inc(x_52); +lean_dec(x_20); +x_53 = lean_ctor_get(x_21, 1); +lean_inc(x_53); +lean_dec(x_21); +x_54 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_55 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_53, x_2, x_5, x_4, x_16, x_17, x_18, x_54, x_7, x_8, x_9, x_10, x_11, x_52); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_ctor_get(x_56, 0); +lean_inc(x_58); +lean_dec(x_56); +x_6 = x_58; +x_12 = x_57; +goto _start; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_60 = lean_ctor_get(x_20, 1); +lean_inc(x_60); +lean_dec(x_20); +x_61 = lean_ctor_get(x_21, 1); +lean_inc(x_61); +lean_dec(x_21); +x_62 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_63 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_61, x_2, x_5, x_4, x_16, x_17, x_18, x_62, x_7, x_8, x_9, x_10, x_11, x_60); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +lean_dec(x_64); +x_6 = x_66; +x_12 = x_65; +goto _start; +} +} +else +{ +lean_object* x_68; +lean_dec(x_5); +lean_dec(x_4); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_6); +lean_ctor_set(x_68, 1, x_12); +return x_68; +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_6, 0); +x_70 = lean_ctor_get(x_14, 0); +x_71 = lean_ctor_get(x_14, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_14); +x_72 = l_Array_isEmpty___rarg(x_71); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +lean_free_object(x_6); +lean_inc(x_71); +x_73 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_71, x_7, x_8, x_9, x_10, x_11, x_12); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_unbox(x_75); +lean_dec(x_75); +if (x_76 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_ctor_get(x_74, 1); +lean_inc(x_78); +lean_dec(x_74); +x_79 = lean_array_get_size(x_71); +x_80 = lean_unsigned_to_nat(0u); +x_81 = lean_nat_dec_lt(x_80, x_79); +x_82 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_78); +if (x_81 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +uint8_t x_102; +x_102 = lean_nat_dec_le(x_79, x_79); +if (x_102 == 0) +{ +lean_dec(x_79); +lean_inc(x_4); +x_83 = x_4; +goto block_101; +} +else +{ +size_t x_103; lean_object* x_104; +x_103 = lean_usize_of_nat(x_79); +lean_dec(x_79); +lean_inc(x_4); +x_104 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_78, x_71, x_2, x_103, x_4); +x_83 = x_104; +goto block_101; +} +} +block_101: +{ +size_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_84 = lean_array_size(x_83); +x_85 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_84, x_2, x_83); +x_86 = lean_array_get_size(x_85); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_sub(x_86, x_87); +lean_dec(x_86); +x_89 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_85, x_80, x_88); +lean_dec(x_88); +x_90 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_89); +lean_dec(x_89); +x_91 = lean_array_size(x_90); +x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_91, x_2, x_90); +lean_inc(x_78); +x_93 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_93, 0, x_78); +lean_ctor_set(x_93, 1, x_92); +lean_ctor_set_uint8(x_93, sizeof(void*)*2, x_82); +x_94 = lean_array_push(x_69, x_93); +x_95 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_96 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_78, x_2, x_5, x_4, x_94, x_70, x_71, x_95, x_7, x_8, x_9, x_10, x_11, x_77); +x_97 = lean_ctor_get(x_96, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +lean_dec(x_96); +x_99 = lean_ctor_get(x_97, 0); +lean_inc(x_99); +lean_dec(x_97); +x_6 = x_99; +x_12 = x_98; +goto _start; +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_105 = lean_ctor_get(x_73, 1); +lean_inc(x_105); +lean_dec(x_73); +x_106 = lean_ctor_get(x_74, 1); +lean_inc(x_106); +lean_dec(x_74); +x_107 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_108 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_106, x_2, x_5, x_4, x_69, x_70, x_71, x_107, x_7, x_8, x_9, x_10, x_11, x_105); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +lean_dec(x_109); +x_6 = x_111; +x_12 = x_110; +goto _start; +} +} +else +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_113 = lean_ctor_get(x_73, 1); +lean_inc(x_113); +lean_dec(x_73); +x_114 = lean_ctor_get(x_74, 1); +lean_inc(x_114); +lean_dec(x_74); +x_115 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_116 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_114, x_2, x_5, x_4, x_69, x_70, x_71, x_115, x_7, x_8, x_9, x_10, x_11, x_113); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_119 = lean_ctor_get(x_117, 0); +lean_inc(x_119); +lean_dec(x_117); +x_6 = x_119; +x_12 = x_118; +goto _start; +} +} +else +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_5); +lean_dec(x_4); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_70); +lean_ctor_set(x_121, 1, x_71); +lean_ctor_set(x_6, 1, x_121); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_6); +lean_ctor_set(x_122, 1, x_12); +return x_122; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_127 = x_123; +} else { + lean_dec_ref(x_123); + x_127 = lean_box(0); +} +x_128 = l_Array_isEmpty___rarg(x_126); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +lean_dec(x_127); +lean_inc(x_126); +x_129 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_126, x_7, x_8, x_9, x_10, x_11, x_12); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_unbox(x_131); +lean_dec(x_131); +if (x_132 == 0) +{ +if (x_1 == 0) +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; uint8_t x_138; lean_object* x_139; +x_133 = lean_ctor_get(x_129, 1); +lean_inc(x_133); +lean_dec(x_129); +x_134 = lean_ctor_get(x_130, 1); +lean_inc(x_134); +lean_dec(x_130); +x_135 = lean_array_get_size(x_126); +x_136 = lean_unsigned_to_nat(0u); +x_137 = lean_nat_dec_lt(x_136, x_135); +x_138 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_3, x_134); +if (x_137 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +uint8_t x_158; +x_158 = lean_nat_dec_le(x_135, x_135); +if (x_158 == 0) +{ +lean_dec(x_135); +lean_inc(x_4); +x_139 = x_4; +goto block_157; +} +else +{ +size_t x_159; lean_object* x_160; +x_159 = lean_usize_of_nat(x_135); +lean_dec(x_135); +lean_inc(x_4); +x_160 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_134, x_126, x_2, x_159, x_4); +x_139 = x_160; +goto block_157; +} +} +block_157: +{ +size_t x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; size_t x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_140 = lean_array_size(x_139); +x_141 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__8(x_140, x_2, x_139); +x_142 = lean_array_get_size(x_141); +x_143 = lean_unsigned_to_nat(1u); +x_144 = lean_nat_sub(x_142, x_143); +lean_dec(x_142); +x_145 = l_Array_qsort_sort___at_Lean_computeStructureResolutionOrder___spec__9(x_141, x_136, x_144); +lean_dec(x_144); +x_146 = l_Array_eraseReps___at_Lean_computeStructureResolutionOrder___spec__10(x_145); +lean_dec(x_145); +x_147 = lean_array_size(x_146); +x_148 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_3, x_147, x_2, x_146); +lean_inc(x_134); +x_149 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_149, 0, x_134); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set_uint8(x_149, sizeof(void*)*2, x_138); +x_150 = lean_array_push(x_124, x_149); +x_151 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_152 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_134, x_2, x_5, x_4, x_150, x_125, x_126, x_151, x_7, x_8, x_9, x_10, x_11, x_133); +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +lean_dec(x_152); +x_155 = lean_ctor_get(x_153, 0); +lean_inc(x_155); +lean_dec(x_153); +x_6 = x_155; +x_12 = x_154; +goto _start; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_161 = lean_ctor_get(x_129, 1); +lean_inc(x_161); +lean_dec(x_129); +x_162 = lean_ctor_get(x_130, 1); +lean_inc(x_162); +lean_dec(x_130); +x_163 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_164 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_162, x_2, x_5, x_4, x_124, x_125, x_126, x_163, x_7, x_8, x_9, x_10, x_11, x_161); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_164, 1); +lean_inc(x_166); +lean_dec(x_164); +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_6 = x_167; +x_12 = x_166; +goto _start; +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_169 = lean_ctor_get(x_129, 1); +lean_inc(x_169); +lean_dec(x_129); +x_170 = lean_ctor_get(x_130, 1); +lean_inc(x_170); +lean_dec(x_130); +x_171 = lean_box(0); +lean_inc(x_4); +lean_inc(x_5); +x_172 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_170, x_2, x_5, x_4, x_124, x_125, x_126, x_171, x_7, x_8, x_9, x_10, x_11, x_169); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); +lean_inc(x_174); +lean_dec(x_172); +x_175 = lean_ctor_get(x_173, 0); +lean_inc(x_175); +lean_dec(x_173); +x_6 = x_175; +x_12 = x_174; +goto _start; +} +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_5); +lean_dec(x_4); +if (lean_is_scalar(x_127)) { + x_177 = lean_alloc_ctor(0, 2, 0); +} else { + x_177 = x_127; +} +lean_ctor_set(x_177, 0, x_125); +lean_ctor_set(x_177, 1, x_126); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_124); +lean_ctor_set(x_178, 1, x_177); +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_179, 1, x_12); +return x_179; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; size_t x_15; lean_object* x_16; uint8_t x_17; +lean_inc(x_2); +x_11 = l_Lean_getStructureParentInfo(x_1, x_2); +x_12 = lean_array_size(x_11); +x_13 = 0; +x_14 = l_Array_mapMUnsafe_map___at_Lean_computeStructureResolutionOrder___spec__1(x_12, x_13, x_11); +x_15 = lean_array_size(x_14); +lean_inc(x_14); +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(x_15, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +x_20 = lean_array_get_size(x_18); +x_21 = lean_unsigned_to_nat(1u); +x_22 = lean_nat_add(x_20, x_21); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_mod(x_23, x_22); +lean_dec(x_22); +lean_inc(x_14); +lean_inc(x_18); +x_25 = lean_array_push(x_18, x_14); +x_26 = l_Array_insertAt_loop___rarg(x_18, x_24, x_25, x_20); +lean_dec(x_24); +lean_dec(x_18); +x_27 = lean_array_get_size(x_26); +x_28 = lean_box(0); +x_29 = lean_nat_dec_lt(x_23, x_27); +lean_inc(x_2); +lean_ctor_set_tag(x_16, 1); +lean_ctor_set(x_16, 1, x_28); +lean_ctor_set(x_16, 0, x_2); +x_30 = lean_array_mk(x_16); +if (x_29 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +lean_dec(x_27); +lean_dec(x_26); +x_31 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_3, x_13, x_14, x_31, x_28, x_33, x_5, x_6, x_7, x_8, x_9, x_19); +lean_dec(x_14); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = lean_ctor_get(x_35, 0); +lean_inc(x_38); +lean_dec(x_35); +x_39 = !lean_is_exclusive(x_36); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = lean_ctor_get(x_36, 0); +x_41 = lean_ctor_get(x_36, 1); +lean_dec(x_41); +lean_inc(x_40); +x_42 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_40, x_5, x_6, x_7, x_8, x_9, x_37); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +lean_ctor_set(x_36, 1, x_38); +lean_ctor_set(x_42, 0, x_36); +return x_42; +} +else +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +lean_ctor_set(x_36, 1, x_38); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_47 = lean_ctor_get(x_36, 0); +lean_inc(x_47); +lean_dec(x_36); +lean_inc(x_47); +x_48 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_47, x_5, x_6, x_7, x_8, x_9, x_37); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_50 = x_48; +} else { + lean_dec_ref(x_48); + x_50 = lean_box(0); +} +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_47); +lean_ctor_set(x_51, 1, x_38); +if (lean_is_scalar(x_50)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_50; +} +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +return x_52; +} +} +else +{ +uint8_t x_53; +x_53 = lean_nat_dec_le(x_27, x_27); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +lean_dec(x_27); +lean_dec(x_26); +x_54 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_30); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_3, x_13, x_14, x_54, x_28, x_56, x_5, x_6, x_7, x_8, x_9, x_19); +lean_dec(x_14); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); +lean_dec(x_57); +x_61 = lean_ctor_get(x_58, 0); +lean_inc(x_61); +lean_dec(x_58); +x_62 = !lean_is_exclusive(x_59); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = lean_ctor_get(x_59, 0); +x_64 = lean_ctor_get(x_59, 1); +lean_dec(x_64); +lean_inc(x_63); +x_65 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_63, x_5, x_6, x_7, x_8, x_9, x_60); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +lean_object* x_67; +x_67 = lean_ctor_get(x_65, 0); +lean_dec(x_67); +lean_ctor_set(x_59, 1, x_61); +lean_ctor_set(x_65, 0, x_59); +return x_65; +} +else +{ +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +lean_ctor_set(x_59, 1, x_61); +x_69 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_69, 0, x_59); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_70 = lean_ctor_get(x_59, 0); +lean_inc(x_70); +lean_dec(x_59); +lean_inc(x_70); +x_71 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_70, x_5, x_6, x_7, x_8, x_9, x_60); +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_73 = x_71; +} else { + lean_dec_ref(x_71); + x_73 = lean_box(0); +} +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_61); +if (lean_is_scalar(x_73)) { + x_75 = lean_alloc_ctor(0, 2, 0); +} else { + x_75 = x_73; +} +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_72); +return x_75; +} +} +else +{ +size_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_76 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_77 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_78 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_26, x_13, x_76, x_77); +lean_dec(x_26); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_30); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_3, x_13, x_14, x_77, x_28, x_80, x_5, x_6, x_7, x_8, x_9, x_19); +lean_dec(x_14); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_82, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_84); +lean_dec(x_81); +x_85 = lean_ctor_get(x_82, 0); +lean_inc(x_85); +lean_dec(x_82); +x_86 = !lean_is_exclusive(x_83); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_87 = lean_ctor_get(x_83, 0); +x_88 = lean_ctor_get(x_83, 1); +lean_dec(x_88); +lean_inc(x_87); +x_89 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_87, x_5, x_6, x_7, x_8, x_9, x_84); +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) +{ +lean_object* x_91; +x_91 = lean_ctor_get(x_89, 0); +lean_dec(x_91); +lean_ctor_set(x_83, 1, x_85); +lean_ctor_set(x_89, 0, x_83); +return x_89; +} +else +{ +lean_object* x_92; lean_object* x_93; +x_92 = lean_ctor_get(x_89, 1); +lean_inc(x_92); +lean_dec(x_89); +lean_ctor_set(x_83, 1, x_85); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_83); +lean_ctor_set(x_93, 1, x_92); +return x_93; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_94 = lean_ctor_get(x_83, 0); +lean_inc(x_94); +lean_dec(x_83); +lean_inc(x_94); +x_95 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_94, x_5, x_6, x_7, x_8, x_9, x_84); +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_97 = x_95; +} else { + lean_dec_ref(x_95); + x_97 = lean_box(0); +} +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_94); +lean_ctor_set(x_98, 1, x_85); +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_97; +} +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_96); +return x_99; +} +} +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; +x_100 = lean_ctor_get(x_16, 0); +x_101 = lean_ctor_get(x_16, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_16); +x_102 = lean_array_get_size(x_100); +x_103 = lean_unsigned_to_nat(1u); +x_104 = lean_nat_add(x_102, x_103); +x_105 = lean_unsigned_to_nat(0u); +x_106 = lean_nat_mod(x_105, x_104); +lean_dec(x_104); +lean_inc(x_14); +lean_inc(x_100); +x_107 = lean_array_push(x_100, x_14); +x_108 = l_Array_insertAt_loop___rarg(x_100, x_106, x_107, x_102); +lean_dec(x_106); +lean_dec(x_100); +x_109 = lean_array_get_size(x_108); +x_110 = lean_box(0); +x_111 = lean_nat_dec_lt(x_105, x_109); +lean_inc(x_2); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_2); +lean_ctor_set(x_112, 1, x_110); +x_113 = lean_array_mk(x_112); +if (x_111 == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_109); +lean_dec(x_108); +x_114 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +x_117 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_3, x_13, x_14, x_114, x_110, x_116, x_5, x_6, x_7, x_8, x_9, x_101); +lean_dec(x_14); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_118, 1); +lean_inc(x_119); +x_120 = lean_ctor_get(x_117, 1); +lean_inc(x_120); +lean_dec(x_117); +x_121 = lean_ctor_get(x_118, 0); +lean_inc(x_121); +lean_dec(x_118); +x_122 = lean_ctor_get(x_119, 0); +lean_inc(x_122); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_123 = x_119; +} else { + lean_dec_ref(x_119); + x_123 = lean_box(0); +} +lean_inc(x_122); +x_124 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_122, x_5, x_6, x_7, x_8, x_9, x_120); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_126 = x_124; +} else { + lean_dec_ref(x_124); + x_126 = lean_box(0); +} +if (lean_is_scalar(x_123)) { + x_127 = lean_alloc_ctor(0, 2, 0); +} else { + x_127 = x_123; +} +lean_ctor_set(x_127, 0, x_122); +lean_ctor_set(x_127, 1, x_121); +if (lean_is_scalar(x_126)) { + x_128 = lean_alloc_ctor(0, 2, 0); +} else { + x_128 = x_126; +} +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_125); +return x_128; +} +else +{ +uint8_t x_129; +x_129 = lean_nat_dec_le(x_109, x_109); +if (x_129 == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_dec(x_109); +lean_dec(x_108); +x_130 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_113); +lean_ctor_set(x_131, 1, x_130); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +x_133 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_3, x_13, x_14, x_130, x_110, x_132, x_5, x_6, x_7, x_8, x_9, x_101); +lean_dec(x_14); +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +x_136 = lean_ctor_get(x_133, 1); +lean_inc(x_136); +lean_dec(x_133); +x_137 = lean_ctor_get(x_134, 0); +lean_inc(x_137); +lean_dec(x_134); +x_138 = lean_ctor_get(x_135, 0); +lean_inc(x_138); +if (lean_is_exclusive(x_135)) { + lean_ctor_release(x_135, 0); + lean_ctor_release(x_135, 1); + x_139 = x_135; +} else { + lean_dec_ref(x_135); + x_139 = lean_box(0); +} +lean_inc(x_138); +x_140 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_138, x_5, x_6, x_7, x_8, x_9, x_136); +x_141 = lean_ctor_get(x_140, 1); +lean_inc(x_141); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_142 = x_140; +} else { + lean_dec_ref(x_140); + x_142 = lean_box(0); +} +if (lean_is_scalar(x_139)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_139; +} +lean_ctor_set(x_143, 0, x_138); +lean_ctor_set(x_143, 1, x_137); +if (lean_is_scalar(x_142)) { + x_144 = lean_alloc_ctor(0, 2, 0); +} else { + x_144 = x_142; +} +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_141); +return x_144; +} +else +{ +size_t x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_145 = lean_usize_of_nat(x_109); +lean_dec(x_109); +x_146 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_147 = l_Array_foldlMUnsafe_fold___at_Lean_computeStructureResolutionOrder___spec__5(x_108, x_13, x_145, x_146); +lean_dec(x_108); +x_148 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_148, 0, x_113); +lean_ctor_set(x_148, 1, x_147); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_146); +lean_ctor_set(x_149, 1, x_148); +x_150 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_3, x_13, x_14, x_146, x_110, x_149, x_5, x_6, x_7, x_8, x_9, x_101); +lean_dec(x_14); +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_151, 1); +lean_inc(x_152); +x_153 = lean_ctor_get(x_150, 1); +lean_inc(x_153); +lean_dec(x_150); +x_154 = lean_ctor_get(x_151, 0); +lean_inc(x_154); +lean_dec(x_151); +x_155 = lean_ctor_get(x_152, 0); +lean_inc(x_155); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_156 = x_152; +} else { + lean_dec_ref(x_152); + x_156 = lean_box(0); +} +lean_inc(x_155); +x_157 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_2, x_155, x_5, x_6, x_7, x_8, x_9, x_153); +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_159 = x_157; +} else { + lean_dec_ref(x_157); + x_159 = lean_box(0); +} +if (lean_is_scalar(x_156)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_156; +} +lean_ctor_set(x_160, 0, x_155); +lean_ctor_set(x_160, 1, x_154); +if (lean_is_scalar(x_159)) { + x_161 = lean_alloc_ctor(0, 2, 0); +} else { + x_161 = x_159; +} +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_158); +return x_161; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(x_13, x_1); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_free_object(x_9); +x_15 = lean_box(0); +x_16 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_13, x_1, x_2, x_15, x_3, x_4, x_5, x_6, x_7, x_12); +lean_dec(x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_13); +lean_dec(x_1); +x_17 = lean_ctor_get(x_14, 0); +lean_inc(x_17); +lean_dec(x_14); +x_18 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_9, 0, x_19); +return x_9; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_9, 0); +x_21 = lean_ctor_get(x_9, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_9); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Structure_0__Lean_getStructureResolutionOrder_x3f(x_22, x_1); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_22, x_1, x_2, x_24, x_3, x_4, x_5, x_6, x_7, x_21); +lean_dec(x_22); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_22); +lean_dec(x_1); +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +lean_dec(x_23); +x_27 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_21); +return x_29; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; uint8_t x_10; +x_8 = 1; +x_9 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_1, x_8, x_2, x_3, x_4, x_5, x_6, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_9; +lean_inc(x_1); +x_8 = l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = l_Array_erase___at_Lean_getAllParentStructures___spec__1(x_10, x_1); +lean_dec(x_1); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +x_14 = l_Array_erase___at_Lean_getAllParentStructures___spec__1(x_12, x_1); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_5, x_4); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_6); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; +lean_dec(x_6); +x_15 = lean_array_uget(x_3, x_5); +x_16 = lean_st_ref_take(x_7, x_12); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_box(0); +x_20 = l_Lean_RBNode_insert___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(x_17, x_15, x_19); +x_21 = lean_st_ref_set(x_7, x_20, x_18); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = 1; +x_24 = lean_usize_add(x_5, x_23); +x_5 = x_24; +x_6 = x_19; +x_12 = x_22; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_unfoldeDefinitionGuarded_x3f(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set(x_9, 0, x_13); +return x_9; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_dec(x_9); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_dec(x_9); +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_18, x_3, x_4, x_5, x_6, x_7, x_17); +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_9); +if (x_20 == 0) +{ +return x_9; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_9, 0); +x_22 = lean_ctor_get(x_9, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_9); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_8) == 4) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_take(x_2, x_7); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_box(0); +lean_inc(x_9); +x_14 = l_Lean_RBNode_insert___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__1(x_11, x_9, x_13); +x_15 = lean_st_ref_set(x_2, x_14, x_12); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_st_ref_get(x_6, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +lean_inc(x_9); +x_21 = l_Lean_isStructure(x_20, x_9); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; +lean_dec(x_9); +x_22 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_19); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; size_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_23 = l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(x_9, x_2, x_3, x_4, x_5, x_6, x_19); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_box(0); +x_27 = lean_array_size(x_24); +x_28 = 0; +x_29 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(x_24, x_26, x_24, x_27, x_28, x_13, x_2, x_3, x_4, x_5, x_6, x_25); +lean_dec(x_24); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_13, x_2, x_3, x_4, x_5, x_6, x_30); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_7); +return x_33; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_11 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__6(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__8(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__9(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__10(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__11(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__12(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__13(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__14(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__15(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +lean_object* x_18 = _args[17]; +lean_object* x_19 = _args[18]; +_start: +{ +lean_object* x_20; +x_20 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_2); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +lean_object* x_18; +x_18 = l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_computeStructureResolutionOrder_selectParent___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__18(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__19(x_7, x_2, x_3, x_8, x_9, x_6); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__20(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__21(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__22(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__23(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; lean_object* x_16; +x_15 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_16 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___lambda__1(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__24(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__26(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_15 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__27(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_computeStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__5(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_getStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_getAllParentStructures___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_14 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_15 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__28(x_1, x_2, x_3, x_13, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_box(0); +x_8 = lean_st_mk_ref(x_7, x_6); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit(x_1, x_9, x_2, x_3, x_4, x_5, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_9, x_12); +lean_dec(x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +uint8_t x_18; +lean_dec(x_9); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +return x_11; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_searchAlias(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_1); +lean_inc(x_3); +lean_inc(x_5); +x_13 = lean_apply_2(x_1, x_5, x_3); +if (lean_obj_tag(x_13) == 0) +{ +if (lean_obj_tag(x_5) == 1) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_5, 0); +lean_inc(x_14); +lean_dec(x_5); +x_5 = x_14; +goto _start; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_12); +return x_17; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_5); +lean_dec(x_1); +x_18 = lean_ctor_get(x_13, 0); +lean_inc(x_18); +lean_dec(x_13); +x_19 = lean_apply_10(x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_2) == 1) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +lean_inc(x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_12 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_2 = x_11; +x_3 = x_15; +x_10 = x_14; +goto _start; +} +else +{ +uint8_t x_17; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) +{ +return x_12; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 0); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_12); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_3); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_10); +return x_23; +} +} +} +static lean_object* _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_3, x_4, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 0); +lean_dec(x_17); +x_18 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_6); +lean_ctor_set(x_14, 0, x_19); +return x_14; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +x_21 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_6); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_15); +if (x_24 == 0) +{ +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_14); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_14, 0); +lean_dec(x_27); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_15); +lean_ctor_set(x_14, 0, x_30); +return x_14; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_31 = lean_ctor_get(x_26, 0); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_26); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_15, 0, x_33); +x_34 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_15); +lean_ctor_set(x_14, 0, x_35); +return x_14; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_36 = lean_ctor_get(x_15, 0); +x_37 = lean_ctor_get(x_14, 1); +lean_inc(x_37); +lean_dec(x_14); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_36)) { + lean_ctor_release(x_36, 0); + lean_ctor_release(x_36, 1); + x_40 = x_36; +} else { + lean_dec_ref(x_36); + x_40 = lean_box(0); +} +if (lean_is_scalar(x_40)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_40; +} +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_15, 0, x_41); +x_42 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_15); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_37); +return x_44; +} +} +else +{ +lean_object* x_45; uint8_t x_46; +x_45 = lean_ctor_get(x_15, 0); +lean_free_object(x_15); +x_46 = !lean_is_exclusive(x_6); +if (x_46 == 0) +{ +uint8_t x_47; +x_47 = !lean_is_exclusive(x_14); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = lean_ctor_get(x_6, 0); +x_49 = lean_ctor_get(x_14, 0); +lean_dec(x_49); +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_48); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_45, 0); +x_53 = lean_ctor_get(x_45, 1); +x_54 = lean_ctor_get(x_48, 0); +x_55 = lean_ctor_get(x_48, 1); +x_56 = l_Lean_Name_isSuffixOf(x_52, x_54); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_53); +lean_dec(x_52); +x_57 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +lean_ctor_set(x_45, 1, x_6); +lean_ctor_set(x_45, 0, x_57); +lean_ctor_set(x_14, 0, x_45); +return x_14; +} +else +{ +lean_object* x_58; +lean_dec(x_55); +lean_dec(x_54); +lean_ctor_set(x_48, 1, x_53); +lean_ctor_set(x_48, 0, x_52); +x_58 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +lean_ctor_set(x_45, 1, x_6); +lean_ctor_set(x_45, 0, x_58); +lean_ctor_set(x_14, 0, x_45); +return x_14; +} +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_59 = lean_ctor_get(x_45, 0); +x_60 = lean_ctor_get(x_45, 1); +x_61 = lean_ctor_get(x_48, 0); +x_62 = lean_ctor_get(x_48, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_48); +x_63 = l_Lean_Name_isSuffixOf(x_59, x_61); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_60); +lean_dec(x_59); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_62); +lean_ctor_set(x_6, 0, x_64); +x_65 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +lean_ctor_set(x_45, 1, x_6); +lean_ctor_set(x_45, 0, x_65); +lean_ctor_set(x_14, 0, x_45); +return x_14; +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +lean_dec(x_61); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_59); +lean_ctor_set(x_66, 1, x_60); +lean_ctor_set(x_6, 0, x_66); +x_67 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +lean_ctor_set(x_45, 1, x_6); +lean_ctor_set(x_45, 0, x_67); +lean_ctor_set(x_14, 0, x_45); +return x_14; +} +} +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_68 = lean_ctor_get(x_45, 0); +x_69 = lean_ctor_get(x_45, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_45); +x_70 = lean_ctor_get(x_48, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_48, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_72 = x_48; +} else { + lean_dec_ref(x_48); + x_72 = lean_box(0); +} +x_73 = l_Lean_Name_isSuffixOf(x_68, x_70); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_69); +lean_dec(x_68); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_72; +} +lean_ctor_set(x_74, 0, x_70); +lean_ctor_set(x_74, 1, x_71); +lean_ctor_set(x_6, 0, x_74); +x_75 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_6); +lean_ctor_set(x_14, 0, x_76); +return x_14; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_71); +lean_dec(x_70); +if (lean_is_scalar(x_72)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_72; +} +lean_ctor_set(x_77, 0, x_68); +lean_ctor_set(x_77, 1, x_69); +lean_ctor_set(x_6, 0, x_77); +x_78 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_6); +lean_ctor_set(x_14, 0, x_79); +return x_14; +} +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_80 = lean_ctor_get(x_6, 0); +x_81 = lean_ctor_get(x_14, 1); +lean_inc(x_81); +lean_dec(x_14); +x_82 = lean_ctor_get(x_45, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_45, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_84 = x_45; +} else { + lean_dec_ref(x_45); + x_84 = lean_box(0); +} +x_85 = lean_ctor_get(x_80, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_80, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_87 = x_80; +} else { + lean_dec_ref(x_80); + x_87 = lean_box(0); +} +x_88 = l_Lean_Name_isSuffixOf(x_82, x_85); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_83); +lean_dec(x_82); +if (lean_is_scalar(x_87)) { + x_89 = lean_alloc_ctor(0, 2, 0); +} else { + x_89 = x_87; +} +lean_ctor_set(x_89, 0, x_85); +lean_ctor_set(x_89, 1, x_86); +lean_ctor_set(x_6, 0, x_89); +x_90 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_84)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_84; +} +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_6); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_81); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_dec(x_86); +lean_dec(x_85); +if (lean_is_scalar(x_87)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_87; +} +lean_ctor_set(x_93, 0, x_82); +lean_ctor_set(x_93, 1, x_83); +lean_ctor_set(x_6, 0, x_93); +x_94 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_84)) { + x_95 = lean_alloc_ctor(0, 2, 0); +} else { + x_95 = x_84; +} +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_6); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_81); +return x_96; +} +} +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_97 = lean_ctor_get(x_6, 0); +lean_inc(x_97); +lean_dec(x_6); +x_98 = lean_ctor_get(x_14, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + x_99 = x_14; +} else { + lean_dec_ref(x_14); + x_99 = lean_box(0); +} +x_100 = lean_ctor_get(x_45, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_45, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_102 = x_45; +} else { + lean_dec_ref(x_45); + x_102 = lean_box(0); +} +x_103 = lean_ctor_get(x_97, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_97, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_97)) { + lean_ctor_release(x_97, 0); + lean_ctor_release(x_97, 1); + x_105 = x_97; +} else { + lean_dec_ref(x_97); + x_105 = lean_box(0); +} +x_106 = l_Lean_Name_isSuffixOf(x_100, x_103); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_101); +lean_dec(x_100); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_103); +lean_ctor_set(x_107, 1, x_104); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +x_109 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_102)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_102; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_108); +if (lean_is_scalar(x_99)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_99; +} +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_98); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_104); +lean_dec(x_103); +if (lean_is_scalar(x_105)) { + x_112 = lean_alloc_ctor(0, 2, 0); +} else { + x_112 = x_105; +} +lean_ctor_set(x_112, 0, x_100); +lean_ctor_set(x_112, 1, x_101); +x_113 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_113, 0, x_112); +x_114 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_102)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_102; +} +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_113); +if (lean_is_scalar(x_99)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_99; +} +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_98); +return x_116; +} +} +} +} +else +{ +lean_object* x_117; +x_117 = lean_ctor_get(x_15, 0); +lean_inc(x_117); +lean_dec(x_15); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_118 = lean_ctor_get(x_14, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + x_119 = x_14; +} else { + lean_dec_ref(x_14); + x_119 = lean_box(0); +} +x_120 = lean_ctor_get(x_117, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_117, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_122 = x_117; +} else { + lean_dec_ref(x_117); + x_122 = lean_box(0); +} +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_122; +} +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +x_124 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_124, 0, x_123); +x_125 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_124); +if (lean_is_scalar(x_119)) { + x_127 = lean_alloc_ctor(0, 2, 0); +} else { + x_127 = x_119; +} +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_118); +return x_127; +} +else +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; +x_128 = lean_ctor_get(x_6, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + x_129 = x_6; +} else { + lean_dec_ref(x_6); + x_129 = lean_box(0); +} +x_130 = lean_ctor_get(x_14, 1); +lean_inc(x_130); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + x_131 = x_14; +} else { + lean_dec_ref(x_14); + x_131 = lean_box(0); +} +x_132 = lean_ctor_get(x_117, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_117, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_134 = x_117; +} else { + lean_dec_ref(x_117); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_128, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_128, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_128)) { + lean_ctor_release(x_128, 0); + lean_ctor_release(x_128, 1); + x_137 = x_128; +} else { + lean_dec_ref(x_128); + x_137 = lean_box(0); +} +x_138 = l_Lean_Name_isSuffixOf(x_132, x_135); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +lean_dec(x_133); +lean_dec(x_132); +if (lean_is_scalar(x_137)) { + x_139 = lean_alloc_ctor(0, 2, 0); +} else { + x_139 = x_137; +} +lean_ctor_set(x_139, 0, x_135); +lean_ctor_set(x_139, 1, x_136); +if (lean_is_scalar(x_129)) { + x_140 = lean_alloc_ctor(1, 1, 0); +} else { + x_140 = x_129; +} +lean_ctor_set(x_140, 0, x_139); +x_141 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_134)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_134; +} +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_140); +if (lean_is_scalar(x_131)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_131; +} +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_130); +return x_143; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_dec(x_136); +lean_dec(x_135); +if (lean_is_scalar(x_137)) { + x_144 = lean_alloc_ctor(0, 2, 0); +} else { + x_144 = x_137; +} +lean_ctor_set(x_144, 0, x_132); +lean_ctor_set(x_144, 1, x_133); +if (lean_is_scalar(x_129)) { + x_145 = lean_alloc_ctor(1, 1, 0); +} else { + x_145 = x_129; +} +lean_ctor_set(x_145, 0, x_144); +x_146 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +if (lean_is_scalar(x_134)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_134; +} +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_145); +if (lean_is_scalar(x_131)) { + x_148 = lean_alloc_ctor(0, 2, 0); +} else { + x_148 = x_131; +} +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_130); +return x_148; +} +} +} +} +} +else +{ +uint8_t x_149; +lean_dec(x_6); +x_149 = !lean_is_exclusive(x_14); +if (x_149 == 0) +{ +return x_14; +} +else +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_14, 0); +x_151 = lean_ctor_get(x_14, 1); +lean_inc(x_151); +lean_inc(x_150); +lean_dec(x_14); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +return x_152; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_3); +lean_dec(x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_8); +lean_ctor_set(x_18, 1, x_10); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +else +{ +lean_object* x_20; +lean_dec(x_8); +x_20 = lean_ctor_get(x_7, 0); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_7, 1); +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = l_List_elem___at_Lean_addAliasEntry___spec__16(x_3, x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_box(0); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_3); +lean_inc(x_2); +x_26 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(x_22, x_2, x_1, x_3, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +if (lean_obj_tag(x_28) == 0) +{ +uint8_t x_29; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_3); +lean_dec(x_2); +x_29 = !lean_is_exclusive(x_26); +if (x_29 == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_ctor_get(x_26, 0); +lean_dec(x_30); +x_31 = !lean_is_exclusive(x_27); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_27, 0); +lean_dec(x_32); +x_33 = lean_ctor_get(x_28, 0); +lean_inc(x_33); +lean_dec(x_28); +lean_ctor_set(x_27, 0, x_33); +return x_26; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_dec(x_27); +x_35 = lean_ctor_get(x_28, 0); +lean_inc(x_35); +lean_dec(x_28); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +lean_ctor_set(x_26, 0, x_36); +return x_26; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = lean_ctor_get(x_26, 1); +lean_inc(x_37); +lean_dec(x_26); +x_38 = lean_ctor_get(x_27, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_39 = x_27; +} else { + lean_dec_ref(x_27); + x_39 = lean_box(0); +} +x_40 = lean_ctor_get(x_28, 0); +lean_inc(x_40); +lean_dec(x_28); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_39; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_37); +return x_42; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_26, 1); +lean_inc(x_43); +lean_dec(x_26); +x_44 = lean_ctor_get(x_27, 1); +lean_inc(x_44); +lean_dec(x_27); +x_45 = lean_ctor_get(x_28, 0); +lean_inc(x_45); +lean_dec(x_28); +x_7 = x_21; +x_8 = x_45; +x_9 = lean_box(0); +x_10 = x_44; +x_17 = x_43; +goto _start; +} +} +else +{ +uint8_t x_47; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_26); +if (x_47 == 0) +{ +return x_26; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_26, 0); +x_49 = lean_ctor_get(x_26, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_26); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +else +{ +lean_object* x_51; +x_51 = lean_box(0); +x_7 = x_21; +x_8 = x_51; +x_9 = lean_box(0); +goto _start; +} +} +else +{ +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_7, 1); +x_54 = lean_box(0); +x_7 = x_53; +x_8 = x_54; +x_9 = lean_box(0); +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_11; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 2); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = lean_apply_9(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_2 = x_16; +x_3 = x_14; +x_10 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = lean_box(0); +x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +x_14 = lean_array_uget(x_2, x_3); +x_15 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__3(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +x_5 = x_17; +x_12 = x_18; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_12); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +lean_inc(x_3); +x_12 = l_Lean_Meta_allowCompletion(x_1, x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_11); +return x_14; +} +else +{ +lean_object* x_15; +x_15 = lean_apply_9(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_12); +x_13 = l_Lean_Server_Completion_getEligibleHeaderDecls(x_12, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_array_get_size(x_16); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_lt(x_18, x_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_17); +lean_dec(x_16); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); +lean_closure_set(x_22, 0, x_12); +lean_closure_set(x_22, 1, x_1); +x_23 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(x_21, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_21); +return x_23; +} +else +{ +uint8_t x_24; +x_24 = lean_nat_dec_le(x_17, x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_17); +lean_dec(x_16); +x_25 = lean_ctor_get(x_12, 1); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); +lean_closure_set(x_27, 0, x_12); +lean_closure_set(x_27, 1, x_1); +x_28 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(x_26, x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_26); +return x_28; +} +else +{ +size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; +x_29 = 0; +x_30 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_31 = lean_box(0); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_32 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6(x_1, x_16, x_29, x_30, x_31, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +lean_dec(x_16); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_ctor_get(x_12, 1); +lean_inc(x_34); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2___lambda__1), 11, 2); +lean_closure_set(x_36, 0, x_12); +lean_closure_set(x_36, 1, x_1); +x_37 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(x_35, x_36, x_2, x_3, x_4, x_5, x_6, x_7, x_33); +lean_dec(x_35); +return x_37; +} +else +{ +uint8_t x_38; +lean_dec(x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = !lean_is_exclusive(x_32); +if (x_38 == 0) +{ +return x_32; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_32, 0); +x_40 = lean_ctor_get(x_32, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_32); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_5, 1); +lean_inc(x_16); +lean_dec(x_5); +lean_inc(x_15); +lean_inc(x_1); +x_17 = l_Lean_Server_Completion_allowCompletion(x_2, x_1, x_15); +if (x_17 == 0) +{ +lean_dec(x_15); +x_5 = x_16; +goto _start; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = l_Lean_Name_getString_x21(x_3); +x_20 = lean_box(0); +x_21 = l_Lean_Name_str___override(x_20, x_19); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_22 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_21, x_15, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_5 = x_16; +x_12 = x_23; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, double x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_5, 1); +lean_inc(x_16); +lean_dec(x_5); +lean_inc(x_15); +lean_inc(x_1); +x_17 = l_Lean_Server_Completion_allowCompletion(x_2, x_1, x_15); +if (x_17 == 0) +{ +lean_dec(x_15); +x_5 = x_16; +goto _start; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = l_Lean_Name_getString_x21(x_3); +x_20 = lean_box(0); +x_21 = l_Lean_Name_str___override(x_20, x_19); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_22 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_21, x_15, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_5 = x_16; +x_12 = x_23; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_11; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 2); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = lean_apply_9(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_2 = x_16; +x_3 = x_14; +x_10 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = lean_array_uget(x_2, x_3); +switch (lean_obj_tag(x_14)) { +case 0: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_17 = lean_apply_10(x_1, x_5, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_3 = x_21; +x_5 = x_18; +x_12 = x_19; +goto _start; +} +else +{ +uint8_t x_23; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +case 1: +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_14, 0); +lean_inc(x_27); +lean_dec(x_14); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_28 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_1, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; size_t x_31; size_t x_32; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = 1; +x_32 = lean_usize_add(x_3, x_31); +x_3 = x_32; +x_5 = x_29; +x_12 = x_30; +goto _start; +} +else +{ +uint8_t x_34; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) +{ +return x_28; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_28, 0); +x_36 = lean_ctor_get(x_28, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_28); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +default: +{ +size_t x_38; size_t x_39; +x_38 = 1; +x_39 = lean_usize_add(x_3, x_38); +x_3 = x_39; +goto _start; +} +} +} +else +{ +lean_object* x_41; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_5); +lean_ctor_set(x_41, 1, x_12); +return x_41; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed), 12, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_array_get_size(x_2); +x_15 = lean_nat_dec_lt(x_5, x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_6); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_array_fget(x_2, x_5); +x_18 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_19 = lean_apply_10(x_1, x_6, x_17, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_5, x_22); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_23; +x_6 = x_20; +x_13 = x_21; +goto _start; +} +else +{ +uint8_t x_25; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_19); +if (x_25 == 0) +{ +return x_19; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_19, 0); +x_27 = lean_ctor_get(x_19, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_19); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed), 13, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = lean_nat_dec_le(x_12, x_12); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_3); +lean_ctor_set(x_17, 1, x_10); +return x_17; +} +else +{ +size_t x_18; size_t x_19; lean_object* x_20; +x_18 = 0; +x_19 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_20 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(x_1, x_11, x_18, x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_20; +} +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_2, 0); +x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_unsigned_to_nat(0u); +x_24 = l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(x_1, x_21, x_22, lean_box(0), x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg___boxed), 10, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = lean_box(0); +x_12 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +x_14 = lean_array_uget(x_2, x_3); +x_15 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__10(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +x_5 = x_17; +x_12 = x_18; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_12); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_10, 1); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_15 = lean_ctor_get(x_1, 1); +x_16 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_16; +} +else +{ +uint8_t x_17; +x_17 = lean_nat_dec_le(x_12, x_12); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_12); +x_18 = lean_ctor_get(x_1, 1); +x_19 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_19; +} +else +{ +size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_20 = 0; +x_21 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_22 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_23 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_ctor_get(x_1, 1); +x_26 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = l_Lean_Name_isPrefixOf(x_1, x_5); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_13); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_box(0); +lean_inc(x_5); +x_18 = l_Lean_Name_replacePrefix(x_5, x_1, x_17); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_2, x_18); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_13); +return x_21; +} +else +{ +lean_object* x_22; double x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_unbox_float(x_22); +lean_dec(x_22); +x_24 = l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8(x_3, x_4, x_5, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_5); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +else +{ +lean_object* x_18; +lean_dec(x_8); +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = lean_ctor_get(x_7, 1); +lean_inc(x_19); +lean_dec(x_7); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_getAliasState(x_2); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_22 = lean_alloc_closure((void*)(l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed), 13, 4); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_1); +lean_closure_set(x_22, 2, x_2); +lean_closure_set(x_22, 3, x_3); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_23 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9(x_21, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_box(0); +x_7 = x_19; +x_8 = x_25; +x_9 = lean_box(0); +x_16 = x_24; +goto _start; +} +else +{ +uint8_t x_27; +lean_dec(x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_7, 1); +lean_inc(x_31); +lean_dec(x_7); +x_32 = lean_ctor_get(x_18, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_18, 1); +lean_inc(x_33); +lean_dec(x_18); +lean_inc(x_33); +lean_inc(x_2); +x_34 = l_Lean_Server_Completion_allowCompletion(x_3, x_2, x_33); +if (x_34 == 0) +{ +lean_object* x_35; +lean_dec(x_33); +lean_dec(x_32); +x_35 = lean_box(0); +x_7 = x_31; +x_8 = x_35; +x_9 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_37; +lean_inc(x_1); +x_37 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_1, x_32); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; +lean_dec(x_33); +lean_dec(x_32); +x_38 = lean_box(0); +x_7 = x_31; +x_8 = x_38; +x_9 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; double x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_37, 0); +lean_inc(x_40); +lean_dec(x_37); +x_41 = l_Lean_Name_getString_x21(x_32); +lean_dec(x_32); +x_42 = lean_box(0); +x_43 = l_Lean_Name_str___override(x_42, x_41); +x_44 = lean_unbox_float(x_40); +lean_dec(x_40); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_45 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItemForDecl(x_43, x_33, x_44, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +x_47 = lean_box(0); +x_7 = x_31; +x_8 = x_47; +x_9 = lean_box(0); +x_16 = x_46; +goto _start; +} +else +{ +uint8_t x_49; +lean_dec(x_31); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_45); +if (x_49 == 0) +{ +return x_45; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_45, 0); +x_51 = lean_ctor_get(x_45, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_45); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_11; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_2); +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 2); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_1); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_15 = lean_apply_9(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_2 = x_16; +x_3 = x_14; +x_10 = x_17; +goto _start; +} +else +{ +uint8_t x_19; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_15); +if (x_19 == 0) +{ +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 0); +x_21 = lean_ctor_get(x_15, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_15); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_2, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_completeNamespaces___spec__4___lambda__1___boxed), 11, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = lean_box(0); +x_12 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Server_Completion_completeNamespaces___spec__6___rarg(x_10, x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_5); +x_14 = lean_array_uget(x_2, x_3); +x_15 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_16 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__19(x_1, x_15, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +x_5 = x_17; +x_12 = x_18; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +else +{ +lean_object* x_26; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_5); +lean_ctor_set(x_26, 1, x_12); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_10, 1); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_15 = lean_ctor_get(x_1, 1); +x_16 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_16; +} +else +{ +uint8_t x_17; +x_17 = lean_nat_dec_le(x_12, x_12); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_12); +x_18 = lean_ctor_get(x_1, 1); +x_19 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_19; +} +else +{ +size_t x_20; size_t x_21; lean_object* x_22; lean_object* x_23; +x_20 = 0; +x_21 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_22 = lean_box(0); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_23 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22(x_2, x_11, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_ctor_get(x_1, 1); +x_26 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_27 = !lean_is_exclusive(x_23); +if (x_27 == 0) +{ +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 0); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_23); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_6, x_5); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_1); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; double x_18; lean_object* x_19; +lean_dec(x_7); +x_17 = lean_array_uget(x_4, x_6); +x_18 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +lean_inc(x_1); +x_19 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_17, x_18); +if (lean_obj_tag(x_19) == 0) +{ +size_t x_20; size_t x_21; lean_object* x_22; +lean_dec(x_17); +x_20 = 1; +x_21 = lean_usize_add(x_6, x_20); +x_22 = lean_box(0); +x_6 = x_21; +x_7 = x_22; +goto _start; +} +else +{ +lean_object* x_24; double x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = lean_unbox_float(x_24); +lean_dec(x_24); +x_26 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem(x_17, x_25, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = 1; +x_29 = lean_usize_add(x_6, x_28); +x_30 = lean_box(0); +x_6 = x_29; +x_7 = x_30; +x_14 = x_27; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; +x_17 = lean_usize_dec_lt(x_8, x_7); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_5); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_array_uget(x_6, x_8); +x_20 = !lean_is_exclusive(x_9); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_9, 1); +x_22 = lean_ctor_get(x_9, 0); +lean_dec(x_22); +lean_inc(x_21); +lean_inc(x_1); +x_23 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_19, x_21, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_19); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +lean_dec(x_5); +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_9, 0, x_27); +lean_ctor_set(x_23, 0, x_9); +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 1); +lean_inc(x_28); +lean_dec(x_23); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_9, 0, x_29); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_9); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; +lean_dec(x_21); +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_dec(x_23); +x_32 = lean_ctor_get(x_24, 0); +lean_inc(x_32); +lean_dec(x_24); +lean_inc(x_5); +lean_ctor_set(x_9, 1, x_32); +lean_ctor_set(x_9, 0, x_5); +x_33 = 1; +x_34 = lean_usize_add(x_8, x_33); +x_8 = x_34; +x_16 = x_31; +goto _start; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_9, 1); +lean_inc(x_36); +lean_dec(x_9); +lean_inc(x_36); +lean_inc(x_1); +x_37 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_19, x_36, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_19); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_5); +lean_dec(x_1); +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_40 = x_37; +} else { + lean_dec_ref(x_37); + x_40 = lean_box(0); +} +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_36); +if (lean_is_scalar(x_40)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_40; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_39); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; +lean_dec(x_36); +x_44 = lean_ctor_get(x_37, 1); +lean_inc(x_44); +lean_dec(x_37); +x_45 = lean_ctor_get(x_38, 0); +lean_inc(x_45); +lean_dec(x_38); +lean_inc(x_5); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_5); +lean_ctor_set(x_46, 1, x_45); +x_47 = 1; +x_48 = lean_usize_add(x_8, x_47); +x_8 = x_48; +x_9 = x_46; +x_16 = x_44; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_7, x_6); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_4); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_array_uget(x_5, x_7); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + x_20 = x_8; +} else { + lean_dec_ref(x_8); + x_20 = lean_box(0); +} +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_29; +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_19); +x_21 = x_29; +x_22 = x_15; +goto block_28; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_19); +x_30 = lean_ctor_get(x_18, 0); +lean_inc(x_30); +lean_dec(x_18); +x_31 = l_Lean_LocalDecl_userName(x_30); +lean_inc(x_1); +x_32 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_1, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_dec(x_31); +lean_dec(x_30); +x_33 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_33; +x_22 = x_15; +goto block_28; +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_32); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; double x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_35 = lean_ctor_get(x_32, 0); +x_36 = l_Lean_LocalDecl_fvarId(x_30); +lean_dec(x_30); +lean_ctor_set(x_32, 0, x_36); +x_37 = 5; +x_38 = lean_unbox_float(x_35); +lean_dec(x_35); +x_39 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_31, x_32, x_37, x_38, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_41; +x_22 = x_40; +goto block_28; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; double x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = lean_ctor_get(x_32, 0); +lean_inc(x_42); +lean_dec(x_32); +x_43 = l_Lean_LocalDecl_fvarId(x_30); +lean_dec(x_30); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = 5; +x_46 = lean_unbox_float(x_42); +lean_dec(x_42); +x_47 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_31, x_44, x_45, x_46, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_49; +x_22 = x_48; +goto block_28; +} +} +} +block_28: +{ +lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_4); +if (lean_is_scalar(x_20)) { + x_24 = lean_alloc_ctor(0, 2, 0); +} else { + x_24 = x_20; +} +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +x_25 = 1; +x_26 = lean_usize_add(x_7, x_25); +x_7 = x_26; +x_8 = x_24; +x_15 = x_22; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_box(0); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_4); +x_16 = lean_array_size(x_12); +x_17 = 0; +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26(x_1, x_2, x_12, x_13, x_14, x_12, x_16, x_17, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_box(0); +x_24 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_22, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +return x_24; +} +else +{ +uint8_t x_25; +lean_dec(x_19); +x_25 = !lean_is_exclusive(x_18); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_18, 0); +lean_dec(x_26); +x_27 = lean_ctor_get(x_20, 0); +lean_inc(x_27); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_27); +return x_18; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_18, 1); +lean_inc(x_28); +lean_dec(x_18); +x_29 = lean_ctor_get(x_20, 0); +lean_inc(x_29); +lean_dec(x_20); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; size_t x_35; size_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_31 = lean_ctor_get(x_3, 0); +x_32 = lean_box(0); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_4); +x_35 = lean_array_size(x_31); +x_36 = 0; +x_37 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27(x_1, x_31, x_32, x_33, x_31, x_35, x_36, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +x_42 = lean_box(0); +x_43 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_41, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_40); +return x_43; +} +else +{ +uint8_t x_44; +lean_dec(x_38); +x_44 = !lean_is_exclusive(x_37); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_37, 0); +lean_dec(x_45); +x_46 = lean_ctor_get(x_39, 0); +lean_inc(x_46); +lean_dec(x_39); +lean_ctor_set(x_37, 0, x_46); +return x_37; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_37, 1); +lean_inc(x_47); +lean_dec(x_37); +x_48 = lean_ctor_get(x_39, 0); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +return x_49; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; +x_16 = lean_usize_dec_lt(x_7, x_6); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_4); +lean_dec(x_1); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_array_uget(x_5, x_7); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + x_20 = x_8; +} else { + lean_dec_ref(x_8); + x_20 = lean_box(0); +} +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_29; +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_19); +x_21 = x_29; +x_22 = x_15; +goto block_28; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_19); +x_30 = lean_ctor_get(x_18, 0); +lean_inc(x_30); +lean_dec(x_18); +x_31 = l_Lean_LocalDecl_userName(x_30); +lean_inc(x_1); +x_32 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_1, x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_dec(x_31); +lean_dec(x_30); +x_33 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_33; +x_22 = x_15; +goto block_28; +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_32); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; double x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_35 = lean_ctor_get(x_32, 0); +x_36 = l_Lean_LocalDecl_fvarId(x_30); +lean_dec(x_30); +lean_ctor_set(x_32, 0, x_36); +x_37 = 5; +x_38 = lean_unbox_float(x_35); +lean_dec(x_35); +x_39 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_31, x_32, x_37, x_38, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_41; +x_22 = x_40; +goto block_28; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; double x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = lean_ctor_get(x_32, 0); +lean_inc(x_42); +lean_dec(x_32); +x_43 = l_Lean_LocalDecl_fvarId(x_30); +lean_dec(x_30); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = 5; +x_46 = lean_unbox_float(x_42); +lean_dec(x_42); +x_47 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_31, x_44, x_45, x_46, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1; +x_21 = x_49; +x_22 = x_48; +goto block_28; +} +} +} +block_28: +{ +lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_4); +if (lean_is_scalar(x_20)) { + x_24 = lean_alloc_ctor(0, 2, 0); +} else { + x_24 = x_20; +} +lean_ctor_set(x_24, 0, x_4); +lean_ctor_set(x_24, 1, x_23); +x_25 = 1; +x_26 = lean_usize_add(x_7, x_25); +x_7 = x_26; +x_8 = x_24; +x_15 = x_22; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +lean_inc(x_1); +x_12 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_3, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_3); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); +lean_dec(x_13); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_ctor_get(x_13, 0); +lean_inc(x_18); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_box(0); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = lean_array_size(x_22); +x_27 = 0; +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28(x_1, x_22, x_23, x_24, x_22, x_26, x_27, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_28, 0); +lean_dec(x_32); +x_33 = lean_ctor_get(x_29, 1); +lean_inc(x_33); +lean_dec(x_29); +lean_ctor_set(x_28, 0, x_33); +return x_28; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +x_35 = lean_ctor_get(x_29, 1); +lean_inc(x_35); +lean_dec(x_29); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +else +{ +uint8_t x_37; +lean_dec(x_29); +x_37 = !lean_is_exclusive(x_28); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_28, 0); +lean_dec(x_38); +x_39 = lean_ctor_get(x_30, 0); +lean_inc(x_39); +lean_dec(x_30); +lean_ctor_set(x_28, 0, x_39); +return x_28; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_28, 1); +lean_inc(x_40); +lean_dec(x_28); +x_41 = lean_ctor_get(x_30, 0); +lean_inc(x_41); +lean_dec(x_30); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(x_4, x_1, x_2, x_3, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_13, 0); +lean_dec(x_16); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_5); +lean_ctor_set(x_13, 0, x_18); +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_13, 1); +lean_inc(x_19); +lean_dec(x_13); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_5); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_14); +if (x_23 == 0) +{ +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_13, 0); +lean_dec(x_26); +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_14); +lean_ctor_set(x_13, 0, x_29); +return x_13; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_30 = lean_ctor_get(x_25, 0); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_25); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_14, 0, x_32); +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_14); +lean_ctor_set(x_13, 0, x_34); +return x_13; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_35 = lean_ctor_get(x_14, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_35)) { + lean_ctor_release(x_35, 0); + lean_ctor_release(x_35, 1); + x_39 = x_35; +} else { + lean_dec_ref(x_35); + x_39 = lean_box(0); +} +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_39; +} +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_38); +lean_ctor_set(x_14, 0, x_40); +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_14); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_36); +return x_43; +} +} +else +{ +lean_object* x_44; uint8_t x_45; +x_44 = lean_ctor_get(x_14, 0); +lean_free_object(x_14); +x_45 = !lean_is_exclusive(x_5); +if (x_45 == 0) +{ +uint8_t x_46; +x_46 = !lean_is_exclusive(x_13); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = lean_ctor_get(x_5, 0); +x_48 = lean_ctor_get(x_13, 0); +lean_dec(x_48); +x_49 = !lean_is_exclusive(x_44); +if (x_49 == 0) +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_47); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_51 = lean_ctor_get(x_44, 0); +x_52 = lean_ctor_get(x_44, 1); +x_53 = lean_ctor_get(x_47, 0); +x_54 = lean_ctor_get(x_47, 1); +x_55 = l_Lean_Name_isSuffixOf(x_51, x_53); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_52); +lean_dec(x_51); +x_56 = lean_box(0); +lean_ctor_set(x_44, 1, x_5); +lean_ctor_set(x_44, 0, x_56); +lean_ctor_set(x_13, 0, x_44); +return x_13; +} +else +{ +lean_object* x_57; +lean_dec(x_54); +lean_dec(x_53); +lean_ctor_set(x_47, 1, x_52); +lean_ctor_set(x_47, 0, x_51); +x_57 = lean_box(0); +lean_ctor_set(x_44, 1, x_5); +lean_ctor_set(x_44, 0, x_57); +lean_ctor_set(x_13, 0, x_44); +return x_13; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_58 = lean_ctor_get(x_44, 0); +x_59 = lean_ctor_get(x_44, 1); +x_60 = lean_ctor_get(x_47, 0); +x_61 = lean_ctor_get(x_47, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_47); +x_62 = l_Lean_Name_isSuffixOf(x_58, x_60); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_59); +lean_dec(x_58); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_61); +lean_ctor_set(x_5, 0, x_63); +x_64 = lean_box(0); +lean_ctor_set(x_44, 1, x_5); +lean_ctor_set(x_44, 0, x_64); +lean_ctor_set(x_13, 0, x_44); +return x_13; +} +else +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_61); +lean_dec(x_60); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_58); +lean_ctor_set(x_65, 1, x_59); +lean_ctor_set(x_5, 0, x_65); +x_66 = lean_box(0); +lean_ctor_set(x_44, 1, x_5); +lean_ctor_set(x_44, 0, x_66); +lean_ctor_set(x_13, 0, x_44); +return x_13; +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_67 = lean_ctor_get(x_44, 0); +x_68 = lean_ctor_get(x_44, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_44); +x_69 = lean_ctor_get(x_47, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_47, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_71 = x_47; +} else { + lean_dec_ref(x_47); + x_71 = lean_box(0); +} +x_72 = l_Lean_Name_isSuffixOf(x_67, x_69); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_68); +lean_dec(x_67); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 2, 0); +} else { + x_73 = x_71; +} +lean_ctor_set(x_73, 0, x_69); +lean_ctor_set(x_73, 1, x_70); +lean_ctor_set(x_5, 0, x_73); +x_74 = lean_box(0); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_5); +lean_ctor_set(x_13, 0, x_75); +return x_13; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_70); +lean_dec(x_69); +if (lean_is_scalar(x_71)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_71; +} +lean_ctor_set(x_76, 0, x_67); +lean_ctor_set(x_76, 1, x_68); +lean_ctor_set(x_5, 0, x_76); +x_77 = lean_box(0); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_5); +lean_ctor_set(x_13, 0, x_78); +return x_13; +} +} +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_79 = lean_ctor_get(x_5, 0); +x_80 = lean_ctor_get(x_13, 1); +lean_inc(x_80); +lean_dec(x_13); +x_81 = lean_ctor_get(x_44, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_44, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_83 = x_44; +} else { + lean_dec_ref(x_44); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_79, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_86 = x_79; +} else { + lean_dec_ref(x_79); + x_86 = lean_box(0); +} +x_87 = l_Lean_Name_isSuffixOf(x_81, x_84); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_82); +lean_dec(x_81); +if (lean_is_scalar(x_86)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_86; +} +lean_ctor_set(x_88, 0, x_84); +lean_ctor_set(x_88, 1, x_85); +lean_ctor_set(x_5, 0, x_88); +x_89 = lean_box(0); +if (lean_is_scalar(x_83)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_83; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_5); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_80); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_85); +lean_dec(x_84); +if (lean_is_scalar(x_86)) { + x_92 = lean_alloc_ctor(0, 2, 0); +} else { + x_92 = x_86; +} +lean_ctor_set(x_92, 0, x_81); +lean_ctor_set(x_92, 1, x_82); +lean_ctor_set(x_5, 0, x_92); +x_93 = lean_box(0); +if (lean_is_scalar(x_83)) { + x_94 = lean_alloc_ctor(0, 2, 0); +} else { + x_94 = x_83; +} +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_5); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_80); +return x_95; +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_96 = lean_ctor_get(x_5, 0); +lean_inc(x_96); +lean_dec(x_5); +x_97 = lean_ctor_get(x_13, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_98 = x_13; +} else { + lean_dec_ref(x_13); + x_98 = lean_box(0); +} +x_99 = lean_ctor_get(x_44, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_44, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_101 = x_44; +} else { + lean_dec_ref(x_44); + x_101 = lean_box(0); +} +x_102 = lean_ctor_get(x_96, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_96, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_104 = x_96; +} else { + lean_dec_ref(x_96); + x_104 = lean_box(0); +} +x_105 = l_Lean_Name_isSuffixOf(x_99, x_102); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_100); +lean_dec(x_99); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(0, 2, 0); +} else { + x_106 = x_104; +} +lean_ctor_set(x_106, 0, x_102); +lean_ctor_set(x_106, 1, x_103); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = lean_box(0); +if (lean_is_scalar(x_101)) { + x_109 = lean_alloc_ctor(0, 2, 0); +} else { + x_109 = x_101; +} +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +if (lean_is_scalar(x_98)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_98; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_97); +return x_110; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_103); +lean_dec(x_102); +if (lean_is_scalar(x_104)) { + x_111 = lean_alloc_ctor(0, 2, 0); +} else { + x_111 = x_104; +} +lean_ctor_set(x_111, 0, x_99); +lean_ctor_set(x_111, 1, x_100); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = lean_box(0); +if (lean_is_scalar(x_101)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_101; +} +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_112); +if (lean_is_scalar(x_98)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_98; +} +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_97); +return x_115; +} +} +} +} +else +{ +lean_object* x_116; +x_116 = lean_ctor_get(x_14, 0); +lean_inc(x_116); +lean_dec(x_14); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_117 = lean_ctor_get(x_13, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_118 = x_13; +} else { + lean_dec_ref(x_13); + x_118 = lean_box(0); +} +x_119 = lean_ctor_get(x_116, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_116, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_121 = x_116; +} else { + lean_dec_ref(x_116); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(0, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +x_123 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_123, 0, x_122); +x_124 = lean_box(0); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_123); +if (lean_is_scalar(x_118)) { + x_126 = lean_alloc_ctor(0, 2, 0); +} else { + x_126 = x_118; +} +lean_ctor_set(x_126, 0, x_125); +lean_ctor_set(x_126, 1, x_117); +return x_126; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_127 = lean_ctor_get(x_5, 0); +lean_inc(x_127); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + x_128 = x_5; +} else { + lean_dec_ref(x_5); + x_128 = lean_box(0); +} +x_129 = lean_ctor_get(x_13, 1); +lean_inc(x_129); +if (lean_is_exclusive(x_13)) { + lean_ctor_release(x_13, 0); + lean_ctor_release(x_13, 1); + x_130 = x_13; +} else { + lean_dec_ref(x_13); + x_130 = lean_box(0); +} +x_131 = lean_ctor_get(x_116, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_116, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_133 = x_116; +} else { + lean_dec_ref(x_116); + x_133 = lean_box(0); +} +x_134 = lean_ctor_get(x_127, 0); +lean_inc(x_134); +x_135 = lean_ctor_get(x_127, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_136 = x_127; +} else { + lean_dec_ref(x_127); + x_136 = lean_box(0); +} +x_137 = l_Lean_Name_isSuffixOf(x_131, x_134); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_132); +lean_dec(x_131); +if (lean_is_scalar(x_136)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_136; +} +lean_ctor_set(x_138, 0, x_134); +lean_ctor_set(x_138, 1, x_135); +if (lean_is_scalar(x_128)) { + x_139 = lean_alloc_ctor(1, 1, 0); +} else { + x_139 = x_128; +} +lean_ctor_set(x_139, 0, x_138); +x_140 = lean_box(0); +if (lean_is_scalar(x_133)) { + x_141 = lean_alloc_ctor(0, 2, 0); +} else { + x_141 = x_133; +} +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_139); +if (lean_is_scalar(x_130)) { + x_142 = lean_alloc_ctor(0, 2, 0); +} else { + x_142 = x_130; +} +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_129); +return x_142; +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +lean_dec(x_135); +lean_dec(x_134); +if (lean_is_scalar(x_136)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_136; +} +lean_ctor_set(x_143, 0, x_131); +lean_ctor_set(x_143, 1, x_132); +if (lean_is_scalar(x_128)) { + x_144 = lean_alloc_ctor(1, 1, 0); +} else { + x_144 = x_128; +} +lean_ctor_set(x_144, 0, x_143); +x_145 = lean_box(0); +if (lean_is_scalar(x_133)) { + x_146 = lean_alloc_ctor(0, 2, 0); +} else { + x_146 = x_133; +} +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +if (lean_is_scalar(x_130)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_130; +} +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_129); +return x_147; +} +} +} +} +} +else +{ +uint8_t x_148; +lean_dec(x_5); +x_148 = !lean_is_exclusive(x_13); +if (x_148 == 0) +{ +return x_13; +} +else +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_149 = lean_ctor_get(x_13, 0); +x_150 = lean_ctor_get(x_13, 1); +lean_inc(x_150); +lean_inc(x_149); +lean_dec(x_13); +x_151 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_151, 0, x_149); +lean_ctor_set(x_151, 1, x_150); +return x_151; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_51; +x_13 = lean_box(x_2); +lean_inc(x_4); +lean_inc(x_1); +x_14 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed), 12, 3); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_13); +lean_closure_set(x_14, 2, x_4); +x_15 = lean_ctor_get(x_3, 0); +lean_inc(x_15); +lean_dec(x_3); +x_16 = lean_ctor_get(x_15, 4); +lean_inc(x_16); +x_17 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_51 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_visitNamespaces(x_14, x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_ctor_get(x_15, 5); +lean_inc(x_55); +lean_dec(x_15); +x_56 = lean_box(0); +x_57 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_1); +x_58 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1(x_2, x_1, x_4, x_55, x_56, x_55, x_55, x_57, lean_box(0), x_54, x_6, x_7, x_8, x_9, x_10, x_11, x_53); +lean_dec(x_55); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_box(0); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +x_63 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(x_62, x_1, x_2, x_4, x_8, x_9, x_10, x_11, x_60); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_18 = x_61; +x_19 = x_65; +goto block_50; +} +else +{ +uint8_t x_66; +x_66 = !lean_is_exclusive(x_64); +if (x_66 == 0) +{ +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_64, 0); +x_68 = lean_ctor_get(x_63, 1); +lean_inc(x_68); +lean_dec(x_63); +x_69 = !lean_is_exclusive(x_67); +if (x_69 == 0) +{ +x_18 = x_64; +x_19 = x_68; +goto block_50; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_67, 0); +x_71 = lean_ctor_get(x_67, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_67); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_64, 0, x_72); +x_18 = x_64; +x_19 = x_68; +goto block_50; +} +} +else +{ +lean_object* x_73; uint8_t x_74; +x_73 = lean_ctor_get(x_64, 0); +lean_free_object(x_64); +x_74 = !lean_is_exclusive(x_61); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_75 = lean_ctor_get(x_61, 0); +x_76 = lean_ctor_get(x_63, 1); +lean_inc(x_76); +lean_dec(x_63); +x_77 = lean_ctor_get(x_73, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_73, 1); +lean_inc(x_78); +lean_dec(x_73); +x_79 = !lean_is_exclusive(x_75); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = lean_ctor_get(x_75, 0); +x_81 = lean_ctor_get(x_75, 1); +x_82 = l_Lean_Name_isSuffixOf(x_77, x_80); +if (x_82 == 0) +{ +lean_dec(x_78); +lean_dec(x_77); +x_18 = x_61; +x_19 = x_76; +goto block_50; +} +else +{ +lean_dec(x_81); +lean_dec(x_80); +lean_ctor_set(x_75, 1, x_78); +lean_ctor_set(x_75, 0, x_77); +x_18 = x_61; +x_19 = x_76; +goto block_50; +} +} +else +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_ctor_get(x_75, 0); +x_84 = lean_ctor_get(x_75, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_75); +x_85 = l_Lean_Name_isSuffixOf(x_77, x_83); +if (x_85 == 0) +{ +lean_object* x_86; +lean_dec(x_78); +lean_dec(x_77); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +lean_ctor_set(x_61, 0, x_86); +x_18 = x_61; +x_19 = x_76; +goto block_50; +} +else +{ +lean_object* x_87; +lean_dec(x_84); +lean_dec(x_83); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_77); +lean_ctor_set(x_87, 1, x_78); +lean_ctor_set(x_61, 0, x_87); +x_18 = x_61; +x_19 = x_76; +goto block_50; +} +} +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_88 = lean_ctor_get(x_61, 0); +lean_inc(x_88); +lean_dec(x_61); +x_89 = lean_ctor_get(x_63, 1); +lean_inc(x_89); +lean_dec(x_63); +x_90 = lean_ctor_get(x_73, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_73, 1); +lean_inc(x_91); +lean_dec(x_73); +x_92 = lean_ctor_get(x_88, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_88, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_94 = x_88; +} else { + lean_dec_ref(x_88); + x_94 = lean_box(0); +} +x_95 = l_Lean_Name_isSuffixOf(x_90, x_92); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; +lean_dec(x_91); +lean_dec(x_90); +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(0, 2, 0); +} else { + x_96 = x_94; +} +lean_ctor_set(x_96, 0, x_92); +lean_ctor_set(x_96, 1, x_93); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +x_18 = x_97; +x_19 = x_89; +goto block_50; +} +else +{ +lean_object* x_98; lean_object* x_99; +lean_dec(x_93); +lean_dec(x_92); +if (lean_is_scalar(x_94)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_94; +} +lean_ctor_set(x_98, 0, x_90); +lean_ctor_set(x_98, 1, x_91); +x_99 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_99, 0, x_98); +x_18 = x_99; +x_19 = x_89; +goto block_50; +} +} +} +} +else +{ +lean_object* x_100; +x_100 = lean_ctor_get(x_64, 0); +lean_inc(x_100); +lean_dec(x_64); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_101 = lean_ctor_get(x_63, 1); +lean_inc(x_101); +lean_dec(x_63); +x_102 = lean_ctor_get(x_100, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_100, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_104 = x_100; +} else { + lean_dec_ref(x_100); + x_104 = lean_box(0); +} +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_104; +} +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_18 = x_106; +x_19 = x_101; +goto block_50; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_107 = lean_ctor_get(x_61, 0); +lean_inc(x_107); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + x_108 = x_61; +} else { + lean_dec_ref(x_61); + x_108 = lean_box(0); +} +x_109 = lean_ctor_get(x_63, 1); +lean_inc(x_109); +lean_dec(x_63); +x_110 = lean_ctor_get(x_100, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_100, 1); +lean_inc(x_111); +lean_dec(x_100); +x_112 = lean_ctor_get(x_107, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_114 = x_107; +} else { + lean_dec_ref(x_107); + x_114 = lean_box(0); +} +x_115 = l_Lean_Name_isSuffixOf(x_110, x_112); +if (x_115 == 0) +{ +lean_object* x_116; lean_object* x_117; +lean_dec(x_111); +lean_dec(x_110); +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_114; +} +lean_ctor_set(x_116, 0, x_112); +lean_ctor_set(x_116, 1, x_113); +if (lean_is_scalar(x_108)) { + x_117 = lean_alloc_ctor(1, 1, 0); +} else { + x_117 = x_108; +} +lean_ctor_set(x_117, 0, x_116); +x_18 = x_117; +x_19 = x_109; +goto block_50; +} +else +{ +lean_object* x_118; lean_object* x_119; +lean_dec(x_113); +lean_dec(x_112); +if (lean_is_scalar(x_114)) { + x_118 = lean_alloc_ctor(0, 2, 0); +} else { + x_118 = x_114; +} +lean_ctor_set(x_118, 0, x_110); +lean_ctor_set(x_118, 1, x_111); +if (lean_is_scalar(x_108)) { + x_119 = lean_alloc_ctor(1, 1, 0); +} else { + x_119 = x_108; +} +lean_ctor_set(x_119, 0, x_118); +x_18 = x_119; +x_19 = x_109; +goto block_50; +} +} +} +} +} +else +{ +uint8_t x_120; +lean_dec(x_61); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +x_120 = !lean_is_exclusive(x_63); +if (x_120 == 0) +{ +return x_63; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_63, 0); +x_122 = lean_ctor_get(x_63, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_63); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_124 = !lean_is_exclusive(x_58); +if (x_124 == 0) +{ +return x_58; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_58, 0); +x_126 = lean_ctor_get(x_58, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_58); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +} +else +{ +uint8_t x_128; +lean_dec(x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_128 = !lean_is_exclusive(x_51); +if (x_128 == 0) +{ +return x_51; +} +else +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_51, 0); +x_130 = lean_ctor_get(x_51, 1); +lean_inc(x_130); +lean_inc(x_129); +lean_dec(x_51); +x_131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +} +block_50: +{ +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_26 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; double x_30; lean_object* x_31; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +lean_ctor_set_tag(x_18, 0); +lean_ctor_set(x_18, 0, x_4); +x_29 = lean_unbox(x_27); +lean_dec(x_27); +x_30 = lean_unbox_float(x_25); +lean_dec(x_25); +x_31 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_24, x_18, x_29, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_28); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_31; +} +else +{ +uint8_t x_32; +lean_dec(x_25); +lean_dec(x_24); +lean_free_object(x_18); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +x_32 = !lean_is_exclusive(x_26); +if (x_32 == 0) +{ +return x_26; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_26, 0); +x_34 = lean_ctor_get(x_26, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_26); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_18, 0); +lean_inc(x_36); +lean_dec(x_18); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_39 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; double x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_39, 1); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_42, 0, x_4); +x_43 = lean_unbox(x_40); +lean_dec(x_40); +x_44 = lean_unbox_float(x_38); +lean_dec(x_38); +x_45 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_37, x_42, x_43, x_44, x_6, x_7, x_8, x_9, x_10, x_11, x_41); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_38); +lean_dec(x_37); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +x_46 = lean_ctor_get(x_39, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_48 = x_39; +} else { + lean_dec_ref(x_39); + x_48 = lean_box(0); +} +if (lean_is_scalar(x_48)) { + x_49 = lean_alloc_ctor(1, 2, 0); +} else { + x_49 = x_48; +} +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_47); +return x_49; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Lean_Name_isPrefixOf(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +lean_dec(x_1); +x_5 = lean_box(0); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_box(0); +x_7 = l_Lean_Name_replacePrefix(x_3, x_2, x_6); +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchAtomic(x_1, x_7); +lean_dec(x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7(x_1, x_2, x_3, x_5, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 4); +lean_inc(x_13); +lean_dec(x_1); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore_searchAlias(x_2, x_3, x_4, x_5, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Server_Completion_completeNamespaces(x_1, x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_st_ref_get(x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_box(x_2); +lean_inc(x_3); +lean_inc(x_1); +x_17 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed), 12, 3); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_16); +lean_closure_set(x_17, 2, x_3); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_18 = l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +lean_inc(x_1); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed), 3, 1); +lean_closure_set(x_20, 0, x_1); +lean_inc(x_15); +x_21 = l_Lean_Server_Completion_getEligibleHeaderDecls(x_15, x_19); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +lean_inc(x_22); +lean_inc(x_15); +x_24 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed), 12, 2); +lean_closure_set(x_24, 0, x_15); +lean_closure_set(x_24, 1, x_22); +x_25 = lean_ctor_get(x_3, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_25, 5); +lean_inc(x_26); +x_27 = lean_box(0); +x_28 = lean_box(0); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_26); +lean_inc(x_15); +lean_inc(x_1); +x_29 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17(x_1, x_15, x_22, x_26, x_27, x_26, x_26, x_28, lean_box(0), x_5, x_6, x_7, x_8, x_9, x_10, x_23); +lean_dec(x_26); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_getAliasState(x_15); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__5), 12, 3); +lean_closure_set(x_32, 0, x_25); +lean_closure_set(x_32, 1, x_20); +lean_closure_set(x_32, 2, x_24); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_33 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18(x_31, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_30); +lean_dec(x_31); +if (lean_obj_tag(x_33) == 0) +{ +if (lean_obj_tag(x_1) == 1) +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_1, 0); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_1, 1); +lean_inc(x_36); +x_37 = l_Lean_Parser_getTokenTable(x_15); +lean_dec(x_15); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Lean_Data_Trie_findPrefix_go___rarg(x_36, x_37, x_38); +x_40 = lean_array_size(x_39); +x_41 = 0; +x_42 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23(x_36, x_39, x_27, x_39, x_40, x_41, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +lean_dec(x_39); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_34); +lean_dec(x_15); +x_45 = lean_ctor_get(x_33, 1); +lean_inc(x_45); +lean_dec(x_33); +x_46 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_45); +return x_46; +} +} +else +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_15); +x_47 = lean_ctor_get(x_33, 1); +lean_inc(x_47); +lean_dec(x_33); +x_48 = l_Lean_Server_Completion_completeNamespaces(x_3, x_1, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_47); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_33); +if (x_49 == 0) +{ +return x_33; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_33, 0); +x_51 = lean_ctor_get(x_33, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_33); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +else +{ +uint8_t x_53; +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_20); +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_53 = !lean_is_exclusive(x_29); +if (x_53 == 0) +{ +return x_29; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_29, 0); +x_55 = lean_ctor_get(x_29, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_29); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; +} +} +} +else +{ +uint8_t x_57; +lean_dec(x_15); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_18); +if (x_57 == 0) +{ +return x_18; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_18, 0); +x_59 = lean_ctor_get(x_18, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_18); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = l_Lean_Name_isAtomic(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_3, x_2, x_1, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_box(0); +lean_inc(x_3); +x_18 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24(x_3, x_16, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_16); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_3, x_2, x_1, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +return x_20; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_ctor_get(x_2, 0); +lean_inc(x_15); +lean_dec(x_2); +x_16 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate(x_4, x_15); +x_17 = 0; +x_18 = lean_box(0); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_17, x_16, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; +x_13 = l_Lean_Name_hasMacroScopes(x_3); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_box(0); +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_4, x_5, x_3, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +else +{ +lean_object* x_16; +x_16 = l_Lean_Syntax_getHeadInfo(x_2); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +x_17 = lean_erase_macro_scopes(x_3); +x_18 = lean_box(0); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_4, x_5, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_20 = lean_box(0); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_12); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___boxed(lean_object** _args) { +lean_object* x_1 = _args[0]; +lean_object* x_2 = _args[1]; +lean_object* x_3 = _args[2]; +lean_object* x_4 = _args[3]; +lean_object* x_5 = _args[4]; +lean_object* x_6 = _args[5]; +lean_object* x_7 = _args[6]; +lean_object* x_8 = _args[7]; +lean_object* x_9 = _args[8]; +lean_object* x_10 = _args[9]; +lean_object* x_11 = _args[10]; +lean_object* x_12 = _args[11]; +lean_object* x_13 = _args[12]; +lean_object* x_14 = _args[13]; +lean_object* x_15 = _args[14]; +lean_object* x_16 = _args[15]; +lean_object* x_17 = _args[16]; +_start: +{ +uint8_t x_18; lean_object* x_19; +x_18 = lean_unbox(x_1); +lean_dec(x_1); +x_19 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__6(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +double x_13; lean_object* x_14; +x_13 = lean_unbox_float(x_4); +lean_dec(x_4); +x_14 = l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__7(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +double x_13; lean_object* x_14; +x_13 = lean_unbox_float(x_4); +lean_dec(x_4); +x_14 = l_List_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__8(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__14___rarg(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__15___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__13___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__16(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +x_17 = l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentHashMap_foldlM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__21(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentHashMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__20(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +size_t x_13; size_t x_14; lean_object* x_15; +x_13 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_14 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_15 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__22(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_SMap_forM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__23(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_18 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__26(x_1, x_2, x_3, x_4, x_5, x_6, x_17, x_18, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__27(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_PersistentArray_forInAux___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__25(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_18 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__28(x_1, x_2, x_3, x_4, x_5, x_16, x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_PersistentArray_forIn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__24(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_2); +lean_dec(x_2); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__1(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_2); +lean_dec(x_2); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__2(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__3(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +double x_13; lean_object* x_14; +x_13 = lean_unbox_float(x_5); +lean_dec(x_5); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__4(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_3); +lean_dec(x_3); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__6(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__7(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__8(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_3); +lean_dec(x_3); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___lambda__9(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_5); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_idCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, uint8_t x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_box(x_8); +lean_inc(x_3); +x_11 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___boxed), 12, 5); +lean_closure_set(x_11, 0, x_3); +lean_closure_set(x_11, 1, x_5); +lean_closure_set(x_11, 2, x_6); +lean_closure_set(x_11, 3, x_7); +lean_closure_set(x_11, 4, x_10); +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_4, x_11, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_idCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_8); +lean_dec(x_8); +x_11 = l_Lean_Server_Completion_idCompletion(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_10, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at_Lean_Server_Completion_dotCompletion___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 2); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_dec(x_1); +lean_inc(x_5); +lean_inc(x_2); +x_8 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_cmpModPrivate(x_2, x_5); +switch (x_8) { +case 0: +{ +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_1 = x_4; +goto _start; +} +case 1: +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_2); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_5); +lean_ctor_set(x_10, 1, x_6); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +default: +{ +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_1 = x_7; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_hasMVar(x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_11 = lean_st_ref_get(x_5, x_8); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_instantiateMVarsCore(x_14, x_1); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_st_ref_take(x_5, x_13); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +lean_ctor_set(x_19, 0, x_17); +x_23 = lean_st_ref_set(x_5, x_19, x_20); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +lean_ctor_set(x_23, 0, x_16); +return x_23; +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_28 = lean_ctor_get(x_19, 1); +x_29 = lean_ctor_get(x_19, 2); +x_30 = lean_ctor_get(x_19, 3); +x_31 = lean_ctor_get(x_19, 4); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_32 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_32, 0, x_17); +lean_ctor_set(x_32, 1, x_28); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_30); +lean_ctor_set(x_32, 4, x_31); +x_33 = lean_st_ref_set(x_5, x_32, x_20); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_35 = x_33; +} else { + lean_dec_ref(x_33); + x_35 = lean_box(0); +} +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_35; +} +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_36, 1, x_34); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_10 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_apply_8(x_2, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_13; +} +else +{ +uint8_t x_14; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +return x_10; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_10 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; double x_18; uint8_t x_19; lean_object* x_20; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_ConstantInfo_name(x_1); +x_14 = l_Lean_Name_getString_x21(x_13); +x_15 = lean_box(0); +x_16 = l_Lean_Name_str___override(x_15, x_14); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_13); +x_18 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_19 = lean_unbox(x_11); +lean_dec(x_11); +x_20 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_16, x_17, x_19, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_20; +} +else +{ +uint8_t x_21; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) +{ +return x_10; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_10); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_6, x_7, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_13 = !lean_is_exclusive(x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_11, 0, x_15); +return x_11; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_dec(x_11); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_ctor_get(x_12, 0); +lean_inc(x_20); +lean_dec(x_12); +x_21 = l_Lean_Name_getPrefix(x_20); +lean_dec(x_20); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_22 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod(x_21, x_2, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +uint8_t x_25; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_22, 0, x_27); +return x_22; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_dec(x_22); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_dec(x_22); +x_32 = lean_box(0); +x_33 = l_Lean_Server_Completion_dotCompletion___lambda__1(x_2, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +return x_33; +} +} +else +{ +uint8_t x_34; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_34 = !lean_is_exclusive(x_22); +if (x_34 == 0) +{ +return x_22; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_22, 0); +x_36 = lean_ctor_get(x_22, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_22); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_Name_getPrefix(x_2); +x_12 = l_Lean_RBNode_findCore___at_Lean_Server_Completion_dotCompletion___spec__1(x_1, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = l_Lean_Server_Completion_dotCompletion___lambda__2(x_2, x_3, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Server_Completion_dotCompletion___lambda__3___boxed), 10, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Server_Completion_dotCompletion___lambda__4(x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames(x_13, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_15) == 0) +{ +return x_15; +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = l_Lean_Exception_isInterrupt(x_17); +if (x_18 == 0) +{ +uint8_t x_19; +x_19 = l_Lean_Exception_isRuntime(x_17); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_17); +x_20 = lean_box(0); +lean_ctor_set_tag(x_15, 0); +lean_ctor_set(x_15, 0, x_20); +return x_15; +} +else +{ +return x_15; +} +} +else +{ +return x_15; +} +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = l_Lean_Exception_isInterrupt(x_21); +if (x_23 == 0) +{ +uint8_t x_24; +x_24 = l_Lean_Exception_isRuntime(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_21); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +return x_26; +} +else +{ +lean_object* x_27; +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_22); +return x_27; +} +} +else +{ +lean_object* x_28; +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_22); +return x_28; +} +} +} +} +else +{ +uint8_t x_29; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_29 = !lean_is_exclusive(x_9); +if (x_29 == 0) +{ +lean_object* x_30; uint8_t x_31; +x_30 = lean_ctor_get(x_9, 0); +x_31 = l_Lean_Exception_isInterrupt(x_30); +if (x_31 == 0) +{ +uint8_t x_32; +x_32 = l_Lean_Exception_isRuntime(x_30); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_30); +x_33 = lean_box(0); +lean_ctor_set_tag(x_9, 0); +lean_ctor_set(x_9, 0, x_33); +return x_9; +} +else +{ +return x_9; +} +} +else +{ +return x_9; +} +} +else +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_9, 0); +x_35 = lean_ctor_get(x_9, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_9); +x_36 = l_Lean_Exception_isInterrupt(x_34); +if (x_36 == 0) +{ +uint8_t x_37; +x_37 = l_Lean_Exception_isRuntime(x_34); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_34); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_35); +return x_39; +} +else +{ +lean_object* x_40; +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_35); +return x_40; +} +} +else +{ +lean_object* x_41; +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_34); +lean_ctor_set(x_41, 1, x_35); +return x_41; +} +} +} +} +} +static lean_object* _init_l_Lean_Server_Completion_dotCompletion___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Server_Completion_dotCompletion___lambda__5), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Lean_Server_Completion_dotCompletion___closed__1; +x_8 = lean_ctor_get(x_4, 3); +lean_inc(x_8); +lean_dec(x_4); +x_9 = lean_alloc_closure((void*)(l_Lean_Server_Completion_dotCompletion___lambda__6___boxed), 8, 1); +lean_closure_set(x_9, 0, x_8); +x_10 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_7); +x_11 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_6, x_10, x_5); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Server_Completion_dotCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Server_Completion_dotCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Server_Completion_dotCompletion___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Server_Completion_dotCompletion___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotCompletion___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Server_Completion_dotCompletion___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = 0; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f(x_1, x_2, x_14, x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_15, 0, x_19); +return x_15; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_16); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; double x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_16, 0); +x_25 = lean_ctor_get(x_15, 1); +lean_inc(x_25); +lean_dec(x_15); +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = l_Lean_ConstantInfo_name(x_4); +lean_ctor_set_tag(x_16, 0); +lean_ctor_set(x_16, 0, x_28); +x_29 = lean_unbox_float(x_27); +lean_dec(x_27); +x_30 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_26, x_16, x_5, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_25); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; double x_37; lean_object* x_38; +x_31 = lean_ctor_get(x_16, 0); +lean_inc(x_31); +lean_dec(x_16); +x_32 = lean_ctor_get(x_15, 1); +lean_inc(x_32); +lean_dec(x_15); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = l_Lean_ConstantInfo_name(x_4); +x_36 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_unbox_float(x_34); +lean_dec(x_34); +x_38 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_33, x_36, x_5, x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_32); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +x_39 = !lean_is_exclusive(x_15); +if (x_39 == 0) +{ +return x_15; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_15, 0); +x_41 = lean_ctor_get(x_15, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_15); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Name_isAnonymous(x_3); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_box(0); +x_18 = lean_unbox(x_14); +lean_dec(x_14); +x_19 = l_Lean_Server_Completion_dotIdCompletion___lambda__1(x_2, x_3, x_4, x_1, x_18, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; double x_25; uint8_t x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_4); +lean_dec(x_3); +x_20 = l_Lean_ConstantInfo_name(x_1); +x_21 = l_Lean_Name_getString_x21(x_20); +x_22 = lean_box(0); +x_23 = l_Lean_Name_str___override(x_22, x_21); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_20); +x_25 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_26 = lean_unbox(x_14); +lean_dec(x_14); +x_27 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem(x_23, x_24, x_26, x_25, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +x_30 = lean_box(0); +lean_ctor_set(x_27, 0, x_30); +return x_27; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_27, 1); +lean_inc(x_31); +lean_dec(x_27); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +} +else +{ +uint8_t x_34; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +x_34 = !lean_is_exclusive(x_13); +if (x_34 == 0) +{ +return x_13; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_13, 0); +x_36 = lean_ctor_get(x_13, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_13); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_normPrivateName_x3f(x_1, x_7, x_8, x_9, x_10, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_12, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_12, 0, x_16); +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); +lean_dec(x_12); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +x_22 = l_Lean_Name_getPrefix(x_21); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_22); +x_23 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotIdCompletionMethod(x_22, x_2, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_unbox(x_24); +lean_dec(x_24); +if (x_25 == 0) +{ +uint8_t x_26; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +x_28 = lean_box(0); +lean_ctor_set(x_23, 0, x_28); +return x_23; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_dec(x_23); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_dec(x_23); +x_33 = lean_box(0); +x_34 = l_Lean_Server_Completion_dotIdCompletion___lambda__2(x_2, x_22, x_3, x_21, x_33, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +lean_dec(x_22); +return x_34; +} +} +else +{ +uint8_t x_35; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_35 = !lean_is_exclusive(x_23); +if (x_35 == 0) +{ +return x_23; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_23, 0); +x_37 = lean_ctor_get(x_23, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_23); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_Lean_Name_getPrefix(x_3); +x_13 = l_Lean_RBNode_findCore___at_Lean_Server_Completion_dotCompletion___spec__1(x_2, x_12); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_1); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_13); +x_16 = lean_box(0); +x_17 = l_Lean_Server_Completion_dotIdCompletion___lambda__3(x_3, x_4, x_1, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Server_Completion_dotIdCompletion___lambda__4___boxed), 11, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_2); +x_11 = l_Lean_Server_Completion_forEligibleDeclsM___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l_Lean_Expr_cleanupAnnotations(x_2); +x_11 = l_Lean_Expr_getAppFn(x_10); +lean_dec(x_10); +x_12 = l_Lean_Expr_cleanupAnnotations(x_11); +if (lean_obj_tag(x_12) == 4) +{ +lean_object* x_13; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_13 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames(x_12, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_15); +return x_16; +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_13); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_ctor_get(x_13, 1); +x_20 = l_Lean_Exception_isInterrupt(x_18); +if (x_20 == 0) +{ +uint8_t x_21; +x_21 = l_Lean_Exception_isRuntime(x_18); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_free_object(x_13); +lean_dec(x_18); +x_22 = lean_box(0); +x_23 = l_Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +return x_23; +} +else +{ +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_13; +} +} +else +{ +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_13; +} +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = l_Lean_Exception_isInterrupt(x_24); +if (x_26 == 0) +{ +uint8_t x_27; +x_27 = l_Lean_Exception_isRuntime(x_24); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l_Lean_Server_Completion_dotIdCompletion___lambda__5(x_1, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_25); +return x_29; +} +else +{ +lean_object* x_30; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_24); +lean_ctor_set(x_30, 1, x_25); +return x_30; +} +} +else +{ +lean_object* x_31; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_24); +lean_ctor_set(x_31, 1, x_25); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_32 = lean_box(0); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_9); +return x_33; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +x_8 = lean_box(0); +x_9 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed), 8, 1); +lean_closure_set(x_9, 0, x_8); +x_10 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_4, x_9, x_7); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_6, 0); +lean_inc(x_11); +lean_dec(x_6); +x_12 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at_Lean_Server_Completion_dotCompletion___spec__2___boxed), 8, 1); +lean_closure_set(x_12, 0, x_11); +x_13 = lean_alloc_closure((void*)(l_Lean_Server_Completion_dotIdCompletion___lambda__6), 9, 1); +lean_closure_set(x_13, 0, x_5); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_4, x_14, x_7); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_ReaderT_pure___at_Lean_Server_Completion_dotIdCompletion___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_5); +lean_dec(x_5); +x_15 = l_Lean_Server_Completion_dotIdCompletion___lambda__1(x_1, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; +x_13 = l_Lean_Server_Completion_dotIdCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Server_Completion_dotIdCompletion___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_dotIdCompletion___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Server_Completion_dotIdCompletion___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_12; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("field", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 4; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = lean_usize_dec_lt(x_6, x_5); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_1); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +else +{ +lean_object* x_17; +lean_dec(x_7); +x_17 = lean_array_uget(x_4, x_6); +if (lean_obj_tag(x_17) == 1) +{ +lean_object* x_18; double x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +lean_inc(x_1); +x_20 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_1, x_18, x_19); +if (lean_obj_tag(x_20) == 0) +{ +size_t x_21; size_t x_22; lean_object* x_23; +lean_dec(x_18); +x_21 = 1; +x_22 = lean_usize_add(x_6, x_21); +x_23 = lean_box(0); +x_6 = x_22; +x_7 = x_23; +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; double x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; lean_object* x_35; +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +lean_dec(x_20); +x_26 = lean_box(0); +x_27 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2; +x_28 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3; +x_29 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_29, 0, x_18); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set(x_29, 2, x_26); +lean_ctor_set(x_29, 3, x_28); +lean_ctor_set(x_29, 4, x_26); +lean_ctor_set(x_29, 5, x_26); +lean_ctor_set(x_29, 6, x_26); +lean_ctor_set(x_29, 7, x_26); +x_30 = lean_unbox_float(x_25); +lean_dec(x_25); +x_31 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addItem(x_29, x_30, x_26, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = 1; +x_34 = lean_usize_add(x_6, x_33); +x_35 = lean_box(0); +x_6 = x_34; +x_7 = x_35; +x_14 = x_32; +goto _start; +} +} +else +{ +size_t x_37; size_t x_38; lean_object* x_39; +lean_dec(x_17); +x_37 = 1; +x_38 = lean_usize_add(x_6, x_37); +x_39 = lean_box(0); +x_6 = x_38; +x_7 = x_39; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_11); +return x_8; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; lean_object* x_13; size_t x_14; size_t x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_11 = 0; +x_12 = l_Lean_getStructureFieldsFlattened(x_3, x_1, x_11); +x_13 = lean_box(0); +x_14 = lean_array_size(x_12); +x_15 = 0; +x_16 = lean_box(0); +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1(x_2, x_12, x_13, x_12, x_14, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_12); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +static lean_object* _init_l_Lean_Server_Completion_fieldIdCompletion___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_14; +x_14 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_8 = x_14; +goto block_13; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +lean_dec(x_5); +x_16 = 1; +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +x_18 = l_Lean_Name_toString(x_15, x_16, x_17); +x_8 = x_18; +goto block_13; +} +block_13: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_alloc_closure((void*)(l_Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed), 10, 2); +lean_closure_set(x_9, 0, x_6); +lean_closure_set(x_9, 1, x_8); +x_10 = l_Lean_Server_Completion_fieldIdCompletion___closed__1; +x_11 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Server_Completion_dotCompletion___spec__3___rarg), 9, 2); +lean_closure_set(x_11, 0, x_10); +lean_closure_set(x_11, 1, x_9); +x_12 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM(x_1, x_2, x_3, x_4, x_11, x_7); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +size_t x_15; size_t x_16; lean_object* x_17; +x_15 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_16 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_17 = l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1(x_1, x_2, x_3, x_4, x_15, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Server_Completion_fieldIdCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_fieldIdCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Lean_Server_Completion_fieldIdCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 9; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("), ", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 32; +x_2 = l_Char_utf8Size(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_10); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_34; lean_object* x_35; lean_object* x_36; double x_37; lean_object* x_38; +x_18 = lean_ctor_get(x_9, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_9, 2); +lean_inc(x_20); +x_21 = lean_ctor_get(x_9, 3); +lean_inc(x_21); +lean_dec(x_9); +lean_inc(x_6); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_22 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18, x_10, x_11, x_12, x_13, x_14, x_15); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + x_25 = x_22; +} else { + lean_dec_ref(x_22); + x_25 = lean_box(0); +} +x_26 = lean_ctor_get(x_23, 0); +lean_inc(x_26); +if (lean_is_exclusive(x_23)) { + lean_ctor_release(x_23, 0); + x_27 = x_23; +} else { + lean_dec_ref(x_23); + x_27 = lean_box(0); +} +x_34 = 1; +x_35 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2; +lean_inc(x_19); +x_36 = l_Lean_Name_toString(x_19, x_34, x_35); +x_37 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2; +lean_inc(x_6); +x_38 = l_Lean_FuzzyMatching_fuzzyMatchScoreWithThreshold_x3f(x_6, x_36, x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_36); +lean_dec(x_20); +lean_dec(x_19); +if (lean_is_scalar(x_27)) { + x_39 = lean_alloc_ctor(1, 1, 0); +} else { + x_39 = x_27; +} +lean_ctor_set(x_39, 0, x_26); +if (lean_is_scalar(x_25)) { + x_40 = lean_alloc_ctor(0, 2, 0); +} else { + x_40 = x_25; +} +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_24); +x_28 = x_40; +goto block_33; +} +else +{ +uint8_t x_41; +x_41 = !lean_is_exclusive(x_38); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_102; lean_object* x_120; lean_object* x_121; +x_42 = lean_ctor_get(x_38, 0); +x_43 = lean_ctor_get(x_4, 0); +x_44 = lean_ctor_get(x_20, 1); +lean_inc(x_44); +x_45 = l_Lean_KVMap_findCore(x_8, x_19); +lean_dec(x_19); +x_46 = lean_ctor_get(x_20, 3); +lean_inc(x_46); +lean_dec(x_20); +x_47 = lean_box(0); +lean_inc(x_2); +lean_inc(x_1); +x_48 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_48, 0, x_1); +lean_ctor_set(x_48, 1, x_2); +lean_ctor_set(x_48, 2, x_47); +x_49 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(x_48); +lean_ctor_set(x_38, 0, x_49); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_139; +x_139 = lean_box(0); +x_50 = x_139; +goto block_101; +} +else +{ +lean_object* x_140; lean_object* x_141; +x_140 = lean_ctor_get(x_43, 0); +x_141 = lean_ctor_get(x_140, 0); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; +x_142 = lean_box(0); +x_50 = x_142; +goto block_101; +} +else +{ +lean_object* x_143; +x_143 = lean_ctor_get(x_141, 0); +if (lean_obj_tag(x_143) == 0) +{ +lean_object* x_144; +x_144 = lean_box(0); +x_50 = x_144; +goto block_101; +} +else +{ +lean_object* x_145; +x_145 = lean_ctor_get(x_143, 0); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; +x_146 = lean_box(0); +x_50 = x_146; +goto block_101; +} +else +{ +lean_object* x_147; uint8_t x_148; +x_147 = lean_ctor_get(x_145, 0); +x_148 = lean_unbox(x_147); +if (x_148 == 0) +{ +lean_object* x_149; +x_149 = lean_box(0); +x_50 = x_149; +goto block_101; +} +else +{ +uint8_t x_150; lean_object* x_151; +lean_dec(x_27); +lean_dec(x_25); +x_150 = 0; +x_151 = l_Lean_Syntax_getRange_x3f(x_5, x_150); +if (lean_obj_tag(x_151) == 0) +{ +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_47; +goto block_119; +} +else +{ +lean_object* x_152; +lean_dec(x_44); +x_152 = lean_ctor_get(x_45, 0); +lean_inc(x_152); +lean_dec(x_45); +x_120 = x_47; +x_121 = x_152; +goto block_138; +} +} +else +{ +uint8_t x_153; +x_153 = !lean_is_exclusive(x_151); +if (x_153 == 0) +{ +lean_object* x_154; uint8_t x_155; +x_154 = lean_ctor_get(x_151, 0); +x_155 = !lean_is_exclusive(x_154); +if (x_155 == 0) +{ +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_156 = lean_ctor_get(x_154, 0); +x_157 = lean_ctor_get(x_154, 1); +x_158 = lean_ctor_get(x_3, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_158, 1); +lean_inc(x_159); +lean_dec(x_158); +lean_inc(x_159); +x_160 = l_Lean_FileMap_utf8PosToLspPos(x_159, x_156); +lean_dec(x_156); +if (x_7 == 0) +{ +lean_object* x_161; lean_object* x_162; +x_161 = l_Lean_FileMap_utf8PosToLspPos(x_159, x_157); +lean_dec(x_157); +lean_ctor_set(x_154, 1, x_161); +lean_ctor_set(x_154, 0, x_160); +lean_inc(x_154); +lean_inc(x_36); +x_162 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_162, 0, x_36); +lean_ctor_set(x_162, 1, x_154); +lean_ctor_set(x_162, 2, x_154); +lean_ctor_set(x_151, 0, x_162); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_151; +goto block_119; +} +else +{ +lean_object* x_163; +lean_dec(x_44); +x_163 = lean_ctor_get(x_45, 0); +lean_inc(x_163); +lean_dec(x_45); +x_120 = x_151; +x_121 = x_163; +goto block_138; +} +} +else +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_164 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4; +x_165 = lean_nat_add(x_157, x_164); +lean_dec(x_157); +x_166 = l_Lean_FileMap_utf8PosToLspPos(x_159, x_165); +lean_dec(x_165); +lean_ctor_set(x_154, 1, x_166); +lean_ctor_set(x_154, 0, x_160); +lean_inc(x_154); +lean_inc(x_36); +x_167 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_167, 0, x_36); +lean_ctor_set(x_167, 1, x_154); +lean_ctor_set(x_167, 2, x_154); +lean_ctor_set(x_151, 0, x_167); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_151; +goto block_119; +} +else +{ +lean_object* x_168; +lean_dec(x_44); +x_168 = lean_ctor_get(x_45, 0); +lean_inc(x_168); +lean_dec(x_45); +x_120 = x_151; +x_121 = x_168; +goto block_138; +} +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_169 = lean_ctor_get(x_154, 0); +x_170 = lean_ctor_get(x_154, 1); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_154); +x_171 = lean_ctor_get(x_3, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_171, 1); +lean_inc(x_172); +lean_dec(x_171); +lean_inc(x_172); +x_173 = l_Lean_FileMap_utf8PosToLspPos(x_172, x_169); +lean_dec(x_169); +if (x_7 == 0) +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = l_Lean_FileMap_utf8PosToLspPos(x_172, x_170); +lean_dec(x_170); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +lean_inc(x_175); +lean_inc(x_36); +x_176 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_176, 0, x_36); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_175); +lean_ctor_set(x_151, 0, x_176); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_151; +goto block_119; +} +else +{ +lean_object* x_177; +lean_dec(x_44); +x_177 = lean_ctor_get(x_45, 0); +lean_inc(x_177); +lean_dec(x_45); +x_120 = x_151; +x_121 = x_177; +goto block_138; +} +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4; +x_179 = lean_nat_add(x_170, x_178); +lean_dec(x_170); +x_180 = l_Lean_FileMap_utf8PosToLspPos(x_172, x_179); +lean_dec(x_179); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_173); +lean_ctor_set(x_181, 1, x_180); +lean_inc(x_181); +lean_inc(x_36); +x_182 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_182, 0, x_36); +lean_ctor_set(x_182, 1, x_181); +lean_ctor_set(x_182, 2, x_181); +lean_ctor_set(x_151, 0, x_182); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_151; +goto block_119; +} +else +{ +lean_object* x_183; +lean_dec(x_44); +x_183 = lean_ctor_get(x_45, 0); +lean_inc(x_183); +lean_dec(x_45); +x_120 = x_151; +x_121 = x_183; +goto block_138; +} +} +} +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_184 = lean_ctor_get(x_151, 0); +lean_inc(x_184); +lean_dec(x_151); +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_187 = x_184; +} else { + lean_dec_ref(x_184); + x_187 = lean_box(0); +} +x_188 = lean_ctor_get(x_3, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_188, 1); +lean_inc(x_189); +lean_dec(x_188); +lean_inc(x_189); +x_190 = l_Lean_FileMap_utf8PosToLspPos(x_189, x_185); +lean_dec(x_185); +if (x_7 == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_191 = l_Lean_FileMap_utf8PosToLspPos(x_189, x_186); +lean_dec(x_186); +if (lean_is_scalar(x_187)) { + x_192 = lean_alloc_ctor(0, 2, 0); +} else { + x_192 = x_187; +} +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, x_191); +lean_inc(x_192); +lean_inc(x_36); +x_193 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_193, 0, x_36); +lean_ctor_set(x_193, 1, x_192); +lean_ctor_set(x_193, 2, x_192); +x_194 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_194, 0, x_193); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_194; +goto block_119; +} +else +{ +lean_object* x_195; +lean_dec(x_44); +x_195 = lean_ctor_get(x_45, 0); +lean_inc(x_195); +lean_dec(x_45); +x_120 = x_194; +x_121 = x_195; +goto block_138; +} +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_196 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4; +x_197 = lean_nat_add(x_186, x_196); +lean_dec(x_186); +x_198 = l_Lean_FileMap_utf8PosToLspPos(x_189, x_197); +lean_dec(x_197); +if (lean_is_scalar(x_187)) { + x_199 = lean_alloc_ctor(0, 2, 0); +} else { + x_199 = x_187; +} +lean_ctor_set(x_199, 0, x_190); +lean_ctor_set(x_199, 1, x_198); +lean_inc(x_199); +lean_inc(x_36); +x_200 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_200, 0, x_36); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_199); +x_201 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_201, 0, x_200); +if (lean_obj_tag(x_45) == 0) +{ +x_102 = x_201; +goto block_119; +} +else +{ +lean_object* x_202; +lean_dec(x_44); +x_202 = lean_ctor_get(x_45, 0); +lean_inc(x_202); +lean_dec(x_45); +x_120 = x_201; +x_121 = x_202; +goto block_138; +} +} +} +} +} +} +} +} +} +block_101: +{ +lean_dec(x_50); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; double x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_51 = lean_data_value_to_string(x_44); +x_52 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_53 = lean_string_append(x_52, x_51); +lean_dec(x_51); +x_54 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_55 = lean_string_append(x_53, x_54); +x_56 = lean_string_append(x_55, x_46); +lean_dec(x_46); +x_57 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_58 = lean_string_append(x_56, x_57); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_60 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_61 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_61, 0, x_36); +lean_ctor_set(x_61, 1, x_59); +lean_ctor_set(x_61, 2, x_47); +lean_ctor_set(x_61, 3, x_60); +lean_ctor_set(x_61, 4, x_47); +lean_ctor_set(x_61, 5, x_47); +lean_ctor_set(x_61, 6, x_38); +lean_ctor_set(x_61, 7, x_47); +x_62 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_62, 0, x_61); +x_63 = lean_unbox_float(x_42); +lean_dec(x_42); +lean_ctor_set_float(x_62, sizeof(void*)*1, x_63); +x_64 = lean_array_push(x_26, x_62); +if (lean_is_scalar(x_27)) { + x_65 = lean_alloc_ctor(1, 1, 0); +} else { + x_65 = x_27; +} +lean_ctor_set(x_65, 0, x_64); +if (lean_is_scalar(x_25)) { + x_66 = lean_alloc_ctor(0, 2, 0); +} else { + x_66 = x_25; +} +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_24); +x_28 = x_66; +goto block_33; +} +else +{ +uint8_t x_67; +lean_dec(x_44); +x_67 = !lean_is_exclusive(x_45); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; double x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_68 = lean_ctor_get(x_45, 0); +x_69 = lean_data_value_to_string(x_68); +x_70 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_71 = lean_string_append(x_70, x_69); +lean_dec(x_69); +x_72 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_73 = lean_string_append(x_71, x_72); +x_74 = lean_string_append(x_73, x_46); +lean_dec(x_46); +x_75 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_76 = lean_string_append(x_74, x_75); +lean_ctor_set(x_45, 0, x_76); +x_77 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_78 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_78, 0, x_36); +lean_ctor_set(x_78, 1, x_45); +lean_ctor_set(x_78, 2, x_47); +lean_ctor_set(x_78, 3, x_77); +lean_ctor_set(x_78, 4, x_47); +lean_ctor_set(x_78, 5, x_47); +lean_ctor_set(x_78, 6, x_38); +lean_ctor_set(x_78, 7, x_47); +x_79 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_79, 0, x_78); +x_80 = lean_unbox_float(x_42); +lean_dec(x_42); +lean_ctor_set_float(x_79, sizeof(void*)*1, x_80); +x_81 = lean_array_push(x_26, x_79); +if (lean_is_scalar(x_27)) { + x_82 = lean_alloc_ctor(1, 1, 0); +} else { + x_82 = x_27; +} +lean_ctor_set(x_82, 0, x_81); +if (lean_is_scalar(x_25)) { + x_83 = lean_alloc_ctor(0, 2, 0); +} else { + x_83 = x_25; +} +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_24); +x_28 = x_83; +goto block_33; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; double x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_84 = lean_ctor_get(x_45, 0); +lean_inc(x_84); +lean_dec(x_45); +x_85 = lean_data_value_to_string(x_84); +x_86 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_87 = lean_string_append(x_86, x_85); +lean_dec(x_85); +x_88 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_89 = lean_string_append(x_87, x_88); +x_90 = lean_string_append(x_89, x_46); +lean_dec(x_46); +x_91 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_92 = lean_string_append(x_90, x_91); +x_93 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_93, 0, x_92); +x_94 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_95 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_95, 0, x_36); +lean_ctor_set(x_95, 1, x_93); +lean_ctor_set(x_95, 2, x_47); +lean_ctor_set(x_95, 3, x_94); +lean_ctor_set(x_95, 4, x_47); +lean_ctor_set(x_95, 5, x_47); +lean_ctor_set(x_95, 6, x_38); +lean_ctor_set(x_95, 7, x_47); +x_96 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_96, 0, x_95); +x_97 = lean_unbox_float(x_42); +lean_dec(x_42); +lean_ctor_set_float(x_96, sizeof(void*)*1, x_97); +x_98 = lean_array_push(x_26, x_96); +if (lean_is_scalar(x_27)) { + x_99 = lean_alloc_ctor(1, 1, 0); +} else { + x_99 = x_27; +} +lean_ctor_set(x_99, 0, x_98); +if (lean_is_scalar(x_25)) { + x_100 = lean_alloc_ctor(0, 2, 0); +} else { + x_100 = x_25; +} +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_24); +x_28 = x_100; +goto block_33; +} +} +} +block_119: +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; double x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_103 = lean_data_value_to_string(x_44); +x_104 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_105 = lean_string_append(x_104, x_103); +lean_dec(x_103); +x_106 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_107 = lean_string_append(x_105, x_106); +x_108 = lean_string_append(x_107, x_46); +lean_dec(x_46); +x_109 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_110 = lean_string_append(x_108, x_109); +x_111 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_111, 0, x_110); +x_112 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_113 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_113, 0, x_36); +lean_ctor_set(x_113, 1, x_111); +lean_ctor_set(x_113, 2, x_47); +lean_ctor_set(x_113, 3, x_112); +lean_ctor_set(x_113, 4, x_102); +lean_ctor_set(x_113, 5, x_47); +lean_ctor_set(x_113, 6, x_38); +lean_ctor_set(x_113, 7, x_47); +x_114 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_114, 0, x_113); +x_115 = lean_unbox_float(x_42); +lean_dec(x_42); +lean_ctor_set_float(x_114, sizeof(void*)*1, x_115); +x_116 = lean_array_push(x_26, x_114); +x_117 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_117, 0, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_117); +lean_ctor_set(x_118, 1, x_24); +x_28 = x_118; +goto block_33; +} +block_138: +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; double x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_122 = lean_data_value_to_string(x_121); +x_123 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_124 = lean_string_append(x_123, x_122); +lean_dec(x_122); +x_125 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_126 = lean_string_append(x_124, x_125); +x_127 = lean_string_append(x_126, x_46); +lean_dec(x_46); +x_128 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_129 = lean_string_append(x_127, x_128); +x_130 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_130, 0, x_129); +x_131 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_132 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_132, 0, x_36); +lean_ctor_set(x_132, 1, x_130); +lean_ctor_set(x_132, 2, x_47); +lean_ctor_set(x_132, 3, x_131); +lean_ctor_set(x_132, 4, x_120); +lean_ctor_set(x_132, 5, x_47); +lean_ctor_set(x_132, 6, x_38); +lean_ctor_set(x_132, 7, x_47); +x_133 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_133, 0, x_132); +x_134 = lean_unbox_float(x_42); +lean_dec(x_42); +lean_ctor_set_float(x_133, sizeof(void*)*1, x_134); +x_135 = lean_array_push(x_26, x_133); +x_136 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_136, 0, x_135); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_24); +x_28 = x_137; +goto block_33; +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_248; lean_object* x_266; lean_object* x_267; +x_203 = lean_ctor_get(x_38, 0); +lean_inc(x_203); +lean_dec(x_38); +x_204 = lean_ctor_get(x_4, 0); +x_205 = lean_ctor_get(x_20, 1); +lean_inc(x_205); +x_206 = l_Lean_KVMap_findCore(x_8, x_19); +lean_dec(x_19); +x_207 = lean_ctor_get(x_20, 3); +lean_inc(x_207); +lean_dec(x_20); +x_208 = lean_box(0); +lean_inc(x_2); +lean_inc(x_1); +x_209 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_209, 0, x_1); +lean_ctor_set(x_209, 1, x_2); +lean_ctor_set(x_209, 2, x_208); +x_210 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(x_209); +x_211 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_211, 0, x_210); +if (lean_obj_tag(x_204) == 0) +{ +lean_object* x_285; +x_285 = lean_box(0); +x_212 = x_285; +goto block_247; +} +else +{ +lean_object* x_286; lean_object* x_287; +x_286 = lean_ctor_get(x_204, 0); +x_287 = lean_ctor_get(x_286, 0); +if (lean_obj_tag(x_287) == 0) +{ +lean_object* x_288; +x_288 = lean_box(0); +x_212 = x_288; +goto block_247; +} +else +{ +lean_object* x_289; +x_289 = lean_ctor_get(x_287, 0); +if (lean_obj_tag(x_289) == 0) +{ +lean_object* x_290; +x_290 = lean_box(0); +x_212 = x_290; +goto block_247; +} +else +{ +lean_object* x_291; +x_291 = lean_ctor_get(x_289, 0); +if (lean_obj_tag(x_291) == 0) +{ +lean_object* x_292; +x_292 = lean_box(0); +x_212 = x_292; +goto block_247; +} +else +{ +lean_object* x_293; uint8_t x_294; +x_293 = lean_ctor_get(x_291, 0); +x_294 = lean_unbox(x_293); +if (x_294 == 0) +{ +lean_object* x_295; +x_295 = lean_box(0); +x_212 = x_295; +goto block_247; +} +else +{ +uint8_t x_296; lean_object* x_297; +lean_dec(x_27); +lean_dec(x_25); +x_296 = 0; +x_297 = l_Lean_Syntax_getRange_x3f(x_5, x_296); +if (lean_obj_tag(x_297) == 0) +{ +if (lean_obj_tag(x_206) == 0) +{ +x_248 = x_208; +goto block_265; +} +else +{ +lean_object* x_298; +lean_dec(x_205); +x_298 = lean_ctor_get(x_206, 0); +lean_inc(x_298); +lean_dec(x_206); +x_266 = x_208; +x_267 = x_298; +goto block_284; +} +} +else +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; +x_299 = lean_ctor_get(x_297, 0); +lean_inc(x_299); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + x_300 = x_297; +} else { + lean_dec_ref(x_297); + x_300 = lean_box(0); +} +x_301 = lean_ctor_get(x_299, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_299, 1); +lean_inc(x_302); +if (lean_is_exclusive(x_299)) { + lean_ctor_release(x_299, 0); + lean_ctor_release(x_299, 1); + x_303 = x_299; +} else { + lean_dec_ref(x_299); + x_303 = lean_box(0); +} +x_304 = lean_ctor_get(x_3, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_304, 1); +lean_inc(x_305); +lean_dec(x_304); +lean_inc(x_305); +x_306 = l_Lean_FileMap_utf8PosToLspPos(x_305, x_301); +lean_dec(x_301); +if (x_7 == 0) +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_307 = l_Lean_FileMap_utf8PosToLspPos(x_305, x_302); +lean_dec(x_302); +if (lean_is_scalar(x_303)) { + x_308 = lean_alloc_ctor(0, 2, 0); +} else { + x_308 = x_303; +} +lean_ctor_set(x_308, 0, x_306); +lean_ctor_set(x_308, 1, x_307); +lean_inc(x_308); +lean_inc(x_36); +x_309 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_309, 0, x_36); +lean_ctor_set(x_309, 1, x_308); +lean_ctor_set(x_309, 2, x_308); +if (lean_is_scalar(x_300)) { + x_310 = lean_alloc_ctor(1, 1, 0); +} else { + x_310 = x_300; +} +lean_ctor_set(x_310, 0, x_309); +if (lean_obj_tag(x_206) == 0) +{ +x_248 = x_310; +goto block_265; +} +else +{ +lean_object* x_311; +lean_dec(x_205); +x_311 = lean_ctor_get(x_206, 0); +lean_inc(x_311); +lean_dec(x_206); +x_266 = x_310; +x_267 = x_311; +goto block_284; +} +} +else +{ +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; +x_312 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4; +x_313 = lean_nat_add(x_302, x_312); +lean_dec(x_302); +x_314 = l_Lean_FileMap_utf8PosToLspPos(x_305, x_313); +lean_dec(x_313); +if (lean_is_scalar(x_303)) { + x_315 = lean_alloc_ctor(0, 2, 0); +} else { + x_315 = x_303; +} +lean_ctor_set(x_315, 0, x_306); +lean_ctor_set(x_315, 1, x_314); +lean_inc(x_315); +lean_inc(x_36); +x_316 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_316, 0, x_36); +lean_ctor_set(x_316, 1, x_315); +lean_ctor_set(x_316, 2, x_315); +if (lean_is_scalar(x_300)) { + x_317 = lean_alloc_ctor(1, 1, 0); +} else { + x_317 = x_300; +} +lean_ctor_set(x_317, 0, x_316); +if (lean_obj_tag(x_206) == 0) +{ +x_248 = x_317; +goto block_265; +} +else +{ +lean_object* x_318; +lean_dec(x_205); +x_318 = lean_ctor_get(x_206, 0); +lean_inc(x_318); +lean_dec(x_206); +x_266 = x_317; +x_267 = x_318; +goto block_284; +} +} +} +} +} +} +} +} +block_247: +{ +lean_dec(x_212); +if (lean_obj_tag(x_206) == 0) +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; double x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_213 = lean_data_value_to_string(x_205); +x_214 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_215 = lean_string_append(x_214, x_213); +lean_dec(x_213); +x_216 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_217 = lean_string_append(x_215, x_216); +x_218 = lean_string_append(x_217, x_207); +lean_dec(x_207); +x_219 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_220 = lean_string_append(x_218, x_219); +x_221 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_221, 0, x_220); +x_222 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_223 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_223, 0, x_36); +lean_ctor_set(x_223, 1, x_221); +lean_ctor_set(x_223, 2, x_208); +lean_ctor_set(x_223, 3, x_222); +lean_ctor_set(x_223, 4, x_208); +lean_ctor_set(x_223, 5, x_208); +lean_ctor_set(x_223, 6, x_211); +lean_ctor_set(x_223, 7, x_208); +x_224 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_224, 0, x_223); +x_225 = lean_unbox_float(x_203); +lean_dec(x_203); +lean_ctor_set_float(x_224, sizeof(void*)*1, x_225); +x_226 = lean_array_push(x_26, x_224); +if (lean_is_scalar(x_27)) { + x_227 = lean_alloc_ctor(1, 1, 0); +} else { + x_227 = x_27; +} +lean_ctor_set(x_227, 0, x_226); +if (lean_is_scalar(x_25)) { + x_228 = lean_alloc_ctor(0, 2, 0); +} else { + x_228 = x_25; +} +lean_ctor_set(x_228, 0, x_227); +lean_ctor_set(x_228, 1, x_24); +x_28 = x_228; +goto block_33; +} +else +{ +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; double x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_205); +x_229 = lean_ctor_get(x_206, 0); +lean_inc(x_229); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + x_230 = x_206; +} else { + lean_dec_ref(x_206); + x_230 = lean_box(0); +} +x_231 = lean_data_value_to_string(x_229); +x_232 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_233 = lean_string_append(x_232, x_231); +lean_dec(x_231); +x_234 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_235 = lean_string_append(x_233, x_234); +x_236 = lean_string_append(x_235, x_207); +lean_dec(x_207); +x_237 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_238 = lean_string_append(x_236, x_237); +if (lean_is_scalar(x_230)) { + x_239 = lean_alloc_ctor(1, 1, 0); +} else { + x_239 = x_230; +} +lean_ctor_set(x_239, 0, x_238); +x_240 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_241 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_241, 0, x_36); +lean_ctor_set(x_241, 1, x_239); +lean_ctor_set(x_241, 2, x_208); +lean_ctor_set(x_241, 3, x_240); +lean_ctor_set(x_241, 4, x_208); +lean_ctor_set(x_241, 5, x_208); +lean_ctor_set(x_241, 6, x_211); +lean_ctor_set(x_241, 7, x_208); +x_242 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_242, 0, x_241); +x_243 = lean_unbox_float(x_203); +lean_dec(x_203); +lean_ctor_set_float(x_242, sizeof(void*)*1, x_243); +x_244 = lean_array_push(x_26, x_242); +if (lean_is_scalar(x_27)) { + x_245 = lean_alloc_ctor(1, 1, 0); +} else { + x_245 = x_27; +} +lean_ctor_set(x_245, 0, x_244); +if (lean_is_scalar(x_25)) { + x_246 = lean_alloc_ctor(0, 2, 0); +} else { + x_246 = x_25; +} +lean_ctor_set(x_246, 0, x_245); +lean_ctor_set(x_246, 1, x_24); +x_28 = x_246; +goto block_33; +} +} +block_265: +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; double x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_249 = lean_data_value_to_string(x_205); +x_250 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_251 = lean_string_append(x_250, x_249); +lean_dec(x_249); +x_252 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_253 = lean_string_append(x_251, x_252); +x_254 = lean_string_append(x_253, x_207); +lean_dec(x_207); +x_255 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_256 = lean_string_append(x_254, x_255); +x_257 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_257, 0, x_256); +x_258 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_259 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_259, 0, x_36); +lean_ctor_set(x_259, 1, x_257); +lean_ctor_set(x_259, 2, x_208); +lean_ctor_set(x_259, 3, x_258); +lean_ctor_set(x_259, 4, x_248); +lean_ctor_set(x_259, 5, x_208); +lean_ctor_set(x_259, 6, x_211); +lean_ctor_set(x_259, 7, x_208); +x_260 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_260, 0, x_259); +x_261 = lean_unbox_float(x_203); +lean_dec(x_203); +lean_ctor_set_float(x_260, sizeof(void*)*1, x_261); +x_262 = lean_array_push(x_26, x_260); +x_263 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_263, 0, x_262); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_24); +x_28 = x_264; +goto block_33; +} +block_284: +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; double x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_268 = lean_data_value_to_string(x_267); +x_269 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2; +x_270 = lean_string_append(x_269, x_268); +lean_dec(x_268); +x_271 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3; +x_272 = lean_string_append(x_270, x_271); +x_273 = lean_string_append(x_272, x_207); +lean_dec(x_207); +x_274 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_275 = lean_string_append(x_273, x_274); +x_276 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_276, 0, x_275); +x_277 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1; +x_278 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_278, 0, x_36); +lean_ctor_set(x_278, 1, x_276); +lean_ctor_set(x_278, 2, x_208); +lean_ctor_set(x_278, 3, x_277); +lean_ctor_set(x_278, 4, x_266); +lean_ctor_set(x_278, 5, x_208); +lean_ctor_set(x_278, 6, x_211); +lean_ctor_set(x_278, 7, x_208); +x_279 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_279, 0, x_278); +x_280 = lean_unbox_float(x_203); +lean_dec(x_203); +lean_ctor_set_float(x_279, sizeof(void*)*1, x_280); +x_281 = lean_array_push(x_26, x_279); +x_282 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_282, 0, x_281); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_24); +x_28 = x_283; +goto block_33; +} +} +} +block_33: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_9 = x_21; +x_10 = x_31; +x_15 = x_30; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_13 = l_Lean_getOptionDecls(x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_10, 2); +x_17 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1; +x_18 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_14, x_17, x_8, x_9, x_10, x_11, x_15); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +lean_dec(x_20); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_18, 0); +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_18); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; +} +} +} +static lean_object* _init_l_Lean_Server_Completion_optionCompletion___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(32u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Server_Completion_optionCompletion___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Server_Completion_optionCompletion___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Server_Completion_optionCompletion___closed__3() { +_start: +{ +size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 5; +x_2 = l_Lean_Server_Completion_optionCompletion___closed__2; +x_3 = l_Lean_Server_Completion_optionCompletion___closed__1; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_3); +lean_ctor_set(x_5, 2, x_4); +lean_ctor_set(x_5, 3, x_4); +lean_ctor_set_usize(x_5, 4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Server_Completion_optionCompletion___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3; +x_2 = l_Lean_Server_Completion_optionCompletion___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_optionCompletion___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; +x_7 = lean_unsigned_to_nat(1u); +x_8 = l_Lean_Syntax_getArg(x_4, x_7); +x_9 = 0; +x_10 = l_Lean_Syntax_getSubstring_x3f(x_8, x_9, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1; +x_12 = lean_box(x_9); +lean_inc(x_3); +x_13 = lean_alloc_closure((void*)(l_Lean_Server_Completion_optionCompletion___lambda__1___boxed), 12, 7); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_3); +lean_closure_set(x_13, 3, x_5); +lean_closure_set(x_13, 4, x_8); +lean_closure_set(x_13, 5, x_11); +lean_closure_set(x_13, 6, x_12); +x_14 = l_Lean_Server_Completion_optionCompletion___closed__4; +x_15 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_14, x_13, x_6); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_16 = lean_ctor_get(x_10, 0); +lean_inc(x_16); +lean_dec(x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 2); +lean_inc(x_18); +x_19 = lean_string_utf8_at_end(x_17, x_18); +if (x_19 == 0) +{ +uint32_t x_20; uint32_t x_21; uint8_t x_22; +x_20 = lean_string_utf8_get(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +x_21 = 46; +x_22 = lean_uint32_dec_eq(x_20, x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_16, 2); +lean_inc(x_25); +lean_dec(x_16); +x_26 = lean_string_utf8_extract(x_23, x_24, x_25); +lean_dec(x_25); +lean_dec(x_24); +lean_dec(x_23); +x_27 = lean_box(x_9); +lean_inc(x_3); +x_28 = lean_alloc_closure((void*)(l_Lean_Server_Completion_optionCompletion___lambda__1___boxed), 12, 7); +lean_closure_set(x_28, 0, x_1); +lean_closure_set(x_28, 1, x_2); +lean_closure_set(x_28, 2, x_3); +lean_closure_set(x_28, 3, x_5); +lean_closure_set(x_28, 4, x_8); +lean_closure_set(x_28, 5, x_26); +lean_closure_set(x_28, 6, x_27); +x_29 = l_Lean_Server_Completion_optionCompletion___closed__4; +x_30 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_29, x_28, x_6); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_31 = lean_ctor_get(x_16, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_16, 2); +lean_inc(x_33); +lean_dec(x_16); +x_34 = lean_string_utf8_extract(x_31, x_32, x_33); +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_31); +x_35 = l_Lean_Server_Completion_optionCompletion___closed__5; +x_36 = lean_string_append(x_34, x_35); +x_37 = 1; +x_38 = lean_box(x_37); +lean_inc(x_3); +x_39 = lean_alloc_closure((void*)(l_Lean_Server_Completion_optionCompletion___lambda__1___boxed), 12, 7); +lean_closure_set(x_39, 0, x_1); +lean_closure_set(x_39, 1, x_2); +lean_closure_set(x_39, 2, x_3); +lean_closure_set(x_39, 3, x_5); +lean_closure_set(x_39, 4, x_8); +lean_closure_set(x_39, 5, x_36); +lean_closure_set(x_39, 6, x_38); +x_40 = l_Lean_Server_Completion_optionCompletion___closed__4; +x_41 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_40, x_39, x_6); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_18); +lean_dec(x_17); +x_42 = lean_ctor_get(x_16, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_16, 1); +lean_inc(x_43); +x_44 = lean_ctor_get(x_16, 2); +lean_inc(x_44); +lean_dec(x_16); +x_45 = lean_string_utf8_extract(x_42, x_43, x_44); +lean_dec(x_44); +lean_dec(x_43); +lean_dec(x_42); +x_46 = lean_box(x_9); +lean_inc(x_3); +x_47 = lean_alloc_closure((void*)(l_Lean_Server_Completion_optionCompletion___lambda__1___boxed), 12, 7); +lean_closure_set(x_47, 0, x_1); +lean_closure_set(x_47, 1, x_2); +lean_closure_set(x_47, 2, x_3); +lean_closure_set(x_47, 3, x_5); +lean_closure_set(x_47, 4, x_8); +lean_closure_set(x_47, 5, x_45); +lean_closure_set(x_47, 6, x_46); +x_48 = l_Lean_Server_Completion_optionCompletion___closed__4; +x_49 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_48, x_47, x_6); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_7); +lean_dec(x_7); +x_14 = l_Lean_Server_Completion_optionCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_13, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_4); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_optionCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Server_Completion_optionCompletion(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_lt(x_4, x_3); +if (x_6 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; +x_7 = lean_array_uget(x_5, x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_array_uset(x_5, x_4, x_8); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +x_11 = lean_box(0); +x_12 = lean_ctor_get(x_7, 3); +lean_inc(x_12); +lean_dec(x_7); +lean_inc(x_2); +lean_inc(x_1); +x_13 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_2); +lean_ctor_set(x_13, 2, x_11); +x_14 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +x_16 = 1; +x_17 = lean_usize_add(x_4, x_16); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_18; lean_object* x_19; double x_20; lean_object* x_21; lean_object* x_22; +x_18 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; +x_19 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_19, 0, x_10); +lean_ctor_set(x_19, 1, x_11); +lean_ctor_set(x_19, 2, x_11); +lean_ctor_set(x_19, 3, x_18); +lean_ctor_set(x_19, 4, x_11); +lean_ctor_set(x_19, 5, x_11); +lean_ctor_set(x_19, 6, x_15); +lean_ctor_set(x_19, 7, x_11); +x_20 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_21 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set_float(x_21, sizeof(void*)*1, x_20); +x_22 = lean_array_uset(x_9, x_4, x_21); +x_4 = x_17; +x_5 = x_22; +goto _start; +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; double x_30; lean_object* x_31; lean_object* x_32; +x_25 = lean_ctor_get(x_12, 0); +x_26 = 1; +x_27 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_26); +lean_ctor_set(x_12, 0, x_27); +x_28 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; +x_29 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_29, 0, x_10); +lean_ctor_set(x_29, 1, x_11); +lean_ctor_set(x_29, 2, x_12); +lean_ctor_set(x_29, 3, x_28); +lean_ctor_set(x_29, 4, x_11); +lean_ctor_set(x_29, 5, x_11); +lean_ctor_set(x_29, 6, x_15); +lean_ctor_set(x_29, 7, x_11); +x_30 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_31 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set_float(x_31, sizeof(void*)*1, x_30); +x_32 = lean_array_uset(x_9, x_4, x_31); +x_4 = x_17; +x_5 = x_32; +goto _start; +} +else +{ +lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; double x_40; lean_object* x_41; lean_object* x_42; +x_34 = lean_ctor_get(x_12, 0); +lean_inc(x_34); +lean_dec(x_12); +x_35 = 1; +x_36 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_35); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3; +x_39 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_39, 0, x_10); +lean_ctor_set(x_39, 1, x_11); +lean_ctor_set(x_39, 2, x_37); +lean_ctor_set(x_39, 3, x_38); +lean_ctor_set(x_39, 4, x_11); +lean_ctor_set(x_39, 5, x_11); +lean_ctor_set(x_39, 6, x_15); +lean_ctor_set(x_39, 7, x_11); +x_40 = l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3; +x_41 = lean_alloc_ctor(0, 1, 8); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set_float(x_41, sizeof(void*)*1, x_40); +x_42 = lean_array_uset(x_9, x_4, x_41); +x_4 = x_17; +x_5 = x_42; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_tacticCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Elab_Tactic_Doc_allTacticDocs(x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_array_size(x_10); +x_12 = 0; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_2, x_11, x_12, x_10); +lean_ctor_set(x_8, 0, x_13); +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +x_16 = lean_array_size(x_14); +x_17 = 0; +x_18 = l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_2, x_16, x_17, x_14); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_15); +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_2); +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_8); +if (x_20 == 0) +{ +return x_8; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_8, 0); +x_22 = lean_ctor_get(x_8, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_8); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_tacticCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_alloc_closure((void*)(l_Lean_Server_Completion_tacticCompletion___lambda__1), 7, 2); +lean_closure_set(x_5, 0, x_1); +lean_closure_set(x_5, 1, x_2); +x_6 = l_Lean_LocalContext_empty; +x_7 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_3, x_6, x_5, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_mapMUnsafe_map___at_Lean_Server_Completion_tacticCompletion___spec__1(x_1, x_2, x_6, x_7, x_5); +return x_8; +} +} +lean_object* initialize_Lean_Data_FuzzyMatching(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_Tactic_Doc(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_CompletionResolution(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_EligibleHeaderDecls(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_CompletionCollectors(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Data_FuzzyMatching(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_Tactic_Doc(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Server_Completion_CompletionResolution(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Server_Completion_EligibleHeaderDecls(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1 = _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__1); +l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2 = _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2(); +lean_mark_persistent(l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__2); +l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3 = _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3(); +lean_mark_persistent(l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__3); +l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4 = _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4(); +lean_mark_persistent(l_Lean_Server_Completion_instInhabitedScoredCompletionItem___closed__4); +l_Lean_Server_Completion_instInhabitedScoredCompletionItem = _init_l_Lean_Server_Completion_instInhabitedScoredCompletionItem(); +lean_mark_persistent(l_Lean_Server_Completion_instInhabitedScoredCompletionItem); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___lambda__3___closed__3); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__3); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__4); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__5); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__6); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__7); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addUnresolvedCompletionItem___closed__8); +l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1 = _init_l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1(); +lean_mark_persistent(l_Lean_isProjectionFn___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__1___closed__1); +l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__1); +l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__2); +l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__3); +l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4 = _init_l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4(); +lean_mark_persistent(l_Lean_getConstInfo___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getCompletionKindForDecl___spec__3___closed__4); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addKeywordCompletionItem___closed__3); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_addNamespaceCompletionItem___closed__3); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_runM___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__2(); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___closed__3(); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_matchDecl_x3f___lambda__2___boxed__const__1); +l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1 = _init_l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___spec__1___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__1); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__3); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_truncate_go___closed__4); +l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1 = _init_l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1(); +lean_mark_persistent(l_Lean_Server_Completion_matchNamespace___closed__1___boxed__const__1); +l_Lean_Server_Completion_matchNamespace___closed__1 = _init_l_Lean_Server_Completion_matchNamespace___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_matchNamespace___closed__1); +l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1 = _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__1); +l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2 = _init_l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2(); +lean_mark_persistent(l_List_forIn_x27_loop___at_Lean_Server_Completion_completeNamespaces___spec__1___closed__2); +l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1 = _init_l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_completeNamespaces___lambda__3___closed__1); +l_Lean_Server_Completion_completeNamespaces___closed__1 = _init_l_Lean_Server_Completion_completeNamespaces___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_completeNamespaces___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___spec__1___closed__2); +l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1 = _init_l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_isDotCompletionMethod___lambda__2___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__16___closed__1); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__1); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__2); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__3); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__4); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__5); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__6); +l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7 = _init_l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7(); +lean_mark_persistent(l___private_Lean_Structure_0__Lean_setStructureResolutionOrder___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_getDotCompletionTypeNames_visit___spec__25___closed__7); +l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1 = _init_l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_List_forIn_x27_loop___at___private_Lean_Server_Completion_CompletionCollectors_0__Lean_Server_Completion_idCompletionCore___spec__1___lambda__1___closed__1); +l_Lean_Server_Completion_dotCompletion___closed__1 = _init_l_Lean_Server_Completion_dotCompletion___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_dotCompletion___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__1); +l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__2); +l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3 = _init_l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Completion_fieldIdCompletion___spec__1___closed__3); +l_Lean_Server_Completion_fieldIdCompletion___closed__1 = _init_l_Lean_Server_Completion_fieldIdCompletion___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_fieldIdCompletion___closed__1); +l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__1); +l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__2); +l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__3); +l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4 = _init_l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4(); +lean_mark_persistent(l_Lean_RBNode_forIn_visit___at_Lean_Server_Completion_optionCompletion___spec__1___closed__4); +l_Lean_Server_Completion_optionCompletion___closed__1 = _init_l_Lean_Server_Completion_optionCompletion___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_optionCompletion___closed__1); +l_Lean_Server_Completion_optionCompletion___closed__2 = _init_l_Lean_Server_Completion_optionCompletion___closed__2(); +lean_mark_persistent(l_Lean_Server_Completion_optionCompletion___closed__2); +l_Lean_Server_Completion_optionCompletion___closed__3 = _init_l_Lean_Server_Completion_optionCompletion___closed__3(); +lean_mark_persistent(l_Lean_Server_Completion_optionCompletion___closed__3); +l_Lean_Server_Completion_optionCompletion___closed__4 = _init_l_Lean_Server_Completion_optionCompletion___closed__4(); +lean_mark_persistent(l_Lean_Server_Completion_optionCompletion___closed__4); +l_Lean_Server_Completion_optionCompletion___closed__5 = _init_l_Lean_Server_Completion_optionCompletion___closed__5(); +lean_mark_persistent(l_Lean_Server_Completion_optionCompletion___closed__5); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionInfoSelection.c b/stage0/stdlib/Lean/Server/Completion/CompletionInfoSelection.c new file mode 100644 index 000000000000..93bfab00222d --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/CompletionInfoSelection.c @@ -0,0 +1,2962 @@ +// Lean compiler output +// Module: Lean.Server.Completion.CompletionInfoSelection +// Imports: Lean.Server.Completion.SyntheticCompletion +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___lambda__1(lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9___boxed(lean_object*, lean_object*); +lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11___boxed(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2___boxed(lean_object*, lean_object*); +uint64_t lean_uint64_of_nat(lean_object*); +uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); +size_t lean_uint64_to_usize(uint64_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__5(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions(lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___boxed(lean_object*); +lean_object* l_Lean_Server_Completion_findSyntheticCompletions(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos(lean_object*); +lean_object* l_Lean_Elab_Info_size_x3f(lean_object*); +lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9; +lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq___boxed(lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(size_t, size_t, lean_object*); +lean_object* l_Lean_Elab_Info_tailPos_x3f(lean_object*); +LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1(lean_object*, lean_object*); +uint8_t lean_expr_eqv(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1; +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2; +lean_object* lean_nat_div(lean_object*, lean_object*); +lean_object* l_instDecidableEqPos___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instHashableBool___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14(lean_object*, size_t, size_t, lean_object*); +lean_object* l_instDecidableEqBool___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1(lean_object*, lean_object*, lean_object*); +uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2852____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2; +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8; +static uint64_t l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1; +lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); +uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_zipWithIndex___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(lean_object*, lean_object*); +lean_object* l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(lean_object*); +lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12; +static lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_instHashableOption___rarg___boxed(lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt___lambda__1(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4; +uint64_t lean_uint64_xor(uint64_t, uint64_t); +static lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1; +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +size_t lean_usize_sub(size_t, size_t); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findPrioritizedCompletionPartitionsAt(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2; +size_t lean_usize_add(size_t, size_t); +uint8_t l_Lean_Syntax_eqWithInfo(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___boxed(lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Server_Completion_findCompletionInfosAt___closed__1; +lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* l_instHashablePos___boxed(lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +uint8_t l_Lean_Elab_Info_occursInOrOnBoundary(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7(lean_object*, lean_object*); +static lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13; +size_t lean_usize_land(size_t, size_t); +static lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3; +uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = l_Lean_Syntax_eqWithInfo(x_6, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_4); +lean_dec(x_3); +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_3, 3); +lean_inc(x_11); +lean_dec(x_3); +x_12 = lean_ctor_get(x_4, 3); +lean_inc(x_12); +lean_dec(x_4); +x_13 = lean_expr_eqv(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_2); +lean_dec(x_1); +x_14 = 0; +return x_14; +} +} +case 3: +{ +if (lean_obj_tag(x_2) == 3) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_1, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_1, 3); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 3); +lean_inc(x_20); +lean_dec(x_2); +x_21 = l_Lean_Syntax_eqWithInfo(x_15, x_18); +if (x_21 == 0) +{ +uint8_t x_22; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_17); +lean_dec(x_16); +x_22 = 0; +return x_22; +} +else +{ +uint8_t x_23; +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2(x_16, x_19); +lean_dec(x_19); +lean_dec(x_16); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_20); +lean_dec(x_17); +x_24 = 0; +return x_24; +} +else +{ +uint8_t x_25; +x_25 = lean_name_eq(x_17, x_20); +lean_dec(x_20); +lean_dec(x_17); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_2); +lean_dec(x_1); +x_26 = 0; +return x_26; +} +} +case 4: +{ +if (lean_obj_tag(x_2) == 4) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_ctor_get(x_2, 0); +lean_inc(x_28); +lean_dec(x_2); +x_29 = l_Lean_Syntax_eqWithInfo(x_27, x_28); +return x_29; +} +else +{ +uint8_t x_30; +lean_dec(x_2); +lean_dec(x_1); +x_30 = 0; +return x_30; +} +} +case 5: +{ +if (lean_obj_tag(x_2) == 5) +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_ctor_get(x_1, 0); +lean_inc(x_31); +lean_dec(x_1); +x_32 = lean_ctor_get(x_2, 0); +lean_inc(x_32); +lean_dec(x_2); +x_33 = l_Lean_Syntax_eqWithInfo(x_31, x_32); +return x_33; +} +else +{ +uint8_t x_34; +lean_dec(x_2); +lean_dec(x_1); +x_34 = 0; +return x_34; +} +} +case 6: +{ +if (lean_obj_tag(x_2) == 6) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_35 = lean_ctor_get(x_1, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_1, 1); +lean_inc(x_36); +lean_dec(x_1); +x_37 = lean_ctor_get(x_2, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_2, 1); +lean_inc(x_38); +lean_dec(x_2); +x_39 = l_Lean_Syntax_eqWithInfo(x_35, x_37); +if (x_39 == 0) +{ +uint8_t x_40; +lean_dec(x_38); +lean_dec(x_36); +x_40 = 0; +return x_40; +} +else +{ +uint8_t x_41; +x_41 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2852____spec__1(x_36, x_38); +lean_dec(x_38); +lean_dec(x_36); +return x_41; +} +} +else +{ +uint8_t x_42; +lean_dec(x_2); +lean_dec(x_1); +x_42 = 0; +return x_42; +} +} +case 7: +{ +if (lean_obj_tag(x_2) == 7) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_1, 0); +lean_inc(x_43); +lean_dec(x_1); +x_44 = lean_ctor_get(x_2, 0); +lean_inc(x_44); +lean_dec(x_2); +x_45 = l_Lean_Syntax_eqWithInfo(x_43, x_44); +return x_45; +} +else +{ +uint8_t x_46; +lean_dec(x_2); +lean_dec(x_1); +x_46 = 0; +return x_46; +} +} +default: +{ +if (lean_obj_tag(x_2) == 1) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_47 = lean_ctor_get(x_1, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_1, 1); +lean_inc(x_48); +lean_dec(x_1); +x_49 = lean_ctor_get(x_2, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_2, 1); +lean_inc(x_50); +lean_dec(x_2); +x_51 = l_Lean_Syntax_eqWithInfo(x_47, x_49); +if (x_51 == 0) +{ +uint8_t x_52; +lean_dec(x_50); +lean_dec(x_48); +x_52 = 0; +return x_52; +} +else +{ +uint8_t x_53; +x_53 = lean_name_eq(x_48, x_50); +lean_dec(x_50); +lean_dec(x_48); +return x_53; +} +} +else +{ +uint8_t x_54; +lean_dec(x_2); +lean_dec(x_1); +x_54 = 0; +return x_54; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_ctor_get(x_6, 2); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos_eq(x_7, x_8); +if (x_9 == 0) +{ +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +goto _start; +} +else +{ +uint8_t x_13; +lean_dec(x_1); +x_13 = 1; +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = 0; +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_array_push(x_2, x_1); +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_array_uget(x_3, x_5); +x_9 = lean_array_get_size(x_6); +x_10 = lean_unsigned_to_nat(0u); +x_11 = lean_nat_dec_lt(x_10, x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; +lean_dec(x_9); +x_12 = lean_box(0); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1(x_8, x_6, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = 1; +x_16 = lean_usize_add(x_5, x_15); +x_5 = x_16; +x_6 = x_14; +goto _start; +} +else +{ +size_t x_18; size_t x_19; uint8_t x_20; +x_18 = 0; +x_19 = lean_usize_of_nat(x_9); +lean_dec(x_9); +lean_inc(x_8); +x_20 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1(x_8, x_6, x_18, x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; +x_21 = lean_box(0); +x_22 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1(x_8, x_6, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = 1; +x_25 = lean_usize_add(x_5, x_24); +x_5 = x_25; +x_6 = x_23; +goto _start; +} +else +{ +size_t x_27; size_t x_28; +lean_dec(x_8); +x_27 = 1; +x_28 = lean_usize_add(x_5, x_27); +x_5 = x_28; +goto _start; +} +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos(lean_object* x_1) { +_start: +{ +lean_object* x_2; size_t x_3; size_t x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_box(0); +x_3 = lean_array_size(x_1); +x_4 = 0; +x_5 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_6 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2(x_1, x_2, x_1, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__1(x_1, x_2, x_5, x_6); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___spec__2(x_1, x_2, x_3, x_7, x_8, x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_box(0); +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_2); +x_7 = lean_array_push(x_3, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +x_7 = lean_array_push(x_4, x_6); +return x_7; +} +} +static lean_object* _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1; +x_2 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Lean_Elab_Info_pos_x3f(x_1); +x_10 = l_Lean_Elab_Info_tailPos_x3f(x_1); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4; +x_37 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_36); +x_11 = x_37; +goto block_35; +} +else +{ +lean_object* x_38; +x_38 = lean_ctor_get(x_9, 0); +lean_inc(x_38); +lean_dec(x_9); +x_11 = x_38; +goto block_35; +} +block_35: +{ +lean_object* x_12; lean_object* x_13; +lean_inc(x_2); +x_12 = l_Lean_FileMap_toPosition(x_2, x_11); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4; +x_33 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_32); +x_13 = x_33; +goto block_31; +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_10, 0); +lean_inc(x_34); +lean_dec(x_10); +x_13 = x_34; +goto block_31; +} +block_31: +{ +uint8_t x_14; +x_14 = lean_nat_dec_lt(x_3, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_11); +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +x_16 = l_Lean_FileMap_toPosition(x_2, x_13); +lean_dec(x_13); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_nat_dec_eq(x_15, x_7); +if (x_18 == 0) +{ +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_5); +lean_dec(x_4); +return x_6; +} +else +{ +uint8_t x_19; +x_19 = lean_nat_dec_eq(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +if (x_19 == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_6; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_box(0); +x_21 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1(x_4, x_5, x_6, x_20); +return x_21; +} +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_22 = lean_nat_sub(x_3, x_11); +lean_dec(x_11); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +lean_dec(x_12); +x_25 = l_Lean_FileMap_toPosition(x_2, x_13); +lean_dec(x_13); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_nat_dec_eq(x_24, x_7); +if (x_27 == 0) +{ +lean_dec(x_26); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_5); +lean_dec(x_4); +return x_6; +} +else +{ +uint8_t x_28; +x_28 = lean_nat_dec_eq(x_24, x_26); +lean_dec(x_26); +lean_dec(x_24); +if (x_28 == 0) +{ +lean_dec(x_23); +lean_dec(x_5); +lean_dec(x_4); +return x_6; +} +else +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_box(0); +x_30 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2(x_23, x_4, x_5, x_6, x_29); +return x_30; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +if (lean_obj_tag(x_5) == 7) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +x_8 = l_Lean_Elab_Info_occursInOrOnBoundary(x_5, x_2); +if (x_8 == 0) +{ +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3(x_5, x_1, x_2, x_4, x_7, x_6, x_3, x_9); +lean_dec(x_5); +return x_10; +} +} +else +{ +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Server_Completion_findCompletionInfosAt_go(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_findCompletionInfosAt___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Server_Completion_findCompletionInfosAt___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +lean_inc(x_1); +x_5 = l_Lean_FileMap_toPosition(x_1, x_2); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +lean_inc(x_2); +lean_inc(x_1); +x_7 = lean_alloc_closure((void*)(l_Lean_Server_Completion_findCompletionInfosAt_go___boxed), 6, 3); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_2); +lean_closure_set(x_7, 2, x_6); +x_8 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +lean_inc(x_4); +x_9 = l_Lean_Elab_InfoTree_foldInfo___rarg(x_7, x_8, x_4); +x_10 = l_Lean_Server_Completion_findCompletionInfosAt___closed__1; +x_11 = l_Array_isEmpty___rarg(x_9); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(0); +x_13 = lean_apply_2(x_10, x_9, x_12); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_9); +x_14 = l_Lean_Server_Completion_findSyntheticCompletions(x_1, x_2, x_3, x_4); +x_15 = lean_box(0); +x_16 = lean_apply_2(x_10, x_14, x_15); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findCompletionInfosAt___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Server_Completion_findCompletionInfosAt___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_4, 0); +x_6 = lean_unbox(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_unbox(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_ctor_get(x_2, 2); +x_11 = lean_ctor_get(x_4, 1); +x_12 = lean_ctor_get(x_1, 1); +x_13 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_11, x_12); +if (x_13 == 0) +{ +x_2 = x_10; +goto _start; +} +else +{ +lean_object* x_15; +lean_inc(x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_9); +return x_15; +} +} +else +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_2, 2); +x_2 = x_16; +goto _start; +} +} +else +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_1, 0); +x_19 = lean_unbox(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_2, 2); +x_2 = x_20; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_ctor_get(x_2, 2); +x_24 = lean_ctor_get(x_4, 1); +x_25 = lean_ctor_get(x_1, 1); +x_26 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_24, x_25); +if (x_26 == 0) +{ +x_2 = x_23; +goto _start; +} +else +{ +lean_object* x_28; +lean_inc(x_22); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_22); +return x_28; +} +} +} +} +} +} +LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_4, 0); +x_6 = lean_unbox(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_unbox(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_2, 2); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_ctor_get(x_1, 1); +x_12 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_10, x_11); +if (x_12 == 0) +{ +x_2 = x_9; +goto _start; +} +else +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +} +else +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_2, 2); +x_2 = x_15; +goto _start; +} +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_unbox(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_2, 2); +x_2 = x_19; +goto _start; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_2, 2); +x_22 = lean_ctor_get(x_4, 1); +x_23 = lean_ctor_get(x_1, 1); +x_24 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_22, x_23); +if (x_24 == 0) +{ +x_2 = x_21; +goto _start; +} +else +{ +uint8_t x_26; +x_26 = 1; +return x_26; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; size_t x_16; size_t x_17; size_t x_18; size_t x_19; size_t x_20; lean_object* x_21; lean_object* x_22; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 2); +x_7 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_5); +x_8 = lean_apply_1(x_1, x_5); +x_9 = lean_unbox_uint64(x_8); +lean_dec(x_8); +x_10 = 32; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = 16; +x_14 = lean_uint64_shift_right(x_12, x_13); +x_15 = lean_uint64_xor(x_12, x_14); +x_16 = lean_uint64_to_usize(x_15); +x_17 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_18 = 1; +x_19 = lean_usize_sub(x_17, x_18); +x_20 = lean_usize_land(x_16, x_19); +x_21 = lean_array_uget(x_2, x_20); +lean_ctor_set(x_3, 2, x_21); +x_22 = lean_array_uset(x_2, x_20, x_3); +x_2 = x_22; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint64_t x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; uint64_t x_34; uint64_t x_35; size_t x_36; size_t x_37; size_t x_38; size_t x_39; size_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +x_26 = lean_ctor_get(x_3, 2); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_27 = lean_array_get_size(x_2); +lean_inc(x_1); +lean_inc(x_24); +x_28 = lean_apply_1(x_1, x_24); +x_29 = lean_unbox_uint64(x_28); +lean_dec(x_28); +x_30 = 32; +x_31 = lean_uint64_shift_right(x_29, x_30); +x_32 = lean_uint64_xor(x_29, x_31); +x_33 = 16; +x_34 = lean_uint64_shift_right(x_32, x_33); +x_35 = lean_uint64_xor(x_32, x_34); +x_36 = lean_uint64_to_usize(x_35); +x_37 = lean_usize_of_nat(x_27); +lean_dec(x_27); +x_38 = 1; +x_39 = lean_usize_sub(x_37, x_38); +x_40 = lean_usize_land(x_36, x_39); +x_41 = lean_array_uget(x_2, x_40); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_24); +lean_ctor_set(x_42, 1, x_25); +lean_ctor_set(x_42, 2, x_41); +x_43 = lean_array_uset(x_2, x_40, x_42); +x_2 = x_43; +x_3 = x_26; +goto _start; +} +} +} +} +static uint64_t _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1() { +_start: +{ +uint64_t x_1; uint64_t x_2; uint64_t x_3; +x_1 = 13; +x_2 = 11; +x_3 = lean_uint64_mix_hash(x_1, x_2); +return x_3; +} +} +static uint64_t _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2() { +_start: +{ +uint64_t x_1; uint64_t x_2; +x_1 = 11; +x_2 = lean_uint64_mix_hash(x_1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; size_t x_10; size_t x_11; size_t x_12; uint64_t x_13; lean_object* x_25; uint8_t x_26; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + x_6 = x_2; +} else { + lean_dec_ref(x_2); + x_6 = lean_box(0); +} +x_7 = lean_array_get_size(x_1); +x_8 = 32; +x_9 = 16; +x_10 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_11 = 1; +x_12 = lean_usize_sub(x_10, x_11); +x_25 = lean_ctor_get(x_3, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_3, 1); +lean_inc(x_27); +if (lean_obj_tag(x_27) == 0) +{ +uint64_t x_28; +x_28 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1; +x_13 = x_28; +goto block_24; +} +else +{ +lean_object* x_29; uint64_t x_30; uint64_t x_31; uint64_t x_32; uint64_t x_33; +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = 13; +x_31 = lean_uint64_of_nat(x_29); +lean_dec(x_29); +x_32 = lean_uint64_mix_hash(x_31, x_30); +x_33 = lean_uint64_mix_hash(x_30, x_32); +x_13 = x_33; +goto block_24; +} +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_3, 1); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +uint64_t x_35; +x_35 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2; +x_13 = x_35; +goto block_24; +} +else +{ +lean_object* x_36; uint64_t x_37; uint64_t x_38; uint64_t x_39; uint64_t x_40; uint64_t x_41; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = 11; +x_38 = lean_uint64_of_nat(x_36); +lean_dec(x_36); +x_39 = 13; +x_40 = lean_uint64_mix_hash(x_38, x_39); +x_41 = lean_uint64_mix_hash(x_37, x_40); +x_13 = x_41; +goto block_24; +} +} +block_24: +{ +uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = lean_uint64_shift_right(x_13, x_8); +x_15 = lean_uint64_xor(x_13, x_14); +x_16 = lean_uint64_shift_right(x_15, x_9); +x_17 = lean_uint64_xor(x_15, x_16); +x_18 = lean_uint64_to_usize(x_17); +x_19 = lean_usize_land(x_18, x_12); +x_20 = lean_array_uget(x_1, x_19); +if (lean_is_scalar(x_6)) { + x_21 = lean_alloc_ctor(1, 3, 0); +} else { + x_21 = x_6; +} +lean_ctor_set(x_21, 0, x_3); +lean_ctor_set(x_21, 1, x_4); +lean_ctor_set(x_21, 2, x_20); +x_22 = lean_array_uset(x_1, x_19, x_21); +x_1 = x_22; +x_2 = x_5; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_array_get_size(x_1); +x_3 = lean_unsigned_to_nat(2u); +x_4 = lean_nat_mul(x_2, x_3); +lean_dec(x_2); +x_5 = lean_box(0); +x_6 = lean_mk_array(x_4, x_5); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Std_DHashMap_Internal_Raw_u2080_expand_go___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__5(x_7, x_1, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_unbox(x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_3, 2); +x_13 = lean_ctor_get(x_3, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_5, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_14, x_15); +lean_dec(x_15); +lean_dec(x_14); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_12); +lean_ctor_set(x_3, 2, x_17); +return x_3; +} +else +{ +lean_dec(x_11); +lean_dec(x_5); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_3, 1); +x_19 = lean_ctor_get(x_3, 2); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_3); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_1, 1); +lean_inc(x_21); +x_22 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_19); +x_24 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_24, 0, x_5); +lean_ctor_set(x_24, 1, x_18); +lean_ctor_set(x_24, 2, x_23); +return x_24; +} +else +{ +lean_object* x_25; +lean_dec(x_18); +lean_dec(x_5); +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_2); +lean_ctor_set(x_25, 2, x_19); +return x_25; +} +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_3); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_3, 2); +x_28 = lean_ctor_get(x_3, 0); +lean_dec(x_28); +x_29 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_27); +lean_ctor_set(x_3, 2, x_29); +return x_3; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_3, 1); +x_31 = lean_ctor_get(x_3, 2); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_3); +x_32 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_31); +x_33 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_33, 0, x_5); +lean_ctor_set(x_33, 1, x_30); +lean_ctor_set(x_33, 2, x_32); +return x_33; +} +} +} +else +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_1, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_3); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_3, 2); +x_38 = lean_ctor_get(x_3, 0); +lean_dec(x_38); +x_39 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_37); +lean_ctor_set(x_3, 2, x_39); +return x_3; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_3, 1); +x_41 = lean_ctor_get(x_3, 2); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_3); +x_42 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_41); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_5); +lean_ctor_set(x_43, 1, x_40); +lean_ctor_set(x_43, 2, x_42); +return x_43; +} +} +else +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_3); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_45 = lean_ctor_get(x_3, 1); +x_46 = lean_ctor_get(x_3, 2); +x_47 = lean_ctor_get(x_3, 0); +lean_dec(x_47); +x_48 = lean_ctor_get(x_5, 1); +lean_inc(x_48); +x_49 = lean_ctor_get(x_1, 1); +lean_inc(x_49); +x_50 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_48, x_49); +lean_dec(x_49); +lean_dec(x_48); +if (x_50 == 0) +{ +lean_object* x_51; +x_51 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_46); +lean_ctor_set(x_3, 2, x_51); +return x_3; +} +else +{ +lean_dec(x_45); +lean_dec(x_5); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_3, 1); +x_53 = lean_ctor_get(x_3, 2); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_3); +x_54 = lean_ctor_get(x_5, 1); +lean_inc(x_54); +x_55 = lean_ctor_get(x_1, 1); +lean_inc(x_55); +x_56 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; +x_57 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_1, x_2, x_53); +x_58 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_58, 0, x_5); +lean_ctor_set(x_58, 1, x_52); +lean_ctor_set(x_58, 2, x_57); +return x_58; +} +else +{ +lean_object* x_59; +lean_dec(x_52); +lean_dec(x_5); +x_59 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_59, 0, x_1); +lean_ctor_set(x_59, 1, x_2); +lean_ctor_set(x_59, 2, x_53); +return x_59; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_unbox(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_unbox(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_ctor_get(x_2, 1); +x_11 = lean_ctor_get(x_2, 2); +x_12 = lean_ctor_get(x_2, 0); +lean_dec(x_12); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_1, 1); +x_15 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_11); +lean_ctor_set(x_2, 2, x_16); +return x_2; +} +else +{ +lean_free_object(x_2); +lean_dec(x_10); +lean_dec(x_4); +return x_11; +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_2, 1); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_2); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_1, 1); +x_21 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_19, x_20); +lean_dec(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_18); +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_4); +lean_ctor_set(x_23, 1, x_17); +lean_ctor_set(x_23, 2, x_22); +return x_23; +} +else +{ +lean_dec(x_17); +lean_dec(x_4); +return x_18; +} +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_2); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_2, 2); +x_26 = lean_ctor_get(x_2, 0); +lean_dec(x_26); +x_27 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_25); +lean_ctor_set(x_2, 2, x_27); +return x_2; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_2, 1); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_2); +x_30 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_29); +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_4); +lean_ctor_set(x_31, 1, x_28); +lean_ctor_set(x_31, 2, x_30); +return x_31; +} +} +} +else +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_1, 0); +x_33 = lean_unbox(x_32); +if (x_33 == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_2); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_2, 2); +x_36 = lean_ctor_get(x_2, 0); +lean_dec(x_36); +x_37 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_35); +lean_ctor_set(x_2, 2, x_37); +return x_2; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_2, 1); +x_39 = lean_ctor_get(x_2, 2); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_2); +x_40 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_39); +x_41 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_41, 0, x_4); +lean_ctor_set(x_41, 1, x_38); +lean_ctor_set(x_41, 2, x_40); +return x_41; +} +} +else +{ +uint8_t x_42; +x_42 = !lean_is_exclusive(x_2); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_43 = lean_ctor_get(x_2, 1); +x_44 = lean_ctor_get(x_2, 2); +x_45 = lean_ctor_get(x_2, 0); +lean_dec(x_45); +x_46 = lean_ctor_get(x_4, 1); +lean_inc(x_46); +x_47 = lean_ctor_get(x_1, 1); +x_48 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_46, x_47); +lean_dec(x_46); +if (x_48 == 0) +{ +lean_object* x_49; +x_49 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_44); +lean_ctor_set(x_2, 2, x_49); +return x_2; +} +else +{ +lean_free_object(x_2); +lean_dec(x_43); +lean_dec(x_4); +return x_44; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_50 = lean_ctor_get(x_2, 1); +x_51 = lean_ctor_get(x_2, 2); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_2); +x_52 = lean_ctor_get(x_4, 1); +lean_inc(x_52); +x_53 = lean_ctor_get(x_1, 1); +x_54 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_System_FilePath_parent___spec__1(x_52, x_53); +lean_dec(x_52); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; +x_55 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_51); +x_56 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_56, 0, x_4); +lean_ctor_set(x_56, 1, x_50); +lean_ctor_set(x_56, 2, x_55); +return x_56; +} +else +{ +lean_dec(x_50); +lean_dec(x_4); +return x_51; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, size_t x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_9, x_8); +if (x_11 == 0) +{ +lean_dec(x_5); +lean_dec(x_1); +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint64_t x_18; uint64_t x_19; size_t x_20; size_t x_21; size_t x_22; uint64_t x_23; lean_object* x_131; uint8_t x_132; +x_12 = lean_array_uget(x_7, x_9); +lean_inc(x_1); +lean_inc(x_12); +x_13 = lean_apply_1(x_1, x_12); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_15); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + x_16 = x_10; +} else { + lean_dec_ref(x_10); + x_16 = lean_box(0); +} +x_17 = lean_array_get_size(x_15); +x_18 = 32; +x_19 = 16; +x_20 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_21 = 1; +x_22 = lean_usize_sub(x_20, x_21); +x_131 = lean_ctor_get(x_13, 0); +lean_inc(x_131); +x_132 = lean_unbox(x_131); +lean_dec(x_131); +if (x_132 == 0) +{ +lean_object* x_133; +x_133 = lean_ctor_get(x_13, 1); +lean_inc(x_133); +if (lean_obj_tag(x_133) == 0) +{ +uint64_t x_134; +x_134 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1; +x_23 = x_134; +goto block_130; +} +else +{ +lean_object* x_135; uint64_t x_136; uint64_t x_137; uint64_t x_138; uint64_t x_139; +x_135 = lean_ctor_get(x_133, 0); +lean_inc(x_135); +lean_dec(x_133); +x_136 = 13; +x_137 = lean_uint64_of_nat(x_135); +lean_dec(x_135); +x_138 = lean_uint64_mix_hash(x_137, x_136); +x_139 = lean_uint64_mix_hash(x_136, x_138); +x_23 = x_139; +goto block_130; +} +} +else +{ +lean_object* x_140; +x_140 = lean_ctor_get(x_13, 1); +lean_inc(x_140); +if (lean_obj_tag(x_140) == 0) +{ +uint64_t x_141; +x_141 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2; +x_23 = x_141; +goto block_130; +} +else +{ +lean_object* x_142; uint64_t x_143; uint64_t x_144; uint64_t x_145; uint64_t x_146; uint64_t x_147; +x_142 = lean_ctor_get(x_140, 0); +lean_inc(x_142); +lean_dec(x_140); +x_143 = 11; +x_144 = lean_uint64_of_nat(x_142); +lean_dec(x_142); +x_145 = 13; +x_146 = lean_uint64_mix_hash(x_144, x_145); +x_147 = lean_uint64_mix_hash(x_143, x_146); +x_23 = x_147; +goto block_130; +} +} +block_130: +{ +uint64_t x_24; uint64_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_uint64_shift_right(x_23, x_18); +x_25 = lean_uint64_xor(x_23, x_24); +x_26 = lean_uint64_shift_right(x_25, x_19); +x_27 = lean_uint64_xor(x_25, x_26); +x_28 = lean_uint64_to_usize(x_27); +x_29 = lean_usize_land(x_28, x_22); +x_30 = lean_array_uget(x_15, x_29); +x_31 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2(x_13, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_33 = lean_array_push(x_32, x_12); +x_34 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(x_13, x_30); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_35 = lean_unsigned_to_nat(1u); +x_36 = lean_nat_add(x_14, x_35); +lean_dec(x_14); +x_37 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_37, 0, x_13); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_30); +x_38 = lean_array_uset(x_15, x_29, x_37); +x_39 = lean_unsigned_to_nat(4u); +x_40 = lean_nat_mul(x_36, x_39); +x_41 = lean_unsigned_to_nat(3u); +x_42 = lean_nat_div(x_40, x_41); +lean_dec(x_40); +x_43 = lean_array_get_size(x_38); +x_44 = lean_nat_dec_le(x_42, x_43); +lean_dec(x_43); +lean_dec(x_42); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; size_t x_47; +x_45 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__4(x_38); +if (lean_is_scalar(x_16)) { + x_46 = lean_alloc_ctor(0, 2, 0); +} else { + x_46 = x_16; +} +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_usize_add(x_9, x_21); +x_9 = x_47; +x_10 = x_46; +goto _start; +} +else +{ +lean_object* x_49; size_t x_50; +if (lean_is_scalar(x_16)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_16; +} +lean_ctor_set(x_49, 0, x_36); +lean_ctor_set(x_49, 1, x_38); +x_50 = lean_usize_add(x_9, x_21); +x_9 = x_50; +x_10 = x_49; +goto _start; +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; size_t x_56; +lean_inc(x_5); +x_52 = lean_array_uset(x_15, x_29, x_5); +x_53 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_13, x_33, x_30); +x_54 = lean_array_uset(x_52, x_29, x_53); +if (lean_is_scalar(x_16)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_16; +} +lean_ctor_set(x_55, 0, x_14); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_usize_add(x_9, x_21); +x_9 = x_56; +x_10 = x_55; +goto _start; +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_122; +x_58 = lean_ctor_get(x_31, 0); +lean_inc(x_58); +lean_dec(x_31); +x_59 = lean_array_push(x_58, x_12); +x_122 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(x_13, x_30); +if (x_122 == 0) +{ +lean_object* x_123; +lean_dec(x_30); +if (lean_is_scalar(x_16)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_16; +} +lean_ctor_set(x_123, 0, x_14); +lean_ctor_set(x_123, 1, x_15); +x_60 = x_123; +goto block_121; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_inc(x_5); +x_124 = lean_array_uset(x_15, x_29, x_5); +x_125 = lean_unsigned_to_nat(1u); +x_126 = lean_nat_sub(x_14, x_125); +lean_dec(x_14); +x_127 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_13, x_30); +x_128 = lean_array_uset(x_124, x_29, x_127); +if (lean_is_scalar(x_16)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_16; +} +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_128); +x_60 = x_129; +goto block_121; +} +block_121: +{ +uint8_t x_61; +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; size_t x_65; size_t x_66; size_t x_67; lean_object* x_68; uint8_t x_69; +x_62 = lean_ctor_get(x_60, 0); +x_63 = lean_ctor_get(x_60, 1); +x_64 = lean_array_get_size(x_63); +x_65 = lean_usize_of_nat(x_64); +lean_dec(x_64); +x_66 = lean_usize_sub(x_65, x_21); +x_67 = lean_usize_land(x_28, x_66); +x_68 = lean_array_uget(x_63, x_67); +x_69 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(x_13, x_68); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_70 = lean_unsigned_to_nat(1u); +x_71 = lean_nat_add(x_62, x_70); +lean_dec(x_62); +x_72 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_72, 0, x_13); +lean_ctor_set(x_72, 1, x_59); +lean_ctor_set(x_72, 2, x_68); +x_73 = lean_array_uset(x_63, x_67, x_72); +x_74 = lean_unsigned_to_nat(4u); +x_75 = lean_nat_mul(x_71, x_74); +x_76 = lean_unsigned_to_nat(3u); +x_77 = lean_nat_div(x_75, x_76); +lean_dec(x_75); +x_78 = lean_array_get_size(x_73); +x_79 = lean_nat_dec_le(x_77, x_78); +lean_dec(x_78); +lean_dec(x_77); +if (x_79 == 0) +{ +lean_object* x_80; size_t x_81; +x_80 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__4(x_73); +lean_ctor_set(x_60, 1, x_80); +lean_ctor_set(x_60, 0, x_71); +x_81 = lean_usize_add(x_9, x_21); +x_9 = x_81; +x_10 = x_60; +goto _start; +} +else +{ +size_t x_83; +lean_ctor_set(x_60, 1, x_73); +lean_ctor_set(x_60, 0, x_71); +x_83 = lean_usize_add(x_9, x_21); +x_9 = x_83; +x_10 = x_60; +goto _start; +} +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; size_t x_88; +lean_inc(x_5); +x_85 = lean_array_uset(x_63, x_67, x_5); +x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_13, x_59, x_68); +x_87 = lean_array_uset(x_85, x_67, x_86); +lean_ctor_set(x_60, 1, x_87); +x_88 = lean_usize_add(x_9, x_21); +x_9 = x_88; +x_10 = x_60; +goto _start; +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; size_t x_93; size_t x_94; size_t x_95; lean_object* x_96; uint8_t x_97; +x_90 = lean_ctor_get(x_60, 0); +x_91 = lean_ctor_get(x_60, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_60); +x_92 = lean_array_get_size(x_91); +x_93 = lean_usize_of_nat(x_92); +lean_dec(x_92); +x_94 = lean_usize_sub(x_93, x_21); +x_95 = lean_usize_land(x_28, x_94); +x_96 = lean_array_uget(x_91, x_95); +x_97 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(x_13, x_96); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_98 = lean_unsigned_to_nat(1u); +x_99 = lean_nat_add(x_90, x_98); +lean_dec(x_90); +x_100 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_100, 0, x_13); +lean_ctor_set(x_100, 1, x_59); +lean_ctor_set(x_100, 2, x_96); +x_101 = lean_array_uset(x_91, x_95, x_100); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_mul(x_99, x_102); +x_104 = lean_unsigned_to_nat(3u); +x_105 = lean_nat_div(x_103, x_104); +lean_dec(x_103); +x_106 = lean_array_get_size(x_101); +x_107 = lean_nat_dec_le(x_105, x_106); +lean_dec(x_106); +lean_dec(x_105); +if (x_107 == 0) +{ +lean_object* x_108; lean_object* x_109; size_t x_110; +x_108 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__4(x_101); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_99); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_usize_add(x_9, x_21); +x_9 = x_110; +x_10 = x_109; +goto _start; +} +else +{ +lean_object* x_112; size_t x_113; +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_99); +lean_ctor_set(x_112, 1, x_101); +x_113 = lean_usize_add(x_9, x_21); +x_9 = x_113; +x_10 = x_112; +goto _start; +} +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; size_t x_119; +lean_inc(x_5); +x_115 = lean_array_uset(x_91, x_95, x_5); +x_116 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__8(x_13, x_59, x_96); +x_117 = lean_array_uset(x_115, x_95, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_90); +lean_ctor_set(x_118, 1, x_117); +x_119 = lean_usize_add(x_9, x_21); +x_9 = x_119; +x_10 = x_118; +goto _start; +} +} +} +} +} +} +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instDecidableEqBool___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instDecidableEqPos___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instHashablePos___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7; +x_2 = lean_alloc_closure((void*)(l_instHashableOption___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instHashableBool___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9; +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; size_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_array_size(x_2); +x_6 = 0; +x_7 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6; +x_8 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10; +x_9 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13; +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10(x_1, x_2, x_7, x_8, x_3, x_4, x_2, x_5, x_6, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_4); +lean_inc(x_3); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_4); +x_7 = lean_array_push(x_1, x_6); +x_1 = x_7; +x_2 = x_5; +goto _start; +} +} +} +LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_3, 1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_5, 1); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_unbox(x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_unbox(x_9); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_13, 1); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +x_15 = 1; +return x_15; +} +else +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_unbox(x_16); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_13, 0); +x_19 = lean_unbox(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_4, 0); +x_21 = lean_ctor_get(x_14, 0); +x_22 = lean_nat_dec_lt(x_20, x_21); +return x_22; +} +else +{ +uint8_t x_23; +x_23 = 1; +return x_23; +} +} +else +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_unbox(x_24); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = 0; +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = lean_ctor_get(x_4, 0); +x_28 = lean_ctor_get(x_14, 0); +x_29 = lean_nat_dec_lt(x_27, x_28); +return x_29; +} +} +} +} +} +} +static lean_object* _init_l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) +{ +lean_dec(x_2); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1; +lean_inc(x_2); +x_6 = l_Array_qpartition___rarg(x_1, x_5, x_2, x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_nat_dec_le(x_3, x_7); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(x_8, x_2, x_7); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_7, x_11); +lean_dec(x_7); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +else +{ +lean_dec(x_7); +lean_dec(x_2); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(size_t x_1, size_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_lt(x_2, x_1); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_5 = lean_array_uget(x_3, x_2); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_array_uset(x_3, x_2, x_6); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = 1; +x_10 = lean_usize_add(x_2, x_9); +x_11 = lean_array_uset(x_7, x_2, x_8); +x_2 = x_10; +x_3 = x_11; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11(x_4, x_6); +lean_dec(x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +x_4 = x_7; +goto _start; +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +lean_dec(x_4); +x_5 = lean_ctor_get(x_3, 2); +lean_inc(x_5); +lean_dec(x_3); +lean_inc(x_5); +x_6 = lean_alloc_ctor(7, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = l_Lean_Elab_Info_size_x3f(x_6); +lean_dec(x_6); +if (lean_obj_tag(x_5) == 1) +{ +uint8_t x_8; lean_object* x_9; +lean_dec(x_5); +x_8 = 1; +x_9 = lean_box(x_8); +lean_ctor_set(x_1, 1, x_7); +lean_ctor_set(x_1, 0, x_9); +return x_1; +} +else +{ +uint8_t x_10; lean_object* x_11; +lean_dec(x_5); +x_10 = 0; +x_11 = lean_box(x_10); +lean_ctor_set(x_1, 1, x_7); +lean_ctor_set(x_1, 0, x_11); +return x_1; +} +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_12, 2); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_13); +x_14 = lean_alloc_ctor(7, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = l_Lean_Elab_Info_size_x3f(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_13) == 1) +{ +uint8_t x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_13); +x_16 = 1; +x_17 = lean_box(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +else +{ +uint8_t x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_13); +x_19 = 0; +x_20 = lean_box(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_15); +return x_21; +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___lambda__1), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_2 = lean_array_get_size(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; size_t x_8; +x_2 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1; +x_3 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1(x_2, x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +x_8 = 0; +if (x_7 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; lean_object* x_13; +lean_dec(x_5); +lean_dec(x_4); +x_9 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_10 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3; +x_11 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(x_9, x_6, x_10); +x_12 = lean_array_size(x_11); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(x_12, x_8, x_11); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = lean_nat_dec_le(x_5, x_5); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; lean_object* x_19; +lean_dec(x_5); +lean_dec(x_4); +x_15 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_16 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3; +x_17 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(x_15, x_6, x_16); +x_18 = lean_array_size(x_17); +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(x_18, x_8, x_17); +return x_19; +} +else +{ +size_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; size_t x_27; lean_object* x_28; +x_20 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_21 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1; +x_22 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14(x_4, x_8, x_20, x_21); +lean_dec(x_4); +x_23 = lean_array_get_size(x_22); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(x_22, x_6, x_25); +lean_dec(x_25); +x_27 = lean_array_size(x_26); +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(x_27, x_8, x_26); +return x_28; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_get_x3f___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_erase___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__9(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__11(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__13(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__14(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findPrioritizedCompletionPartitionsAt(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_Server_Completion_findCompletionInfosAt(x_1, x_2, x_3, x_4); +x_6 = l_Array_zipWithIndex___rarg(x_5); +lean_dec(x_5); +x_7 = l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions(x_6); +lean_dec(x_6); +return x_7; +} +} +lean_object* initialize_Lean_Server_Completion_SyntheticCompletion(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_CompletionInfoSelection(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Server_Completion_SyntheticCompletion(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1 = _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_filterDuplicateCompletionInfos___closed__1); +l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1 = _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__1); +l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2 = _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__2); +l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3 = _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__3); +l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4 = _init_l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Server_Completion_findCompletionInfosAt_go___lambda__3___closed__4); +l_Lean_Server_Completion_findCompletionInfosAt___closed__1 = _init_l_Lean_Server_Completion_findCompletionInfosAt___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_findCompletionInfosAt___closed__1); +l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__1(); +l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2 = _init_l_Std_DHashMap_Internal_AssocList_foldlM___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__6___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__7___closed__2(); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__1); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__2); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__3); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__4); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__5); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__6); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__7); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__8); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__9); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__10); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__11); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__12); +l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13 = _init_l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13(); +lean_mark_persistent(l_Array_groupByKey___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__1___closed__13); +l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1 = _init_l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1(); +lean_mark_persistent(l_Array_qsort_sort___at___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___spec__12___closed__1); +l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1 = _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__1); +l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2 = _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__2); +l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3 = _init_l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionInfoSelection_0__Lean_Server_Completion_computePrioritizedCompletionPartitions___closed__3); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionItemData.c b/stage0/stdlib/Lean/Server/Completion/CompletionItemData.c new file mode 100644 index 000000000000..5023ea039362 --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/CompletionItemData.c @@ -0,0 +1,528 @@ +// Lean compiler output +// Module: Lean.Server.Completion.CompletionItemData +// Imports: Lean.Server.FileSource +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(lean_object*); +static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; +lean_object* l_Lean_Json_mkObj(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instFileSourceCompletionItem; +static lean_object* l_Lean_Lsp_instFileSourceCompletionItem___closed__1; +lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6; +lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(lean_object*); +static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instToJsonCompletionItemData___closed__1; +lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13; +static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; +extern lean_object* l_String_instInhabited; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7; +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14; +static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1; +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15_(lean_object*); +static lean_object* l_Lean_Lsp_instFromJsonCompletionItemData___closed__1; +LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemData; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3; +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4; +LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemData; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9; +static lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Lean_Json_getObjValD(x_1, x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("params", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("CompletionItemData", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3; +x_3 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14; +x_10 = lean_string_append(x_9, x_8); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +return x_3; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_3, 0); +lean_inc(x_13); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemData___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemData() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFromJsonCompletionItemData___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(x_1); +x_3 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +x_5 = lean_box(0); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1; +x_9 = l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(x_7, x_8); +x_10 = l_Lean_Json_mkObj(x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemData___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemData() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonCompletionItemData___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_String_instInhabited; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("no data param on completion item ", 33, 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Server.Completion.CompletionItemData", 41, 41); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Lsp.CompletionItem.getFileSource!", 38, 38); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 6); +lean_inc(x_2); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1; +x_5 = lean_string_append(x_4, x_3); +lean_dec(x_3); +x_6 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2; +x_7 = lean_string_append(x_5, x_6); +x_8 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; +x_9 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; +x_10 = lean_unsigned_to_nat(34u); +x_11 = lean_unsigned_to_nat(22u); +x_12 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_8, x_9, x_10, x_11, x_7); +lean_dec(x_7); +x_13 = l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(x_12); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_1); +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +lean_dec(x_2); +x_15 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15_(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; +x_18 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; +x_19 = lean_unsigned_to_nat(34u); +x_20 = lean_unsigned_to_nat(22u); +x_21 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_17, x_18, x_19, x_20, x_16); +lean_dec(x_16); +x_22 = l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec(x_15); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +return x_24; +} +} +} +} +static lean_object* _init_l_Lean_Lsp_instFileSourceCompletionItem___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_getFileSource_x21), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFileSourceCompletionItem() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFileSourceCompletionItem___closed__1; +return x_1; +} +} +lean_object* initialize_Lean_Server_FileSource(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_CompletionItemData(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Server_FileSource(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__1); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__2); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__3); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__4); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__5); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__6); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__7); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__8); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__9); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__10); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__11); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__12); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__13); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____closed__14); +l_Lean_Lsp_instFromJsonCompletionItemData___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItemData___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemData___closed__1); +l_Lean_Lsp_instFromJsonCompletionItemData = _init_l_Lean_Lsp_instFromJsonCompletionItemData(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemData); +l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1 = _init_l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82____closed__1); +l_Lean_Lsp_instToJsonCompletionItemData___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionItemData___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemData___closed__1); +l_Lean_Lsp_instToJsonCompletionItemData = _init_l_Lean_Lsp_instToJsonCompletionItemData(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemData); +l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1); +l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2); +l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3); +l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4); +l_Lean_Lsp_instFileSourceCompletionItem___closed__1 = _init_l_Lean_Lsp_instFileSourceCompletionItem___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFileSourceCompletionItem___closed__1); +l_Lean_Lsp_instFileSourceCompletionItem = _init_l_Lean_Lsp_instFileSourceCompletionItem(); +lean_mark_persistent(l_Lean_Lsp_instFileSourceCompletionItem); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c b/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c new file mode 100644 index 000000000000..2524cd1fde0f --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/CompletionResolution.c @@ -0,0 +1,2352 @@ +// Lean compiler output +// Module: Lean.Server.Completion.CompletionResolution +// Imports: Lean.Server.Completion.CompletionItemData Lean.Server.Completion.CompletionInfoSelection +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Lean_JsonNumber_fromNat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(lean_object*); +lean_object* l_Lean_Meta_ppExpr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2; +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +lean_object* l_Lean_Json_mkObj(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17; +lean_object* l_Lean_ConstantInfo_type(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6; +lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14; +uint8_t l_Lean_Name_isAnonymous(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7; +static lean_object* l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1; +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157_(lean_object*); +static lean_object* l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1; +lean_object* lean_environment_find(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7; +extern lean_object* l_Lean_instInhabitedJson; +lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_findCompletionInfosAt(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5; +lean_object* l_Lean_Json_getStr_x3f(lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_local_ctx_find(lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23_(lean_object*); +lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2; +static lean_object* l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1; +lean_object* l_Lean_Elab_CompletionInfo_lctx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonResolvableCompletionItemData; +lean_object* l_outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionIdentifier; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4; +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3; +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13; +extern lean_object* l_Std_Format_defWidth; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_CompletionItem_resolve___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19; +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5; +lean_object* l_String_toName(lean_object*); +lean_object* l_Lean_LocalDecl_type(lean_object*); +lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6; +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Position_0__Lean_fromJsonPosition____x40_Lean_Data_Position___hyg_289____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___boxed(lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1; +LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonResolvableCompletionItemData; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1(lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(lean_object*); +uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_404_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1; +LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionIdentifier; +lean_object* lean_expr_instantiate1(lean_object*, lean_object*); +static lean_object* l_Lean_Lsp_CompletionItem_resolve___closed__2; +lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4; +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9; +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18; +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("no inductive constructor matched", 32, 32); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2; +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("declName", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("const", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[anonymous]", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("expected a `Name`, got '", 24, 24); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_3 = l_Except_orElseLazy___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_4 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2; +x_5 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_1); +x_6 = lean_array_mk(x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3; +x_9 = lean_unsigned_to_nat(1u); +x_10 = l_Lean_Json_parseTagged(x_2, x_8, x_9, x_7); +lean_dec(x_7); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_58; +x_58 = !lean_is_exclusive(x_10); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +x_59 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_60 = l_Except_orElseLazy___rarg(x_10, x_59); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_10, 0); +lean_inc(x_61); +lean_dec(x_10); +x_62 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_62, 0, x_61); +x_63 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_64 = l_Except_orElseLazy___rarg(x_62, x_63); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_65 = lean_ctor_get(x_10, 0); +lean_inc(x_65); +lean_dec(x_10); +x_66 = lean_array_get_size(x_65); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +lean_dec(x_66); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; +lean_dec(x_65); +x_69 = l_Lean_instInhabitedJson; +x_70 = l_outOfBounds___rarg(x_69); +x_11 = x_70; +goto block_57; +} +else +{ +lean_object* x_71; +x_71 = lean_array_fget(x_65, x_67); +lean_dec(x_65); +x_11 = x_71; +goto block_57; +} +} +block_57: +{ +lean_object* x_12; +lean_inc(x_11); +x_12 = l_Lean_Json_getStr_x3f(x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_11); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_15 = l_Except_orElseLazy___rarg(x_12, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_19 = l_Except_orElseLazy___rarg(x_17, x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5; +x_23 = lean_string_dec_eq(x_21, x_22); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; +x_24 = l_String_toName(x_21); +x_25 = l_Lean_Name_isAnonymous(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_11); +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_12, 0, x_26); +x_27 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_28 = l_Except_orElseLazy___rarg(x_12, x_27); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_dec(x_24); +x_29 = lean_unsigned_to_nat(80u); +x_30 = l_Lean_Json_pretty(x_11, x_29); +x_31 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6; +x_32 = lean_string_append(x_31, x_30); +lean_dec(x_30); +x_33 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7; +x_34 = lean_string_append(x_32, x_33); +lean_ctor_set_tag(x_12, 0); +lean_ctor_set(x_12, 0, x_34); +x_35 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_36 = l_Except_orElseLazy___rarg(x_12, x_35); +return x_36; +} +} +else +{ +lean_object* x_37; +lean_free_object(x_12); +lean_dec(x_21); +lean_dec(x_11); +x_37 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10; +return x_37; +} +} +else +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = lean_ctor_get(x_12, 0); +lean_inc(x_38); +lean_dec(x_12); +x_39 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5; +x_40 = lean_string_dec_eq(x_38, x_39); +if (x_40 == 0) +{ +lean_object* x_41; uint8_t x_42; +x_41 = l_String_toName(x_38); +x_42 = l_Lean_Name_isAnonymous(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_11); +x_43 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_43, 0, x_41); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_46 = l_Except_orElseLazy___rarg(x_44, x_45); +return x_46; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_41); +x_47 = lean_unsigned_to_nat(80u); +x_48 = l_Lean_Json_pretty(x_11, x_47); +x_49 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6; +x_50 = lean_string_append(x_49, x_48); +lean_dec(x_48); +x_51 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7; +x_52 = lean_string_append(x_50, x_51); +x_53 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_53, 0, x_52); +x_54 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4; +x_55 = l_Except_orElseLazy___rarg(x_53, x_54); +return x_55; +} +} +else +{ +lean_object* x_56; +lean_dec(x_38); +lean_dec(x_11); +x_56 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10; +return x_56; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("fvar", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_box(0); +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6; +x_4 = lean_unsigned_to_nat(1u); +x_5 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5; +lean_inc(x_1); +x_6 = l_Lean_Json_parseTagged(x_1, x_3, x_4, x_5); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___boxed), 3, 2); +lean_closure_set(x_7, 0, x_2); +lean_closure_set(x_7, 1, x_1); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_6); +if (x_51 == 0) +{ +lean_object* x_52; +x_52 = l_Except_orElseLazy___rarg(x_6, x_7); +return x_52; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_6, 0); +lean_inc(x_53); +lean_dec(x_6); +x_54 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = l_Except_orElseLazy___rarg(x_54, x_7); +return x_55; +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_56 = lean_ctor_get(x_6, 0); +lean_inc(x_56); +lean_dec(x_6); +x_57 = lean_array_get_size(x_56); +x_58 = lean_unsigned_to_nat(0u); +x_59 = lean_nat_dec_lt(x_58, x_57); +lean_dec(x_57); +if (x_59 == 0) +{ +lean_object* x_60; lean_object* x_61; +lean_dec(x_56); +x_60 = l_Lean_instInhabitedJson; +x_61 = l_outOfBounds___rarg(x_60); +x_8 = x_61; +goto block_50; +} +else +{ +lean_object* x_62; +x_62 = lean_array_fget(x_56, x_58); +lean_dec(x_56); +x_8 = x_62; +goto block_50; +} +} +block_50: +{ +lean_object* x_9; +lean_inc(x_8); +x_9 = l_Lean_Json_getStr_x3f(x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +lean_dec(x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Except_orElseLazy___rarg(x_13, x_7); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5; +x_18 = lean_string_dec_eq(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = l_String_toName(x_16); +x_20 = l_Lean_Name_isAnonymous(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_8); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_9, 0, x_21); +x_22 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_19); +x_23 = lean_unsigned_to_nat(80u); +x_24 = l_Lean_Json_pretty(x_8, x_23); +x_25 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6; +x_26 = lean_string_append(x_25, x_24); +lean_dec(x_24); +x_27 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7; +x_28 = lean_string_append(x_26, x_27); +lean_ctor_set_tag(x_9, 0); +lean_ctor_set(x_9, 0, x_28); +x_29 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; +lean_free_object(x_9); +lean_dec(x_16); +lean_dec(x_8); +x_30 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8; +x_31 = l_Except_orElseLazy___rarg(x_30, x_7); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = lean_ctor_get(x_9, 0); +lean_inc(x_32); +lean_dec(x_9); +x_33 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5; +x_34 = lean_string_dec_eq(x_32, x_33); +if (x_34 == 0) +{ +lean_object* x_35; uint8_t x_36; +x_35 = l_String_toName(x_32); +x_36 = l_Lean_Name_isAnonymous(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_8); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_35); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = l_Except_orElseLazy___rarg(x_38, x_7); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_35); +x_40 = lean_unsigned_to_nat(80u); +x_41 = l_Lean_Json_pretty(x_8, x_40); +x_42 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6; +x_43 = lean_string_append(x_42, x_41); +lean_dec(x_41); +x_44 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7; +x_45 = lean_string_append(x_43, x_44); +x_46 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = l_Except_orElseLazy___rarg(x_46, x_7); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_32); +lean_dec(x_8); +x_48 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8; +x_49 = l_Except_orElseLazy___rarg(x_48, x_7); +return x_49; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__2(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionIdentifier() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157_(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_1, 0); +x_4 = 1; +x_5 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_6 = l_Lean_Name_toString(x_3, x_4, x_5); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_6); +x_7 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_1); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +x_11 = l_Lean_Json_mkObj(x_10); +x_12 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3; +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_9); +x_15 = l_Lean_Json_mkObj(x_14); +return x_15; +} +else +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = 1; +x_18 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_19 = l_Lean_Name_toString(x_16, x_17, x_18); +x_20 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Json_mkObj(x_24); +x_26 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_23); +x_29 = l_Lean_Json_mkObj(x_28); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_1); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_31 = lean_ctor_get(x_1, 0); +x_32 = 1; +x_33 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_34 = l_Lean_Name_toString(x_31, x_32, x_33); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_34); +x_35 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_1); +x_37 = lean_box(0); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Json_mkObj(x_38); +x_40 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6; +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_37); +x_43 = l_Lean_Json_mkObj(x_42); +return x_43; +} +else +{ +lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_44 = lean_ctor_get(x_1, 0); +lean_inc(x_44); +lean_dec(x_1); +x_45 = 1; +x_46 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_47 = l_Lean_Name_toString(x_44, x_45, x_46); +x_48 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Lean_Json_mkObj(x_52); +x_54 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6; +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_51); +x_57 = l_Lean_Json_mkObj(x_56); +return x_57; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonCompletionIdentifier() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_getObjValD(x_1, x_2); +switch (lean_obj_tag(x_3)) { +case 0: +{ +lean_object* x_4; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1; +return x_4; +} +case 1: +{ +lean_object* x_5; +x_5 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23_(x_3); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +return x_5; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_5); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_5, 0, x_11); +return x_5; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_5, 0); +lean_inc(x_12); +lean_dec(x_5); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +default: +{ +lean_object* x_15; uint8_t x_16; +lean_inc(x_3); +x_15 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23_(x_3); +x_16 = !lean_is_exclusive(x_3); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_3, 0); +lean_dec(x_17); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_18; +lean_free_object(x_3); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_15, 0); +lean_inc(x_19); +lean_dec(x_15); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_19); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_15); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_15, 0); +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_22); +lean_ctor_set(x_15, 0, x_3); +return x_15; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_15, 0); +lean_inc(x_23); +lean_dec(x_15); +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 0, x_23); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_3); +return x_24; +} +} +} +else +{ +lean_dec(x_3); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_15, 0); +lean_inc(x_25); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + x_26 = x_15; +} else { + lean_dec_ref(x_15); + x_26 = lean_box(0); +} +if (lean_is_scalar(x_26)) { + x_27 = lean_alloc_ctor(0, 1, 0); +} else { + x_27 = x_26; +} +lean_ctor_set(x_27, 0, x_25); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_15, 0); +lean_inc(x_28); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + x_29 = x_15; +} else { + lean_dec_ref(x_15); + x_29 = lean_box(0); +} +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_28); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(1, 1, 0); +} else { + x_31 = x_29; +} +lean_ctor_set(x_31, 0, x_30); +return x_31; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("params", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lsp", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ResolvableCompletionItemData", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(": ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cPos", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id\?", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20; +x_2 = 1; +x_3 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1; +x_4 = l_Lean_Name_toString(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1; +lean_inc(x_1); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_15____spec__1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13; +x_10 = lean_string_append(x_9, x_8); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +lean_dec(x_3); +x_13 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14; +lean_inc(x_1); +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Position_0__Lean_fromJsonPosition____x40_Lean_Data_Position___hyg_289____spec__1(x_1, x_13); +if (lean_obj_tag(x_14) == 0) +{ +uint8_t x_15; +lean_dec(x_12); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18; +x_18 = lean_string_append(x_17, x_16); +lean_dec(x_16); +lean_ctor_set(x_14, 0, x_18); +return x_14; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_14, 0); +lean_inc(x_19); +lean_dec(x_14); +x_20 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18; +x_21 = lean_string_append(x_20, x_19); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_14, 0); +lean_inc(x_23); +lean_dec(x_14); +x_24 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1(x_1, x_24); +if (lean_obj_tag(x_25) == 0) +{ +uint8_t x_26; +lean_dec(x_23); +lean_dec(x_12); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +lean_ctor_set(x_25, 0, x_29); +return x_25; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_25, 0); +lean_inc(x_30); +lean_dec(x_25); +x_31 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23; +x_32 = lean_string_append(x_31, x_30); +lean_dec(x_30); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_32); +return x_33; +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_25); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_25, 0); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_12); +lean_ctor_set(x_36, 1, x_23); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_25, 0, x_36); +return x_25; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_25, 0); +lean_inc(x_37); +lean_dec(x_25); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_12); +lean_ctor_set(x_38, 1, x_23); +lean_ctor_set(x_38, 2, x_37); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +return x_39; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instFromJsonResolvableCompletionItemData() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_1); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157_(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3049_(x_2); +x_4 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +x_6 = lean_box(0); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = l_Lean_JsonNumber_fromNat(x_8); +x_10 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_6); +x_14 = lean_ctor_get(x_1, 2); +lean_inc(x_14); +lean_dec(x_1); +x_15 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1; +x_16 = l_Lean_Json_opt___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____spec__1(x_15, x_14); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_6); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_13); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_7); +lean_ctor_set(x_19, 1, x_18); +x_20 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1; +x_21 = l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(x_19, x_20); +x_22 = l_Lean_Json_mkObj(x_21); +return x_22; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406_), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_instToJsonResolvableCompletionItemData() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_expr_instantiate1(x_1, x_3); +x_10 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg(x_9, x_2, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +if (lean_obj_tag(x_1) == 7) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 2); +lean_inc(x_10); +x_11 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 8); +x_12 = 1; +x_13 = l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_404_(x_11, x_12); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +x_14 = lean_apply_6(x_2, x_1, x_3, x_4, x_5, x_6, x_7); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; +lean_dec(x_1); +x_15 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1___boxed), 8, 2); +lean_closure_set(x_15, 0, x_10); +lean_closure_set(x_15, 1, x_2); +x_16 = 0; +x_17 = l_Lean_Meta_withLocalDecl___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__3___rarg(x_8, x_11, x_9, x_15, x_16, x_3, x_4, x_5, x_6, x_7); +return x_17; +} +} +else +{ +lean_object* x_18; +x_18 = lean_apply_6(x_2, x_1, x_3, x_4, x_5, x_6, x_7); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg), 7, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_ppExpr(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_7, 0); +x_10 = l_Std_Format_defWidth; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_format_pretty(x_9, x_10, x_11, x_11); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_13 = lean_ctor_get(x_7, 0); +x_14 = lean_ctor_get(x_7, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_7); +x_15 = l_Std_Format_defWidth; +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_format_pretty(x_13, x_15, x_16, x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_14); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_7); +if (x_19 == 0) +{ +return x_7; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_7, 0); +x_21 = lean_ctor_get(x_7, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_7); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_resolve___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed), 7, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Lsp_CompletionItem_resolve___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_76; +x_8 = lean_st_ref_get(x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +x_13 = l_Lean_Lsp_CompletionItem_resolve___closed__1; +x_76 = lean_ctor_get(x_1, 1); +lean_inc(x_76); +if (lean_obj_tag(x_76) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_12); +x_77 = lean_ctor_get(x_2, 0); +lean_inc(x_77); +lean_dec(x_2); +x_78 = lean_environment_find(x_11, x_77); +if (lean_obj_tag(x_78) == 0) +{ +lean_object* x_79; +x_79 = lean_box(0); +x_14 = x_79; +goto block_75; +} +else +{ +uint8_t x_80; +x_80 = !lean_is_exclusive(x_78); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_78, 0); +x_82 = l_Lean_ConstantInfo_type(x_81); +lean_dec(x_81); +lean_ctor_set(x_78, 0, x_82); +x_14 = x_78; +goto block_75; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_78, 0); +lean_inc(x_83); +lean_dec(x_78); +x_84 = l_Lean_ConstantInfo_type(x_83); +lean_dec(x_83); +x_85 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_14 = x_85; +goto block_75; +} +} +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_11); +x_86 = lean_ctor_get(x_2, 0); +lean_inc(x_86); +lean_dec(x_2); +x_87 = lean_local_ctx_find(x_12, x_86); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; +x_88 = lean_box(0); +x_14 = x_88; +goto block_75; +} +else +{ +uint8_t x_89; +x_89 = !lean_is_exclusive(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; +x_90 = lean_ctor_get(x_87, 0); +x_91 = l_Lean_LocalDecl_type(x_90); +lean_dec(x_90); +lean_ctor_set(x_87, 0, x_91); +x_14 = x_87; +goto block_75; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_87, 0); +lean_inc(x_92); +lean_dec(x_87); +x_93 = l_Lean_LocalDecl_type(x_92); +lean_dec(x_92); +x_94 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_14 = x_94; +goto block_75; +} +} +} +} +else +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_76); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +x_95 = lean_box(0); +x_96 = lean_apply_7(x_13, x_1, x_95, x_3, x_4, x_5, x_6, x_10); +return x_96; +} +block_75: +{ +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_box(0); +x_16 = !lean_is_exclusive(x_1); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_1, 1); +lean_dec(x_17); +lean_ctor_set(x_1, 1, x_15); +x_18 = lean_box(0); +x_19 = lean_apply_7(x_13, x_1, x_18, x_3, x_4, x_5, x_6, x_10); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_ctor_get(x_1, 2); +x_22 = lean_ctor_get(x_1, 3); +x_23 = lean_ctor_get(x_1, 4); +x_24 = lean_ctor_get(x_1, 5); +x_25 = lean_ctor_get(x_1, 6); +x_26 = lean_ctor_get(x_1, 7); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_1); +x_27 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_15); +lean_ctor_set(x_27, 2, x_21); +lean_ctor_set(x_27, 3, x_22); +lean_ctor_set(x_27, 4, x_23); +lean_ctor_set(x_27, 5, x_24); +lean_ctor_set(x_27, 6, x_25); +lean_ctor_set(x_27, 7, x_26); +x_28 = lean_box(0); +x_29 = lean_apply_7(x_13, x_27, x_28, x_3, x_4, x_5, x_6, x_10); +return x_29; +} +} +else +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_14); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_14, 0); +x_32 = l_Lean_Lsp_CompletionItem_resolve___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_33 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg(x_31, x_32, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +lean_ctor_set(x_14, 0, x_34); +x_36 = !lean_is_exclusive(x_1); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_1, 1); +lean_dec(x_37); +lean_ctor_set(x_1, 1, x_14); +x_38 = lean_box(0); +x_39 = lean_apply_7(x_13, x_1, x_38, x_3, x_4, x_5, x_6, x_35); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_40 = lean_ctor_get(x_1, 0); +x_41 = lean_ctor_get(x_1, 2); +x_42 = lean_ctor_get(x_1, 3); +x_43 = lean_ctor_get(x_1, 4); +x_44 = lean_ctor_get(x_1, 5); +x_45 = lean_ctor_get(x_1, 6); +x_46 = lean_ctor_get(x_1, 7); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_1); +x_47 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_47, 0, x_40); +lean_ctor_set(x_47, 1, x_14); +lean_ctor_set(x_47, 2, x_41); +lean_ctor_set(x_47, 3, x_42); +lean_ctor_set(x_47, 4, x_43); +lean_ctor_set(x_47, 5, x_44); +lean_ctor_set(x_47, 6, x_45); +lean_ctor_set(x_47, 7, x_46); +x_48 = lean_box(0); +x_49 = lean_apply_7(x_13, x_47, x_48, x_3, x_4, x_5, x_6, x_35); +return x_49; +} +} +else +{ +uint8_t x_50; +lean_free_object(x_14); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_50 = !lean_is_exclusive(x_33); +if (x_50 == 0) +{ +return x_33; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_33, 0); +x_52 = lean_ctor_get(x_33, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_33); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_14, 0); +lean_inc(x_54); +lean_dec(x_14); +x_55 = l_Lean_Lsp_CompletionItem_resolve___closed__2; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_56 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_consumeImplicitPrefix___rarg(x_54, x_55, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_57); +x_60 = lean_ctor_get(x_1, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_1, 2); +lean_inc(x_61); +x_62 = lean_ctor_get(x_1, 3); +lean_inc(x_62); +x_63 = lean_ctor_get(x_1, 4); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 5); +lean_inc(x_64); +x_65 = lean_ctor_get(x_1, 6); +lean_inc(x_65); +x_66 = lean_ctor_get(x_1, 7); +lean_inc(x_66); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + x_67 = x_1; +} else { + lean_dec_ref(x_1); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(0, 8, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_60); +lean_ctor_set(x_68, 1, x_59); +lean_ctor_set(x_68, 2, x_61); +lean_ctor_set(x_68, 3, x_62); +lean_ctor_set(x_68, 4, x_63); +lean_ctor_set(x_68, 5, x_64); +lean_ctor_set(x_68, 6, x_65); +lean_ctor_set(x_68, 7, x_66); +x_69 = lean_box(0); +x_70 = lean_apply_7(x_13, x_68, x_69, x_3, x_4, x_5, x_6, x_58); +return x_70; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_71 = lean_ctor_get(x_56, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_56, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + x_73 = x_56; +} else { + lean_dec_ref(x_56); + x_73 = lean_box(0); +} +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_73; +} +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Lsp_CompletionItem_resolve___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_resolve___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Lsp_CompletionItem_resolve___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_Server_Completion_findCompletionInfosAt(x_1, x_2, x_3, x_4); +x_10 = l_Array_get_x3f___rarg(x_9, x_7); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +lean_dec(x_6); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_5); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 2); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_CompletionInfo_lctx(x_14); +lean_dec(x_14); +x_16 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_resolve), 7, 2); +lean_closure_set(x_16, 0, x_5); +lean_closure_set(x_16, 1, x_6); +x_17 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_13, x_15, x_16, x_8); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Server_Completion_resolveCompletionItem_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +return x_9; +} +} +lean_object* initialize_Lean_Server_Completion_CompletionItemData(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_CompletionInfoSelection(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_CompletionResolution(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Server_Completion_CompletionItemData(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Server_Completion_CompletionInfoSelection(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__1); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__1___closed__2); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__1); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__2); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__3); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__4); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__5); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__6); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__7); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__8); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__9); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____lambda__3___closed__10); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__1); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__2); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__3); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__4); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__5); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__6); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__7); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_23____closed__8); +l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionIdentifier___closed__1); +l_Lean_Lsp_instFromJsonCompletionIdentifier = _init_l_Lean_Lsp_instFromJsonCompletionIdentifier(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionIdentifier); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonCompletionIdentifier____x40_Lean_Server_Completion_CompletionResolution___hyg_157____closed__1); +l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionIdentifier___closed__1); +l_Lean_Lsp_instToJsonCompletionIdentifier = _init_l_Lean_Lsp_instToJsonCompletionIdentifier(); +lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionIdentifier); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____spec__1___closed__1); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__1); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__2); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__3); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__4); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__5); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__6); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__7); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__8); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__9); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__10); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__11); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__12); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__13); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__14); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__15); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__16); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__17); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__18); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__19); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__20); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__21); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__22); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261____closed__23); +l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1 = _init_l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonResolvableCompletionItemData___closed__1); +l_Lean_Lsp_instFromJsonResolvableCompletionItemData = _init_l_Lean_Lsp_instFromJsonResolvableCompletionItemData(); +lean_mark_persistent(l_Lean_Lsp_instFromJsonResolvableCompletionItemData); +l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1 = _init_l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_toJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_406____closed__1); +l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1 = _init_l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1(); +lean_mark_persistent(l_Lean_Lsp_instToJsonResolvableCompletionItemData___closed__1); +l_Lean_Lsp_instToJsonResolvableCompletionItemData = _init_l_Lean_Lsp_instToJsonResolvableCompletionItemData(); +lean_mark_persistent(l_Lean_Lsp_instToJsonResolvableCompletionItemData); +l_Lean_Lsp_CompletionItem_resolve___closed__1 = _init_l_Lean_Lsp_CompletionItem_resolve___closed__1(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_resolve___closed__1); +l_Lean_Lsp_CompletionItem_resolve___closed__2 = _init_l_Lean_Lsp_CompletionItem_resolve___closed__2(); +lean_mark_persistent(l_Lean_Lsp_CompletionItem_resolve___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/Completion/CompletionUtils.c b/stage0/stdlib/Lean/Server/Completion/CompletionUtils.c new file mode 100644 index 000000000000..34cc2f471927 --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/CompletionUtils.c @@ -0,0 +1,33 @@ +// Lean compiler output +// Module: Lean.Server.Completion.CompletionUtils +// Imports: Init.Prelude Lean.Elab.InfoTree.Types +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Init_Prelude(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_InfoTree_Types(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_CompletionUtils(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Prelude(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Elab_InfoTree_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c b/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c new file mode 100644 index 000000000000..334222cfdbe5 --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/EligibleHeaderDecls.c @@ -0,0 +1,1205 @@ +// Lean compiler output +// Module: Lean.Server.Completion.EligibleHeaderDecls +// Imports: Lean.Meta.CompletionName +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Server_Completion_allowCompletion(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_uint64_to_usize(uint64_t); +lean_object* l_Lean_PersistentHashMap_foldlMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_array(lean_object*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_allowCompletion___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_eligibleHeaderDeclsRef; +static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2; +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_shift_right(uint64_t, uint64_t); +lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); +lean_object* lean_st_mk_ref(lean_object*, lean_object*); +uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_allowCompletion(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion_EligibleHeaderDecls___hyg_10_(lean_object*); +uint64_t l_Lean_Name_hash___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM(lean_object*); +static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +uint64_t lean_uint64_xor(uint64_t, uint64_t); +lean_object* l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); +size_t lean_usize_sub(size_t, size_t); +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*); +static lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3; +lean_object* lean_array_get_size(lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion_EligibleHeaderDecls___hyg_10_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_box(0); +x_3 = lean_st_mk_ref(x_2, x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_3); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_5; +lean_dec(x_1); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_2); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +else +{ +uint8_t x_6; +lean_dec(x_2); +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +x_9 = lean_ctor_get(x_3, 2); +lean_inc(x_7); +lean_inc(x_1); +x_10 = l_Lean_Meta_allowCompletion(x_1, x_7); +if (x_10 == 0) +{ +lean_object* x_11; +lean_free_object(x_3); +lean_dec(x_8); +lean_dec(x_7); +x_11 = lean_box(0); +x_2 = x_11; +x_3 = x_9; +goto _start; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint64_t x_17; uint64_t x_18; uint64_t x_19; uint64_t x_20; uint64_t x_21; uint64_t x_22; uint64_t x_23; size_t x_24; size_t x_25; size_t x_26; size_t x_27; size_t x_28; lean_object* x_29; uint8_t x_30; +x_14 = lean_ctor_get(x_4, 0); +x_15 = lean_ctor_get(x_4, 1); +x_16 = lean_array_get_size(x_15); +x_17 = l_Lean_Name_hash___override(x_7); +x_18 = 32; +x_19 = lean_uint64_shift_right(x_17, x_18); +x_20 = lean_uint64_xor(x_17, x_19); +x_21 = 16; +x_22 = lean_uint64_shift_right(x_20, x_21); +x_23 = lean_uint64_xor(x_20, x_22); +x_24 = lean_uint64_to_usize(x_23); +x_25 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_26 = 1; +x_27 = lean_usize_sub(x_25, x_26); +x_28 = lean_usize_land(x_24, x_27); +x_29 = lean_array_uget(x_15, x_28); +x_30 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_add(x_14, x_31); +lean_dec(x_14); +lean_ctor_set(x_3, 2, x_29); +x_33 = lean_array_uset(x_15, x_28, x_3); +x_34 = lean_unsigned_to_nat(4u); +x_35 = lean_nat_mul(x_32, x_34); +x_36 = lean_unsigned_to_nat(3u); +x_37 = lean_nat_div(x_35, x_36); +lean_dec(x_35); +x_38 = lean_array_get_size(x_33); +x_39 = lean_nat_dec_le(x_37, x_38); +lean_dec(x_38); +lean_dec(x_37); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; +x_40 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_33); +lean_ctor_set(x_4, 1, x_40); +lean_ctor_set(x_4, 0, x_32); +x_41 = lean_box(0); +x_2 = x_41; +x_3 = x_9; +goto _start; +} +else +{ +lean_object* x_43; +lean_ctor_set(x_4, 1, x_33); +lean_ctor_set(x_4, 0, x_32); +x_43 = lean_box(0); +x_2 = x_43; +x_3 = x_9; +goto _start; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_free_object(x_3); +x_45 = lean_box(0); +x_46 = lean_array_uset(x_15, x_28, x_45); +x_47 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_29); +x_48 = lean_array_uset(x_46, x_28, x_47); +lean_ctor_set(x_4, 1, x_48); +x_49 = lean_box(0); +x_2 = x_49; +x_3 = x_9; +goto _start; +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint64_t x_54; uint64_t x_55; uint64_t x_56; uint64_t x_57; uint64_t x_58; uint64_t x_59; uint64_t x_60; size_t x_61; size_t x_62; size_t x_63; size_t x_64; size_t x_65; lean_object* x_66; uint8_t x_67; +x_51 = lean_ctor_get(x_4, 0); +x_52 = lean_ctor_get(x_4, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_4); +x_53 = lean_array_get_size(x_52); +x_54 = l_Lean_Name_hash___override(x_7); +x_55 = 32; +x_56 = lean_uint64_shift_right(x_54, x_55); +x_57 = lean_uint64_xor(x_54, x_56); +x_58 = 16; +x_59 = lean_uint64_shift_right(x_57, x_58); +x_60 = lean_uint64_xor(x_57, x_59); +x_61 = lean_uint64_to_usize(x_60); +x_62 = lean_usize_of_nat(x_53); +lean_dec(x_53); +x_63 = 1; +x_64 = lean_usize_sub(x_62, x_63); +x_65 = lean_usize_land(x_61, x_64); +x_66 = lean_array_uget(x_52, x_65); +x_67 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_7, x_66); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_68 = lean_unsigned_to_nat(1u); +x_69 = lean_nat_add(x_51, x_68); +lean_dec(x_51); +lean_ctor_set(x_3, 2, x_66); +x_70 = lean_array_uset(x_52, x_65, x_3); +x_71 = lean_unsigned_to_nat(4u); +x_72 = lean_nat_mul(x_69, x_71); +x_73 = lean_unsigned_to_nat(3u); +x_74 = lean_nat_div(x_72, x_73); +lean_dec(x_72); +x_75 = lean_array_get_size(x_70); +x_76 = lean_nat_dec_le(x_74, x_75); +lean_dec(x_75); +lean_dec(x_74); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_70); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_69); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_box(0); +x_2 = x_79; +x_3 = x_9; +x_4 = x_78; +goto _start; +} +else +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_69); +lean_ctor_set(x_81, 1, x_70); +x_82 = lean_box(0); +x_2 = x_82; +x_3 = x_9; +x_4 = x_81; +goto _start; +} +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_free_object(x_3); +x_84 = lean_box(0); +x_85 = lean_array_uset(x_52, x_65, x_84); +x_86 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_7, x_8, x_66); +x_87 = lean_array_uset(x_85, x_65, x_86); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_51); +lean_ctor_set(x_88, 1, x_87); +x_89 = lean_box(0); +x_2 = x_89; +x_3 = x_9; +x_4 = x_88; +goto _start; +} +} +} +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_91 = lean_ctor_get(x_3, 0); +x_92 = lean_ctor_get(x_3, 1); +x_93 = lean_ctor_get(x_3, 2); +lean_inc(x_93); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_3); +lean_inc(x_91); +lean_inc(x_1); +x_94 = l_Lean_Meta_allowCompletion(x_1, x_91); +if (x_94 == 0) +{ +lean_object* x_95; +lean_dec(x_92); +lean_dec(x_91); +x_95 = lean_box(0); +x_2 = x_95; +x_3 = x_93; +goto _start; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint64_t x_101; uint64_t x_102; uint64_t x_103; uint64_t x_104; uint64_t x_105; uint64_t x_106; uint64_t x_107; size_t x_108; size_t x_109; size_t x_110; size_t x_111; size_t x_112; lean_object* x_113; uint8_t x_114; +x_97 = lean_ctor_get(x_4, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_4, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_4)) { + lean_ctor_release(x_4, 0); + lean_ctor_release(x_4, 1); + x_99 = x_4; +} else { + lean_dec_ref(x_4); + x_99 = lean_box(0); +} +x_100 = lean_array_get_size(x_98); +x_101 = l_Lean_Name_hash___override(x_91); +x_102 = 32; +x_103 = lean_uint64_shift_right(x_101, x_102); +x_104 = lean_uint64_xor(x_101, x_103); +x_105 = 16; +x_106 = lean_uint64_shift_right(x_104, x_105); +x_107 = lean_uint64_xor(x_104, x_106); +x_108 = lean_uint64_to_usize(x_107); +x_109 = lean_usize_of_nat(x_100); +lean_dec(x_100); +x_110 = 1; +x_111 = lean_usize_sub(x_109, x_110); +x_112 = lean_usize_land(x_108, x_111); +x_113 = lean_array_uget(x_98, x_112); +x_114 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_91, x_113); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_115 = lean_unsigned_to_nat(1u); +x_116 = lean_nat_add(x_97, x_115); +lean_dec(x_97); +x_117 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_117, 0, x_91); +lean_ctor_set(x_117, 1, x_92); +lean_ctor_set(x_117, 2, x_113); +x_118 = lean_array_uset(x_98, x_112, x_117); +x_119 = lean_unsigned_to_nat(4u); +x_120 = lean_nat_mul(x_116, x_119); +x_121 = lean_unsigned_to_nat(3u); +x_122 = lean_nat_div(x_120, x_121); +lean_dec(x_120); +x_123 = lean_array_get_size(x_118); +x_124 = lean_nat_dec_le(x_122, x_123); +lean_dec(x_123); +lean_dec(x_122); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = l_Std_DHashMap_Internal_Raw_u2080_expand___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(x_118); +if (lean_is_scalar(x_99)) { + x_126 = lean_alloc_ctor(0, 2, 0); +} else { + x_126 = x_99; +} +lean_ctor_set(x_126, 0, x_116); +lean_ctor_set(x_126, 1, x_125); +x_127 = lean_box(0); +x_2 = x_127; +x_3 = x_93; +x_4 = x_126; +goto _start; +} +else +{ +lean_object* x_129; lean_object* x_130; +if (lean_is_scalar(x_99)) { + x_129 = lean_alloc_ctor(0, 2, 0); +} else { + x_129 = x_99; +} +lean_ctor_set(x_129, 0, x_116); +lean_ctor_set(x_129, 1, x_118); +x_130 = lean_box(0); +x_2 = x_130; +x_3 = x_93; +x_4 = x_129; +goto _start; +} +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_132 = lean_box(0); +x_133 = lean_array_uset(x_98, x_112, x_132); +x_134 = l_Std_DHashMap_Internal_AssocList_replace___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(x_91, x_92, x_113); +x_135 = lean_array_uset(x_133, x_112, x_134); +if (lean_is_scalar(x_99)) { + x_136 = lean_alloc_ctor(0, 2, 0); +} else { + x_136 = x_99; +} +lean_ctor_set(x_136, 0, x_97); +lean_ctor_set(x_136, 1, x_135); +x_137 = lean_box(0); +x_2 = x_137; +x_3 = x_93; +x_4 = x_136; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_3, x_4); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; +lean_dec(x_5); +x_8 = lean_array_uget(x_2, x_3); +x_9 = lean_box(0); +lean_inc(x_1); +x_10 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__1(x_1, x_9, x_8, x_6); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_3 = x_14; +x_5 = x_11; +x_6 = x_12; +goto _start; +} +else +{ +lean_object* x_16; +lean_dec(x_1); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_5); +lean_ctor_set(x_16, 1, x_6); +return x_16; +} +} +} +static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Server_Completion_eligibleHeaderDeclsRef; +return x_1; +} +} +static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_unsigned_to_nat(1u); +x_3 = l_Nat_nextPowerOfTwo_go(x_1, x_2, lean_box(0)); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2; +x_3 = lean_mk_array(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_getEligibleHeaderDecls(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1; +x_4 = lean_st_ref_take(x_3, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_array_get_size(x_9); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_13 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; +x_14 = lean_st_ref_set(x_3, x_13, x_6); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_14, 0); +lean_dec(x_16); +x_17 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +lean_ctor_set(x_14, 0, x_17); +return x_14; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_14, 1); +lean_inc(x_18); +lean_dec(x_14); +x_19 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = lean_nat_dec_le(x_10, x_10); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_1); +x_22 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5; +x_23 = lean_st_ref_set(x_3, x_22, x_6); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +x_26 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +lean_ctor_set(x_23, 0, x_26); +return x_23; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_30 = 0; +x_31 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_32 = lean_box(0); +x_33 = l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4; +x_34 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(x_1, x_9, x_30, x_31, x_32, x_33); +lean_dec(x_9); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +lean_inc(x_35); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_st_ref_set(x_3, x_36, x_6); +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +lean_ctor_set(x_37, 0, x_35); +return x_37; +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_35); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +} +else +{ +lean_object* x_42; uint8_t x_43; +lean_dec(x_1); +x_42 = lean_ctor_get(x_4, 1); +lean_inc(x_42); +lean_dec(x_4); +x_43 = !lean_is_exclusive(x_5); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = lean_ctor_get(x_5, 0); +lean_inc(x_44); +x_45 = lean_st_ref_set(x_3, x_5, x_42); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +lean_object* x_47; +x_47 = lean_ctor_get(x_45, 0); +lean_dec(x_47); +lean_ctor_set(x_45, 0, x_44); +return x_45; +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_dec(x_45); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_44); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_50 = lean_ctor_get(x_5, 0); +lean_inc(x_50); +lean_dec(x_5); +lean_inc(x_50); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = lean_st_ref_set(x_3, x_51, x_42); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_54 = x_52; +} else { + lean_dec_ref(x_52); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_50); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_getEligibleHeaderDecls___spec__2(x_1, x_2, x_7, x_8, x_5, x_6); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(x_1, x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_apply_2(x_6, lean_box(0), x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_3); +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_4, 2); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_inc(x_2); +x_12 = lean_apply_2(x_2, x_8, x_9); +x_13 = lean_alloc_closure((void*)(l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg___lambda__1), 4, 3); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_2); +lean_closure_set(x_13, 2, x_10); +x_14 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_12, x_13); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_PersistentHashMap_foldlMAux___rarg(x_1, lean_box(0), lean_box(0), lean_box(0), x_3, x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__3___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_apply_2(x_1, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed), 4, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = lean_box(0); +x_6 = l_Lean_PersistentHashMap_foldlMAux___rarg(x_1, lean_box(0), lean_box(0), lean_box(0), x_4, x_2, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = 1; +x_8 = lean_usize_add(x_1, x_7); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_2, x_3, x_4, x_8, x_5, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +x_7 = lean_usize_dec_eq(x_4, x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +x_9 = lean_array_uget(x_3, x_4); +x_10 = lean_box(0); +lean_inc(x_2); +lean_inc(x_1); +x_11 = l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Completion_forEligibleDeclsM___spec__1___rarg(x_1, x_2, x_10, x_9); +x_12 = lean_box_usize(x_4); +x_13 = lean_box_usize(x_5); +x_14 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_1); +lean_closure_set(x_14, 2, x_2); +lean_closure_set(x_14, 3, x_3); +lean_closure_set(x_14, 4, x_13); +x_15 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_11, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_apply_2(x_17, lean_box(0), x_6); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +lean_inc(x_4); +x_6 = l_Lean_Meta_allowCompletion(x_1, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +lean_dec(x_2); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_box(0); +x_10 = lean_apply_2(x_8, lean_box(0), x_9); +return x_10; +} +else +{ +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_apply_2(x_3, x_4, x_5); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +lean_inc(x_2); +x_7 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__1), 5, 3); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_2); +lean_closure_set(x_7, 2, x_3); +x_8 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg(x_2, x_6, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +lean_inc(x_3); +lean_inc(x_2); +x_10 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed), 4, 3); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_2); +lean_closure_set(x_10, 2, x_3); +if (x_9 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +lean_dec(x_2); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_box(0); +x_14 = lean_apply_2(x_12, lean_box(0), x_13); +x_15 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_14, x_10); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = lean_nat_dec_le(x_7, x_7); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_17 = lean_ctor_get(x_2, 0); +lean_inc(x_17); +lean_dec(x_2); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_box(0); +x_20 = lean_apply_2(x_18, lean_box(0), x_19); +x_21 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_20, x_10); +return x_21; +} +else +{ +size_t x_22; size_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_22 = 0; +x_23 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_24 = lean_box(0); +x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_2, x_3, x_6, x_22, x_23, x_24); +x_26 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_25, x_10); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_5); +x_6 = lean_alloc_closure((void*)(l_Lean_Server_Completion_getEligibleHeaderDecls), 2, 1); +lean_closure_set(x_6, 0, x_5); +x_7 = lean_apply_2(x_1, lean_box(0), x_6); +lean_inc(x_4); +x_8 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__3), 5, 4); +lean_closure_set(x_8, 0, x_5); +lean_closure_set(x_8, 1, x_2); +lean_closure_set(x_8, 2, x_3); +lean_closure_set(x_8, 3, x_4); +x_9 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_2, 0); +lean_inc(x_7); +lean_dec(x_2); +lean_inc(x_6); +x_8 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__4), 5, 4); +lean_closure_set(x_8, 0, x_4); +lean_closure_set(x_8, 1, x_1); +lean_closure_set(x_8, 2, x_5); +lean_closure_set(x_8, 3, x_6); +x_9 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Server_Completion_forEligibleDeclsM___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_PersistentHashMap_forM___at_Lean_Server_Completion_forEligibleDeclsM___spec__2___rarg___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___lambda__1(x_7, x_2, x_3, x_4, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +size_t x_7; size_t x_8; lean_object* x_9; +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Server_Completion_forEligibleDeclsM___spec__4___rarg(x_1, x_2, x_3, x_7, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Server_Completion_forEligibleDeclsM___rarg___lambda__2(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_forEligibleDeclsM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Lean_Server_Completion_forEligibleDeclsM___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_3); +return x_6; +} +} +LEAN_EXPORT uint8_t l_Lean_Server_Completion_allowCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; uint8_t x_19; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_array_get_size(x_4); +x_6 = l_Lean_Name_hash___override(x_3); +x_7 = 32; +x_8 = lean_uint64_shift_right(x_6, x_7); +x_9 = lean_uint64_xor(x_6, x_8); +x_10 = 16; +x_11 = lean_uint64_shift_right(x_9, x_10); +x_12 = lean_uint64_xor(x_9, x_11); +x_13 = lean_uint64_to_usize(x_12); +x_14 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_15 = 1; +x_16 = lean_usize_sub(x_14, x_15); +x_17 = lean_usize_land(x_13, x_16); +x_18 = lean_array_uget(x_4, x_17); +x_19 = l_Std_DHashMap_Internal_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__6(x_3, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__2(x_21, x_3); +if (x_22 == 0) +{ +uint8_t x_23; +lean_dec(x_3); +lean_dec(x_2); +x_23 = 0; +return x_23; +} +else +{ +uint8_t x_24; +x_24 = l_Lean_Meta_allowCompletion(x_2, x_3); +return x_24; +} +} +else +{ +uint8_t x_25; +lean_dec(x_3); +lean_dec(x_2); +x_25 = 1; +return x_25; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_allowCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Lean_Server_Completion_allowCompletion(x_1, x_2, x_3); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +lean_object* initialize_Lean_Meta_CompletionName(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_EligibleHeaderDecls(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Meta_CompletionName(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +if (builtin) {res = l_Lean_Server_Completion_initFn____x40_Lean_Server_Completion_EligibleHeaderDecls___hyg_10_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_Server_Completion_eligibleHeaderDeclsRef = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_Server_Completion_eligibleHeaderDeclsRef); +lean_dec_ref(res); +}l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__1); +l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2(); +lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__2); +l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3(); +lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__3); +l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4(); +lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__4); +l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5 = _init_l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5(); +lean_mark_persistent(l_Lean_Server_Completion_getEligibleHeaderDecls___closed__5); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/ImportCompletion.c b/stage0/stdlib/Lean/Server/Completion/ImportCompletion.c similarity index 98% rename from stage0/stdlib/Lean/Server/ImportCompletion.c rename to stage0/stdlib/Lean/Server/Completion/ImportCompletion.c index 5f4ac935b446..da748f3fd3ca 100644 --- a/stage0/stdlib/Lean/Server/ImportCompletion.c +++ b/stage0/stdlib/Lean/Server/Completion/ImportCompletion.c @@ -1,6 +1,6 @@ // Lean compiler output -// Module: Lean.Server.ImportCompletion -// Imports: Lean.Data.Name Lean.Data.NameTrie Lean.Data.Lsp.Utf16 Lean.Data.Lsp.LanguageFeatures Lean.Util.Paths Lean.Util.LakePath Lean.Server.CompletionItemData +// Module: Lean.Server.Completion.ImportCompletion +// Imports: Lean.Data.NameTrie Lean.Util.Paths Lean.Util.LakePath Lean.Server.Completion.CompletionItemData #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -20,6 +20,7 @@ lean_object* l_Lean_initSrcSearchPath(lean_object*, lean_object*); static lean_object* l_ImportCompletion_collectAvailableImportsFromLake___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_ImportCompletion_find___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_ImportCompletion_collectAvailableImportsFromLake___closed__5; +lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_ImportCompletion_collectAvailableImportsFromSrcSearchPath___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_ImportCompletion_computePartialImportCompletions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_NameTrie_toArray___rarg(lean_object*); @@ -56,6 +57,7 @@ lean_object* lean_io_process_child_wait(lean_object*, lean_object*, lean_object* static lean_object* l_ImportCompletion_computePartialImportCompletions___closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at_ImportCompletion_collectAvailableImportsFromSrcSearchPath___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_ImportCompletion_collectAvailableImportsFromSrcSearchPath___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); static lean_object* l_ImportCompletion_computePartialImportCompletions___closed__6; @@ -98,7 +100,6 @@ LEAN_EXPORT lean_object* l_ImportCompletion_determinePartialHeaderCompletions___ static lean_object* l_ImportCompletion_computePartialImportCompletions___closed__4; static lean_object* l_ImportCompletion_computePartialImportCompletions___closed__5; static lean_object* l_Array_mapMUnsafe_map___at_ImportCompletion_collectAvailableImportsFromLake___spec__1___closed__3; -lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_(lean_object*); LEAN_EXPORT lean_object* l_ImportCompletion_computePartialImportCompletions(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_ImportCompletion_computePartialImportCompletions___spec__3___lambda__1(lean_object*); lean_object* lean_string_length(lean_object*); @@ -133,7 +134,6 @@ lean_object* lean_array_mk(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_ImportCompletion_find___spec__2___closed__1; uint8_t l_Substring_beq(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(lean_object*, lean_object*); static lean_object* l_ImportCompletion_collectAvailableImportsFromLake___closed__2; lean_object* l_Lean_NameTrie_matchingToArray___rarg(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); @@ -2248,7 +2248,7 @@ x_29 = lean_ctor_get(x_24, 0); lean_dec(x_29); x_30 = l_System_FilePath_extension(x_23); x_31 = l_Array_forIn_x27Unsafe_loop___at_ImportCompletion_collectAvailableImportsFromSrcSearchPath___spec__2___closed__2; -x_32 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(x_30, x_31); +x_32 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -2342,7 +2342,7 @@ lean_inc(x_52); lean_dec(x_24); x_53 = l_System_FilePath_extension(x_23); x_54 = l_Array_forIn_x27Unsafe_loop___at_ImportCompletion_collectAvailableImportsFromSrcSearchPath___spec__2___closed__2; -x_55 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__3(x_53, x_54); +x_55 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2487____spec__1(x_53, x_54); lean_dec(x_53); if (x_55 == 0) { @@ -3178,7 +3178,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_1 x_10 = lean_ctor_get(x_6, 6); lean_dec(x_10); lean_inc(x_1); -x_11 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_(x_1); +x_11 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(x_1); x_12 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_6, 6, x_12); @@ -3208,7 +3208,7 @@ lean_inc(x_18); lean_inc(x_17); lean_dec(x_6); lean_inc(x_1); -x_24 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_(x_1); +x_24 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(x_1); x_25 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_25, 0, x_24); x_26 = lean_alloc_ctor(0, 8, 0); @@ -3530,37 +3530,25 @@ lean_dec(x_1); return x_5; } } -lean_object* initialize_Lean_Data_Name(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Data_NameTrie(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Lsp_Utf16(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Data_Lsp_LanguageFeatures(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_Paths(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_LakePath(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Server_CompletionItemData(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_CompletionItemData(uint8_t builtin, lean_object*); static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Server_ImportCompletion(uint8_t builtin, lean_object* w) { +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_ImportCompletion(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; -res = initialize_Lean_Data_Name(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Data_NameTrie(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Data_Lsp_Utf16(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -res = initialize_Lean_Data_Lsp_LanguageFeatures(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); res = initialize_Lean_Util_Paths(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); res = initialize_Lean_Util_LakePath(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Server_CompletionItemData(builtin, lean_io_mk_world()); +res = initialize_Lean_Server_Completion_CompletionItemData(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_ImportCompletion_AvailableImports_toImportTrie___closed__1 = _init_l_ImportCompletion_AvailableImports_toImportTrie___closed__1(); diff --git a/stage0/stdlib/Lean/Server/Completion/SyntheticCompletion.c b/stage0/stdlib/Lean/Server/Completion/SyntheticCompletion.c new file mode 100644 index 000000000000..37a9168fb9d9 --- /dev/null +++ b/stage0/stdlib/Lean/Server/Completion/SyntheticCompletion.c @@ -0,0 +1,5561 @@ +// Lean compiler output +// Module: Lean.Server.Completion.SyntheticCompletion +// Imports: Lean.Server.InfoUtils Lean.Server.Completion.CompletionUtils +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint32_t lean_string_utf8_get(lean_object*, lean_object*); +static lean_object* l_Lean_Server_Completion_findSyntheticCompletions___closed__1; +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1; +lean_object* l_Lean_Syntax_getId(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1___boxed(lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArgs(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1; +lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findSyntheticCompletions(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_head_x3f___rarg(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_LocalContext_empty; +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2; +lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(lean_object*); +uint8_t l_List_any___rarg(lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isAtom(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +lean_object* l_Lean_Elab_InfoTree_smallestInfo_x3f(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed(lean_object*); +lean_object* l_Lean_Elab_Info_tailPos_x3f(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(lean_object*, lean_object*); +lean_object* lean_string_utf8_next(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +static lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Info_updateContext_x3f(lean_object*, lean_object*); +lean_object* l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getKind(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1; +uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed(lean_object*); +lean_object* l_Lean_FileMap_lineStart(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; +uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed(lean_object*, lean_object*); +lean_object* l_Array_zipWithIndex___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken(lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getTrailingTailPos_x3f(lean_object*, uint8_t); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7; +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2(lean_object*); +uint8_t l_Lean_Syntax_isToken(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_hasArgs(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_findStack_x3f(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1___boxed(lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(lean_object*, lean_object*); +lean_object* lean_nat_mod(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2; +uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace___boxed(lean_object*, lean_object*); +extern lean_object* l_Id_instMonad; +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1; +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(lean_object*, size_t, size_t); +lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f(lean_object*); +uint8_t l_Lean_isStructure(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5; +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* l_List_reverse___rarg(lean_object*); +lean_object* lean_array_mk(lean_object*); +lean_object* l_List_dropWhile___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1(lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Elab_Info_isSmaller(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* l_Lean_Elab_Info_lctx(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4; +lean_object* l_Lean_Elab_Info_pos_x3f(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(lean_object*, lean_object*, size_t, size_t); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l_Lean_Elab_Info_stx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1(lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Elab_Info_occursInOrOnBoundary(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +uint8_t lean_local_ctx_is_empty(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2; +static lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2; +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion(lean_object*, lean_object*, lean_object*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_4; lean_object* x_5; +lean_dec(x_3); +lean_dec(x_1); +x_4 = 0; +x_5 = lean_box(x_4); +return x_5; +} +else +{ +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = 1; +x_7 = lean_box(x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_apply_2(x_1, x_8, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_1, 0); +lean_dec(x_7); +x_8 = lean_box(0); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_8); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_1, 1); +x_16 = lean_ctor_get(x_1, 0); +lean_dec(x_16); +x_17 = lean_ctor_get(x_4, 0); +lean_inc(x_17); +lean_dec(x_4); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_17); +{ +lean_object* _tmp_0 = x_15; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_1, 1); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); +lean_dec(x_4); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_2); +x_1 = x_19; +x_2 = x_21; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_2, x_4); +x_7 = lean_unbox(x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_dec(x_2); +x_2 = x_4; +x_3 = x_5; +goto _start; +} +else +{ +lean_dec(x_4); +x_3 = x_5; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_box(0); +x_8 = lean_box(0); +x_9 = l_List_mapTR_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__1___rarg(x_6, x_8); +lean_inc(x_1); +x_10 = l_List_foldl___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___spec__2___rarg(x_1, x_7, x_9); +x_11 = lean_apply_3(x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_11) == 0) +{ +lean_dec(x_1); +return x_10; +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +lean_inc(x_10); +lean_inc(x_11); +x_13 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_11, x_10); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_11); +return x_10; +} +else +{ +lean_dec(x_10); +return x_11; +} +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_11, 0); +lean_inc(x_15); +lean_dec(x_11); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +lean_inc(x_10); +lean_inc(x_16); +x_17 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_isBetter___rarg(x_1, x_16, x_10); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_dec(x_16); +return x_10; +} +else +{ +lean_dec(x_10); +return x_16; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___rarg), 6, 0); +return x_2; +} +} +static lean_object* _init_l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Id_instMonad; +x_3 = l_instInhabitedOfMonad___rarg(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_6; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_6 = l_List_reverse___rarg(x_5); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_4); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_10 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_1, x_2, x_3, x_8); +lean_ctor_set(x_4, 1, x_5); +lean_ctor_set(x_4, 0, x_10); +{ +lean_object* _tmp_3 = x_9; +lean_object* _tmp_4 = x_4; +x_4 = _tmp_3; +x_5 = _tmp_4; +} +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_4, 0); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_14 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_1, x_2, x_3, x_12); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_5); +x_4 = x_13; +x_5 = x_15; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg), 5, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Server.InfoUtils", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean.Elab.InfoTree.visitM.go", 28, 28); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("unexpected context-free info tree node", 38, 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1; +x_2 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2; +x_3 = lean_unsigned_to_nat(62u); +x_4 = lean_unsigned_to_nat(21u); +x_5 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +switch (lean_obj_tag(x_4)) { +case 0: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_5, x_3); +x_3 = x_7; +x_4 = x_6; +goto _start; +} +case 1: +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_9 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4; +x_10 = l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg(x_9); +return x_10; +} +default: +{ +lean_object* x_11; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box(0); +return x_11; +} +} +} +else +{ +switch (lean_obj_tag(x_4)) { +case 0: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); +lean_dec(x_4); +x_14 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_12, x_3); +x_3 = x_14; +x_4 = x_13; +goto _start; +} +case 1: +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_3); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_17 = lean_ctor_get(x_3, 0); +x_18 = lean_ctor_get(x_4, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +lean_dec(x_4); +lean_inc(x_1); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_20 = lean_apply_3(x_1, x_17, x_18, x_19); +x_21 = lean_unbox(x_20); +lean_dec(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_1); +x_22 = lean_box(0); +x_23 = lean_apply_4(x_2, x_17, x_18, x_19, x_22); +lean_ctor_set(x_3, 0, x_23); +return x_3; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_inc(x_17); +x_24 = l_Lean_Elab_Info_updateContext_x3f(x_3, x_18); +x_25 = l_Lean_PersistentArray_toList___rarg(x_19); +x_26 = lean_box(0); +lean_inc(x_2); +x_27 = l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(x_1, x_2, x_24, x_25, x_26); +x_28 = lean_apply_4(x_2, x_17, x_18, x_19, x_27); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_30 = lean_ctor_get(x_3, 0); +lean_inc(x_30); +lean_dec(x_3); +x_31 = lean_ctor_get(x_4, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_4, 1); +lean_inc(x_32); +lean_dec(x_4); +lean_inc(x_1); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +x_33 = lean_apply_3(x_1, x_30, x_31, x_32); +x_34 = lean_unbox(x_33); +lean_dec(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_1); +x_35 = lean_box(0); +x_36 = lean_apply_4(x_2, x_30, x_31, x_32, x_35); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_inc(x_30); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_30); +x_39 = l_Lean_Elab_Info_updateContext_x3f(x_38, x_31); +x_40 = l_Lean_PersistentArray_toList___rarg(x_32); +x_41 = lean_box(0); +lean_inc(x_2); +x_42 = l_List_mapM_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__3___rarg(x_1, x_2, x_39, x_40, x_41); +x_43 = lean_apply_4(x_2, x_30, x_31, x_32, x_42); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +return x_44; +} +} +} +default: +{ +lean_object* x_45; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_45 = lean_box(0); +return x_45; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f_choose___rarg), 6, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_3); +x_5 = lean_box(0); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1; +x_7 = l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg(x_6, x_4, x_5, x_1); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; +x_8 = lean_box(0); +return x_8; +} +else +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_Elab_Info_lctx(x_3); +x_6 = lean_local_ctx_is_empty(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = l_Lean_Elab_Info_lctx(x_4); +x_8 = lean_local_ctx_is_empty(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = l_Lean_Elab_Info_isSmaller(x_3, x_4); +return x_9; +} +else +{ +uint8_t x_10; +x_10 = 1; +return x_10; +} +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = l_Lean_Elab_Info_lctx(x_4); +x_12 = lean_local_ctx_is_empty(x_11); +if (x_12 == 0) +{ +uint8_t x_13; +x_13 = 0; +return x_13; +} +else +{ +uint8_t x_14; +x_14 = l_Lean_Elab_Info_isSmaller(x_3, x_4); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = l_Lean_Elab_Info_occursInOrOnBoundary(x_3, x_1); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_3); +lean_dec(x_2); +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_2); +lean_ctor_set(x_7, 1, x_3); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f_isBetter___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed), 4, 1); +lean_closure_set(x_3, 0, x_1); +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg(x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = 0; +x_4 = l_Lean_Syntax_getRange_x3f(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; uint8_t x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = 1; +x_8 = l_String_Range_contains(x_6, x_1, x_7); +lean_dec(x_6); +return x_8; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_Lean_Syntax_hasArgs(x_1); +if (x_2 == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ident", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("completion", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; +lean_inc(x_2); +x_4 = l_Lean_Syntax_isOfKind(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7; +lean_inc(x_2); +x_6 = l_Lean_Syntax_isOfKind(x_2, x_5); +if (x_6 == 0) +{ +uint8_t x_7; +lean_dec(x_2); +x_7 = 1; +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_Syntax_getArg(x_2, x_8); +lean_dec(x_2); +x_10 = l_Lean_Syntax_isOfKind(x_9, x_3); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +else +{ +uint8_t x_13; +lean_dec(x_2); +x_13 = 0; +return x_13; +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dotIdent", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2; +lean_inc(x_2); +x_4 = l_Lean_Syntax_isOfKind(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +lean_dec(x_2); +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_unsigned_to_nat(1u); +x_7 = l_Lean_Syntax_getArg(x_2, x_6); +lean_dec(x_2); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; +x_9 = l_Lean_Syntax_isOfKind(x_7, x_8); +return x_9; +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Init.Data.Option.BasicAux", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Option.get!", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("value is none", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_43; uint8_t x_44; +x_43 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2; +lean_inc(x_1); +x_44 = l_Lean_Syntax_isOfKind(x_1, x_43); +if (x_44 == 0) +{ +lean_object* x_45; uint8_t x_46; +x_45 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7; +lean_inc(x_1); +x_46 = l_Lean_Syntax_isOfKind(x_1, x_45); +if (x_46 == 0) +{ +lean_object* x_47; +lean_dec(x_4); +lean_dec(x_1); +x_47 = lean_box(0); +return x_47; +} +else +{ +lean_object* x_48; lean_object* x_49; uint8_t x_50; +x_48 = lean_unsigned_to_nat(0u); +x_49 = l_Lean_Syntax_getArg(x_1, x_48); +lean_inc(x_49); +x_50 = l_Lean_Syntax_isOfKind(x_49, x_43); +if (x_50 == 0) +{ +lean_object* x_51; +lean_dec(x_49); +lean_dec(x_4); +lean_dec(x_1); +x_51 = lean_box(0); +return x_51; +} +else +{ +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_52 = l_Lean_Syntax_getId(x_49); +lean_dec(x_49); +x_53 = 1; +x_54 = lean_box(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +x_6 = x_55; +goto block_42; +} +} +} +else +{ +lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; +x_56 = l_Lean_Syntax_getId(x_1); +x_57 = 0; +x_58 = lean_box(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_58); +x_6 = x_59; +goto block_42; +} +block_42: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = 0; +x_10 = l_Lean_Syntax_getTailPos_x3f(x_1, x_9); +x_11 = l_Lean_Elab_Info_lctx(x_2); +x_12 = lean_box(0); +x_13 = lean_alloc_ctor(1, 4, 1); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_7); +lean_ctor_set(x_13, 2, x_11); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_unbox(x_8); +lean_dec(x_8); +lean_ctor_set_uint8(x_13, sizeof(void*)*4, x_14); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_16 = l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(x_15); +x_17 = lean_nat_dec_lt(x_3, x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_16); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_4); +lean_ctor_set(x_19, 2, x_13); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_nat_sub(x_16, x_3); +lean_dec(x_16); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_4); +lean_ctor_set(x_23, 2, x_13); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +else +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_10); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_10, 0); +x_27 = lean_nat_dec_lt(x_3, x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_26); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_4); +lean_ctor_set(x_29, 2, x_13); +lean_ctor_set(x_10, 0, x_29); +return x_10; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_nat_sub(x_26, x_3); +lean_dec(x_26); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_4); +lean_ctor_set(x_32, 2, x_13); +lean_ctor_set(x_10, 0, x_32); +return x_10; +} +} +else +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_10, 0); +lean_inc(x_33); +lean_dec(x_10); +x_34 = lean_nat_dec_lt(x_3, x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_33); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_4); +lean_ctor_set(x_36, 2, x_13); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_nat_sub(x_33, x_3); +lean_dec(x_33); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_4); +lean_ctor_set(x_40, 2, x_13); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +lean_inc(x_1); +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l_Lean_Elab_Info_stx(x_7); +lean_inc(x_1); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed), 2, 1); +lean_closure_set(x_9, 0, x_1); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1; +x_11 = l_Lean_Syntax_findStack_x3f(x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2; +x_15 = l_List_dropWhile___rarg(x_14, x_13); +x_16 = l_List_head_x3f___rarg(x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +lean_dec(x_15); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_17 = lean_box(0); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3; +x_21 = l_List_any___rarg(x_15, x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(x_19, x_7, x_1, x_6, x_22); +lean_dec(x_1); +lean_dec(x_7); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_19); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_24 = lean_box(0); +return x_24; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__2(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_string_utf8_next(x_5, x_6); +lean_dec(x_6); +lean_ctor_set(x_2, 1, x_7); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); +x_12 = lean_string_utf8_next(x_10, x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_10); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +static lean_object* _init_l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = lean_nat_dec_lt(x_6, x_1); +lean_dec(x_6); +if (x_7 == 0) +{ +return x_2; +} +else +{ +lean_object* x_8; uint32_t x_9; lean_object* x_10; lean_object* x_11; uint32_t x_12; uint8_t x_13; +x_8 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; +x_9 = 32; +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_5, 1); +lean_inc(x_11); +x_12 = lean_string_utf8_get(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = lean_uint32_dec_eq(x_12, x_9); +if (x_13 == 0) +{ +uint32_t x_14; uint8_t x_15; +x_14 = 9; +x_15 = lean_uint32_dec_eq(x_12, x_14); +if (x_15 == 0) +{ +return x_2; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_free_object(x_2); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_add(x_4, x_16); +lean_dec(x_4); +x_18 = lean_box(0); +x_19 = lean_apply_3(x_8, x_17, x_5, x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_dec(x_19); +return x_20; +} +else +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_2 = x_21; +goto _start; +} +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_free_object(x_2); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_4, x_23); +lean_dec(x_4); +x_25 = lean_box(0); +x_26 = lean_apply_3(x_8, x_24, x_5, x_25); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +lean_dec(x_26); +return x_27; +} +else +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_2 = x_28; +goto _start; +} +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_ctor_get(x_2, 0); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_2); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +x_33 = lean_nat_dec_lt(x_32, x_1); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_30); +lean_ctor_set(x_34, 1, x_31); +return x_34; +} +else +{ +lean_object* x_35; uint32_t x_36; lean_object* x_37; lean_object* x_38; uint32_t x_39; uint8_t x_40; +x_35 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1; +x_36 = 32; +x_37 = lean_ctor_get(x_31, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +x_39 = lean_string_utf8_get(x_37, x_38); +lean_dec(x_38); +lean_dec(x_37); +x_40 = lean_uint32_dec_eq(x_39, x_36); +if (x_40 == 0) +{ +uint32_t x_41; uint8_t x_42; +x_41 = 9; +x_42 = lean_uint32_dec_eq(x_39, x_41); +if (x_42 == 0) +{ +lean_object* x_43; +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_30); +lean_ctor_set(x_43, 1, x_31); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_add(x_30, x_44); +lean_dec(x_30); +x_46 = lean_box(0); +x_47 = lean_apply_3(x_35, x_45, x_31, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +lean_dec(x_47); +return x_48; +} +else +{ +lean_object* x_49; +x_49 = lean_ctor_get(x_47, 0); +lean_inc(x_49); +lean_dec(x_47); +x_2 = x_49; +goto _start; +} +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_51 = lean_unsigned_to_nat(1u); +x_52 = lean_nat_add(x_30, x_51); +lean_dec(x_30); +x_53 = lean_box(0); +x_54 = lean_apply_3(x_35, x_52, x_31, x_53); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +lean_dec(x_54); +return x_55; +} +else +{ +lean_object* x_56; +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +lean_dec(x_54); +x_2 = x_56; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = l_Lean_FileMap_lineStart(x_1, x_2); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_add(x_2, x_4); +x_6 = l_Lean_FileMap_lineStart(x_1, x_5); +lean_dec(x_5); +x_7 = !lean_is_exclusive(x_1); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_1, 1); +lean_dec(x_8); +lean_ctor_set(x_1, 1, x_3); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_1); +x_11 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_6, x_10); +lean_dec(x_6); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_6, x_16); +lean_dec(x_6); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_string_utf8_at_end(x_3, x_2); +if (x_4 == 0) +{ +uint32_t x_5; uint32_t x_6; uint8_t x_7; +x_5 = lean_string_utf8_get(x_3, x_2); +x_6 = 32; +x_7 = lean_uint32_dec_eq(x_5, x_6); +if (x_7 == 0) +{ +uint32_t x_8; uint8_t x_9; +x_8 = 9; +x_9 = lean_uint32_dec_eq(x_5, x_8); +if (x_9 == 0) +{ +uint32_t x_10; uint8_t x_11; +x_10 = 13; +x_11 = lean_uint32_dec_eq(x_5, x_10); +if (x_11 == 0) +{ +uint32_t x_12; uint8_t x_13; +x_12 = 10; +x_13 = lean_uint32_dec_eq(x_5, x_12); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = 1; +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = 1; +return x_17; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_20; +x_3 = lean_ctor_get(x_1, 0); +x_20 = lean_string_utf8_at_end(x_3, x_2); +if (x_20 == 0) +{ +uint32_t x_21; uint32_t x_22; uint8_t x_23; +x_21 = lean_string_utf8_get(x_3, x_2); +x_22 = 32; +x_23 = lean_uint32_dec_eq(x_21, x_22); +if (x_23 == 0) +{ +uint32_t x_24; uint8_t x_25; +x_24 = 9; +x_25 = lean_uint32_dec_eq(x_21, x_24); +if (x_25 == 0) +{ +uint32_t x_26; uint8_t x_27; +x_26 = 13; +x_27 = lean_uint32_dec_eq(x_21, x_26); +if (x_27 == 0) +{ +uint32_t x_28; uint8_t x_29; +x_28 = 10; +x_29 = lean_uint32_dec_eq(x_21, x_28); +if (x_29 == 0) +{ +uint8_t x_30; +x_30 = 0; +return x_30; +} +else +{ +lean_object* x_31; +x_31 = lean_box(0); +x_4 = x_31; +goto block_19; +} +} +else +{ +lean_object* x_32; +x_32 = lean_box(0); +x_4 = x_32; +goto block_19; +} +} +else +{ +lean_object* x_33; +x_33 = lean_box(0); +x_4 = x_33; +goto block_19; +} +} +else +{ +lean_object* x_34; +x_34 = lean_box(0); +x_4 = x_34; +goto block_19; +} +} +else +{ +lean_object* x_35; +x_35 = lean_box(0); +x_4 = x_35; +goto block_19; +} +block_19: +{ +lean_object* x_5; lean_object* x_6; uint32_t x_7; uint32_t x_8; uint8_t x_9; +lean_dec(x_4); +x_5 = lean_unsigned_to_nat(1u); +x_6 = lean_nat_sub(x_2, x_5); +x_7 = lean_string_utf8_get(x_3, x_6); +lean_dec(x_6); +x_8 = 32; +x_9 = lean_uint32_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +uint32_t x_10; uint8_t x_11; +x_10 = 9; +x_11 = lean_uint32_dec_eq(x_7, x_10); +if (x_11 == 0) +{ +uint32_t x_12; uint8_t x_13; +x_12 = 13; +x_13 = lean_uint32_dec_eq(x_7, x_12); +if (x_13 == 0) +{ +uint32_t x_14; uint8_t x_15; +x_14 = 10; +x_15 = lean_uint32_dec_eq(x_7, x_14); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = 1; +return x_17; +} +} +else +{ +uint8_t x_18; +x_18 = 1; +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeqBracketed", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +lean_inc(x_1); +x_2 = l_Lean_Syntax_getKind(x_1); +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; +x_4 = lean_name_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; +x_6 = lean_name_eq(x_2, x_5); +lean_dec(x_2); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_1); +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_Syntax_getArg(x_1, x_11); +lean_dec(x_1); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(x_1); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +lean_dec(x_2); +x_7 = 0; +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +lean_dec(x_6); +x_9 = 0; +x_10 = l_Lean_Syntax_getPos_x3f(x_8, x_9); +lean_dec(x_8); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_2); +x_11 = 0; +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +lean_inc(x_2); +x_13 = l_Lean_FileMap_toPosition(x_2, x_12); +lean_dec(x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +lean_inc(x_2); +x_15 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(x_2, x_14); +lean_dec(x_14); +x_16 = lean_nat_dec_eq(x_3, x_15); +lean_dec(x_15); +x_17 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(x_2, x_4); +lean_dec(x_2); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +return x_16; +} +} +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_3, 1); +x_7 = lean_nat_dec_le(x_6, x_4); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_5); +lean_dec(x_1); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(x_5, x_1, x_4, x_2, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_7 = lean_box(x_6); +return x_7; +} +} +static lean_object* _init_l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(";", 1, 1); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_6 = lean_array_uget(x_2, x_3); +x_7 = 0; +x_8 = l_Lean_Syntax_getTailPos_x3f(x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +size_t x_9; size_t x_10; +lean_dec(x_6); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +x_3 = x_10; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_8, 0); +lean_inc(x_12); +lean_dec(x_8); +x_13 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1; +x_14 = l_Lean_Syntax_isToken(x_13, x_6); +if (x_14 == 0) +{ +size_t x_15; size_t x_16; +lean_dec(x_12); +lean_dec(x_6); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +goto _start; +} +else +{ +uint8_t x_18; +x_18 = lean_nat_dec_le(x_12, x_1); +if (x_18 == 0) +{ +size_t x_19; size_t x_20; +lean_dec(x_12); +lean_dec(x_6); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +goto _start; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = l_Lean_Syntax_getTrailingSize(x_6); +lean_dec(x_6); +x_23 = lean_nat_add(x_12, x_22); +lean_dec(x_22); +lean_dec(x_12); +x_24 = lean_nat_dec_le(x_1, x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +size_t x_25; size_t x_26; +x_25 = 1; +x_26 = lean_usize_add(x_3, x_25); +x_3 = x_26; +goto _start; +} +else +{ +uint8_t x_28; +x_28 = 1; +return x_28; +} +} +} +} +} +else +{ +uint8_t x_29; +x_29 = 0; +return x_29; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f(x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Lean_Syntax_getArgs(x_6); +lean_dec(x_6); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace(x_1, x_2); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_7); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_array_get_size(x_7); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_10); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_10); +lean_dec(x_7); +x_13 = 0; +return x_13; +} +else +{ +size_t x_14; size_t x_15; uint8_t x_16; +x_14 = 0; +x_15 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_16 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(x_2, x_7, x_14, x_15); +lean_dec(x_7); +return x_16; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(lean_object* x_1, size_t x_2, size_t x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_uget(x_1, x_2); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +else +{ +size_t x_8; size_t x_9; +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +x_2 = x_9; +goto _start; +} +} +else +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +case 1: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 2); +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +uint8_t x_7; +lean_dec(x_4); +x_7 = 1; +return x_7; +} +else +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(x_3, x_8, x_9); +if (x_10 == 0) +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +else +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +} +} +default: +{ +uint8_t x_13; +x_13 = 0; +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; size_t x_5; uint8_t x_6; lean_object* x_7; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___spec__1(x_1, x_4, x_5); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +lean_inc(x_1); +x_2 = l_Lean_Syntax_getKind(x_1); +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2; +x_4 = lean_name_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3; +x_6 = lean_name_eq(x_2, x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; +x_8 = lean_name_eq(x_2, x_7); +lean_dec(x_2); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_unsigned_to_nat(1u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +lean_dec(x_1); +x_12 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_11); +lean_dec(x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; +x_15 = lean_name_eq(x_2, x_14); +lean_dec(x_2); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_1); +x_16 = 0; +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_unsigned_to_nat(1u); +x_18 = l_Lean_Syntax_getArg(x_1, x_17); +lean_dec(x_1); +x_19 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_18); +lean_dec(x_18); +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_2); +lean_dec(x_1); +x_20 = 1; +return x_20; +} +} +} +else +{ +uint8_t x_21; +x_21 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_1); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5; +x_23 = lean_name_eq(x_2, x_22); +lean_dec(x_2); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_1); +x_24 = 0; +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_unsigned_to_nat(1u); +x_26 = l_Lean_Syntax_getArg(x_1, x_25); +lean_dec(x_1); +x_27 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmpty(x_26); +lean_dec(x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +lean_dec(x_2); +lean_dec(x_1); +x_28 = 1; +return x_28; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(x_1, x_2); +if (x_4 == 0) +{ +uint8_t x_5; +lean_dec(x_3); +x_5 = 0; +return x_5; +} +else +{ +uint8_t x_6; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock(x_3); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_Syntax_getTrailingSize(x_1); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +static lean_object* _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, size_t x_9, size_t x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_10, x_9); +if (x_12 == 0) +{ +lean_dec(x_7); +lean_dec(x_1); +return x_11; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_array_uget(x_8, x_10); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_11, 1); +x_16 = lean_ctor_get(x_11, 0); +lean_dec(x_16); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_1); +x_17 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_13, x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; +lean_free_object(x_11); +x_18 = lean_box(0); +lean_inc(x_7); +x_19 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_13, x_7, x_15, x_18); +lean_dec(x_15); +lean_dec(x_13); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_dec(x_19); +x_21 = 1; +x_22 = lean_usize_add(x_10, x_21); +x_10 = x_22; +x_11 = x_20; +goto _start; +} +else +{ +lean_object* x_24; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_1); +x_24 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1; +lean_ctor_set(x_11, 0, x_24); +return x_11; +} +} +else +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_ctor_get(x_11, 1); +lean_inc(x_25); +lean_dec(x_11); +lean_inc(x_25); +lean_inc(x_13); +lean_inc(x_1); +x_26 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_13, x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; +x_27 = lean_box(0); +lean_inc(x_7); +x_28 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_13, x_7, x_25, x_27); +lean_dec(x_25); +lean_dec(x_13); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +lean_dec(x_28); +x_30 = 1; +x_31 = lean_usize_add(x_10, x_30); +x_10 = x_31; +x_11 = x_29; +goto _start; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_1); +x_33 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_25); +return x_34; +} +} +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; +lean_inc(x_3); +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_3); +if (x_7 == 0) +{ +uint8_t x_8; +lean_inc(x_3); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon(x_1, x_2, x_3); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionOnTacticBlockIndentation(x_1, x_2, x_4, x_5, x_3); +return x_9; +} +else +{ +uint8_t x_10; +lean_dec(x_3); +lean_dec(x_1); +x_10 = 1; +return x_10; +} +} +else +{ +uint8_t x_11; +lean_dec(x_3); +lean_dec(x_1); +x_11 = 1; +return x_11; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; +x_8 = l_Lean_Syntax_getArgs(x_1); +x_9 = lean_box(0); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_2); +x_12 = lean_array_size(x_8); +x_13 = 0; +lean_inc(x_3); +x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(x_3, x_4, x_5, x_6, x_8, x_9, x_10, x_8, x_12, x_13, x_11); +lean_dec(x_8); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; uint8_t x_17; +x_16 = lean_box(0); +x_17 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(x_3, x_4, x_1, x_5, x_6, x_16); +return x_17; +} +else +{ +lean_object* x_18; uint8_t x_19; +lean_dec(x_3); +lean_dec(x_1); +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +return x_19; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = 0; +x_8 = l_Lean_Syntax_getPos_x3f(x_5, x_7); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +lean_dec(x_6); +x_9 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_5); +lean_dec(x_1); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Syntax_getTailPos_x3f(x_5, x_7); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +lean_dec(x_10); +lean_dec(x_6); +x_12 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionInEmptyTacticBlock(x_1, x_2, x_5); +lean_dec(x_1); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_nat_sub(x_10, x_6); +lean_dec(x_10); +x_15 = lean_nat_dec_le(x_14, x_2); +lean_dec(x_14); +if (x_15 == 0) +{ +uint8_t x_16; +lean_dec(x_13); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_16 = 0; +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = l_Lean_Syntax_getTrailingSize(x_5); +x_18 = lean_nat_add(x_13, x_17); +lean_dec(x_17); +lean_dec(x_13); +x_19 = lean_nat_dec_le(x_2, x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +x_20 = 0; +return x_20; +} +else +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_box(0); +x_22 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(x_5, x_6, x_1, x_2, x_3, x_4, x_21); +return x_22; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +size_t x_12; size_t x_13; lean_object* x_14; +x_12 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_13 = lean_unbox_usize(x_10); +lean_dec(x_10); +x_14 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12, x_13, x_11); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_14; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go(x_1, x_2, x_3, x_5, x_4, x_7); +return x_8; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Id_instMonad; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +lean_dec(x_2); +x_4 = lean_box(0); +x_5 = lean_apply_2(x_3, lean_box(0), x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +lean_inc(x_1); +x_4 = l_Lean_FileMap_toPosition(x_1, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_inc(x_1); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(x_1, x_5); +lean_dec(x_5); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); +x_8 = lean_nat_dec_lt(x_7, x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_7); +x_9 = l_Id_instMonad; +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +x_11 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed), 6, 5); +lean_closure_set(x_11, 0, x_1); +lean_closure_set(x_11, 1, x_2); +lean_closure_set(x_11, 2, x_4); +lean_closure_set(x_11, 3, x_3); +lean_closure_set(x_11, 4, x_6); +x_12 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; +x_13 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_12, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_6); +x_14 = l_Id_instMonad; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed), 6, 5); +lean_closure_set(x_16, 0, x_1); +lean_closure_set(x_16, 1, x_2); +lean_closure_set(x_16, 2, x_4); +lean_closure_set(x_16, 3, x_3); +lean_closure_set(x_16, 4, x_7); +x_17 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1; +x_18 = lean_apply_4(x_15, lean_box(0), lean_box(0), x_17, x_16); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_inc(x_7); +return x_7; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_9); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 0) +{ +size_t x_11; size_t x_12; +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +{ +size_t _tmp_5 = x_12; +lean_object* _tmp_6 = x_3; +x_6 = _tmp_5; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_10); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_inc(x_7); +return x_7; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_9); +if (lean_obj_tag(x_10) == 0) +{ +size_t x_11; size_t x_12; +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +{ +size_t _tmp_5 = x_12; +lean_object* _tmp_6 = x_3; +x_6 = _tmp_5; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_10); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +static lean_object* _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = lean_box(0); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2; +x_2 = lean_box(0); +x_3 = lean_apply_1(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; size_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_box(0); +x_4 = lean_array_size(x_2); +x_5 = 0; +x_6 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; +x_7 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(x_2, x_3, x_6, x_2, x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; +x_9 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3; +return x_9; +} +else +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +return x_10; +} +} +else +{ +lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_box(0); +x_13 = lean_array_size(x_11); +x_14 = 0; +x_15 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; +x_16 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(x_11, x_12, x_15, x_11, x_13, x_14, x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +x_18 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3; +return x_18; +} +else +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +return x_19; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_6, x_5); +if (x_8 == 0) +{ +lean_inc(x_7); +return x_7; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_array_uget(x_4, x_6); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_9); +if (lean_obj_tag(x_10) == 0) +{ +size_t x_11; size_t x_12; +x_11 = 1; +x_12 = lean_usize_add(x_6, x_11); +{ +size_t _tmp_5 = x_12; +lean_object* _tmp_6 = x_3; +x_6 = _tmp_5; +x_7 = _tmp_6; +} +goto _start; +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_10); +x_16 = lean_box(0); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_box(0); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_box(0); +x_6 = lean_array_size(x_4); +x_7 = 0; +x_8 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1; +x_9 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(x_4, x_5, x_8, x_4, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +x_11 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3; +return x_11; +} +else +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_3); +if (x_13 == 0) +{ +return x_3; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_3, 0); +lean_inc(x_14); +lean_dec(x_3); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(lean_object* x_1) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 1); +lean_dec(x_4); +x_5 = lean_ctor_get(x_1, 0); +lean_dec(x_5); +x_6 = !lean_is_exclusive(x_2); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_box(0); +lean_ctor_set(x_1, 1, x_8); +lean_ctor_set(x_1, 0, x_7); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_2, 0); +lean_inc(x_9); +lean_dec(x_2); +x_10 = lean_box(0); +lean_ctor_set(x_1, 1, x_10); +lean_ctor_set(x_1, 0, x_9); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_1); +return x_11; +} +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_12 = lean_ctor_get(x_2, 0); +lean_inc(x_12); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_13 = x_2; +} else { + lean_dec_ref(x_2); + x_13 = lean_box(0); +} +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_12); +lean_ctor_set(x_15, 1, x_14); +if (lean_is_scalar(x_13)) { + x_16 = lean_alloc_ctor(1, 1, 0); +} else { + x_16 = x_13; + lean_ctor_set_tag(x_16, 1); +} +lean_ctor_set(x_16, 0, x_15); +return x_16; +} +} +else +{ +lean_object* x_17; +lean_dec(x_2); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +lean_dec(x_1); +x_1 = x_17; +goto _start; +} +} +case 1: +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_1, 1); +lean_inc(x_19); +lean_dec(x_1); +x_20 = l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(x_19); +lean_dec(x_19); +return x_20; +} +default: +{ +lean_object* x_21; +lean_dec(x_1); +x_21 = lean_box(0); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__3(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__4(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__5(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PersistentArray_findSomeM_x3f___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(7, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_box(0); +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1; +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_4); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go(x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion(x_1, x_2, x_3); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_7); +x_10 = lean_box(0); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(x_7, x_11); +return x_12; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Elab_Info_pos_x3f(x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l_Lean_Elab_Info_tailPos_x3f(x_2); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +lean_dec(x_5); +x_7 = 0; +return x_7; +} +else +{ +if (lean_obj_tag(x_2) == 1) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_ctor_get(x_8, 2); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +lean_dec(x_6); +lean_dec(x_5); +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_6, 0); +lean_inc(x_11); +lean_dec(x_6); +x_12 = lean_nat_dec_le(x_5, x_1); +lean_dec(x_5); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_11); +x_13 = 0; +return x_13; +} +else +{ +uint8_t x_14; +x_14 = lean_nat_dec_le(x_1, x_11); +lean_dec(x_11); +return x_14; +} +} +} +else +{ +uint8_t x_15; +lean_dec(x_6); +lean_dec(x_5); +x_15 = 0; +return x_15; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1___boxed), 2, 1); +lean_closure_set(x_3, 0, x_2); +x_4 = l_Lean_Elab_InfoTree_smallestInfo_x3f(x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +if (lean_obj_tag(x_8) == 1) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_7); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_7, 1); +lean_dec(x_12); +x_13 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_14 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_13); +lean_ctor_set(x_7, 1, x_14); +return x_4; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +lean_dec(x_7); +x_16 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_17 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_15); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_4, 0, x_18); +return x_4; +} +} +else +{ +uint8_t x_19; +lean_free_object(x_4); +x_19 = !lean_is_exclusive(x_7); +if (x_19 == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_7, 1); +lean_dec(x_20); +x_21 = !lean_is_exclusive(x_10); +if (x_21 == 0) +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_10, 0); +lean_ctor_set(x_7, 1, x_22); +lean_ctor_set(x_10, 0, x_7); +return x_10; +} +else +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_10, 0); +lean_inc(x_23); +lean_dec(x_10); +lean_ctor_set(x_7, 1, x_23); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_7); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_25 = lean_ctor_get(x_7, 0); +lean_inc(x_25); +lean_dec(x_7); +x_26 = lean_ctor_get(x_10, 0); +lean_inc(x_26); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + x_27 = x_10; +} else { + lean_dec_ref(x_10); + x_27 = lean_box(0); +} +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_26); +if (lean_is_scalar(x_27)) { + x_29 = lean_alloc_ctor(1, 1, 0); +} else { + x_29 = x_27; +} +lean_ctor_set(x_29, 0, x_28); +return x_29; +} +} +} +else +{ +lean_object* x_30; +lean_dec(x_8); +lean_free_object(x_4); +lean_dec(x_7); +x_30 = lean_box(0); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_4, 0); +lean_inc(x_31); +lean_dec(x_4); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 1) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +lean_dec(x_32); +x_34 = lean_ctor_get(x_33, 2); +lean_inc(x_34); +lean_dec(x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_31, 0); +lean_inc(x_35); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_36 = x_31; +} else { + lean_dec_ref(x_31); + x_36 = lean_box(0); +} +x_37 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4; +x_38 = l_panic___at_Lean_Expr_appFn_x21___spec__1(x_37); +if (lean_is_scalar(x_36)) { + x_39 = lean_alloc_ctor(0, 2, 0); +} else { + x_39 = x_36; +} +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_39); +return x_40; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_41 = lean_ctor_get(x_31, 0); +lean_inc(x_41); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + x_42 = x_31; +} else { + lean_dec_ref(x_31); + x_42 = lean_box(0); +} +x_43 = lean_ctor_get(x_34, 0); +lean_inc(x_43); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + x_44 = x_34; +} else { + lean_dec_ref(x_34); + x_44 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_42; +} +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_43); +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(1, 1, 0); +} else { + x_46 = x_44; +} +lean_ctor_set(x_46, 0, x_45); +return x_46; +} +} +else +{ +lean_object* x_47; +lean_dec(x_32); +lean_dec(x_31); +x_47 = lean_box(0); +return x_47; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, size_t x_7, size_t x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_8, x_7); +if (x_10 == 0) +{ +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_array_uget(x_6, x_8); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_9, 0); +x_15 = lean_ctor_get(x_9, 1); +lean_dec(x_15); +lean_inc(x_3); +lean_inc(x_1); +x_16 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_1, x_2, x_3, x_14, x_11); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_ctor_set(x_9, 0, x_18); +x_19 = 1; +x_20 = lean_usize_add(x_8, x_19); +x_8 = x_20; +goto _start; +} +else +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +size_t x_24; size_t x_25; +lean_ctor_set(x_9, 1, x_17); +lean_ctor_set(x_9, 0, x_22); +x_24 = 1; +x_25 = lean_usize_add(x_8, x_24); +x_8 = x_25; +goto _start; +} +else +{ +lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; +x_27 = lean_ctor_get(x_17, 0); +lean_inc(x_27); +lean_dec(x_17); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_9, 1, x_28); +lean_ctor_set(x_9, 0, x_22); +x_29 = 1; +x_30 = lean_usize_add(x_8, x_29); +x_8 = x_30; +goto _start; +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_9, 0); +lean_inc(x_32); +lean_dec(x_9); +lean_inc(x_3); +lean_inc(x_1); +x_33 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_1, x_2, x_3, x_32, x_11); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_12); +x_37 = 1; +x_38 = lean_usize_add(x_8, x_37); +x_8 = x_38; +x_9 = x_36; +goto _start; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_dec(x_33); +x_41 = lean_ctor_get(x_34, 0); +lean_inc(x_41); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + x_42 = x_34; +} else { + lean_dec_ref(x_34); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(1, 1, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +x_45 = 1; +x_46 = lean_usize_add(x_8, x_45); +x_8 = x_46; +x_9 = x_44; +goto _start; +} +} +} +else +{ +uint8_t x_48; +x_48 = !lean_is_exclusive(x_9); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_49 = lean_ctor_get(x_9, 0); +x_50 = lean_ctor_get(x_9, 1); +lean_dec(x_50); +x_51 = lean_ctor_get(x_12, 0); +lean_inc(x_51); +x_52 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_52, 0, x_51); +lean_inc(x_1); +x_53 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_1, x_2, x_52, x_49, x_11); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; size_t x_56; size_t x_57; +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +lean_ctor_set(x_9, 0, x_55); +x_56 = 1; +x_57 = lean_usize_add(x_8, x_56); +x_8 = x_57; +goto _start; +} +else +{ +lean_object* x_59; uint8_t x_60; +lean_dec(x_12); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +lean_dec(x_53); +x_60 = !lean_is_exclusive(x_54); +if (x_60 == 0) +{ +size_t x_61; size_t x_62; +lean_ctor_set(x_9, 1, x_54); +lean_ctor_set(x_9, 0, x_59); +x_61 = 1; +x_62 = lean_usize_add(x_8, x_61); +x_8 = x_62; +goto _start; +} +else +{ +lean_object* x_64; lean_object* x_65; size_t x_66; size_t x_67; +x_64 = lean_ctor_get(x_54, 0); +lean_inc(x_64); +lean_dec(x_54); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_9, 1, x_65); +lean_ctor_set(x_9, 0, x_59); +x_66 = 1; +x_67 = lean_usize_add(x_8, x_66); +x_8 = x_67; +goto _start; +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_69 = lean_ctor_get(x_9, 0); +lean_inc(x_69); +lean_dec(x_9); +x_70 = lean_ctor_get(x_12, 0); +lean_inc(x_70); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +lean_inc(x_1); +x_72 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_1, x_2, x_71, x_69, x_11); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; size_t x_76; size_t x_77; +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +lean_dec(x_72); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_12); +x_76 = 1; +x_77 = lean_usize_add(x_8, x_76); +x_8 = x_77; +x_9 = x_75; +goto _start; +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; size_t x_84; size_t x_85; +lean_dec(x_12); +x_79 = lean_ctor_get(x_72, 1); +lean_inc(x_79); +lean_dec(x_72); +x_80 = lean_ctor_get(x_73, 0); +lean_inc(x_80); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + x_81 = x_73; +} else { + lean_dec_ref(x_73); + x_81 = lean_box(0); +} +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(1, 1, 0); +} else { + x_82 = x_81; +} +lean_ctor_set(x_82, 0, x_80); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_79); +lean_ctor_set(x_83, 1, x_82); +x_84 = 1; +x_85 = lean_usize_add(x_8, x_84); +x_8 = x_85; +x_9 = x_83; +goto _start; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg___boxed), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +lean_inc(x_1); +lean_inc(x_5); +lean_inc(x_3); +x_6 = lean_apply_3(x_1, x_4, x_3, x_5); +switch (lean_obj_tag(x_5)) { +case 0: +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +lean_dec(x_1); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +case 1: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; uint8_t x_16; +x_9 = lean_ctor_get(x_5, 2); +lean_inc(x_9); +lean_dec(x_5); +x_10 = lean_box(0); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_6); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_array_size(x_9); +x_14 = 0; +x_15 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg(x_1, x_2, x_3, x_9, x_11, x_9, x_13, x_14, x_12); +lean_dec(x_9); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_ctor_get(x_15, 1); +lean_ctor_set(x_15, 1, x_17); +lean_ctor_set(x_15, 0, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; +} +} +default: +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_3); +lean_dec(x_1); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_5); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_6); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_11, x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_box(0); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken_go___rarg(x_2, x_1, x_5, x_3, x_4); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_5; uint8_t x_6; +lean_inc(x_4); +x_5 = lean_apply_2(x_1, x_3, x_4); +x_6 = lean_unbox(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_4); +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_4); +return x_8; +} +} +else +{ +uint8_t x_9; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +return x_2; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +lean_dec(x_2); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(0); +x_4 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f___lambda__1), 4, 1); +lean_closure_set(x_4, 0, x_1); +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_foldWithLeadingToken___rarg(x_3, x_4, x_3, x_2); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = 0; +x_5 = l_Lean_Syntax_getTailPos_x3f(x_1, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_nat_dec_le(x_7, x_2); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_7); +x_9 = 0; +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = l_Lean_Syntax_getTrailingSize(x_1); +x_11 = lean_nat_add(x_7, x_10); +lean_dec(x_10); +lean_dec(x_7); +x_12 = lean_nat_dec_le(x_2, x_11); +lean_dec(x_11); +return x_12; +} +} +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_array_uget(x_2, x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_unsigned_to_nat(2u); +x_10 = lean_nat_mod(x_8, x_9); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_eq(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +uint8_t x_13; +x_13 = l_Lean_Syntax_isAtom(x_7); +if (x_13 == 0) +{ +size_t x_14; size_t x_15; +lean_dec(x_7); +x_14 = 1; +x_15 = lean_usize_add(x_3, x_14); +x_3 = x_15; +goto _start; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_box(0); +x_18 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1(x_7, x_1, x_17); +lean_dec(x_7); +if (x_18 == 0) +{ +size_t x_19; size_t x_20; +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +goto _start; +} +else +{ +uint8_t x_22; +x_22 = 1; +return x_22; +} +} +} +else +{ +size_t x_23; size_t x_24; +lean_dec(x_7); +x_23 = 1; +x_24 = lean_usize_add(x_3, x_23); +x_3 = x_24; +goto _start; +} +} +else +{ +uint8_t x_26; +x_26 = 0; +return x_26; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = 0; +x_6 = l_Lean_Syntax_getPos_x3f(x_1, x_5); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +lean_dec(x_2); +x_7 = 0; +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +lean_dec(x_6); +lean_inc(x_2); +x_9 = l_Lean_FileMap_toPosition(x_2, x_8); +lean_dec(x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(x_2, x_10); +lean_dec(x_10); +x_12 = lean_nat_dec_eq(x_3, x_11); +lean_dec(x_11); +return x_12; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_4, 1); +x_7 = lean_nat_dec_le(x_6, x_3); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_2); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1(x_1, x_2, x_3, x_9); +return x_10; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +if (x_1 == 0) +{ +uint8_t x_7; +lean_dec(x_3); +x_7 = 0; +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_box(0); +x_9 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2(x_2, x_3, x_4, x_5, x_8); +return x_9; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = l_Array_zipWithIndex___rarg(x_1); +x_10 = lean_array_get_size(x_9); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_10); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +lean_dec(x_10); +lean_dec(x_9); +x_13 = lean_box(0); +x_14 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3(x_2, x_3, x_4, x_5, x_6, x_13); +return x_14; +} +else +{ +size_t x_15; size_t x_16; uint8_t x_17; +x_15 = 0; +x_16 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_17 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1(x_7, x_9, x_15, x_16); +lean_dec(x_9); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +x_18 = lean_box(0); +x_19 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3(x_2, x_3, x_4, x_5, x_6, x_18); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_4); +x_20 = 1; +return x_20; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_Lean_Syntax_getArg(x_1, x_9); +x_11 = l_Lean_Syntax_getArgs(x_10); +lean_dec(x_10); +x_12 = 1; +x_13 = l_Lean_Syntax_getTailPos_x3f(x_7, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +lean_dec(x_11); +lean_dec(x_3); +x_14 = 0; +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Syntax_getTrailingTailPos_x3f(x_1, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; +x_17 = l_Lean_Syntax_getTrailingTailPos_x3f(x_7, x_12); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_15); +lean_dec(x_11); +lean_dec(x_3); +x_18 = 0; +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_15); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Array_isEmpty___rarg(x_11); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +lean_dec(x_20); +x_22 = lean_box(0); +x_23 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(x_11, x_2, x_1, x_3, x_4, x_5, x_6, x_22); +lean_dec(x_11); +return x_23; +} +else +{ +uint8_t x_24; +x_24 = l_String_Range_contains(x_20, x_6, x_12); +lean_dec(x_20); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_box(0); +x_26 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(x_11, x_2, x_1, x_3, x_4, x_5, x_6, x_25); +lean_dec(x_11); +return x_26; +} +else +{ +uint8_t x_27; +lean_dec(x_11); +lean_dec(x_3); +x_27 = 1; +return x_27; +} +} +} +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_28 = lean_ctor_get(x_16, 0); +lean_inc(x_28); +lean_dec(x_16); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_15); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Array_isEmpty___rarg(x_11); +if (x_30 == 0) +{ +lean_object* x_31; uint8_t x_32; +lean_dec(x_29); +x_31 = lean_box(0); +x_32 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(x_11, x_2, x_1, x_3, x_4, x_5, x_6, x_31); +lean_dec(x_11); +return x_32; +} +else +{ +uint8_t x_33; +x_33 = l_String_Range_contains(x_29, x_6, x_12); +lean_dec(x_29); +if (x_33 == 0) +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_box(0); +x_35 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(x_11, x_2, x_1, x_3, x_4, x_5, x_6, x_34); +lean_dec(x_11); +return x_35; +} +else +{ +uint8_t x_36; +lean_dec(x_11); +lean_dec(x_3); +x_36 = 1; +return x_36; +} +} +} +} +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3; +x_2 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4; +x_3 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_8; +lean_dec(x_7); +lean_dec(x_2); +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_10 = l_Lean_Syntax_getKind(x_7); +x_11 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2; +x_12 = lean_name_eq(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +uint8_t x_13; +lean_dec(x_7); +lean_dec(x_2); +x_13 = 0; +return x_13; +} +else +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_box(0); +x_15 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5(x_7, x_1, x_2, x_3, x_4, x_5, x_9, x_14); +lean_dec(x_7); +return x_15; +} +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_box(x_1); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___boxed), 7, 5); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_2); +lean_closure_set(x_9, 2, x_6); +lean_closure_set(x_9, 3, x_3); +lean_closure_set(x_9, 4, x_4); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findWithLeadingToken_x3f(x_9, x_5); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = 0; +return x_11; +} +else +{ +uint8_t x_12; +lean_dec(x_10); +x_12 = 1; +return x_12; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_inc(x_1); +x_6 = l_Lean_FileMap_toPosition(x_1, x_2); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_inc(x_1); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount(x_1, x_7); +lean_dec(x_7); +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +x_10 = lean_nat_dec_lt(x_9, x_8); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +lean_dec(x_9); +x_11 = lean_box(0); +x_12 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7(x_3, x_1, x_6, x_2, x_4, x_8, x_11); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +lean_dec(x_8); +x_13 = lean_box(0); +x_14 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7(x_3, x_1, x_6, x_2, x_4, x_9, x_13); +return x_14; +} +} +} +LEAN_EXPORT uint8_t l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorOnWhitespace(x_1, x_2); +if (x_4 == 0) +{ +uint8_t x_5; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = 0; +return x_5; +} +else +{ +uint8_t x_6; lean_object* x_7; uint8_t x_8; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isCursorInProperWhitespace(x_1, x_2); +x_7 = lean_box(0); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8(x_1, x_2, x_6, x_3, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___spec__1(x_1, x_2, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; uint8_t x_8; lean_object* x_9; +x_7 = lean_unbox(x_1); +lean_dec(x_1); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__3(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__4(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_11 = lean_box(x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; uint8_t x_10; lean_object* x_11; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__5(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_11 = lean_box(x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; uint8_t x_9; lean_object* x_10; +x_8 = lean_unbox(x_1); +lean_dec(x_1); +x_9 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_10 = lean_box(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; uint8_t x_9; lean_object* x_10; +x_8 = lean_unbox(x_1); +lean_dec(x_1); +x_9 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__7(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_7); +x_10 = lean_box(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_6 = lean_unbox(x_3); +lean_dec(x_3); +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__8(x_1, x_2, x_6, x_4, x_5); +lean_dec(x_5); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_box(0); +x_5 = lean_box(0); +x_6 = l_Lean_LocalContext_empty; +x_7 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_4); +lean_ctor_set(x_7, 2, x_6); +lean_ctor_set(x_7, 3, x_1); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_2); +lean_ctor_set(x_9, 2, x_7); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findExpectedTypeAt(x_1, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_Expr_getAppFn(x_8); +lean_dec(x_8); +if (lean_obj_tag(x_9) == 4) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +lean_inc(x_10); +x_13 = l_Lean_isStructure(x_12, x_10); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_10); +lean_dec(x_7); +x_14 = lean_box(0); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1(x_10, x_7, x_15); +return x_16; +} +} +else +{ +lean_object* x_17; +lean_dec(x_9); +lean_dec(x_7); +x_17 = lean_box(0); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +lean_inc(x_2); +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion(x_1, x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_4); +lean_dec(x_2); +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2(x_4, x_2, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f___lambda__2(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Server_Completion_findSyntheticCompletions___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_Completion_findSyntheticCompletions(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_5 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; +lean_inc(x_4); +lean_inc(x_2); +x_6 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticFieldCompletion_x3f(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; +x_7 = l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f(x_2, x_4); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; +x_8 = l_Lean_Server_Completion_findSyntheticCompletions___closed__1; +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_array_mk(x_11); +return x_12; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_4); +lean_dec(x_2); +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +lean_dec(x_6); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_array_mk(x_15); +return x_16; +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_ctor_get(x_5, 0); +lean_inc(x_17); +lean_dec(x_5); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_array_mk(x_19); +return x_20; +} +} +} +lean_object* initialize_Lean_Server_InfoUtils(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_CompletionUtils(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Server_Completion_SyntheticCompletion(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Lean_Server_InfoUtils(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Server_Completion_CompletionUtils(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1 = _init_l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1(); +lean_mark_persistent(l_panic___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__2___rarg___closed__1); +l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__1); +l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__2); +l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__3); +l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4 = _init_l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___spec__1___rarg___closed__4); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findBest_x3f___rarg___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findClosestInfoWithLocalContextAt_x3f___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__2); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__3); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__4); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__5); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__6); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__3___closed__7); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__4___closed__2); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__2); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__3); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___lambda__5___closed__4); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__2); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticIdentifierCompletion_x3f___closed__3); +l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1 = _init_l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Loop_forIn_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_getIndentationAmount___spec__1___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__2); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__3); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__4); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_getTacticsNode_x3f___closed__5); +l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1 = _init_l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isCompletionAfterSemicolon___spec__1___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_isEmptyTacticBlock___closed__2); +l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1 = _init_l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1(); +lean_mark_persistent(l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion_go___spec__1___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticTacticCompletion___closed__1); +l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1 = _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1(); +lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__1); +l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2 = _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2(); +lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__2); +l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3 = _init_l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3(); +lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findOutermostContextInfo_x3f_go___spec__2___closed__3); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_findSyntheticTacticCompletion_x3f___lambda__1___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__1); +l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2 = _init_l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2(); +lean_mark_persistent(l___private_Lean_Server_Completion_SyntheticCompletion_0__Lean_Server_Completion_isSyntheticStructFieldCompletion___lambda__6___closed__2); +l_Lean_Server_Completion_findSyntheticCompletions___closed__1 = _init_l_Lean_Server_Completion_findSyntheticCompletions___closed__1(); +lean_mark_persistent(l_Lean_Server_Completion_findSyntheticCompletions___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Server/CompletionItemData.c b/stage0/stdlib/Lean/Server/CompletionItemData.c deleted file mode 100644 index 76668c813afe..000000000000 --- a/stage0/stdlib/Lean/Server/CompletionItemData.c +++ /dev/null @@ -1,528 +0,0 @@ -// Lean compiler output -// Module: Lean.Server.CompletionItemData -// Imports: Lean.Server.FileSource -#include -#if defined(__clang__) -#pragma clang diagnostic ignored "-Wunused-parameter" -#pragma clang diagnostic ignored "-Wunused-label" -#elif defined(__GNUC__) && !defined(__CLANG__) -#pragma GCC diagnostic ignored "-Wunused-parameter" -#pragma GCC diagnostic ignored "-Wunused-label" -#pragma GCC diagnostic ignored "-Wunused-but-set-variable" -#endif -#ifdef __cplusplus -extern "C" { -#endif -LEAN_EXPORT lean_object* l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(lean_object*); -static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; -lean_object* l_Lean_Json_mkObj(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFileSourceCompletionItem; -static lean_object* l_Lean_Lsp_instFileSourceCompletionItem___closed__1; -lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13; -lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3; -static lean_object* l_Lean_Lsp_instToJsonCompletionItemData___closed__1; -lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6; -static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1; -extern lean_object* l_String_instInhabited; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1; -lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_(lean_object*); -static lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15_(lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7; -lean_object* lean_panic_fn(lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_instFromJsonCompletionItemData___closed__1; -LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemData; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14; -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1(lean_object*); -lean_object* lean_array_mk(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemData; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2; -lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11; -static lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(x_3); -return x_4; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1(lean_object* x_1) { -_start: -{ -uint8_t x_2; -x_2 = 0; -return x_2; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("params", 6, 6); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean", 4, 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lsp", 3, 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("CompletionItemData", 18, 18); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3; -x_3 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1___boxed), 1, 0); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5; -x_2 = 1; -x_3 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(".", 1, 1); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10; -x_2 = 1; -x_3 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6; -x_4 = l_Lean_Name_toString(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked(": ", 2, 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13; -x_3 = lean_string_append(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(x_1, x_2); -if (lean_obj_tag(x_3) == 0) -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_3); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14; -x_7 = lean_string_append(x_6, x_5); -lean_dec(x_5); -lean_ctor_set(x_3, 0, x_7); -return x_3; -} -else -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_3, 0); -lean_inc(x_8); -lean_dec(x_3); -x_9 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14; -x_10 = lean_string_append(x_9, x_8); -lean_dec(x_8); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -return x_11; -} -} -else -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_3); -if (x_12 == 0) -{ -return x_3; -} -else -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_3, 0); -lean_inc(x_13); -lean_dec(x_3); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -return x_14; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____lambda__1(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemData___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItemData() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFromJsonCompletionItemData___closed__1; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2476_(x_1); -x_3 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -x_5 = lean_box(0); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_4); -lean_ctor_set(x_6, 1, x_5); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_5); -x_8 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1; -x_9 = l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Lsp_toJsonRpcRef____x40_Lean_Server_Rpc_Basic___hyg_173____spec__1(x_7, x_8); -x_10 = l_Lean_Json_mkObj(x_9); -return x_10; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemData___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82_), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItemData() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instToJsonCompletionItemData___closed__1; -return x_1; -} -} -LEAN_EXPORT lean_object* l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_String_instInhabited; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("no data param on completion item ", 33, 33); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("", 0, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Server.CompletionItemData", 30, 30); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_unchecked("Lean.Lsp.CompletionItem.getFileSource!", 38, 38); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_ctor_get(x_1, 6); -lean_inc(x_2); -if (lean_obj_tag(x_2) == 0) -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -lean_dec(x_1); -x_4 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1; -x_5 = lean_string_append(x_4, x_3); -lean_dec(x_3); -x_6 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2; -x_7 = lean_string_append(x_5, x_6); -x_8 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; -x_9 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; -x_10 = lean_unsigned_to_nat(34u); -x_11 = lean_unsigned_to_nat(22u); -x_12 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_8, x_9, x_10, x_11, x_7); -lean_dec(x_7); -x_13 = l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(x_12); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; -lean_dec(x_1); -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -lean_dec(x_2); -x_15 = l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15_(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3; -x_18 = l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4; -x_19 = lean_unsigned_to_nat(34u); -x_20 = lean_unsigned_to_nat(22u); -x_21 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_17, x_18, x_19, x_20, x_16); -lean_dec(x_16); -x_22 = l_panic___at_Lean_Lsp_CompletionItem_getFileSource_x21___spec__1(x_21); -return x_22; -} -else -{ -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_15, 0); -lean_inc(x_23); -lean_dec(x_15); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -lean_dec(x_23); -return x_24; -} -} -} -} -static lean_object* _init_l_Lean_Lsp_instFileSourceCompletionItem___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Lsp_CompletionItem_getFileSource_x21), 1, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_instFileSourceCompletionItem() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Lsp_instFileSourceCompletionItem___closed__1; -return x_1; -} -} -lean_object* initialize_Lean_Server_FileSource(uint8_t builtin, lean_object*); -static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Server_CompletionItemData(uint8_t builtin, lean_object* w) { -lean_object * res; -if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); -_G_initialized = true; -res = initialize_Lean_Server_FileSource(builtin, lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -lean_dec_ref(res); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__1); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__2); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__3); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__4); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__5); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__6); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__7); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__8); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__9); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__10); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__11); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__12); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__13); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_fromJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_15____closed__14); -l_Lean_Lsp_instFromJsonCompletionItemData___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItemData___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemData___closed__1); -l_Lean_Lsp_instFromJsonCompletionItemData = _init_l_Lean_Lsp_instFromJsonCompletionItemData(); -lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItemData); -l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1 = _init_l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1(); -lean_mark_persistent(l___private_Lean_Server_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_CompletionItemData___hyg_82____closed__1); -l_Lean_Lsp_instToJsonCompletionItemData___closed__1 = _init_l_Lean_Lsp_instToJsonCompletionItemData___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemData___closed__1); -l_Lean_Lsp_instToJsonCompletionItemData = _init_l_Lean_Lsp_instToJsonCompletionItemData(); -lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionItemData); -l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__1); -l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__2); -l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__3); -l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4 = _init_l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4(); -lean_mark_persistent(l_Lean_Lsp_CompletionItem_getFileSource_x21___closed__4); -l_Lean_Lsp_instFileSourceCompletionItem___closed__1 = _init_l_Lean_Lsp_instFileSourceCompletionItem___closed__1(); -lean_mark_persistent(l_Lean_Lsp_instFileSourceCompletionItem___closed__1); -l_Lean_Lsp_instFileSourceCompletionItem = _init_l_Lean_Lsp_instFileSourceCompletionItem(); -lean_mark_persistent(l_Lean_Lsp_instFileSourceCompletionItem); -return lean_io_result_mk_ok(lean_box(0)); -} -#ifdef __cplusplus -} -#endif diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index 2f1566013177..a5fe979dca38 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Server.FileWorker -// Imports: Init.System.IO Init.Data.Channel Lean.Data.RBMap Lean.Environment Lean.Data.Lsp Lean.Data.Json.FromToJson Lean.Util.FileSetupInfo Lean.LoadDynlib Lean.Language.Lean Lean.Server.Utils Lean.Server.AsyncList Lean.Server.References Lean.Server.FileWorker.Utils Lean.Server.FileWorker.RequestHandling Lean.Server.FileWorker.WidgetRequests Lean.Server.FileWorker.SetupFile Lean.Server.Rpc.Basic Lean.Widget.InteractiveDiagnostic Lean.Server.ImportCompletion +// Imports: Init.System.IO Init.Data.Channel Lean.Data.RBMap Lean.Environment Lean.Data.Lsp Lean.Data.Json.FromToJson Lean.Util.FileSetupInfo Lean.LoadDynlib Lean.Language.Lean Lean.Server.Utils Lean.Server.AsyncList Lean.Server.References Lean.Server.FileWorker.Utils Lean.Server.FileWorker.RequestHandling Lean.Server.FileWorker.WidgetRequests Lean.Server.FileWorker.SetupFile Lean.Server.Rpc.Basic Lean.Widget.InteractiveDiagnostic Lean.Server.Completion.ImportCompletion #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -89,6 +89,7 @@ static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__5(lean_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_mainLoop___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_instInhabitedReportSnapshotsState___closed__1; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__55; @@ -104,7 +105,6 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleNotification(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_Server_FileWorker_handleDidChange___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcRelease___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updatePendingRequests___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__21; @@ -140,6 +140,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleGetInteractiveDiagnosticsRequest___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_RBNode_balRight___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_promise_result(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__37; lean_object* l_Prod_map___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Server_FileWorker_mainLoop___spec__2___boxed(lean_object*, lean_object*); @@ -150,7 +151,6 @@ LEAN_EXPORT lean_object* lean_server_worker_main(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__23; static lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots___lambda__6___closed__1; static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Server_FileWorker_runRefreshTask___spec__3___closed__1; -lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker_getImportClosure_x3f(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__64; @@ -424,6 +424,7 @@ static lean_object* l_Lean_Server_FileWorker_sendServerRequest___rarg___closed__ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleRequest___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_(lean_object*); lean_object* l_Lean_Server_foldDocumentChanges(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1850____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__1; static lean_object* l_Lean_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___closed__5; @@ -565,7 +566,6 @@ lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializePar lean_object* l_Lean_Server_maybeTee(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1850____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__11; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(lean_object*); static lean_object* l_Lean_Server_FileWorker_initAndRunWorker_writeErrorDiag___closed__1; lean_object* l___private_Init_Dynamic_0__Dynamic_get_x3fImpl___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_mkIleanInfoUpdateNotification___boxed(lean_object*, lean_object*, lean_object*); @@ -14827,7 +14827,7 @@ x_14 = l_ImportCompletion_find(x_2, x_13, x_3, x_6); x_15 = lean_ctor_get(x_4, 0); lean_inc(x_15); lean_dec(x_4); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(x_14); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(x_14); lean_ctor_set_tag(x_8, 2); lean_ctor_set(x_8, 1, x_16); lean_ctor_set(x_8, 0, x_5); @@ -14873,7 +14873,7 @@ x_28 = l_ImportCompletion_find(x_2, x_27, x_3, x_6); x_29 = lean_ctor_get(x_4, 0); lean_inc(x_29); lean_dec(x_4); -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(x_28); +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(x_28); x_31 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_31, 0, x_5); lean_ctor_set(x_31, 1, x_30); @@ -14912,7 +14912,7 @@ x_13 = l_ImportCompletion_find(x_2, x_12, x_3, x_7); x_14 = lean_ctor_get(x_4, 0); lean_inc(x_14); lean_dec(x_4); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(x_13); +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(x_13); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_5); lean_ctor_set(x_16, 1, x_15); @@ -15282,7 +15282,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_F { lean_object* x_5; lean_inc(x_1); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(x_1); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(x_1); if (lean_obj_tag(x_5) == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; @@ -17025,7 +17025,7 @@ x_85 = lean_ctor_get(x_70, 0); x_86 = lean_ctor_get(x_62, 2); lean_inc(x_86); lean_dec(x_62); -x_87 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_(x_86); +x_87 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_(x_86); x_88 = l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__3(x_87, x_63); if (lean_obj_tag(x_88) == 0) { @@ -17199,7 +17199,7 @@ lean_dec(x_70); x_144 = lean_ctor_get(x_62, 2); lean_inc(x_144); lean_dec(x_62); -x_145 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_(x_144); +x_145 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_(x_144); x_146 = l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__3(x_145, x_63); if (lean_obj_tag(x_146) == 0) { @@ -28081,7 +28081,7 @@ lean_object* initialize_Lean_Server_FileWorker_WidgetRequests(uint8_t builtin, l lean_object* initialize_Lean_Server_FileWorker_SetupFile(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Server_Rpc_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Widget_InteractiveDiagnostic(uint8_t builtin, lean_object*); -lean_object* initialize_Lean_Server_ImportCompletion(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Server_Completion_ImportCompletion(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Server_FileWorker(uint8_t builtin, lean_object* w) { lean_object * res; @@ -28141,7 +28141,7 @@ lean_dec_ref(res); res = initialize_Lean_Widget_InteractiveDiagnostic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -res = initialize_Lean_Server_ImportCompletion(builtin, lean_io_mk_world()); +res = initialize_Lean_Server_Completion_ImportCompletion(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_mkIleanInfoUpdateNotification___closed__1 = _init_l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_mkIleanInfoUpdateNotification___closed__1(); diff --git a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index 18f5a92f78df..50e5e05455b2 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -14,149 +14,150 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283_(lean_object*); lean_object* l_Lean_Server_RequestM_mapTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14; lean_object* l_Lean_Name_getNumParts(lean_object*); static lean_object* l_Array_groupByKey___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__8; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21(lean_object*, lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_compress(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__1(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1; static lean_object* l_Array_groupByKey___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover(lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024_(lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_infoTree(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__10___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Term_identProjKind; lean_object* l_Lean_Widget_diffInteractiveGoals(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_157_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__5; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_NamespaceEntry_finish___lambda__1(lean_object*); +lean_object* l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(lean_object*); +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; static lean_object* l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition(uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); lean_object* l_Lean_Server_documentUriFromModule(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__2(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131_(lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__2___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__14(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__1(lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__1(lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__1___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_List_getLastD___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_FileWorker_handleSemanticTokens_run___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1(size_t, size_t, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1; lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldrTR___at_Lean_Server_FileWorker_NamespaceEntry_finish___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__5; lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__3(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_instBEqAbsoluteLspSemanticToken___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__3___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__2; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1(lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__5; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__4; uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__8; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_Range_includes(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__2; extern lean_object* l_Lean_Lsp_instOrdPosition; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,115 +166,113 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_bindTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__4(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensRange(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__1(lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_groupByKey___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__22(lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__3(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Widget_InteractiveTermGoal_pretty(lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__2(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__10; static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2; uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__3; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__2; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15(lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912_(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__38(lean_object*); static lean_object* l_Lean_Server_FileWorker_instToJsonAbsoluteLspSemanticToken___closed__1; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Nat_nextPowerOfTwo_go(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Server_requestHandlers; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__5; uint8_t l_instDecidableRelLt___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__1; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__9(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704_(lean_object*); LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__14(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__2; static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__28(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__8; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols_popStack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__1(lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__8___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__1(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1057_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_findInfoTreeAtPosWithTrailingWhitespace(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_replace___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__3; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__4; static lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__1___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196_(lean_object*); extern lean_object* l_Lean_projectionFnInfoExt; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9; LEAN_EXPORT lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_initializing(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_fmtHover_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___lambda__1(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___rarg(lean_object*, lean_object*); lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1; size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_pop(lean_object*); static lean_object* l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_findCompletionCmdDataAtPos___lambda__1___boxed(lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__9; @@ -285,6 +284,7 @@ static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___c static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__4; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__8; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__5; +LEAN_EXPORT uint64_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal(lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__5(lean_object*, lean_object*, size_t, size_t); @@ -292,292 +292,273 @@ static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymb static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__5; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2; -lean_object* l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__11(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_filterDuplicateSemanticTokens(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22; LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__2___boxed(lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_shift_right(uint64_t, uint64_t); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__2(lean_object*, lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___boxed(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__3; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_1342_(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__4; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16(size_t, size_t, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__6; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__22(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__6; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqPosition____x40_Lean_Data_Lsp_Basic___hyg_226_(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__28(lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__9; lean_object* l_Except_map___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20(size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__8(lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_env(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10; lean_object* l_Lean_Widget_InteractiveGoal_pretty(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7; lean_object* l_Option_map___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__12; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___boxed(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__18; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_findCompletionCmdDataAtPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_sleep(uint32_t, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__1___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__2(lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855_(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__19; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8; lean_object* l_Lean_Server_RequestM_findCmdDataAtPos(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__3; LEAN_EXPORT lean_object* l_Option_bindM___at_Lean_Server_FileWorker_getInteractiveGoals___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__2; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691_(lean_object*); lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1___boxed(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970_(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleHover___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap; lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__25(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__1(lean_object*); lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485_(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol(lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__5(lean_object*); +lean_object* l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__2(lean_object*, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_asTask___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__11(size_t, size_t, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_(lean_object*); LEAN_EXPORT lean_object* l_Option_bindM___at_Lean_Server_FileWorker_getInteractiveTermGoal___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_findCompletionCmdDataAtPos___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__1; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForDiagnosticsParams____x40_Lean_Data_Lsp_Extra___hyg_370_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___closed__7; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Task_Priority_default; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1690_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_componentsRev(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Server_FileWorker_NamespaceEntry_finish___spec__2(lean_object*, size_t, size_t, lean_object*); static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__8(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1340_(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321_(uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23___boxed(lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Server_GoTo_0__Lean_Server_beqGoToKind____x40_Lean_Server_GoTo___hyg_9_(uint8_t, uint8_t); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543_(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand_go___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instToJsonAbsoluteLspSemanticToken; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__10; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__3; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleHover___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731_(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__11; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__1(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11; -static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5; uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__3; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_FileWorker_handleSemanticTokens_run___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18; static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics_waitLoop(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_computeDeltaLspSemanticTokens(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__2; extern lean_object* l_Std_Format_defWidth; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__1(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__4; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__12(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instBEqAbsoluteLspSemanticToken; extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_String_Range_toLspRange(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18; +uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894_(uint8_t, uint8_t); lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Server_ModuleRefs_toLspModuleRefs(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4; lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__11; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2; lean_object* l_IO_AsyncList_waitUntil___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__2(lean_object*); lean_object* l_Lean_Syntax_reprint(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestM_readDoc___at_Lean_Server_RequestM_withWaitFindSnapAtPos___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8889_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683_(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__1(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__13; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__3; LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__15; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__2(size_t, size_t, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__2___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__8; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__6; @@ -585,104 +566,120 @@ uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDefinition___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27; +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2; static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges_popRanges(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__9; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558_(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_noHighlightKinds; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15; static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1; lean_object* lean_string_length(lean_object*); lean_object* l_IO_AsyncList_waitAll___rarg___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__18(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256_(lean_object*); uint8_t l_Lean_Syntax_hasArgs(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRangeFromSyntax(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_takeTR_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_drop___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2; lean_object* l_Lean_Syntax_findStack_x3f(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Server_FileWorker_handleDocumentSymbol___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_erase___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__8(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8836_(lean_object*); static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__3; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__5; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3; LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Server_FileWorker_handleSemanticTokens_run___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550_(lean_object*); lean_object* l_id___rarg___boxed(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__5(lean_object*, lean_object*, size_t, size_t); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__35(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2(lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131_(lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleFoldingRange_isImport(lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_instDecidableRelLe___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDefinition___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22; uint64_t lean_uint64_xor(uint64_t, uint64_t); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__10; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15; LEAN_EXPORT lean_object* l_List_span_loop___at_Lean_Server_FileWorker_handleFoldingRange_addRanges___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1406_(lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893_(uint8_t); lean_object* lean_task_map(lean_object*, lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Completion_find_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23; lean_object* l_Lean_Expr_getAppFn(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33(size_t, size_t, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320_(uint8_t); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__5(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainTermGoal___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__10(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__7; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___lambda__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__5; -LEAN_EXPORT uint64_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__18(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__31(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_range_x3f(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg(lean_object*, lean_object*); @@ -691,182 +688,179 @@ static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___c LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Server_FileWorker_NamespaceEntry_finish___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Lsp_CompletionItem_getFileSource_x21(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645_(lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_handleWaitForDiagnostics___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__2; lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__2(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__1(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058_(lean_object*); lean_object* l_String_intercalate(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2; lean_object* lean_array_mk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__5; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18; static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__14; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__11___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleDocumentHighlight___spec__4(size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__16; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___closed__6; static lean_object* l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__2; static size_t l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__2; LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__4___boxed(lean_object*, lean_object*); static lean_object* l_Array_groupByKey___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__1___closed__3; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__4(lean_object*); static lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap___closed__3; lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_lctx(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_array_size(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_getInteractiveGoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Server_FileWorker_locationLinksOfInfo___spec__3(lean_object*, lean_object*, size_t, size_t); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_findInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal(lean_object*, lean_object*, lean_object*); -uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339_(uint8_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Server_Completion_resolveCompletionItem_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_AsyncList_getFinishedPrefix___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__35(lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__6(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__7; -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_groupByKey___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__3(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addCommandRange___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Server_registerLspRequestHandler___spec__6(lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokensFull___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__10(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Server_FileWorker_handlePlainGoal___closed__3; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__2(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__6; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_collectInfoBasedSemanticTokens___closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__5; -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__38(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handlePlainGoal___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__6___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_computeDeltaLspSemanticTokens___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t); lean_object* l_Lean_Elab_Info_stx(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2; -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRanges___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__9; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22; static lean_object* l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Info_toElabInfo_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleHover___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleFoldingRange_isImport___closed__4; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5; -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapM_loop___at_Lean_Server_FileWorker_getInteractiveGoals___spec__2___closed__1; static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleFoldingRange_addRange(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__8(lean_object*); +static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__31(lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5; LEAN_EXPORT uint8_t l_Array_contains___at_Lean_Server_FileWorker_collectSyntaxBasedSemanticTokens___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__25(lean_object*); static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__12; +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13; +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_keywordSemanticTokenMap___spec__2(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__11(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1; -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_handleDefinition___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals(lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2; -static lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17; +static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19; lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_FileWorker_findCompletionCmdDataAtPos___lambda__1(lean_object* x_1, lean_object* x_2) { _start: @@ -923,205 +917,107 @@ return x_4; static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 1; -x_2 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_unchecked("-", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1; -x_3 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -lean_ctor_set(x_3, 2, x_1); -lean_ctor_set(x_3, 3, x_1); -lean_ctor_set(x_3, 4, x_1); -lean_ctor_set(x_3, 5, x_1); -lean_ctor_set(x_3, 6, x_1); -lean_ctor_set(x_3, 7, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3; -x_2 = lean_array_mk(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 1; -x_2 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4; -x_3 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed), 3, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { if (lean_obj_tag(x_5) == 0) { -lean_object* x_8; lean_object* x_9; -lean_dec(x_6); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_8 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_7); -return x_9; +x_8 = lean_box(0); +x_9 = l___private_Lean_Server_Completion_CompletionItemData_0__Lean_Lsp_toJsonCompletionItemData____x40_Lean_Server_Completion_CompletionItemData___hyg_82_(x_1); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_12 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +lean_ctor_set(x_12, 2, x_8); +lean_ctor_set(x_12, 3, x_8); +lean_ctor_set(x_12, 4, x_8); +lean_ctor_set(x_12, 5, x_8); +lean_ctor_set(x_12, 6, x_10); +lean_ctor_set(x_12, 7, x_8); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_array_mk(x_14); +x_16 = 1; +x_17 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_7); +return x_18; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_5, 0); -lean_inc(x_10); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_5, 0); +lean_inc(x_19); lean_dec(x_5); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Server_Completion_find_x3f(x_1, x_2, x_3, x_11, x_12, x_4, x_7); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6; -x_17 = lean_box(0); -x_18 = lean_apply_3(x_16, x_17, x_6, x_15); -return x_18; -} -else +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Server_Completion_find_x3f(x_1, x_2, x_3, x_20, x_21, x_4, x_7); +if (lean_obj_tag(x_22) == 0) { -uint8_t x_19; -lean_dec(x_6); -x_19 = !lean_is_exclusive(x_13); -if (x_19 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_13, 0); -lean_dec(x_20); -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -lean_ctor_set(x_13, 0, x_21); -return x_13; +return x_22; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_dec(x_13); -x_23 = lean_ctor_get(x_14, 0); -lean_inc(x_23); -lean_dec(x_14); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_22); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } else { -uint8_t x_25; -lean_dec(x_6); -x_25 = !lean_is_exclusive(x_13); -if (x_25 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_13, 0); -x_27 = l_Lean_Server_RequestError_ofIoError(x_26); -lean_ctor_set(x_13, 0, x_27); -return x_13; +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_22, 0); +x_29 = l_Lean_Server_RequestError_ofIoError(x_28); +lean_ctor_set(x_22, 0, x_29); +return x_22; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_13, 0); -x_29 = lean_ctor_get(x_13, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_13); -x_30 = l_Lean_Server_RequestError_ofIoError(x_28); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -return x_31; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_22, 0); +x_31 = lean_ctor_get(x_22, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_22); +x_32 = l_Lean_Server_RequestError_ofIoError(x_30); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } } @@ -1155,7 +1051,7 @@ lean_inc(x_13); lean_dec(x_12); lean_inc(x_11); x_14 = l_Lean_Server_FileWorker_findCompletionCmdDataAtPos(x_5, x_11); -x_15 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion___lambda__2), 7, 4); +x_15 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed), 7, 4); lean_closure_set(x_15, 0, x_1); lean_closure_set(x_15, 1, x_9); lean_closure_set(x_15, 2, x_11); @@ -1164,89 +1060,89 @@ x_16 = l_Lean_Server_RequestM_mapTask___rarg(x_14, x_15, x_2, x_6); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletion___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_4; -x_4 = l_Lean_Server_FileWorker_handleCompletion___lambda__1(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; +lean_object* x_8; +x_8 = l_Lean_Server_FileWorker_handleCompletion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_6) == 0) { -lean_object* x_8; +lean_object* x_9; +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_7); -return x_8; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_8); +return x_9; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -lean_dec(x_5); -x_10 = lean_ctor_get(x_9, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_6, 0); lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_dec(x_6); +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Server_Completion_resolveCompletionItem_x3f(x_2, x_3, x_10, x_11, x_1, x_4, x_7); -if (lean_obj_tag(x_12) == 0) +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_2, 1); +x_14 = l_Lean_Server_Completion_resolveCompletionItem_x3f(x_3, x_4, x_11, x_12, x_1, x_5, x_13, x_8); +if (lean_obj_tag(x_14) == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -return x_12; +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 0); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_14); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } else { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_14); +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_12, 0); -x_19 = l_Lean_Server_RequestError_ofIoError(x_18); -lean_ctor_set(x_12, 0, x_19); -return x_12; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_14, 0); +x_21 = l_Lean_Server_RequestError_ofIoError(x_20); +lean_ctor_set(x_14, 0, x_21); +return x_14; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_12, 0); -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_12); -x_22 = l_Lean_Server_RequestError_ofIoError(x_20); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_14, 0); +x_23 = lean_ctor_get(x_14, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_14); +x_24 = l_Lean_Server_RequestError_ofIoError(x_22); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } } @@ -1310,7 +1206,7 @@ lean_dec(x_18); x_20 = lean_ctor_get(x_19, 2); lean_inc(x_20); lean_dec(x_19); -x_21 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_(x_17); +x_21 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_(x_17); if (lean_obj_tag(x_21) == 0) { uint8_t x_22; @@ -1348,7 +1244,7 @@ if (x_27 == 0) { lean_object* x_28; lean_object* x_29; x_28 = lean_ctor_get(x_21, 0); -x_29 = lean_ctor_get(x_28, 1); +x_29 = lean_ctor_get(x_28, 2); lean_inc(x_29); if (lean_obj_tag(x_29) == 0) { @@ -1372,18 +1268,18 @@ lean_inc(x_31); lean_dec(x_29); x_32 = lean_ctor_get(x_28, 0); lean_inc(x_32); -lean_dec(x_28); x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); x_34 = l_Lean_FileMap_lspPosToUtf8Pos(x_20, x_33); lean_inc(x_34); x_35 = l_Lean_Server_FileWorker_findCompletionCmdDataAtPos(x_15, x_34); -x_36 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 7, 4); +x_36 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 8, 5); lean_closure_set(x_36, 0, x_1); -lean_closure_set(x_36, 1, x_20); -lean_closure_set(x_36, 2, x_34); -lean_closure_set(x_36, 3, x_31); +lean_closure_set(x_36, 1, x_28); +lean_closure_set(x_36, 2, x_20); +lean_closure_set(x_36, 3, x_34); +lean_closure_set(x_36, 4, x_31); x_37 = l_Lean_Server_RequestM_mapTask___rarg(x_35, x_36, x_2, x_16); return x_37; } @@ -1394,7 +1290,7 @@ lean_object* x_38; lean_object* x_39; x_38 = lean_ctor_get(x_21, 0); lean_inc(x_38); lean_dec(x_21); -x_39 = lean_ctor_get(x_38, 1); +x_39 = lean_ctor_get(x_38, 2); lean_inc(x_39); if (lean_obj_tag(x_39) == 0) { @@ -1418,18 +1314,18 @@ lean_inc(x_42); lean_dec(x_39); x_43 = lean_ctor_get(x_38, 0); lean_inc(x_43); -lean_dec(x_38); x_44 = lean_ctor_get(x_43, 1); lean_inc(x_44); lean_dec(x_43); x_45 = l_Lean_FileMap_lspPosToUtf8Pos(x_20, x_44); lean_inc(x_45); x_46 = l_Lean_Server_FileWorker_findCompletionCmdDataAtPos(x_15, x_45); -x_47 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 7, 4); +x_47 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 8, 5); lean_closure_set(x_47, 0, x_1); -lean_closure_set(x_47, 1, x_20); -lean_closure_set(x_47, 2, x_45); -lean_closure_set(x_47, 3, x_42); +lean_closure_set(x_47, 1, x_38); +lean_closure_set(x_47, 2, x_20); +lean_closure_set(x_47, 3, x_45); +lean_closure_set(x_47, 4, x_42); x_48 = l_Lean_Server_RequestM_mapTask___rarg(x_46, x_47, x_2, x_16); return x_48; } @@ -1455,7 +1351,7 @@ lean_dec(x_52); x_54 = lean_ctor_get(x_53, 2); lean_inc(x_54); lean_dec(x_53); -x_55 = l___private_Lean_Server_Completion_0__Lean_Lsp_fromJsonCompletionItemDataWithId____x40_Lean_Server_Completion___hyg_366_(x_51); +x_55 = l___private_Lean_Server_Completion_CompletionResolution_0__Lean_Lsp_fromJsonResolvableCompletionItemData____x40_Lean_Server_Completion_CompletionResolution___hyg_261_(x_51); if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; @@ -1494,7 +1390,7 @@ if (lean_is_exclusive(x_55)) { lean_dec_ref(x_55); x_61 = lean_box(0); } -x_62 = lean_ctor_get(x_60, 1); +x_62 = lean_ctor_get(x_60, 2); lean_inc(x_62); if (lean_obj_tag(x_62) == 0) { @@ -1524,18 +1420,18 @@ lean_inc(x_66); lean_dec(x_62); x_67 = lean_ctor_get(x_60, 0); lean_inc(x_67); -lean_dec(x_60); x_68 = lean_ctor_get(x_67, 1); lean_inc(x_68); lean_dec(x_67); x_69 = l_Lean_FileMap_lspPosToUtf8Pos(x_54, x_68); lean_inc(x_69); x_70 = l_Lean_Server_FileWorker_findCompletionCmdDataAtPos(x_49, x_69); -x_71 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 7, 4); +x_71 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed), 8, 5); lean_closure_set(x_71, 0, x_1); -lean_closure_set(x_71, 1, x_54); -lean_closure_set(x_71, 2, x_69); -lean_closure_set(x_71, 3, x_66); +lean_closure_set(x_71, 1, x_60); +lean_closure_set(x_71, 2, x_54); +lean_closure_set(x_71, 3, x_69); +lean_closure_set(x_71, 4, x_66); x_72 = l_Lean_Server_RequestM_mapTask___rarg(x_70, x_71, x_2, x_50); return x_72; } @@ -1544,13 +1440,14 @@ return x_72; } } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; -x_8 = l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -return x_8; +lean_object* x_9; +x_9 = l_Lean_Server_FileWorker_handleCompletionItemResolve___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_2); +return x_9; } } LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -2994,6 +2891,15 @@ x_7 = l_Lean_Meta_isInstance(x_1, x_4, x_5, x_6); return x_7; } } +static lean_object* _init_l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -3026,7 +2932,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; x_13 = lean_ctor_get(x_9, 0); lean_dec(x_13); -x_14 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_14 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_ctor_set(x_9, 0, x_14); return x_9; } @@ -3036,7 +2942,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; x_15 = lean_ctor_get(x_9, 1); lean_inc(x_15); lean_dec(x_9); -x_16 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_16 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -3209,7 +3115,7 @@ return x_55; lean_object* x_56; lean_object* x_57; lean_dec(x_3); lean_dec(x_2); -x_56 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_56 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_57 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_57, 0, x_56); lean_ctor_set(x_57, 1, x_5); @@ -3316,7 +3222,7 @@ x_17 = 1; return x_17; } } -case 3: +case 4: { uint8_t x_18; lean_dec(x_12); @@ -3453,7 +3359,7 @@ x_17 = 1; return x_17; } } -case 3: +case 4: { uint8_t x_18; lean_dec(x_12); @@ -3726,7 +3632,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__ _start: { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_7 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -3994,7 +3900,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__ _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_4 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -7041,7 +6947,7 @@ return x_57; } } } -case 2: +case 3: { lean_object* x_58; lean_object* x_59; lean_dec(x_3); @@ -7131,7 +7037,7 @@ x_77 = l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__5(x_12, x_1, x_13, return x_77; } } -case 4: +case 5: { lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_dec(x_3); @@ -7346,7 +7252,7 @@ return x_127; } } } -case 5: +case 6: { lean_object* x_128; uint8_t x_129; uint8_t x_130; lean_dec(x_3); @@ -8224,7 +8130,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_handleDefinition___closed__1( _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_1 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Server_FileWorker_handleHover___spec__1___rarg___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -8256,7 +8162,7 @@ lean_dec(x_10); lean_inc(x_12); x_13 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDefinition___lambda__1___boxed), 2, 1); lean_closure_set(x_13, 0, x_12); -x_14 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_14 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_15 = lean_box(x_1); x_16 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDefinition___lambda__2___boxed), 6, 3); lean_closure_set(x_16, 0, x_12); @@ -9902,7 +9808,7 @@ if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_ctor_get(x_10, 0); -x_13 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_13 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_14 = l_List_foldl___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3(x_13, x_12); lean_dec(x_12); x_15 = lean_alloc_ctor(1, 1, 0); @@ -9918,7 +9824,7 @@ x_17 = lean_ctor_get(x_10, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_10); -x_18 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_18 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_19 = l_List_foldl___at_Lean_Server_FileWorker_getInteractiveGoals___spec__3(x_18, x_16); lean_dec(x_16); x_20 = lean_alloc_ctor(1, 1, 0); @@ -10141,7 +10047,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_handlePlainGoal___lambda__1__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Server_FileWorker_handlePlainGoal___lambda__1___closed__2; -x_2 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_2 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12801,7 +12707,7 @@ lean_inc(x_11); x_12 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDefinition___lambda__1___boxed), 2, 1); lean_closure_set(x_12, 0, x_11); x_13 = lean_box(0); -x_14 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_14 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_15 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentHighlight___lambda__3___boxed), 9, 6); lean_closure_set(x_15, 0, x_14); lean_closure_set(x_15, 1, x_9); @@ -16019,7 +15925,7 @@ lean_ctor_set(x_32, 2, x_7); lean_ctor_set(x_32, 3, x_3); lean_ctor_set(x_2, 1, x_4); lean_ctor_set(x_2, 0, x_32); -x_33 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_33 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; { lean_object* _tmp_1 = x_8; lean_object* _tmp_2 = x_33; @@ -16045,7 +15951,7 @@ lean_ctor_set(x_38, 2, x_35); lean_ctor_set(x_38, 3, x_3); lean_ctor_set(x_2, 1, x_4); lean_ctor_set(x_2, 0, x_38); -x_39 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_39 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; { lean_object* _tmp_1 = x_8; lean_object* _tmp_2 = x_39; @@ -16072,7 +15978,7 @@ lean_ctor_set(x_45, 2, x_42); lean_ctor_set(x_45, 3, x_3); lean_ctor_set(x_2, 1, x_4); lean_ctor_set(x_2, 0, x_45); -x_46 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_46 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; { lean_object* _tmp_1 = x_8; lean_object* _tmp_2 = x_46; @@ -16175,7 +16081,7 @@ lean_ctor_set(x_73, 3, x_3); x_74 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_74, 0, x_73); lean_ctor_set(x_74, 1, x_4); -x_75 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_75 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_2 = x_49; x_3 = x_75; x_4 = x_74; @@ -16197,7 +16103,7 @@ lean_ctor_set(x_80, 3, x_3); x_81 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_4); -x_82 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_82 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_2 = x_49; x_3 = x_82; x_4 = x_81; @@ -16220,7 +16126,7 @@ lean_ctor_set(x_88, 3, x_3); x_89 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_4); -x_90 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_90 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_2 = x_49; x_3 = x_90; x_4 = x_89; @@ -16306,7 +16212,7 @@ x_28 = lean_ctor_get(x_10, 0); lean_dec(x_28); x_29 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_29, 0, x_2); -x_30 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_30 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_inc(x_4); lean_inc(x_12); x_31 = l_List_takeTR_go___rarg(x_12, x_12, x_4, x_30); @@ -16334,7 +16240,7 @@ lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean lean_dec(x_10); x_36 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_36, 0, x_2); -x_37 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_37 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_inc(x_4); lean_inc(x_12); x_38 = l_List_takeTR_go___rarg(x_12, x_12, x_4, x_37); @@ -16440,7 +16346,7 @@ if (lean_is_exclusive(x_47)) { } x_62 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_62, 0, x_2); -x_63 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_63 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_inc(x_4); lean_inc(x_49); x_64 = l_List_takeTR_go___rarg(x_49, x_49, x_4, x_63); @@ -16623,7 +16529,7 @@ lean_dec(x_1); x_11 = lean_ctor_get(x_10, 2); lean_inc(x_11); lean_dec(x_10); -x_12 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_12 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_13 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols(x_11, x_9, x_12, x_8); lean_ctor_set(x_2, 1, x_4); lean_ctor_set(x_2, 0, x_13); @@ -16643,7 +16549,7 @@ lean_dec(x_1); x_18 = lean_ctor_get(x_17, 2); lean_inc(x_18); lean_dec(x_17); -x_19 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_19 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_20 = l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols(x_18, x_16, x_19, x_15); x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -19392,7 +19298,7 @@ x_5 = l_Lean_RBNode_insert___at_Lean_Server_FileWorker_keywordSemanticTokenMap__ return x_5; } } -LEAN_EXPORT uint8_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; @@ -19422,17 +19328,17 @@ return x_12; else { uint8_t x_13; -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8321_(x_5, x_8); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8894_(x_5, x_8); return x_13; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_1, x_2); +x_3 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -19443,7 +19349,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_instBEqAbsoluteLspSemanticTok _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551____boxed), 2, 0); return x_1; } } @@ -19455,7 +19361,7 @@ x_1 = l_Lean_Server_FileWorker_instBEqAbsoluteLspSemanticToken___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; @@ -19467,16 +19373,16 @@ x_6 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_ x_7 = lean_uint64_mix_hash(x_5, x_6); x_8 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashPosition____x40_Lean_Data_Lsp_Basic___hyg_365_(x_3); x_9 = lean_uint64_mix_hash(x_7, x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8339_(x_4); +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8912_(x_4); x_11 = lean_uint64_mix_hash(x_9, x_10); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(x_1); +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -19486,7 +19392,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_instHashableAbsoluteLspSemant _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641____boxed), 1, 0); return x_1; } } @@ -19498,16 +19404,16 @@ x_1 = l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7485_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8058_(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1() { _start: { lean_object* x_1; @@ -19515,7 +19421,7 @@ x_1 = lean_mk_string_unchecked("pos", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2() { _start: { lean_object* x_1; @@ -19523,7 +19429,7 @@ x_1 = lean_mk_string_unchecked("Server", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3() { _start: { lean_object* x_1; @@ -19531,7 +19437,7 @@ x_1 = lean_mk_string_unchecked("FileWorker", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4() { _start: { lean_object* x_1; @@ -19539,30 +19445,30 @@ x_1 = lean_mk_string_unchecked("AbsoluteLspSemanticToken", 24, 24); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__6; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2; -x_3 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3; -x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2; +x_3 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3; +x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5; x_2 = 1; x_3 = l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7() { _start: { lean_object* x_1; @@ -19570,48 +19476,48 @@ x_1 = lean_mk_string_unchecked(".", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9; x_2 = 1; x_3 = l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12() { _start: { lean_object* x_1; @@ -19619,17 +19525,17 @@ x_1 = lean_mk_string_unchecked(": ", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14() { _start: { lean_object* x_1; @@ -19637,48 +19543,48 @@ x_1 = lean_mk_string_unchecked("tailPos", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15; x_2 = 1; x_3 = l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19688,42 +19594,42 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19; x_2 = 1; x_3 = l_Lean_Server_FileWorker_NamespaceEntry_finish___closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22() { +static lean_object* _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12; +x_1 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1; +x_2 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -19735,7 +19641,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13; +x_6 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -19747,7 +19653,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13; +x_9 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -19761,7 +19667,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14; +x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14; lean_inc(x_1); x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) @@ -19774,7 +19680,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18; +x_17 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -19786,7 +19692,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18; +x_20 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -19801,7 +19707,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l_Lean_Server_FileWorker_noHighlightKinds___closed__3; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -19812,7 +19718,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22; +x_28 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -19824,7 +19730,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22; +x_31 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -19870,11 +19776,11 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -19883,7 +19789,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemant _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691_), 1, 0); return x_1; } } @@ -19895,14 +19801,14 @@ x_1 = l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8889_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8836_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_2); -x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1; +x_4 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -19913,7 +19819,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(x_8); -x_10 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14; +x_10 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -19922,7 +19828,7 @@ lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_6); x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); lean_dec(x_1); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7320_(x_13); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7893_(x_13); x_15 = l_Lean_Server_FileWorker_noHighlightKinds___closed__3; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -19939,7 +19845,7 @@ lean_ctor_set(x_19, 1, x_18); x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_7); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_21 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_22 = l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(x_20, x_21); x_23 = l_Lean_Json_mkObj(x_22); return x_23; @@ -19949,7 +19855,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_instToJsonAbsoluteLspSemantic _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8889_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_toJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8836_), 1, 0); return x_1; } } @@ -20083,7 +19989,7 @@ if (x_7 == 0) { lean_object* x_8; lean_dec(x_1); -x_8 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_8 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_8; } else @@ -20096,7 +20002,7 @@ if (x_10 == 0) { lean_object* x_11; lean_dec(x_1); -x_11 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_11 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_11; } else @@ -20104,7 +20010,7 @@ else size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; x_12 = lean_usize_of_nat(x_5); x_13 = lean_usize_of_nat(x_6); -x_14 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_14 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_computeAbsoluteLspSemanticTokens___spec__2(x_1, x_2, x_3, x_4, x_12, x_13, x_14); return x_15; } @@ -20176,7 +20082,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_ctor_get(x_2, 2); -x_7 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_4, x_1); +x_7 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_4, x_1); if (x_7 == 0) { x_2 = x_6; @@ -20207,7 +20113,7 @@ else lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_4, x_1); +x_6 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_4, x_1); if (x_6 == 0) { x_2 = x_5; @@ -20239,7 +20145,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8 x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); x_6 = lean_array_get_size(x_1); -x_7 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(x_4); +x_7 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(x_4); x_8 = 32; x_9 = lean_uint64_shift_right(x_7, x_8); x_10 = lean_uint64_xor(x_7, x_9); @@ -20270,7 +20176,7 @@ lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); x_25 = lean_array_get_size(x_1); -x_26 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(x_22); +x_26 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(x_22); x_27 = 32; x_28 = lean_uint64_shift_right(x_26, x_27); x_29 = lean_uint64_xor(x_26, x_28); @@ -20362,7 +20268,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); -x_9 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_6, x_1); +x_9 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_6, x_1); if (x_9 == 0) { lean_object* x_10; @@ -20389,7 +20295,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_3); -x_14 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_11, x_1); +x_14 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -20434,7 +20340,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); x_7 = lean_ctor_get(x_2, 2); -x_8 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_5, x_1); +x_8 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_5, x_1); if (x_8 == 0) { lean_object* x_9; @@ -20460,7 +20366,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_dec(x_2); -x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8604_(x_10, x_1); +x_13 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_beqAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8551_(x_10, x_1); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -20506,7 +20412,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_13 = lean_ctor_get(x_8, 0); x_14 = lean_ctor_get(x_8, 1); x_15 = lean_array_get_size(x_14); -x_16 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(x_11); +x_16 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(x_11); x_17 = 32; x_18 = lean_uint64_shift_right(x_16, x_17); x_19 = lean_uint64_xor(x_16, x_18); @@ -20524,7 +20430,7 @@ x_29 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_FileWorker_fil if (lean_obj_tag(x_29) == 0) { lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_30 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_31 = lean_array_push(x_30, x_10); x_32 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__3(x_11, x_28); if (x_32 == 0) @@ -20768,7 +20674,7 @@ lean_inc(x_124); lean_inc(x_123); lean_dec(x_8); x_125 = lean_array_get_size(x_124); -x_126 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8694_(x_11); +x_126 = l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_hashAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8641_(x_11); x_127 = 32; x_128 = lean_uint64_shift_right(x_126, x_127); x_129 = lean_uint64_xor(x_126, x_128); @@ -20786,7 +20692,7 @@ x_139 = l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_FileWorker_fi if (lean_obj_tag(x_139) == 0) { lean_object* x_140; lean_object* x_141; uint8_t x_142; -x_140 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_140 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_141 = lean_array_push(x_140, x_10); x_142 = l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__3(x_11, x_138); if (x_142 == 0) @@ -21114,7 +21020,7 @@ static size_t _init_l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___clo _start: { lean_object* x_1; size_t x_2; -x_1 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_1 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_2 = lean_array_size(x_1); return x_2; } @@ -21138,7 +21044,7 @@ size_t x_9; lean_object* x_10; lean_object* x_11; lean_dec(x_5); lean_dec(x_4); x_9 = l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__2; -x_10 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_10 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_11 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__11(x_9, x_8, x_10); return x_11; } @@ -21152,7 +21058,7 @@ size_t x_13; lean_object* x_14; lean_object* x_15; lean_dec(x_5); lean_dec(x_4); x_13 = l_Lean_Server_FileWorker_filterDuplicateSemanticTokens___closed__2; -x_14 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_14 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_15 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__11(x_13, x_8, x_14); return x_15; } @@ -21161,7 +21067,7 @@ else size_t x_16; lean_object* x_17; lean_object* x_18; size_t x_19; lean_object* x_20; x_16 = lean_usize_of_nat(x_5); lean_dec(x_5); -x_17 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_17 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_filterDuplicateSemanticTokens___spec__12(x_4, x_8, x_16, x_17); lean_dec(x_4); x_19 = lean_array_size(x_18); @@ -22587,7 +22493,7 @@ else { lean_object* x_11; lean_dec(x_1); -x_11 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_11 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_11; } } @@ -22622,7 +22528,7 @@ else lean_object* x_23; lean_dec(x_13); lean_dec(x_1); -x_23 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_23 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_23; } } @@ -22653,7 +22559,7 @@ else lean_object* x_32; lean_dec(x_13); lean_dec(x_1); -x_32 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_32 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_32; } } @@ -22703,7 +22609,7 @@ else lean_object* x_48; lean_dec(x_38); lean_dec(x_1); -x_48 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_48 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; return x_48; } } @@ -23156,7 +23062,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleSemanticTokens_run(lean_ { lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; x_7 = lean_box(0); -x_8 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_8 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; lean_inc(x_4); x_9 = l_List_forIn_x27_loop___at_Lean_Server_FileWorker_handleSemanticTokens_run___spec__1(x_1, x_4, x_7, x_4, x_4, x_8, lean_box(0), x_5, x_6); lean_dec(x_4); @@ -24906,7 +24812,7 @@ lean_dec(x_1); x_9 = lean_ctor_get(x_8, 2); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1; +x_10 = l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1; x_11 = l_Lean_Server_FileWorker_handleFoldingRange_addRanges(x_9, x_6, x_7, x_10, x_3, x_4); x_12 = !lean_is_exclusive(x_11); if (x_12 == 0) @@ -25252,7 +25158,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1() { _start: { lean_object* x_1; @@ -25260,7 +25166,7 @@ x_1 = lean_mk_string_unchecked("Cannot parse request params: ", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2() { _start: { lean_object* x_1; @@ -25268,7 +25174,7 @@ x_1 = lean_mk_string_unchecked("\n", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2(lean_object* x_1) { _start: { lean_object* x_2; @@ -25283,10 +25189,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -25306,10 +25212,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -25346,7 +25252,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -25371,11 +25277,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -25425,7 +25331,7 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -25434,38 +25340,38 @@ x_2 = l_Lean_Json_mkObj(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1; +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1; return x_2; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -25484,7 +25390,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -25499,7 +25405,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -25558,7 +25464,7 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -25566,21 +25472,21 @@ x_1 = l_Lean_Server_requestHandlers; return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -25588,7 +25494,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -25620,7 +25526,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -25649,7 +25555,7 @@ return x_27; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -25657,7 +25563,7 @@ x_1 = lean_mk_string_unchecked("Failed to register LSP request handler for '", 4 return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2() { _start: { lean_object* x_1; @@ -25665,11 +25571,11 @@ x_1 = lean_mk_string_unchecked("': already registered", 21, 21); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -25683,17 +25589,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -25715,17 +25621,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -25737,7 +25643,7 @@ return x_28; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1() { _start: { lean_object* x_1; @@ -25745,7 +25651,7 @@ x_1 = lean_mk_string_unchecked("': only possible during initialization", 38, 38) return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -25764,10 +25670,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -25781,10 +25687,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -25801,17 +25707,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__5(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2370_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2943_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -25821,10 +25727,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -25844,10 +25750,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -25884,7 +25790,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -25909,11 +25815,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__5(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__5(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -25963,30 +25869,30 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2306_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2879_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__5(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__5(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -26005,7 +25911,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -26020,7 +25926,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -26079,21 +25985,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -26101,7 +26007,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -26133,7 +26039,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -26162,11 +26068,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -26180,17 +26086,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -26212,17 +26118,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -26234,7 +26140,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -26253,10 +26159,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -26270,10 +26176,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -26290,17 +26196,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__8(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__8(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1756_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2064_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26310,10 +26216,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -26333,10 +26239,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -26373,7 +26279,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26398,11 +26304,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__8(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__8(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26448,30 +26354,30 @@ return x_11; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2096_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2404_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__8(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__8(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -26490,7 +26396,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -26505,7 +26411,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -26564,21 +26470,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -26586,7 +26492,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -26618,7 +26524,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -26647,11 +26553,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -26665,17 +26571,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -26697,17 +26603,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -26719,7 +26625,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -26738,10 +26644,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -26755,10 +26661,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -26775,17 +26681,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__11(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__11(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2710_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3283_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26795,10 +26701,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -26818,10 +26724,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -26858,7 +26764,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26883,11 +26789,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__11(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__11(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -26937,7 +26843,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -26952,35 +26858,35 @@ lean_object* x_3; lean_object* x_4; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); lean_dec(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2550_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3123_(x_3); return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__11(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__11(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -26999,7 +26905,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -27014,7 +26920,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -27073,21 +26979,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -27095,7 +27001,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -27127,7 +27033,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -27156,11 +27062,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -27174,17 +27080,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -27206,17 +27112,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -27228,7 +27134,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -27247,10 +27153,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -27264,10 +27170,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -27284,12 +27190,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__14(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__14(lean_object* x_1) { _start: { lean_object* x_2; @@ -27304,10 +27210,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -27327,10 +27233,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -27367,7 +27273,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -27392,7 +27298,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -27417,11 +27323,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__14(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__14(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -27471,42 +27377,42 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__2(lean_object* x_1) { _start: { size_t x_2; size_t x_3; lean_object* x_4; lean_object* x_5; x_2 = lean_array_size(x_1); x_3 = 0; -x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16(x_2, x_3, x_1); +x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16(x_2, x_3, x_1); x_5 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_5, 0, x_4); return x_5; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__14(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__14(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -27525,7 +27431,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -27540,7 +27446,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -27599,21 +27505,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -27621,7 +27527,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -27653,7 +27559,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -27682,11 +27588,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -27700,17 +27606,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -27732,17 +27638,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -27754,7 +27660,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -27773,10 +27679,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -27790,10 +27696,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -27810,17 +27716,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__18(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__18(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3855_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4428_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -27830,10 +27736,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -27853,10 +27759,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -27893,7 +27799,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -27918,7 +27824,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -27933,7 +27839,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4072_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4645_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -27943,11 +27849,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__18(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__18(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -27997,42 +27903,42 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__2(lean_object* x_1) { _start: { size_t x_2; size_t x_3; lean_object* x_4; lean_object* x_5; x_2 = lean_array_size(x_1); x_3 = 0; -x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20(x_2, x_3, x_1); +x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20(x_2, x_3, x_1); x_5 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_5, 0, x_4); return x_5; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__18(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__18(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -28051,7 +27957,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -28066,7 +27972,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -28125,21 +28031,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -28147,7 +28053,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -28179,7 +28085,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -28208,11 +28114,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -28226,17 +28132,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -28258,17 +28164,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -28280,7 +28186,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -28299,10 +28205,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -28316,10 +28222,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -28336,17 +28242,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__22(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__22(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4131_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4704_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -28356,10 +28262,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -28379,10 +28285,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -28419,7 +28325,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -28444,11 +28350,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__22(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__22(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -28489,7 +28395,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__2(lean_object* x_1) { _start: { size_t x_2; size_t x_3; lean_object* x_4; lean_object* x_5; @@ -28501,30 +28407,30 @@ lean_ctor_set(x_5, 0, x_4); return x_5; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__22(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__22(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -28543,7 +28449,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -28558,7 +28464,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -28617,21 +28523,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -28639,7 +28545,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -28671,7 +28577,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -28700,11 +28606,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -28718,17 +28624,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -28750,17 +28656,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -28772,7 +28678,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -28791,10 +28697,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -28808,10 +28714,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -28828,17 +28734,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__25(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__25(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9558_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10131_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -28848,10 +28754,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -28871,10 +28777,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -28911,7 +28817,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -28936,11 +28842,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__25(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__25(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -28981,30 +28887,30 @@ return x_8; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9970_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10543_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__25(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__25(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -29023,7 +28929,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -29038,7 +28944,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -29097,21 +29003,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -29119,7 +29025,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -29151,7 +29057,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -29180,11 +29086,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -29198,17 +29104,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -29230,17 +29136,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -29252,7 +29158,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -29271,10 +29177,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -29288,10 +29194,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -29308,17 +29214,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__28(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__28(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9683_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10256_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -29328,10 +29234,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -29351,10 +29257,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -29391,7 +29297,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -29416,11 +29322,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__28(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__28(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -29470,12 +29376,12 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__28(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__28(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -29494,7 +29400,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -29509,7 +29415,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -29568,21 +29474,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__2), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -29590,7 +29496,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -29622,7 +29528,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -29651,11 +29557,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -29669,17 +29575,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -29701,17 +29607,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -29723,7 +29629,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -29742,10 +29648,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -29759,10 +29665,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -29779,17 +29685,17 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__31(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__31(lean_object* x_1) { _start: { lean_object* x_2; lean_inc(x_1); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10024_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10597_(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -29799,10 +29705,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -29822,10 +29728,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -29862,7 +29768,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -29887,7 +29793,7 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -29902,7 +29808,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10196_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10769_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -29912,11 +29818,11 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__31(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__31(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -29957,42 +29863,42 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__2(lean_object* x_1) { _start: { size_t x_2; size_t x_3; lean_object* x_4; lean_object* x_5; x_2 = lean_array_size(x_1); x_3 = 0; -x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33(x_2, x_3, x_1); +x_4 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33(x_2, x_3, x_1); x_5 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_5, 0, x_4); return x_5; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__31(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__31(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -30011,7 +29917,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -30026,7 +29932,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -30085,21 +29991,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -30107,7 +30013,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -30139,7 +30045,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -30168,11 +30074,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -30186,17 +30092,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -30218,17 +30124,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -30240,7 +30146,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -30259,10 +30165,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -30276,10 +30182,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -30296,12 +30202,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__35(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__35(lean_object* x_1) { _start: { lean_object* x_2; @@ -30316,10 +30222,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -30339,10 +30245,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -30379,7 +30285,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -30404,11 +30310,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__35(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__35(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -30458,7 +30364,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -30478,30 +30384,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__35(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__35(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -30520,7 +30426,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -30535,7 +30441,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -30594,21 +30500,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -30616,7 +30522,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -30648,7 +30554,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -30677,11 +30583,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -30695,17 +30601,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -30727,17 +30633,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -30749,7 +30655,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -30768,10 +30674,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -30785,10 +30691,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -30805,12 +30711,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__38(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__38(lean_object* x_1) { _start: { lean_object* x_2; @@ -30825,10 +30731,10 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; x_4 = lean_ctor_get(x_2, 0); x_5 = l_Lean_Json_compress(x_1); -x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_6 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); -x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_8 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_9 = lean_string_append(x_7, x_8); x_10 = lean_string_append(x_9, x_4); lean_dec(x_4); @@ -30848,10 +30754,10 @@ x_15 = lean_ctor_get(x_2, 0); lean_inc(x_15); lean_dec(x_2); x_16 = l_Lean_Json_compress(x_1); -x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1; +x_17 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); -x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2; +x_19 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2; x_20 = lean_string_append(x_18, x_19); x_21 = lean_string_append(x_20, x_15); lean_dec(x_15); @@ -30888,7 +30794,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -30913,11 +30819,11 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__38(x_1); +x_2 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__38(x_1); if (lean_obj_tag(x_2) == 0) { uint8_t x_3; @@ -30967,7 +30873,7 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__2(lean_object* x_1) { _start: { if (lean_obj_tag(x_1) == 0) @@ -30987,30 +30893,30 @@ return x_4; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1; +x_1 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1; x_2 = lean_alloc_closure((void*)(l_Except_map___rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__38(x_2); -x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39(x_5, x_3, x_4); +x_5 = l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__38(x_2); +x_6 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39(x_5, x_3, x_4); lean_dec(x_5); if (lean_obj_tag(x_6) == 0) { @@ -31029,7 +30935,7 @@ if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2; x_13 = l_Task_Priority_default; x_14 = 0; x_15 = lean_task_map(x_12, x_11, x_13, x_14); @@ -31044,7 +30950,7 @@ x_17 = lean_ctor_get(x_9, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_9); -x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2; +x_18 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2; x_19 = l_Task_Priority_default; x_20 = 0; x_21 = lean_task_map(x_18, x_16, x_19, x_20); @@ -31103,21 +31009,21 @@ return x_30; } } } -static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3), 4, 1); +x_5 = lean_alloc_closure((void*)(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3), 4, 1); lean_closure_set(x_5, 0, x_1); -x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_6 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) @@ -31125,7 +31031,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1; lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 0, x_11); x_12 = l_Lean_PersistentHashMap_insert___at_Lean_Server_registerLspRequestHandler___spec__2(x_9, x_2, x_7); @@ -31157,7 +31063,7 @@ x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_7); -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_5); @@ -31186,11 +31092,11 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1; +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1; x_6 = lean_st_ref_get(x_5, x_4); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -31204,17 +31110,17 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_free_object(x_6); x_11 = lean_box(0); -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4(x_1, x_2, x_11, x_9); +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4(x_1, x_2, x_11, x_9); return x_12; } else { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_1); -x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_13 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_14 = lean_string_append(x_13, x_2); lean_dec(x_2); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_16 = lean_string_append(x_14, x_15); x_17 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_17, 0, x_16); @@ -31236,17 +31142,17 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_box(0); -x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4(x_1, x_2, x_21, x_19); +x_22 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4(x_1, x_2, x_21, x_19); return x_22; } else { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_dec(x_1); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_24 = lean_string_append(x_23, x_2); lean_dec(x_2); -x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2; +x_25 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -31258,7 +31164,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; uint8_t x_6; @@ -31277,10 +31183,10 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_8 = lean_ctor_get(x_4, 0); lean_dec(x_8); -x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_9 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_10 = lean_string_append(x_9, x_1); lean_dec(x_1); -x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_11 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_12 = lean_string_append(x_10, x_11); x_13 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_13, 0, x_12); @@ -31294,10 +31200,10 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); lean_dec(x_4); -x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1; +x_15 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1; x_16 = lean_string_append(x_15, x_1); lean_dec(x_1); -x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1; +x_17 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -31314,12 +31220,12 @@ x_21 = lean_ctor_get(x_4, 1); lean_inc(x_21); lean_dec(x_4); x_22 = lean_box(0); -x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5(x_2, x_1, x_22, x_21); +x_23 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5(x_2, x_1, x_22, x_21); return x_23; } } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1() { _start: { lean_object* x_1; @@ -31327,7 +31233,7 @@ x_1 = lean_mk_string_unchecked("textDocument/waitForDiagnostics", 31, 31); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2() { _start: { lean_object* x_1; @@ -31335,7 +31241,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleWaitForDiagnosti return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3() { _start: { lean_object* x_1; @@ -31343,7 +31249,7 @@ x_1 = lean_mk_string_unchecked("textDocument/completion", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4() { _start: { lean_object* x_1; @@ -31351,7 +31257,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletion), 3, return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5() { _start: { lean_object* x_1; @@ -31359,7 +31265,7 @@ x_1 = lean_mk_string_unchecked("completionItem/resolve", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6() { _start: { lean_object* x_1; @@ -31367,7 +31273,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleCompletionItemRe return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7() { _start: { lean_object* x_1; @@ -31375,7 +31281,7 @@ x_1 = lean_mk_string_unchecked("textDocument/hover", 18, 18); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8() { _start: { lean_object* x_1; @@ -31383,7 +31289,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleHover), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -31394,7 +31300,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10() { _start: { lean_object* x_1; @@ -31402,7 +31308,7 @@ x_1 = lean_mk_string_unchecked("textDocument/declaration", 24, 24); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -31413,7 +31319,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12() { _start: { lean_object* x_1; @@ -31421,7 +31327,7 @@ x_1 = lean_mk_string_unchecked("textDocument/definition", 23, 23); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -31432,7 +31338,7 @@ lean_closure_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14() { _start: { lean_object* x_1; @@ -31440,7 +31346,7 @@ x_1 = lean_mk_string_unchecked("textDocument/typeDefinition", 27, 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15() { _start: { lean_object* x_1; @@ -31448,7 +31354,7 @@ x_1 = lean_mk_string_unchecked("textDocument/documentHighlight", 30, 30); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16() { _start: { lean_object* x_1; @@ -31456,7 +31362,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentHighligh return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17() { _start: { lean_object* x_1; @@ -31464,7 +31370,7 @@ x_1 = lean_mk_string_unchecked("textDocument/documentSymbol", 27, 27); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18() { _start: { lean_object* x_1; @@ -31472,7 +31378,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleDocumentSymbol__ return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19() { _start: { lean_object* x_1; @@ -31480,7 +31386,7 @@ x_1 = lean_mk_string_unchecked("textDocument/semanticTokens/full", 32, 32); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20() { _start: { lean_object* x_1; @@ -31488,7 +31394,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensFu return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21() { _start: { lean_object* x_1; @@ -31496,7 +31402,7 @@ x_1 = lean_mk_string_unchecked("textDocument/semanticTokens/range", 33, 33); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22() { _start: { lean_object* x_1; @@ -31504,7 +31410,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleSemanticTokensRa return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23() { _start: { lean_object* x_1; @@ -31512,7 +31418,7 @@ x_1 = lean_mk_string_unchecked("textDocument/foldingRange", 25, 25); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24() { _start: { lean_object* x_1; @@ -31520,7 +31426,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handleFoldingRange___b return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25() { _start: { lean_object* x_1; @@ -31528,7 +31434,7 @@ x_1 = lean_mk_string_unchecked("$/lean/plainGoal", 16, 16); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26() { _start: { lean_object* x_1; @@ -31536,7 +31442,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainGoal), 3, 0 return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27() { _start: { lean_object* x_1; @@ -31544,7 +31450,7 @@ x_1 = lean_mk_string_unchecked("$/lean/plainTermGoal", 20, 20); return x_1; } } -static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28() { +static lean_object* _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28() { _start: { lean_object* x_1; @@ -31552,130 +31458,130 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_handlePlainTermGoal), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1; -x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2; -x_4 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1; +x_3 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2; +x_4 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_5 = lean_ctor_get(x_4, 1); lean_inc(x_5); lean_dec(x_4); -x_6 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3; -x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4; -x_8 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4(x_6, x_7, x_5); +x_6 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3; +x_7 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4; +x_8 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4(x_6, x_7, x_5); if (lean_obj_tag(x_8) == 0) { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5; -x_11 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6; -x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7(x_10, x_11, x_9); +x_10 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5; +x_11 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6; +x_12 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7(x_10, x_11, x_9); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7; -x_15 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8; -x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10(x_14, x_15, x_13); +x_14 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7; +x_15 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8; +x_16 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10(x_14, x_15, x_13); if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10; -x_19 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9; -x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13(x_18, x_19, x_17); +x_18 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10; +x_19 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9; +x_20 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13(x_18, x_19, x_17); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12; -x_23 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11; -x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13(x_22, x_23, x_21); +x_22 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12; +x_23 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11; +x_24 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13(x_22, x_23, x_21); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_25 = lean_ctor_get(x_24, 1); lean_inc(x_25); lean_dec(x_24); -x_26 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14; -x_27 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13; -x_28 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13(x_26, x_27, x_25); +x_26 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14; +x_27 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13; +x_28 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13(x_26, x_27, x_25); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; x_29 = lean_ctor_get(x_28, 1); lean_inc(x_29); lean_dec(x_28); -x_30 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15; -x_31 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16; -x_32 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17(x_30, x_31, x_29); +x_30 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15; +x_31 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16; +x_32 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17(x_30, x_31, x_29); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); lean_dec(x_32); -x_34 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17; -x_35 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18; -x_36 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21(x_34, x_35, x_33); +x_34 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17; +x_35 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18; +x_36 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21(x_34, x_35, x_33); if (lean_obj_tag(x_36) == 0) { lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); -x_38 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19; -x_39 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20; -x_40 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24(x_38, x_39, x_37); +x_38 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19; +x_39 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20; +x_40 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24(x_38, x_39, x_37); if (lean_obj_tag(x_40) == 0) { lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_40, 1); lean_inc(x_41); lean_dec(x_40); -x_42 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21; -x_43 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22; -x_44 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27(x_42, x_43, x_41); +x_42 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21; +x_43 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22; +x_44 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27(x_42, x_43, x_41); if (lean_obj_tag(x_44) == 0) { lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; x_45 = lean_ctor_get(x_44, 1); lean_inc(x_45); lean_dec(x_44); -x_46 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23; -x_47 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24; -x_48 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30(x_46, x_47, x_45); +x_46 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23; +x_47 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24; +x_48 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30(x_46, x_47, x_45); if (lean_obj_tag(x_48) == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25; -x_51 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26; -x_52 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34(x_50, x_51, x_49); +x_50 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25; +x_51 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26; +x_52 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34(x_50, x_51, x_49); if (lean_obj_tag(x_52) == 0) { lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; x_53 = lean_ctor_get(x_52, 1); lean_inc(x_53); lean_dec(x_52); -x_54 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27; -x_55 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28; -x_56 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37(x_54, x_55, x_53); +x_54 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27; +x_55 = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28; +x_56 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37(x_54, x_55, x_53); return x_56; } else @@ -31978,138 +31884,138 @@ return x_108; } } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__3(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__3(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2(x_1); +x_2 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__6(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__6(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__9(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__9(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__12(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__12(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__15(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__15(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -32117,39 +32023,39 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__16(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__16(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__19(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__19(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -32157,123 +32063,123 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__20(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__20(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__23(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__23(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__26(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__26(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__29(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__29(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__32(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__32(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -32281,80 +32187,80 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__33(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__33(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__36(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__36(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__39(x_1, x_2, x_3); +x_4 = l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__39(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__5(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -32410,20 +32316,6 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__1); -l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__1___closed__2); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__1); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__2); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__3); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__4); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__5); -l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6 = _init_l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Server_FileWorker_handleCompletion___lambda__2___closed__6); l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__1 = _init_l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__1(); lean_mark_persistent(l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__1); l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__2 = _init_l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__2(); @@ -32434,6 +32326,8 @@ l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed_ lean_mark_persistent(l_List_findSomeM_x3f___at_Lean_Server_FileWorker_handleHover___spec__2___closed__4); l_Lean_Server_FileWorker_handleHover___lambda__7___closed__1 = _init_l_Lean_Server_FileWorker_handleHover___lambda__7___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleHover___lambda__7___closed__1); +l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1 = _init_l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_locationLinksOfInfo_extractInstances___closed__1); l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1 = _init_l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__1); l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__2 = _init_l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__2___closed__2(); @@ -32650,50 +32544,50 @@ l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1 = _ini lean_mark_persistent(l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken___closed__1); l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken = _init_l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken(); lean_mark_persistent(l_Lean_Server_FileWorker_instHashableAbsoluteLspSemanticToken); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__1); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__2); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__3); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__4); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__5); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__6); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__7); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__8); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__9); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__10); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__11); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__12); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__13); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__14); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__15); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__16); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__17); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__18); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__19); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__20); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__21); -l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8744____closed__22); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__1); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__2); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__3); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__4); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__5); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__6); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__7); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__8); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__9); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__10); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__11); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__12); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__13); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__14); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__15); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__16); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__17); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__18); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__19); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__20); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__21); +l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22 = _init_l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_RequestHandling_0__Lean_Server_FileWorker_fromJsonAbsoluteLspSemanticToken____x40_Lean_Server_FileWorker_RequestHandling___hyg_8691____closed__22); l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken___closed__1 = _init_l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken___closed__1); l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken = _init_l_Lean_Server_FileWorker_instFromJsonAbsoluteLspSemanticToken(); @@ -32753,145 +32647,145 @@ l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1 = _ini lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___closed__1); l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1 = _init_l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__1); -l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__4___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___lambda__5___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__1___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__4___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__7___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__10___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__13___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__17___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__21___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__2___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__24___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__27___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__30___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__34___lambda__4___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__1); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__3___closed__2); -l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____spec__37___lambda__4___closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__1); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__2); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__3); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__4); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__5); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__6); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__7); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__8); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__9); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__10); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__11); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__12); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__13); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__14); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__15); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__16); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__17); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__18); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__19); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__20); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__21); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__22); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__23); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__24); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__25); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__26); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__27); -l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28(); -lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731____closed__28); -if (builtin) {res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12731_(lean_io_mk_world()); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__1); +l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2 = _init_l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__4___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___lambda__5___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__1___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__4___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__7___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__10___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__13___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__17___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__21___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__2___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__24___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__27___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__30___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__34___lambda__4___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__1); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__3___closed__2); +l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1 = _init_l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Server_registerLspRequestHandler___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____spec__37___lambda__4___closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__1); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__2); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__3); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__4); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__5); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__6); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__7); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__8); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__9); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__10); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__11); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__12); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__13); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__14); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__15); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__16); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__17); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__18); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__19); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__20); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__21); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__22); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__23); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__24); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__25); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__26); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__27); +l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28 = _init_l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28(); +lean_mark_persistent(l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680____closed__28); +if (builtin) {res = l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12680_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c index 128d1eb8ff7d..686691349bc1 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c +++ b/stage0/stdlib/Lean/Server/FileWorker/WidgetRequests.c @@ -14,308 +14,308 @@ extern "C" { #endif lean_object* l_Lean_JsonNumber_fromNat(lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__9; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_makePopup___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__18; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520_(lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__9; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369_(lean_object*); extern lean_object* l_Lean_Server_builtinRpcProcedures; static lean_object* l_Lean_Widget_instRpcEncodableInfoPopup___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_5_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__18; lean_object* l_Lean_Widget_ppExprTagged(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_201_; +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389_(lean_object*); static lean_object* l_Lean_Widget_instInhabitedMsgToInteractive___closed__1; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2___closed__2; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Server_registerBuiltinRpcProcedure___spec__5(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__12; LEAN_EXPORT lean_object* l_Lean_Widget_lazyTraceChildrenToInteractive___boxed__const__1; lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo_enc____x40_Lean_Widget_InteractiveCode___hyg_291_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6; lean_object* l_Lean_Server_instRpcEncodableOption___rarg(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58_(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__32; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204____closed__1; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4; static lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2; +static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__27; static lean_object* l_Lean_Widget_instRpcEncodableInfoPopup___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730_; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_253_; -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__22; LEAN_EXPORT lean_object* l_Lean_Widget_makePopup___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_makePopup(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__24; -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_instRpcEncodableWithRpcRefOfTypeName_rpcDecode___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_makePopup_ppExprTaggedWithoutTopLevelHighlight(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedGetInteractiveDiagnosticsParams; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___closed__3; static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____closed__2; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__28; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__29; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__19; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__16; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__17; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__31; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1; static lean_object* l_Lean_Widget_instRpcEncodableInfoPopup_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_448____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_725_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_788____closed__1; extern lean_object* l_Lean_Lsp_instFromJsonPlainTermGoalParams; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__7; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInfoPopup_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_448_(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1; static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_253____closed__1; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_569____spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387_(lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive___closed__2; lean_object* l_Lean_initializing(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__11; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__5; +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1; lean_object* l_Lean_Json_getStr_x3f(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__8; -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1733_(lean_object*); -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__10; +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409_(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInfoPopup_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_448____closed__1; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__8; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_lazyTraceChildrenToInteractive___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2; lean_object* l_Lean_Widget_TaggedText_instRpcEncodable___rarg(lean_object*); lean_object* l___private_Lean_Widget_TaggedText_0__Lean_Widget_toJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_620____at_Lean_Widget_TaggedText_instRpcEncodable___spec__3(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_instRpcEncodableArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__7; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2836_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Widget_instTypeNameLazyTraceChildren; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241_(lean_object*); lean_object* l_Lean_Elab_Info_type_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_5_(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6; +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_788_; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_(lean_object*, lean_object*); +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedMsgToInteractive; static lean_object* l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__3; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__11; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784_; lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5; +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1; lean_object* l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__30; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245_(lean_object*); lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4; +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2(lean_object*, lean_object*); extern lean_object* l_Lean_Lsp_instFromJsonLocationLink; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__6; extern lean_object* l_Lean_Widget_instRpcEncodableInteractiveTermGoal; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___closed__2; lean_object* l___private_Lean_Server_GoTo_0__Lean_Server_toJsonGoToKind____x40_Lean_Server_GoTo___hyg_27_(uint8_t); lean_object* l_Lean_Server_RequestM_asTask___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constName_x3f(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1735_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Server_locationLinksFromDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2; +LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732_; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInfoPopup_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_448_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__2; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__33; lean_object* l_Lean_Server_instRpcEncodableWithRpcRefOfTypeName_rpcEncode___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__26; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__12; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9; LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__8; +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1; extern lean_object* l_Lean_Widget_instTypeNameMessageData; extern lean_object* l_Lean_Lsp_instToJsonPlainTermGoalParams; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2; lean_object* l___private_Lean_Widget_TaggedText_0__Lean_Widget_fromJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_406____at_Lean_Widget_TaggedText_instRpcEncodable___spec__5(lean_object*); lean_object* l_Lean_Elab_Info_docString_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_makePopup___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_lazyTraceChildrenToInteractive___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3; +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587_(lean_object*); static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_201____closed__1; static lean_object* l_Lean_Widget_instInhabitedInfoPopup___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_204_(lean_object*); lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__14; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589_(lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLineRange____x40_Lean_Data_Lsp_Extra___hyg_2942_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782_; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2; +static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__13; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Lsp_instToJsonPlainGoalParams; lean_object* l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_makePopup___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive___closed__1; -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__14; lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_291_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519_(lean_object*); extern lean_object* l_Lean_Lsp_instFromJsonPlainGoalParams; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_enc____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__13; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; LEAN_EXPORT lean_object* l_Lean_Widget_makePopup_ppExprTaggedWithoutTopLevelHighlight___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__13; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_instRpcEncodableWithRpcRefOfTypeName___rarg(lean_object*); static lean_object* l_Lean_Widget_instInhabitedMsgToInteractive___closed__2; extern lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed; +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_(lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3; lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____closed__1; static lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3; static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInfoPopup___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_725____boxed(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__2; lean_object* l_Lean_Json_getNat_x3f(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__6; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInfoPopup; LEAN_EXPORT lean_object* l_Lean_Widget_makePopup___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2; lean_object* l_Lean_Widget_msgToInteractive(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__16; lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__10; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2; lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Server_registerBuiltinRpcProcedure___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableMsgToInteractive___closed__3; lean_object* l___private_Lean_Server_GoTo_0__Lean_Server_fromJsonGoToKind____x40_Lean_Server_GoTo___hyg_66_(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Elab_Info_lctx(lean_object*); size_t lean_array_size(lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__13; +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___closed__1; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2; lean_object* l_Lean_Server_FileWorker_getInteractiveTermGoal(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_lazyTraceChildrenToInteractive(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__15; lean_object* lean_string_append(lean_object*, lean_object*); extern lean_object* l_Lean_Lsp_instToJsonLocationLink; static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11; -static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__15; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__21; extern lean_object* l_Lean_Widget_instTypeNameInfoWithCtx; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__23; -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__17; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__5; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5; static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__25; -static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1; lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_722_; +static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11; static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____closed__4; -static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1; -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2; static lean_object* l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925_(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_722____closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedInfoPopup; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923_(lean_object*); static lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_520____closed__3; lean_object* l_Lean_Server_FileWorker_getInteractiveGoals(lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); @@ -2792,7 +2792,7 @@ return x_19; } } } -case 5: +case 6: { lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_20 = lean_ctor_get(x_1, 0); @@ -2810,7 +2810,7 @@ lean_ctor_set(x_26, 0, x_25); x_27 = l_Lean_Widget_makePopup___lambda__1(x_1, x_2, x_26, x_3, x_4, x_5, x_6, x_7); return x_27; } -case 11: +case 12: { lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; x_28 = lean_ctor_get(x_1, 0); @@ -3133,7 +3133,7 @@ lean_dec(x_4); return x_6; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3142,11 +3142,11 @@ x_2 = l_Lean_Server_instRpcEncodableWithRpcRefOfTypeName___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1; +x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1; x_4 = l_Lean_Widget_instRpcEncodableInfoPopup; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); @@ -3158,12 +3158,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -3193,7 +3193,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -3211,7 +3211,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -3243,7 +3243,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -3265,7 +3265,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -3331,13 +3331,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1() { _start: { lean_object* x_1; @@ -3345,19 +3345,19 @@ x_1 = lean_mk_string_unchecked("infoToInteractive", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____closed__1; -x_4 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1; +x_4 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3() { _start: { lean_object* x_1; @@ -3365,36 +3365,36 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Widget_makePopup), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3404,7 +3404,7 @@ x_3 = l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -3413,12 +3413,12 @@ x_2 = l_Lean_Server_instRpcEncodableOption___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1; -x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2; +x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1; +x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); lean_closure_set(x_5, 1, lean_box(0)); @@ -3429,12 +3429,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -3464,7 +3464,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -3482,7 +3482,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -3514,7 +3514,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -3536,7 +3536,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -3602,13 +3602,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1() { _start: { lean_object* x_1; @@ -3616,18 +3616,18 @@ x_1 = lean_mk_string_unchecked("getInteractiveGoals", 19, 19); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3() { _start: { lean_object* x_1; @@ -3635,36 +3635,36 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_getInteractiveGoals), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3674,7 +3674,7 @@ x_3 = l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -3683,12 +3683,12 @@ x_2 = l_Lean_Server_instRpcEncodableOption___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1; -x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2; +x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1; +x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); lean_closure_set(x_5, 1, lean_box(0)); @@ -3699,12 +3699,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -3734,7 +3734,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -3752,7 +3752,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -3784,7 +3784,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -3806,7 +3806,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -3872,13 +3872,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1() { _start: { lean_object* x_1; @@ -3886,18 +3886,18 @@ x_1 = lean_mk_string_unchecked("getInteractiveTermGoal", 22, 22); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3() { _start: { lean_object* x_1; @@ -3905,30 +3905,30 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_getInteractiveTermGoal return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; @@ -3942,7 +3942,7 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3952,7 +3952,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -3961,7 +3961,7 @@ switch (lean_obj_tag(x_3)) { case 0: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1; +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1; return x_4; } case 1: @@ -4122,7 +4122,7 @@ return x_31; } } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1() { _start: { lean_object* x_1; @@ -4130,7 +4130,7 @@ x_1 = lean_mk_string_unchecked("lineRange", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2() { _start: { lean_object* x_1; @@ -4138,39 +4138,39 @@ x_1 = lean_mk_string_unchecked("GetInteractiveDiagnosticsParams", 31, 31); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; -x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2; +x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3; x_2 = 1; x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6() { _start: { lean_object* x_1; @@ -4178,53 +4178,53 @@ x_1 = lean_mk_string_unchecked("lineRange\?", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7; x_2 = 1; x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__27; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1(x_1, x_2); +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -4233,7 +4233,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10; +x_6 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -4245,7 +4245,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10; +x_9 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -4274,11 +4274,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4287,7 +4287,7 @@ static lean_object* _init_l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsPar _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452_), 1, 0); return x_1; } } @@ -4299,7 +4299,7 @@ x_1 = l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4327,12 +4327,12 @@ return x_8; } } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1; -x_3 = l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517____spec__1(x_2, x_1); +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1; +x_3 = l_Lean_Json_opt___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519____spec__1(x_2, x_1); x_4 = lean_box(0); x_5 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_5, 0, x_3); @@ -4347,7 +4347,7 @@ static lean_object* _init_l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParam _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1517_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1519_), 1, 0); return x_1; } } @@ -4359,7 +4359,7 @@ x_1 = l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1() { _start: { lean_object* x_1; @@ -4367,79 +4367,79 @@ x_1 = lean_mk_string_unchecked("kind", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__18; -x_2 = lean_unsigned_to_nat(1545u); +x_2 = lean_unsigned_to_nat(1547u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2; x_2 = 1; x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5; x_2 = 1; x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__8() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__27; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9() { _start: { lean_object* x_1; @@ -4447,58 +4447,58 @@ x_1 = lean_mk_string_unchecked("info", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10; x_2 = 1; x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__20; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__13() { +static lean_object* _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12; +x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__27; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1; +x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_2); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); lean_dec(x_3); -x_5 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9; +x_5 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9; x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_5); x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) @@ -4526,23 +4526,23 @@ return x_12; } } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730_() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1; +x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1733_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1735_(lean_object* x_1) { _start: { uint8_t x_2; @@ -4552,14 +4552,14 @@ if (x_2 == 0) lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); -x_5 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1; +x_5 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1; lean_ctor_set(x_1, 1, x_3); lean_ctor_set(x_1, 0, x_5); x_6 = lean_box(0); x_7 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_7, 0, x_1); lean_ctor_set(x_7, 1, x_6); -x_8 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9; +x_8 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_4); @@ -4585,7 +4585,7 @@ x_17 = lean_ctor_get(x_1, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_1); -x_18 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1; +x_18 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1; x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_16); @@ -4593,7 +4593,7 @@ x_20 = lean_box(0); x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); -x_22 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9; +x_22 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9; x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_17); @@ -4613,23 +4613,23 @@ return x_29; } } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1733_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1735_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782_() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1; +x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -4648,7 +4648,7 @@ x_9 = lean_ctor_get(x_7, 0); x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_4); lean_ctor_set(x_10, 1, x_9); -x_11 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1733_(x_10); +x_11 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1735_(x_10); lean_ctor_set(x_7, 0, x_11); return x_7; } @@ -4663,7 +4663,7 @@ lean_dec(x_7); x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_4); lean_ctor_set(x_14, 1, x_12); -x_15 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1733_(x_14); +x_15 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1735_(x_14); x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_13); @@ -4671,7 +4671,7 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -4714,18 +4714,18 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587_(x_1); +x_3 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589_(x_1); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); lean_dec(x_3); x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); x_6 = l___private_Lean_Server_GoTo_0__Lean_Server_fromJsonGoToKind____x40_Lean_Server_GoTo___hyg_66_(x_5); -x_7 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1(x_6, x_2); +x_7 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1(x_6, x_2); if (lean_obj_tag(x_7) == 0) { uint8_t x_8; @@ -4813,11 +4813,11 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545____spec__1(x_1, x_2); +x_3 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -4826,7 +4826,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodableGetGoToLocationParams___ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetGoToLocationParams_enc____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_), 2, 0); return x_1; } } @@ -4834,7 +4834,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodableGetGoToLocationParams___ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1545_), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetGoToLocationParams_dec____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1547_), 2, 0); return x_1; } } @@ -4858,7 +4858,7 @@ x_1 = l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__3; return x_1; } } -LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -4868,7 +4868,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4878,21 +4878,21 @@ x_3 = l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1; +x_1 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1; x_2 = l_Lean_Server_instRpcEncodableArray___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; x_3 = l_Lean_Widget_instRpcEncodableGetGoToLocationParams; -x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2; +x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); lean_closure_set(x_5, 1, lean_box(0)); @@ -4903,12 +4903,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -4938,7 +4938,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -4956,7 +4956,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -4988,7 +4988,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -5010,7 +5010,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -5076,13 +5076,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -5229,7 +5229,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; @@ -5260,7 +5260,7 @@ lean_object* x_12; lean_object* x_13; lean_free_object(x_7); lean_dec(x_9); x_12 = lean_box(0); -x_13 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1(x_2, x_3, x_6, x_12, x_4, x_10); +x_13 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1(x_2, x_3, x_6, x_12, x_4, x_10); lean_dec(x_4); return x_13; } @@ -5290,7 +5290,7 @@ else lean_object* x_18; lean_object* x_19; lean_dec(x_14); x_18 = lean_box(0); -x_19 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1(x_2, x_3, x_6, x_18, x_4, x_15); +x_19 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1(x_2, x_3, x_6, x_18, x_4, x_15); lean_dec(x_4); return x_19; } @@ -5323,15 +5323,15 @@ return x_23; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__1), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -5340,10 +5340,10 @@ x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); lean_dec(x_1); x_6 = lean_box(x_4); -x_7 = lean_alloc_closure((void*)(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2___boxed), 5, 2); +x_7 = lean_alloc_closure((void*)(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2___boxed), 5, 2); lean_closure_set(x_7, 0, x_6); lean_closure_set(x_7, 1, x_5); -x_8 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1; +x_8 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1; x_9 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Server_wrapRpcProcedure___elambda__1___spec__2___rarg), 4, 2); lean_closure_set(x_9, 0, x_8); lean_closure_set(x_9, 1, x_7); @@ -5351,7 +5351,7 @@ x_10 = l_Lean_Server_RequestM_asTask___rarg(x_9, x_2, x_3); return x_10; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1() { _start: { lean_object* x_1; @@ -5359,71 +5359,71 @@ x_1 = lean_mk_string_unchecked("getGoToLocation", 15, 15); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__2___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; lean_object* x_7; x_6 = lean_unbox(x_1); lean_dec(x_1); -x_7 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__2(x_6, x_2, x_3, x_4, x_5); +x_7 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__2(x_6, x_2, x_3, x_4, x_5); return x_7; } } @@ -5541,7 +5541,7 @@ lean_dec(x_5); return x_9; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -5550,7 +5550,7 @@ x_2 = l_Lean_Server_instRpcEncodableWithRpcRefOfTypeName___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -5559,12 +5559,12 @@ x_2 = l_Lean_Server_instRpcEncodableArray___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1; -x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2; +x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1; +x_4 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); lean_closure_set(x_5, 1, lean_box(0)); @@ -5575,12 +5575,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_394____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -5610,7 +5610,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -5628,7 +5628,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -5660,7 +5660,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -5682,7 +5682,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -5748,13 +5748,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1() { _start: { lean_object* x_1; @@ -5762,18 +5762,18 @@ x_1 = lean_mk_string_unchecked("lazyTraceChildrenToInteractive", 30, 30); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__2; x_2 = l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_58____closed__4; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3() { _start: { lean_object* x_1; @@ -5781,30 +5781,30 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Widget_lazyTraceChildrenToInteractive), return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; @@ -6014,67 +6014,67 @@ l_Lean_Widget_instRpcEncodableInfoPopup___closed__3 = _init_l_Lean_Widget_instRp lean_mark_persistent(l_Lean_Widget_instRpcEncodableInfoPopup___closed__3); l_Lean_Widget_instRpcEncodableInfoPopup = _init_l_Lean_Widget_instRpcEncodableInfoPopup(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableInfoPopup); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____spec__2___closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1367_(lean_io_mk_world()); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____spec__2___closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1369_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__1); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____spec__2___closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1387_(lean_io_mk_world()); +}l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__1); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____spec__2___closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1389_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__1); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____spec__2___closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1407_(lean_io_mk_world()); +}l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__1); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____spec__2___closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1409_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Widget_instInhabitedGetInteractiveDiagnosticsParams = _init_l_Lean_Widget_instInhabitedGetInteractiveDiagnosticsParams(); lean_mark_persistent(l_Lean_Widget_instInhabitedGetInteractiveDiagnosticsParams); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____spec__1___closed__1); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__1); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__2); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__3); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__4); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__5); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__6); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__7); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__8); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__9); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1450____closed__10); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____spec__1___closed__1); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__1); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__2); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__3); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__4); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__5); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__6); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__7); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__8); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__9); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonGetInteractiveDiagnosticsParams____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1452____closed__10); l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams___closed__1 = _init_l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams___closed__1(); lean_mark_persistent(l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams___closed__1); l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams = _init_l_Lean_Widget_instFromJsonGetInteractiveDiagnosticsParams(); @@ -6083,40 +6083,40 @@ l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams___closed__1 = _init_l_Le lean_mark_persistent(l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams___closed__1); l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams = _init_l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams(); lean_mark_persistent(l_Lean_Widget_instToJsonGetInteractiveDiagnosticsParams); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__1); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__2); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__3); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__4); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__5); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__6); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__7); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__8 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__8(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__8); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__9); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__10); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__11); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__12); -l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__13 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__13(); -lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1587____closed__13); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730____closed__1); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730_(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1730_); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782____closed__1); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782_(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1782_); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__1); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__2); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__3); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__4); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__5); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__6); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__7); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__8 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__8(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__8); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__9); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__10); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__11); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__12); +l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__13 = _init_l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__13(); +lean_mark_persistent(l___private_Lean_Server_FileWorker_WidgetRequests_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1589____closed__13); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732____closed__1); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732_(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1732_); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784____closed__1); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784_(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1784_); l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__1 = _init_l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__1(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__1); l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__2 = _init_l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__2(); @@ -6125,34 +6125,34 @@ l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__3 = _init_l_Lean_W lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetGoToLocationParams___closed__3); l_Lean_Widget_instRpcEncodableGetGoToLocationParams = _init_l_Lean_Widget_instRpcEncodableGetGoToLocationParams(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetGoToLocationParams); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__1); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____spec__3___closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____lambda__3___closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1923_(lean_io_mk_world()); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__1); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____spec__3___closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____lambda__3___closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_1925_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Widget_lazyTraceChildrenToInteractive___boxed__const__1 = _init_l_Lean_Widget_lazyTraceChildrenToInteractive___boxed__const__1(); lean_mark_persistent(l_Lean_Widget_lazyTraceChildrenToInteractive___boxed__const__1); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__1); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____spec__2___closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__1); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__2); -l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2241_(lean_io_mk_world()); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__1); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____spec__2___closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__1); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__2); +l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Server_FileWorker_WidgetRequests___hyg_2245_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/InfoUtils.c b/stage0/stdlib/Lean/Server/InfoUtils.c index 01e0c6c57405..a88615777fe5 100644 --- a/stage0/stdlib/Lean/Server/InfoUtils.c +++ b/stage0/stdlib/Lean/Server/InfoUtils.c @@ -38,6 +38,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__17(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_smallestInfo_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__21___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__17___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_isTerm___boxed(lean_object*); @@ -60,6 +61,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__8___boxed(lean_object*, lean_object*); uint8_t l_Lean_Exception_isInterrupt(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__2(lean_object*); @@ -117,6 +119,7 @@ lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___closed__2; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__17(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_InfoTree_visitM_x27___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__22___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_List_mapTR_loop___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__1___closed__3; @@ -163,6 +166,7 @@ lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__5(lean_object*); lean_object* l_Lean_Elab_CompletionInfo_stx(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__11(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); @@ -199,6 +203,7 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__7___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__52(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__11(lean_object*); @@ -208,6 +213,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__25___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_tailPos_x3f(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_foldInfoTree_go(lean_object*); uint8_t l_Lean_Expr_isSort(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Elab_InfoTree_foldInfo_go___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -259,6 +265,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__22(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__20(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__18___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -269,7 +276,9 @@ LEAN_EXPORT lean_object* l_lexOrd___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_isSmaller___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Elab_CompletionInfo_lctx(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__10(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__54(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -295,6 +304,7 @@ uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* l_Lean_getOptionDecls(lean_object*); lean_object* lean_array_to_list(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_getCompletionInfos___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__25___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -304,6 +314,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtModule_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_String_instInhabitedRange; lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isMissing(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__21(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,6 +348,7 @@ static lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtModule_x3f___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Elab_InfoTree_visitM_x27___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64(lean_object*, lean_object*, size_t, size_t, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -354,6 +366,7 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__50(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__15(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -369,6 +382,8 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_InfoTree_smallestInfo_ LEAN_EXPORT uint8_t l_compareOn___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_lexOrd___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__8; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__3(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_InfoTree_collectNodesBottomUp___spec__4___rarg(lean_object*); @@ -405,6 +420,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry lean_object* l_panic___at_Lean_Parser_ParserState_mkUnexpectedTokenErrors___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_smallestInfo_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_lexOrd___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__8___closed__1; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_docString_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__47(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__33(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -428,6 +444,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Elab_InfoTree_ static lean_object* l_Lean_Elab_InfoTree_visitM_go___rarg___closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__28(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_lexOrd___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__8___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__26(lean_object*, lean_object*, lean_object*); @@ -451,6 +468,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_isEmptyBy___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_159____at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__17(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__9___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__43(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -572,6 +590,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_x27___rarg___lambda__1___bo LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__22___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f_fmtModule_x3f___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -612,6 +631,7 @@ static lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f___lambda__2___closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__24___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Int_negOfNat(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfoTree_go___spec__8(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -5506,7 +5526,7 @@ return x_3; LEAN_EXPORT uint8_t l_Lean_Elab_Info_isCompletion(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 6) +if (lean_obj_tag(x_1) == 7) { uint8_t x_2; x_2 = 1; @@ -5533,7 +5553,7 @@ return x_3; LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_getCompletionInfos___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -if (lean_obj_tag(x_2) == 6) +if (lean_obj_tag(x_2) == 7) { lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_2, 0); @@ -5600,67 +5620,76 @@ x_7 = lean_ctor_get(x_6, 1); lean_inc(x_7); return x_7; } -case 4: +case 2: { -lean_object* x_8; lean_object* x_9; +lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_ctor_get(x_1, 0); x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -return x_9; +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +return x_10; } case 5: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 0); -x_11 = lean_ctor_get(x_10, 4); -lean_inc(x_11); -return x_11; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +return x_12; } case 6: { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_1, 0); -x_13 = l_Lean_Elab_CompletionInfo_stx(x_12); -return x_13; +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_13, 4); +lean_inc(x_14); +return x_14; } -case 8: +case 7: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_1, 0); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -return x_15; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_1, 0); +x_16 = l_Lean_Elab_CompletionInfo_stx(x_15); +return x_16; } case 9: { -lean_object* x_16; -x_16 = lean_box(0); -return x_16; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +return x_18; } case 10: { -lean_object* x_17; -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -return x_17; +lean_object* x_19; +x_19 = lean_box(0); +return x_19; } case 11: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_1, 0); -x_19 = lean_ctor_get(x_18, 0); -x_20 = lean_ctor_get(x_19, 0); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -return x_21; +lean_object* x_20; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +return x_20; +} +case 12: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_ctor_get(x_21, 0); +x_23 = lean_ctor_get(x_22, 0); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +return x_24; } default: { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_1, 0); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -return x_23; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_1, 0); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +return x_26; } } } @@ -5686,7 +5715,7 @@ x_3 = lean_ctor_get(x_2, 1); lean_inc(x_3); return x_3; } -case 3: +case 4: { lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_1, 0); @@ -5694,7 +5723,7 @@ x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); return x_5; } -case 5: +case 6: { lean_object* x_6; lean_object* x_7; x_6 = lean_ctor_get(x_1, 0); @@ -5702,14 +5731,14 @@ x_7 = lean_ctor_get(x_6, 2); lean_inc(x_7); return x_7; } -case 6: +case 7: { lean_object* x_8; lean_object* x_9; x_8 = lean_ctor_get(x_1, 0); x_9 = l_Lean_Elab_CompletionInfo_lctx(x_8); return x_9; } -case 11: +case 12: { lean_object* x_10; lean_object* x_11; lean_object* x_12; x_10 = lean_ctor_get(x_1, 0); @@ -7778,14 +7807,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hoverableInfoAt_x3f___lambda__3(le { lean_object* x_9; lean_object* x_15; switch (lean_obj_tag(x_1)) { -case 4: +case 5: { lean_object* x_20; x_20 = lean_box(0); x_15 = x_20; goto block_19; } -case 5: +case 6: { lean_object* x_21; x_21 = lean_box(0); @@ -8673,7 +8702,7 @@ return x_31; } } } -case 5: +case 6: { uint8_t x_32; x_32 = !lean_is_exclusive(x_1); @@ -8800,7 +8829,7 @@ return x_56; } } } -case 11: +case 12: { uint8_t x_57; x_57 = !lean_is_exclusive(x_1); @@ -9231,7 +9260,7 @@ return x_22; } } } -case 4: +case 5: { lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; x_23 = lean_ctor_get(x_7, 0); @@ -9322,7 +9351,7 @@ return x_43; } } } -case 5: +case 6: { lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; uint8_t x_51; lean_dec(x_5); @@ -9365,7 +9394,7 @@ lean_ctor_set(x_54, 1, x_53); return x_54; } } -case 11: +case 12: { uint8_t x_55; lean_dec(x_5); @@ -10867,7 +10896,7 @@ return x_23; } } } -case 5: +case 6: { uint8_t x_24; x_24 = !lean_is_exclusive(x_1); @@ -12179,7 +12208,7 @@ x_12 = 0; return x_12; } } -case 3: +case 4: { lean_object* x_13; uint8_t x_14; x_13 = lean_ctor_get(x_4, 1); @@ -18811,131 +18840,1105 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hasSorry_go(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { _start: { -switch (lean_obj_tag(x_2)) { -case 0: +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -lean_dec(x_2); -x_6 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_4, x_1); -x_1 = x_6; -x_2 = x_5; -goto _start; -} -case 1: +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57(x_1, x_7, x_5); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) { -if (lean_obj_tag(x_1) == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__1(x_1, x_8, x_3); +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); lean_dec(x_8); -return x_9; +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; } else { -lean_object* x_10; -x_10 = lean_ctor_get(x_2, 0); -lean_inc(x_10); -switch (lean_obj_tag(x_10)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_10); -x_11 = lean_ctor_get(x_2, 1); -lean_inc(x_11); -lean_dec(x_2); -x_12 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__6(x_1, x_11, x_3); -lean_dec(x_11); -return x_12; -} -case 1: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_2); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); +uint8_t x_15; lean_dec(x_1); -x_14 = lean_ctor_get(x_10, 0); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_ctor_get(x_14, 3); -lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); -lean_closure_set(x_16, 0, x_15); -x_17 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_14, x_13, x_16, x_3); -if (lean_obj_tag(x_17) == 0) -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) { -lean_object* x_19; uint8_t x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = l_Lean_Expr_hasSorry(x_19); -lean_dec(x_19); -x_21 = lean_box(x_20); -lean_ctor_set(x_17, 0, x_21); -return x_17; +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; } else { -lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_17, 0); -x_23 = lean_ctor_get(x_17, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_17); -x_24 = l_Lean_Expr_hasSorry(x_22); -lean_dec(x_22); -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_23); -return x_26; +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} } } else { -uint8_t x_27; -x_27 = !lean_is_exclusive(x_17); -if (x_27 == 0) +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) { -return x_17; +return x_8; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_17, 0); -x_29 = lean_ctor_get(x_17, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_17); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } -case 2: +else { -lean_object* x_31; lean_object* x_32; -lean_dec(x_10); -x_31 = lean_ctor_get(x_2, 1); -lean_inc(x_31); -lean_dec(x_2); -x_32 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__11(x_1, x_31, x_3); -lean_dec(x_31); -return x_32; +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_Elab_InfoTree_hasSorry_go(x_1, x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) +{ +return x_8; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_1); +x_8 = 0; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58(x_1, x_4, x_11, x_12, x_3); +return x_13; +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_array_get_size(x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_lt(x_16, x_15); +if (x_17 == 0) +{ +uint8_t x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +lean_dec(x_1); +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; +} +else +{ +size_t x_21; size_t x_22; lean_object* x_23; +x_21 = 0; +x_22 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_23 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59(x_1, x_14, x_21, x_22, x_3); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_Elab_InfoTree_hasSorry_go(x_1, x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) +{ +return x_8; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_1); +x_5 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57(x_1, x_4, x_3); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_unbox(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_6); +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_9 = lean_ctor_get(x_5, 1); +x_10 = lean_ctor_get(x_5, 0); +lean_dec(x_10); +x_11 = lean_ctor_get(x_2, 1); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +uint8_t x_15; lean_object* x_16; +lean_dec(x_12); +lean_dec(x_1); +x_15 = 0; +x_16 = lean_box(x_15); +lean_ctor_set(x_5, 0, x_16); +return x_5; +} +else +{ +size_t x_17; size_t x_18; lean_object* x_19; +lean_free_object(x_5); +x_17 = 0; +x_18 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_19 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60(x_1, x_11, x_17, x_18, x_9); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_dec(x_5); +x_21 = lean_ctor_get(x_2, 1); +x_22 = lean_array_get_size(x_21); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_lt(x_23, x_22); +if (x_24 == 0) +{ +uint8_t x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_22); +lean_dec(x_1); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_20); +return x_27; +} +else +{ +size_t x_28; size_t x_29; lean_object* x_30; +x_28 = 0; +x_29 = lean_usize_of_nat(x_22); +lean_dec(x_22); +x_30 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60(x_1, x_21, x_28, x_29, x_20); +return x_30; +} +} +} +else +{ +uint8_t x_31; +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_5); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_5, 0); +lean_dec(x_32); +return x_5; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_dec(x_5); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_6); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_5); +if (x_35 == 0) +{ +return x_5; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_5, 0); +x_37 = lean_ctor_get(x_5, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_5); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62(x_1, x_7, x_5); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) +{ +return x_8; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_Elab_InfoTree_hasSorry_go(x_1, x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) +{ +return x_8; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_5); +lean_dec(x_1); +x_8 = 0; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63(x_1, x_4, x_11, x_12, x_3); +return x_13; +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_array_get_size(x_14); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_lt(x_16, x_15); +if (x_17 == 0) +{ +uint8_t x_18; lean_object* x_19; lean_object* x_20; +lean_dec(x_15); +lean_dec(x_1); +x_18 = 0; +x_19 = lean_box(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_3); +return x_20; +} +else +{ +size_t x_21; size_t x_22; lean_object* x_23; +x_21 = 0; +x_22 = lean_usize_of_nat(x_15); +lean_dec(x_15); +x_23 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64(x_1, x_14, x_21, x_22, x_3); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = l_Lean_Elab_InfoTree_hasSorry_go(x_1, x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; size_t x_12; size_t x_13; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; +x_5 = x_11; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_8, 0, x_18); +return x_8; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_dec(x_8); +x_20 = 1; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +} +} +else +{ +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_8); +if (x_23 == 0) +{ +return x_8; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_8, 0); +x_25 = lean_ctor_get(x_8, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_8); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_5); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_1); +x_5 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62(x_1, x_4, x_3); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_unbox(x_6); +if (x_7 == 0) +{ +uint8_t x_8; +lean_dec(x_6); +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_9 = lean_ctor_get(x_5, 1); +x_10 = lean_ctor_get(x_5, 0); +lean_dec(x_10); +x_11 = lean_ctor_get(x_2, 1); +x_12 = lean_array_get_size(x_11); +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_lt(x_13, x_12); +if (x_14 == 0) +{ +uint8_t x_15; lean_object* x_16; +lean_dec(x_12); +lean_dec(x_1); +x_15 = 0; +x_16 = lean_box(x_15); +lean_ctor_set(x_5, 0, x_16); +return x_5; +} +else +{ +size_t x_17; size_t x_18; lean_object* x_19; +lean_free_object(x_5); +x_17 = 0; +x_18 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_19 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65(x_1, x_11, x_17, x_18, x_9); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_dec(x_5); +x_21 = lean_ctor_get(x_2, 1); +x_22 = lean_array_get_size(x_21); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_nat_dec_lt(x_23, x_22); +if (x_24 == 0) +{ +uint8_t x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_22); +lean_dec(x_1); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_20); +return x_27; +} +else +{ +size_t x_28; size_t x_29; lean_object* x_30; +x_28 = 0; +x_29 = lean_usize_of_nat(x_22); +lean_dec(x_22); +x_30 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65(x_1, x_21, x_28, x_29, x_20); +return x_30; +} +} +} +else +{ +uint8_t x_31; +lean_dec(x_1); +x_31 = !lean_is_exclusive(x_5); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_5, 0); +lean_dec(x_32); +return x_5; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_dec(x_5); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_6); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +else +{ +uint8_t x_35; +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_5); +if (x_35 == 0) +{ +return x_5; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_5, 0); +x_37 = lean_ctor_get(x_5, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_5); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hasSorry_go(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +lean_dec(x_2); +x_6 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f(x_4, x_1); +x_1 = x_6; +x_2 = x_5; +goto _start; +} +case 1: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__1(x_1, x_8, x_3); +lean_dec(x_8); +return x_9; +} +else +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__6(x_1, x_11, x_3); +lean_dec(x_11); +return x_12; +} +case 1: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_2); +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_ctor_get(x_14, 3); +lean_inc(x_15); +x_16 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); +lean_closure_set(x_16, 0, x_15); +x_17 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_14, x_13, x_16, x_3); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = l_Lean_Expr_hasSorry(x_19); +lean_dec(x_19); +x_21 = lean_box(x_20); +lean_ctor_set(x_17, 0, x_21); +return x_17; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_22 = lean_ctor_get(x_17, 0); +x_23 = lean_ctor_get(x_17, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_17); +x_24 = l_Lean_Expr_hasSorry(x_22); +lean_dec(x_22); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +return x_26; +} +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_17); +if (x_27 == 0) +{ +return x_17; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_17, 0); +x_29 = lean_ctor_get(x_17, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_17); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +case 2: +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_10); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_32 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__11(x_1, x_31, x_3); +lean_dec(x_31); +return x_32; } case 3: { @@ -19025,92 +20028,114 @@ x_48 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__5 lean_dec(x_47); return x_48; } -default: +case 11: { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_2); -x_49 = lean_ctor_get(x_10, 0); -lean_inc(x_49); +lean_object* x_49; lean_object* x_50; lean_dec(x_10); -x_50 = lean_ctor_get(x_1, 0); -lean_inc(x_50); -lean_dec(x_1); -x_51 = lean_ctor_get(x_49, 0); -lean_inc(x_51); +x_49 = lean_ctor_get(x_2, 1); +lean_inc(x_49); +lean_dec(x_2); +x_50 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56(x_1, x_49, x_3); lean_dec(x_49); -x_52 = lean_ctor_get(x_51, 3); +return x_50; +} +case 12: +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_2); +x_51 = lean_ctor_get(x_10, 0); +lean_inc(x_51); +lean_dec(x_10); +x_52 = lean_ctor_get(x_1, 0); lean_inc(x_52); -x_53 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); -lean_closure_set(x_53, 0, x_52); -x_54 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_51, x_50, x_53, x_3); -if (lean_obj_tag(x_54) == 0) +lean_dec(x_1); +x_53 = lean_ctor_get(x_51, 0); +lean_inc(x_53); +lean_dec(x_51); +x_54 = lean_ctor_get(x_53, 3); +lean_inc(x_54); +x_55 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); +lean_closure_set(x_55, 0, x_54); +x_56 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_53, x_52, x_55, x_3); +if (lean_obj_tag(x_56) == 0) { -uint8_t x_55; -x_55 = !lean_is_exclusive(x_54); -if (x_55 == 0) +uint8_t x_57; +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_56; uint8_t x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_54, 0); -x_57 = l_Lean_Expr_hasSorry(x_56); -lean_dec(x_56); -x_58 = lean_box(x_57); -lean_ctor_set(x_54, 0, x_58); -return x_54; +lean_object* x_58; uint8_t x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_56, 0); +x_59 = l_Lean_Expr_hasSorry(x_58); +lean_dec(x_58); +x_60 = lean_box(x_59); +lean_ctor_set(x_56, 0, x_60); +return x_56; } else { -lean_object* x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; -x_59 = lean_ctor_get(x_54, 0); -x_60 = lean_ctor_get(x_54, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_54); -x_61 = l_Lean_Expr_hasSorry(x_59); -lean_dec(x_59); -x_62 = lean_box(x_61); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_60); -return x_63; +lean_object* x_61; lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; +x_61 = lean_ctor_get(x_56, 0); +x_62 = lean_ctor_get(x_56, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_56); +x_63 = l_Lean_Expr_hasSorry(x_61); +lean_dec(x_61); +x_64 = lean_box(x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; } } else { -uint8_t x_64; -x_64 = !lean_is_exclusive(x_54); -if (x_64 == 0) +uint8_t x_66; +x_66 = !lean_is_exclusive(x_56); +if (x_66 == 0) { -return x_54; +return x_56; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_54, 0); -x_66 = lean_ctor_get(x_54, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_54); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_56, 0); +x_68 = lean_ctor_get(x_56, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_56); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} } } +default: +{ +lean_object* x_70; lean_object* x_71; +lean_dec(x_10); +x_70 = lean_ctor_get(x_2, 1); +lean_inc(x_70); +lean_dec(x_2); +x_71 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61(x_1, x_70, x_3); +lean_dec(x_70); +return x_71; } } } } default: { -uint8_t x_68; lean_object* x_69; lean_object* x_70; +uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_dec(x_2); lean_dec(x_1); -x_68 = 0; -x_69 = lean_box(x_68); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_3); -return x_70; +x_72 = 0; +x_73 = lean_box(x_72); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_3); +return x_74; } } } @@ -19742,6 +20767,120 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__58(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__59(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__57(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__60(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__56(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__63(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__64(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_hasSorry_go___spec__62(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__65(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentArray_anyM___at_Lean_Elab_InfoTree_hasSorry_go___spec__61(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_hasSorry(lean_object* x_1, lean_object* x_2) { _start: { diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index b1d757ddca11..83a67e4bc534 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -11154,7 +11154,7 @@ return x_77; } } } -case 4: +case 5: { lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; x_78 = lean_ctor_get(x_2, 0); @@ -11224,7 +11224,7 @@ return x_103; } } } -case 5: +case 6: { lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_104 = lean_ctor_get(x_2, 0); @@ -15614,7 +15614,7 @@ return x_13; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_combineIdents_buildIdMap___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -if (lean_obj_tag(x_2) == 9) +if (lean_obj_tag(x_2) == 10) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; x_5 = lean_ctor_get(x_2, 0); diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index ecfda4d43352..3555a0a60e60 100644 --- a/stage0/stdlib/Lean/Server/Rpc/Deriving.c +++ b/stage0/stdlib/Lean/Server/Rpc/Deriving.c @@ -44,8 +44,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__8(size_t, size_t, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*); @@ -61,6 +60,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_D lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__7; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__33; @@ -84,7 +84,6 @@ lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_ob lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___lambda__1___closed__5; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__5; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; @@ -104,7 +103,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -135,6 +133,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__3; LEAN_EXPORT uint8_t l_Lean_Server_RpcEncodable_isOptField___lambda__1(lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__33; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; @@ -149,6 +148,7 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__54; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__11___closed__1; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; lean_object* l_Lean_Elab_registerDerivingHandler(lean_object*, lean_object*, lean_object*); uint8_t l_instDecidableNot___rarg(uint8_t); @@ -165,7 +165,6 @@ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__25; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1; lean_object* l_Lean_Meta_isType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2___closed__1; @@ -173,6 +172,7 @@ lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_ lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_isOptField___lambda__1___boxed(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12; lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; @@ -202,9 +202,11 @@ static lean_object* l_Lean_Server_RpcEncodable_isOptField___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8; static lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2___closed__2; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__59; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; +static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; @@ -212,6 +214,7 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__64; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_RpcEncodable_isOptField___closed__5; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__45; @@ -254,10 +257,8 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__49; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__50; @@ -292,7 +293,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__81; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__29; @@ -307,14 +307,12 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__6(size_t, size_t, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8; lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___lambda__1___closed__1; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -331,7 +329,6 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_ lean_object* lean_string_length(lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11; lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); @@ -339,29 +336,28 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; lean_object* lean_environment_main_module(lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; LEAN_EXPORT uint8_t l_Lean_Server_RpcEncodable_isOptField(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__37; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__5___closed__4; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_RpcEncodable_isOptField___closed__2; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11; lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__49; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__11; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__52; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Substring_prevn(lean_object*, lean_object*, lean_object*); @@ -377,14 +373,15 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15; uint8_t l_Lean_isStructure(lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_instCoeTSyntaxConsSyntaxNodeKindStrNumAnonymousOfNatNatNilMkStr4___boxed(lean_object*); static lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__1___closed__2; static lean_object* l_Lean_Server_RpcEncodable_isOptField___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__35; static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__1___closed__2; @@ -393,20 +390,22 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__48; lean_object* lean_array_mk(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; uint8_t l_Substring_beq(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__6; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___lambda__1___closed__6; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__4___closed__2; lean_object* lean_array_uget(lean_object*, size_t); size_t lean_array_size(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___lambda__1___closed__2; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; @@ -415,6 +414,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__28; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15; static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_isOptField___boxed(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14___lambda__1___closed__4; @@ -439,6 +439,7 @@ uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; lean_object* l_Lean_throwError___at_Lean_withSetOptionIn___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; @@ -481,6 +482,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_instCoeTSyntaxConsSyntaxNodeKindStrNumAnonymousOfNatNatNilMkStr4(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87; lean_object* l_Lean_Parser_termParser(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837_(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; LEAN_EXPORT uint8_t l_Lean_Server_RpcEncodable_isOptField___lambda__1(lean_object* x_1) { @@ -3532,30 +3534,50 @@ static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_Rp _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("enc", 3, 3); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; +x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; +x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__69; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("enc", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__69; +x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__69; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74() { _start: { lean_object* x_1; @@ -3563,26 +3585,26 @@ x_1 = lean_mk_string_unchecked("dec", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77() { _start: { lean_object* x_1; @@ -3590,19 +3612,19 @@ x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79() { _start: { lean_object* x_1; @@ -3610,7 +3632,7 @@ x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80() { _start: { lean_object* x_1; @@ -3618,7 +3640,7 @@ x_1 = lean_mk_string_unchecked("Termination", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__81() { _start: { lean_object* x_1; @@ -3626,19 +3648,19 @@ x_1 = lean_mk_string_unchecked("suffix", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; -x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__81; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__81() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__83() { _start: { lean_object* x_1; @@ -3646,19 +3668,19 @@ x_1 = lean_mk_string_unchecked("whereDecls", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__81; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__83; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__83() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__85() { _start: { lean_object* x_1; @@ -3666,19 +3688,19 @@ x_1 = lean_mk_string_unchecked("letRecDecl", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__83; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__85; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__85() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87() { _start: { lean_object* x_1; @@ -3686,19 +3708,19 @@ x_1 = lean_mk_string_unchecked("letDecl", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__85; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__89() { _start: { lean_object* x_1; @@ -3706,19 +3728,19 @@ x_1 = lean_mk_string_unchecked("letIdDecl", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__89; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__89() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91() { _start: { lean_object* x_1; @@ -3726,19 +3748,19 @@ x_1 = lean_mk_string_unchecked("termReturn", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__89; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93() { _start: { lean_object* x_1; @@ -3746,7 +3768,7 @@ x_1 = lean_mk_string_unchecked("return", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94() { _start: { lean_object* x_1; @@ -3754,61 +3776,61 @@ x_1 = lean_mk_string_unchecked("toJson", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__33; -x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3820,7 +3842,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101() { _start: { lean_object* x_1; @@ -3828,26 +3850,26 @@ x_1 = lean_mk_string_unchecked("j", 1, 1); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104() { _start: { lean_object* x_1; @@ -3855,19 +3877,19 @@ x_1 = lean_mk_string_unchecked("do", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__106() { _start: { lean_object* x_1; @@ -3875,19 +3897,19 @@ x_1 = lean_mk_string_unchecked("doSeqIndent", 11, 11); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__106; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__106() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108() { _start: { lean_object* x_1; @@ -3895,19 +3917,19 @@ x_1 = lean_mk_string_unchecked("doSeqItem", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__106; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110() { _start: { lean_object* x_1; @@ -3915,19 +3937,19 @@ x_1 = lean_mk_string_unchecked("doLetArrow", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112() { _start: { lean_object* x_1; @@ -3935,7 +3957,7 @@ x_1 = lean_mk_string_unchecked("let", 3, 3); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__113() { _start: { lean_object* x_1; @@ -3943,19 +3965,19 @@ x_1 = lean_mk_string_unchecked("doIdDecl", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__113; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__113() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115() { _start: { lean_object* x_1; @@ -3963,19 +3985,19 @@ x_1 = lean_mk_string_unchecked("doExpr", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__113; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117() { _start: { lean_object* x_1; @@ -3983,61 +4005,61 @@ x_1 = lean_mk_string_unchecked("fromJson\?", 9, 9); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; +x_1 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__24; -x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; +x_3 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__121; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__121() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123() { _start: { lean_object* x_1; @@ -4045,14 +4067,14 @@ x_1 = lean_mk_string_unchecked("doReturn", 8, 8); return x_1; } } -static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122() { +static lean_object* _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__10; x_3 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; -x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__121; +x_4 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -4123,7 +4145,7 @@ x_42 = lean_st_ref_get(x_11, x_37); x_43 = !lean_is_exclusive(x_42); if (x_43 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; size_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; size_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; size_t x_183; lean_object* x_184; size_t x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; size_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; size_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; size_t x_185; lean_object* x_186; size_t x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; x_44 = lean_ctor_get(x_42, 0); x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); @@ -4360,11 +4382,11 @@ x_143 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__ lean_inc(x_49); lean_inc(x_40); x_144 = l_Lean_Syntax_node2(x_40, x_143, x_142, x_49); -x_145 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; +x_145 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; lean_inc(x_41); lean_inc(x_46); x_146 = l_Lean_addMacroScope(x_46, x_145, x_41); -x_147 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +x_147 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; lean_inc(x_40); x_148 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_148, 0, x_40); @@ -4391,11 +4413,11 @@ lean_ctor_set(x_155, 3, x_154); lean_inc(x_49); lean_inc(x_40); x_156 = l_Lean_Syntax_node2(x_40, x_143, x_155, x_49); -x_157 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; +x_157 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; lean_inc(x_41); lean_inc(x_46); x_158 = l_Lean_addMacroScope(x_46, x_157, x_41); -x_159 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +x_159 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; lean_inc(x_40); x_160 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_160, 0, x_40); @@ -4408,741 +4430,755 @@ lean_inc(x_40); x_161 = l_Lean_Syntax_node3(x_40, x_149, x_156, x_135, x_160); lean_inc(x_40); x_162 = l_Lean_Syntax_node3(x_40, x_47, x_150, x_18, x_161); -x_163 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +x_163 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_40); +x_164 = l_Lean_Syntax_node1(x_40, x_163, x_162); +x_165 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; lean_inc(x_49); lean_inc(x_40); -x_164 = l_Lean_Syntax_node1(x_40, x_163, x_49); -x_165 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; +x_166 = l_Lean_Syntax_node1(x_40, x_165, x_49); +x_167 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; lean_inc(x_40); -x_166 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_166, 0, x_40); -lean_ctor_set(x_166, 1, x_165); -x_167 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +x_168 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_168, 0, x_40); +lean_ctor_set(x_168, 1, x_167); +x_169 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_168); lean_inc(x_166); -lean_inc(x_164); lean_inc_n(x_49, 2); lean_inc(x_137); lean_inc(x_40); -x_168 = l_Lean_Syntax_node6(x_40, x_167, x_137, x_49, x_162, x_164, x_49, x_166); -x_169 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; +x_170 = l_Lean_Syntax_node6(x_40, x_169, x_137, x_49, x_164, x_166, x_49, x_168); +x_171 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; lean_inc_n(x_49, 2); lean_inc(x_40); -x_170 = l_Lean_Syntax_node2(x_40, x_169, x_49, x_49); -x_171 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +x_172 = l_Lean_Syntax_node2(x_40, x_171, x_49, x_49); +x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; lean_inc(x_41); lean_inc(x_46); -x_172 = l_Lean_addMacroScope(x_46, x_171, x_41); -x_173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +x_174 = l_Lean_addMacroScope(x_46, x_173, x_41); +x_175 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; lean_inc(x_40); -x_174 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_174, 0, x_40); -lean_ctor_set(x_174, 1, x_173); -lean_ctor_set(x_174, 2, x_172); -lean_ctor_set(x_174, 3, x_13); -lean_inc(x_174); +x_176 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_176, 0, x_40); +lean_ctor_set(x_176, 1, x_175); +lean_ctor_set(x_176, 2, x_174); +lean_ctor_set(x_176, 3, x_13); +lean_inc(x_176); lean_inc(x_40); -x_175 = l_Lean_Syntax_node1(x_40, x_47, x_174); -x_176 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_177 = l_Lean_Syntax_node1(x_40, x_47, x_176); +x_178 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_40); -x_177 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_177, 0, x_40); -lean_ctor_set(x_177, 1, x_176); -x_178 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_179 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_179, 0, x_40); +lean_ctor_set(x_179, 1, x_178); +x_180 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_41); lean_inc(x_46); -x_179 = l_Lean_addMacroScope(x_46, x_178, x_41); -x_180 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_181 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +x_181 = l_Lean_addMacroScope(x_46, x_180, x_41); +x_182 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_183 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_40); -x_182 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_182, 0, x_40); -lean_ctor_set(x_182, 1, x_180); -lean_ctor_set(x_182, 2, x_179); -lean_ctor_set(x_182, 3, x_181); -x_183 = lean_array_size(x_29); -x_184 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_149, x_183, x_16, x_29); -x_185 = lean_array_size(x_184); -x_186 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_185, x_16, x_184); -x_187 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_188 = l_Lean_mkSepArray(x_186, x_187); -lean_dec(x_186); -x_189 = l_Array_append___rarg(x_48, x_188); +x_184 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_184, 0, x_40); +lean_ctor_set(x_184, 1, x_182); +lean_ctor_set(x_184, 2, x_181); +lean_ctor_set(x_184, 3, x_183); +x_185 = lean_array_size(x_29); +x_186 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_149, x_185, x_16, x_29); +x_187 = lean_array_size(x_186); +x_188 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_187, x_16, x_186); +x_189 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_190 = l_Lean_mkSepArray(x_188, x_189); lean_dec(x_188); +x_191 = l_Array_append___rarg(x_48, x_190); +lean_dec(x_190); lean_inc(x_40); -x_190 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_190, 0, x_40); -lean_ctor_set(x_190, 1, x_47); -lean_ctor_set(x_190, 2, x_189); +x_192 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_192, 0, x_40); +lean_ctor_set(x_192, 1, x_47); +lean_ctor_set(x_192, 2, x_191); +lean_inc(x_40); +x_193 = l_Lean_Syntax_node1(x_40, x_163, x_192); lean_inc(x_58); lean_inc(x_103); lean_inc(x_40); -x_191 = l_Lean_Syntax_node2(x_40, x_47, x_103, x_58); +x_194 = l_Lean_Syntax_node2(x_40, x_47, x_103, x_58); +lean_inc(x_168); lean_inc(x_166); -lean_inc(x_164); lean_inc(x_49); lean_inc(x_137); lean_inc(x_40); -x_192 = l_Lean_Syntax_node6(x_40, x_167, x_137, x_49, x_190, x_164, x_191, x_166); +x_195 = l_Lean_Syntax_node6(x_40, x_169, x_137, x_49, x_193, x_166, x_194, x_168); lean_inc(x_40); -x_193 = l_Lean_Syntax_node1(x_40, x_47, x_192); +x_196 = l_Lean_Syntax_node1(x_40, x_47, x_195); lean_inc(x_40); -x_194 = l_Lean_Syntax_node2(x_40, x_122, x_182, x_193); +x_197 = l_Lean_Syntax_node2(x_40, x_122, x_184, x_196); lean_inc(x_40); -x_195 = l_Lean_Syntax_node1(x_40, x_47, x_194); -x_196 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_177); +x_198 = l_Lean_Syntax_node1(x_40, x_47, x_197); +x_199 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_179); lean_inc(x_40); -x_197 = l_Lean_Syntax_node2(x_40, x_196, x_177, x_195); -x_198 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +x_200 = l_Lean_Syntax_node2(x_40, x_199, x_179, x_198); +x_201 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; lean_inc(x_135); lean_inc(x_49); lean_inc(x_40); -x_199 = l_Lean_Syntax_node5(x_40, x_198, x_148, x_175, x_49, x_135, x_197); -x_200 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +x_202 = l_Lean_Syntax_node5(x_40, x_201, x_148, x_177, x_49, x_135, x_200); +x_203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; lean_inc(x_40); -x_201 = l_Lean_Syntax_node1(x_40, x_200, x_199); -x_202 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_170); +x_204 = l_Lean_Syntax_node1(x_40, x_203, x_202); +x_205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_172); lean_inc_n(x_49, 2); lean_inc(x_40); -x_203 = l_Lean_Syntax_node4(x_40, x_202, x_49, x_49, x_201, x_170); -x_204 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; +x_206 = l_Lean_Syntax_node4(x_40, x_205, x_49, x_49, x_204, x_172); +x_207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_41); lean_inc(x_46); -x_205 = l_Lean_addMacroScope(x_46, x_204, x_41); -x_206 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_40); -x_207 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_207, 0, x_40); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_205); -lean_ctor_set(x_207, 3, x_13); -lean_inc(x_40); -x_208 = l_Lean_Syntax_node1(x_40, x_47, x_207); +x_208 = l_Lean_addMacroScope(x_46, x_207, x_41); x_209 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; lean_inc(x_40); -x_210 = lean_alloc_ctor(2, 2, 0); +x_210 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_210, 0, x_40); lean_ctor_set(x_210, 1, x_209); -x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +lean_ctor_set(x_210, 2, x_208); +lean_ctor_set(x_210, 3, x_13); lean_inc(x_40); -x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_40); -lean_ctor_set(x_212, 1, x_211); +x_211 = l_Lean_Syntax_node1(x_40, x_47, x_210); +x_212 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; lean_inc(x_40); -x_213 = l_Lean_Syntax_node2(x_40, x_130, x_103, x_58); +x_213 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_213, 0, x_40); +lean_ctor_set(x_213, 1, x_212); +x_214 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; lean_inc(x_40); -x_214 = l_Lean_Syntax_node1(x_40, x_47, x_213); -x_215 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_215 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_215, 0, x_40); +lean_ctor_set(x_215, 1, x_214); lean_inc(x_40); -x_216 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_216, 0, x_40); -lean_ctor_set(x_216, 1, x_215); -x_217 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_218 = l_Lean_addMacroScope(x_46, x_217, x_41); -x_219 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; +x_216 = l_Lean_Syntax_node2(x_40, x_130, x_103, x_58); lean_inc(x_40); -x_221 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_221, 0, x_40); -lean_ctor_set(x_221, 1, x_219); -lean_ctor_set(x_221, 2, x_218); -lean_ctor_set(x_221, 3, x_220); -lean_inc(x_208); +x_217 = l_Lean_Syntax_node1(x_40, x_47, x_216); +x_218 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_40); -x_222 = l_Lean_Syntax_node2(x_40, x_122, x_221, x_208); -x_223 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_40); +lean_ctor_set(x_219, 1, x_218); +x_220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_221 = l_Lean_addMacroScope(x_46, x_220, x_41); +x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_223 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_40); +x_224 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_224, 0, x_40); +lean_ctor_set(x_224, 1, x_222); +lean_ctor_set(x_224, 2, x_221); +lean_ctor_set(x_224, 3, x_223); +lean_inc(x_211); +lean_inc(x_40); +x_225 = l_Lean_Syntax_node2(x_40, x_122, x_224, x_211); +x_226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; lean_inc(x_40); -x_224 = l_Lean_Syntax_node1(x_40, x_223, x_222); -x_225 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +x_227 = l_Lean_Syntax_node1(x_40, x_226, x_225); +x_228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; lean_inc(x_40); -x_226 = l_Lean_Syntax_node4(x_40, x_225, x_174, x_214, x_216, x_224); -x_227 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +x_229 = l_Lean_Syntax_node4(x_40, x_228, x_176, x_217, x_219, x_227); +x_230 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; lean_inc(x_49); lean_inc(x_40); -x_228 = l_Lean_Syntax_node3(x_40, x_227, x_212, x_49, x_226); -x_229 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +x_231 = l_Lean_Syntax_node3(x_40, x_230, x_215, x_49, x_229); +x_232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; lean_inc(x_49); lean_inc(x_40); -x_230 = l_Lean_Syntax_node2(x_40, x_229, x_228, x_49); -x_231 = l_Lean_Syntax_SepArray_ofElems(x_76, x_26); +x_233 = l_Lean_Syntax_node2(x_40, x_232, x_231, x_49); +x_234 = l_Lean_Syntax_SepArray_ofElems(x_76, x_26); lean_dec(x_26); -x_232 = l_Array_append___rarg(x_48, x_231); -lean_dec(x_231); +x_235 = l_Array_append___rarg(x_48, x_234); +lean_dec(x_234); +lean_inc(x_40); +x_236 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_236, 0, x_40); +lean_ctor_set(x_236, 1, x_47); +lean_ctor_set(x_236, 2, x_235); lean_inc(x_40); -x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_40); -lean_ctor_set(x_233, 1, x_47); -lean_ctor_set(x_233, 2, x_232); +x_237 = l_Lean_Syntax_node1(x_40, x_163, x_236); lean_inc_n(x_49, 2); lean_inc(x_40); -x_234 = l_Lean_Syntax_node6(x_40, x_167, x_137, x_49, x_233, x_164, x_49, x_166); +x_238 = l_Lean_Syntax_node6(x_40, x_169, x_137, x_49, x_237, x_166, x_49, x_168); lean_inc(x_40); -x_235 = l_Lean_Syntax_node1(x_40, x_47, x_234); -x_236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +x_239 = l_Lean_Syntax_node1(x_40, x_47, x_238); +x_240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; lean_inc(x_40); -x_237 = l_Lean_Syntax_node2(x_40, x_236, x_177, x_235); +x_241 = l_Lean_Syntax_node2(x_40, x_240, x_179, x_239); lean_inc(x_49); lean_inc(x_40); -x_238 = l_Lean_Syntax_node2(x_40, x_229, x_237, x_49); +x_242 = l_Lean_Syntax_node2(x_40, x_232, x_241, x_49); lean_inc(x_40); -x_239 = l_Lean_Syntax_node2(x_40, x_47, x_230, x_238); -x_240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_243 = l_Lean_Syntax_node2(x_40, x_47, x_233, x_242); +x_244 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_40); -x_241 = l_Lean_Syntax_node1(x_40, x_240, x_239); -x_242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_245 = l_Lean_Syntax_node1(x_40, x_244, x_243); +x_246 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_40); -x_243 = l_Lean_Syntax_node2(x_40, x_242, x_210, x_241); +x_247 = l_Lean_Syntax_node2(x_40, x_246, x_213, x_245); lean_inc(x_135); lean_inc(x_49); lean_inc(x_40); -x_244 = l_Lean_Syntax_node5(x_40, x_198, x_160, x_208, x_49, x_135, x_243); +x_248 = l_Lean_Syntax_node5(x_40, x_201, x_160, x_211, x_49, x_135, x_247); lean_inc(x_40); -x_245 = l_Lean_Syntax_node1(x_40, x_200, x_244); -lean_inc(x_170); +x_249 = l_Lean_Syntax_node1(x_40, x_203, x_248); +lean_inc(x_172); lean_inc_n(x_49, 2); lean_inc(x_40); -x_246 = l_Lean_Syntax_node4(x_40, x_202, x_49, x_49, x_245, x_170); +x_250 = l_Lean_Syntax_node4(x_40, x_205, x_49, x_49, x_249, x_172); lean_inc(x_49); lean_inc(x_40); -x_247 = l_Lean_Syntax_node3(x_40, x_47, x_203, x_49, x_246); -x_248 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_251 = l_Lean_Syntax_node3(x_40, x_47, x_206, x_49, x_250); +x_252 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_40); -x_249 = l_Lean_Syntax_node2(x_40, x_248, x_20, x_247); +x_253 = l_Lean_Syntax_node2(x_40, x_252, x_20, x_251); lean_inc(x_40); -x_250 = l_Lean_Syntax_node1(x_40, x_47, x_249); -x_251 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +x_254 = l_Lean_Syntax_node1(x_40, x_47, x_253); +x_255 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; lean_inc(x_40); -x_252 = l_Lean_Syntax_node4(x_40, x_251, x_135, x_168, x_170, x_250); -x_253 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +x_256 = l_Lean_Syntax_node4(x_40, x_255, x_135, x_170, x_172, x_254); +x_257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; lean_inc(x_49); lean_inc(x_40); -x_254 = l_Lean_Syntax_node6(x_40, x_253, x_99, x_101, x_49, x_49, x_133, x_252); +x_258 = l_Lean_Syntax_node6(x_40, x_257, x_99, x_101, x_49, x_49, x_133, x_256); lean_inc(x_40); -x_255 = l_Lean_Syntax_node2(x_40, x_88, x_51, x_254); -x_256 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +x_259 = l_Lean_Syntax_node2(x_40, x_88, x_51, x_258); +x_260 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; lean_inc(x_40); -x_257 = l_Lean_Syntax_node3(x_40, x_256, x_95, x_97, x_255); -x_258 = l_Lean_Syntax_node2(x_40, x_47, x_89, x_257); -lean_ctor_set(x_42, 0, x_258); +x_261 = l_Lean_Syntax_node3(x_40, x_260, x_95, x_97, x_259); +x_262 = l_Lean_Syntax_node2(x_40, x_47, x_89, x_261); +lean_ctor_set(x_42, 0, x_262); return x_42; } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; size_t x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; size_t x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; size_t x_399; lean_object* x_400; size_t x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; -x_259 = lean_ctor_get(x_42, 0); -x_260 = lean_ctor_get(x_42, 1); -lean_inc(x_260); -lean_inc(x_259); +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; size_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; size_t x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; size_t x_405; lean_object* x_406; size_t x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; +x_263 = lean_ctor_get(x_42, 0); +x_264 = lean_ctor_get(x_42, 1); +lean_inc(x_264); +lean_inc(x_263); lean_dec(x_42); -x_261 = lean_ctor_get(x_259, 0); -lean_inc(x_261); -lean_dec(x_259); -x_262 = lean_environment_main_module(x_261); -x_263 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -x_264 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +x_265 = lean_ctor_get(x_263, 0); +lean_inc(x_265); +lean_dec(x_263); +x_266 = lean_environment_main_module(x_265); +x_267 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +x_268 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_40); -x_265 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_265, 0, x_40); -lean_ctor_set(x_265, 1, x_263); -lean_ctor_set(x_265, 2, x_264); -x_266 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_265, 6); +x_269 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_269, 0, x_40); +lean_ctor_set(x_269, 1, x_267); +lean_ctor_set(x_269, 2, x_268); +x_270 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_269, 6); lean_inc(x_40); -x_267 = l_Lean_Syntax_node6(x_40, x_266, x_265, x_265, x_265, x_265, x_265, x_265); -x_268 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; +x_271 = l_Lean_Syntax_node6(x_40, x_270, x_269, x_269, x_269, x_269, x_269, x_269); +x_272 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; lean_inc(x_40); lean_ctor_set_tag(x_21, 2); -lean_ctor_set(x_21, 1, x_268); +lean_ctor_set(x_21, 1, x_272); lean_ctor_set(x_21, 0, x_40); -x_269 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; +x_273 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; lean_inc(x_40); -x_270 = l_Lean_Syntax_node1(x_40, x_269, x_21); -x_271 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +x_274 = l_Lean_Syntax_node1(x_40, x_273, x_21); +x_275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; lean_inc(x_41); -lean_inc(x_262); -x_272 = l_Lean_addMacroScope(x_262, x_271, x_41); -x_273 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_266); +x_276 = l_Lean_addMacroScope(x_266, x_275, x_41); +x_277 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; lean_inc(x_40); -x_274 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_274, 0, x_40); -lean_ctor_set(x_274, 1, x_273); -lean_ctor_set(x_274, 2, x_272); -lean_ctor_set(x_274, 3, x_13); -x_275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_265); -lean_inc(x_274); +x_278 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_278, 0, x_40); +lean_ctor_set(x_278, 1, x_277); +lean_ctor_set(x_278, 2, x_276); +lean_ctor_set(x_278, 3, x_13); +x_279 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_269); +lean_inc(x_278); lean_inc(x_40); -x_276 = l_Lean_Syntax_node2(x_40, x_275, x_274, x_265); -x_277 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +x_280 = l_Lean_Syntax_node2(x_40, x_279, x_278, x_269); +x_281 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; lean_inc(x_40); lean_ctor_set_tag(x_20, 2); -lean_ctor_set(x_20, 1, x_277); +lean_ctor_set(x_20, 1, x_281); lean_ctor_set(x_20, 0, x_40); -x_278 = l_Array_zip___rarg(x_32, x_33); +x_282 = l_Array_zip___rarg(x_32, x_33); lean_dec(x_33); lean_dec(x_32); -x_279 = lean_array_size(x_278); -lean_inc(x_267); -lean_inc(x_265); +x_283 = lean_array_size(x_282); +lean_inc(x_271); +lean_inc(x_269); lean_inc(x_40); -x_280 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_40, x_263, x_265, x_267, x_279, x_16, x_278); -x_281 = l_Array_append___rarg(x_264, x_280); -lean_dec(x_280); +x_284 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_40, x_267, x_269, x_271, x_283, x_16, x_282); +x_285 = l_Array_append___rarg(x_268, x_284); +lean_dec(x_284); lean_inc(x_40); -x_282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_282, 0, x_40); -lean_ctor_set(x_282, 1, x_263); -lean_ctor_set(x_282, 2, x_281); -x_283 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; +x_286 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_286, 0, x_40); +lean_ctor_set(x_286, 1, x_267); +lean_ctor_set(x_286, 2, x_285); +x_287 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; lean_inc(x_40); -x_284 = l_Lean_Syntax_node1(x_40, x_283, x_282); -lean_inc(x_265); +x_288 = l_Lean_Syntax_node1(x_40, x_287, x_286); +lean_inc(x_269); lean_inc(x_20); lean_inc(x_40); -x_285 = l_Lean_Syntax_node3(x_40, x_263, x_20, x_265, x_284); -x_286 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +x_289 = l_Lean_Syntax_node3(x_40, x_267, x_20, x_269, x_288); +x_290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; lean_inc(x_40); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_286); +lean_ctor_set(x_19, 1, x_290); lean_ctor_set(x_19, 0, x_40); -x_287 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +x_291 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; lean_inc(x_41); -lean_inc(x_262); -x_288 = l_Lean_addMacroScope(x_262, x_287, x_41); -x_289 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_266); +x_292 = l_Lean_addMacroScope(x_266, x_291, x_41); +x_293 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_294 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; lean_inc(x_40); -x_291 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_291, 0, x_40); -lean_ctor_set(x_291, 1, x_289); -lean_ctor_set(x_291, 2, x_288); -lean_ctor_set(x_291, 3, x_290); -x_292 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +x_295 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_295, 0, x_40); +lean_ctor_set(x_295, 1, x_293); +lean_ctor_set(x_295, 2, x_292); +lean_ctor_set(x_295, 3, x_294); +x_296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; lean_inc(x_40); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_292); +lean_ctor_set(x_18, 1, x_296); lean_ctor_set(x_18, 0, x_40); -x_293 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +x_297 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; lean_inc(x_41); -lean_inc(x_262); -x_294 = l_Lean_addMacroScope(x_262, x_293, x_41); -x_295 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_266); +x_298 = l_Lean_addMacroScope(x_266, x_297, x_41); +x_299 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; lean_inc(x_40); -x_297 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_297, 0, x_40); -lean_ctor_set(x_297, 1, x_295); -lean_ctor_set(x_297, 2, x_294); -lean_ctor_set(x_297, 3, x_296); +x_301 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_301, 0, x_40); +lean_ctor_set(x_301, 1, x_299); +lean_ctor_set(x_301, 2, x_298); +lean_ctor_set(x_301, 3, x_300); lean_inc(x_18); lean_inc(x_40); -x_298 = l_Lean_Syntax_node3(x_40, x_263, x_291, x_18, x_297); -lean_inc(x_40); -x_299 = l_Lean_Syntax_node2(x_40, x_263, x_19, x_298); -x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +x_302 = l_Lean_Syntax_node3(x_40, x_267, x_295, x_18, x_301); lean_inc(x_40); -x_301 = l_Lean_Syntax_node1(x_40, x_300, x_299); -x_302 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; -lean_inc_n(x_265, 3); +x_303 = l_Lean_Syntax_node2(x_40, x_267, x_19, x_302); +x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; lean_inc(x_40); -x_303 = l_Lean_Syntax_node7(x_40, x_302, x_270, x_276, x_265, x_265, x_265, x_285, x_301); -x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_267); +x_305 = l_Lean_Syntax_node1(x_40, x_304, x_303); +x_306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; +lean_inc_n(x_269, 3); lean_inc(x_40); -x_305 = l_Lean_Syntax_node2(x_40, x_304, x_267, x_303); -x_306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +x_307 = l_Lean_Syntax_node7(x_40, x_306, x_274, x_280, x_269, x_269, x_269, x_289, x_305); +x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_271); lean_inc(x_40); -x_307 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_307, 0, x_40); -lean_ctor_set(x_307, 1, x_306); -x_308 = l_Array_append___rarg(x_264, x_3); +x_309 = l_Lean_Syntax_node2(x_40, x_308, x_271, x_307); +x_310 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; lean_inc(x_40); -x_309 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_309, 0, x_40); -lean_ctor_set(x_309, 1, x_263); -lean_ctor_set(x_309, 2, x_308); -x_310 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +x_311 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_311, 0, x_40); +lean_ctor_set(x_311, 1, x_310); +x_312 = l_Array_append___rarg(x_268, x_3); lean_inc(x_40); -x_311 = l_Lean_Syntax_node2(x_40, x_310, x_307, x_309); -x_312 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_40); -x_313 = lean_alloc_ctor(2, 2, 0); +x_313 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_313, 0, x_40); -lean_ctor_set(x_313, 1, x_312); -x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_265); +lean_ctor_set(x_313, 1, x_267); +lean_ctor_set(x_313, 2, x_312); +x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; lean_inc(x_40); -x_315 = l_Lean_Syntax_node1(x_40, x_314, x_265); -x_316 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +x_315 = l_Lean_Syntax_node2(x_40, x_314, x_311, x_313); +x_316 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; lean_inc(x_40); x_317 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_317, 0, x_40); lean_ctor_set(x_317, 1, x_316); -x_318 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +x_318 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_269); lean_inc(x_40); -x_319 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_319, 0, x_40); -lean_ctor_set(x_319, 1, x_318); -x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_41); -lean_inc(x_262); -x_321 = l_Lean_addMacroScope(x_262, x_320, x_41); -x_322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_323 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +x_319 = l_Lean_Syntax_node1(x_40, x_318, x_269); +x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; lean_inc(x_40); -x_324 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_324, 0, x_40); -lean_ctor_set(x_324, 1, x_322); -lean_ctor_set(x_324, 2, x_321); -lean_ctor_set(x_324, 3, x_323); -x_325 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; +x_321 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_321, 0, x_40); +lean_ctor_set(x_321, 1, x_320); +x_322 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; lean_inc(x_40); -x_326 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_326, 0, x_40); -lean_ctor_set(x_326, 1, x_325); -x_327 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +x_323 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_323, 0, x_40); +lean_ctor_set(x_323, 1, x_322); +x_324 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_41); +lean_inc(x_266); +x_325 = l_Lean_addMacroScope(x_266, x_324, x_41); +x_326 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_327 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; lean_inc(x_40); -x_328 = lean_alloc_ctor(2, 2, 0); +x_328 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_328, 0, x_40); -lean_ctor_set(x_328, 1, x_327); -x_329 = lean_box(0); -x_330 = l_Lean_mkCIdentFrom(x_329, x_4, x_39); -x_331 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_ctor_set(x_328, 1, x_326); +lean_ctor_set(x_328, 2, x_325); +lean_ctor_set(x_328, 3, x_327); +x_329 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; lean_inc(x_40); -x_332 = l_Lean_Syntax_node2(x_40, x_331, x_328, x_330); -x_333 = lean_array_size(x_36); -x_334 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; -x_335 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_334, x_333, x_16, x_36); -x_336 = l_Array_append___rarg(x_264, x_335); -lean_dec(x_335); +x_330 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_330, 0, x_40); +lean_ctor_set(x_330, 1, x_329); +x_331 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; lean_inc(x_40); -x_337 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_337, 0, x_40); -lean_ctor_set(x_337, 1, x_263); -lean_ctor_set(x_337, 2, x_336); -x_338 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +x_332 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_332, 0, x_40); +lean_ctor_set(x_332, 1, x_331); +x_333 = lean_box(0); +x_334 = l_Lean_mkCIdentFrom(x_333, x_4, x_39); +x_335 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; lean_inc(x_40); -x_339 = l_Lean_Syntax_node2(x_40, x_338, x_332, x_337); -x_340 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; +x_336 = l_Lean_Syntax_node2(x_40, x_335, x_332, x_334); +x_337 = lean_array_size(x_36); +x_338 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; +x_339 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_338, x_337, x_16, x_36); +x_340 = l_Array_append___rarg(x_268, x_339); +lean_dec(x_339); lean_inc(x_40); -x_341 = lean_alloc_ctor(2, 2, 0); +x_341 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_341, 0, x_40); -lean_ctor_set(x_341, 1, x_340); -x_342 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; +lean_ctor_set(x_341, 1, x_267); +lean_ctor_set(x_341, 2, x_340); +x_342 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; lean_inc(x_40); -x_343 = l_Lean_Syntax_node3(x_40, x_342, x_326, x_339, x_341); +x_343 = l_Lean_Syntax_node2(x_40, x_342, x_336, x_341); +x_344 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; lean_inc(x_40); -x_344 = l_Lean_Syntax_node1(x_40, x_263, x_343); +x_345 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_345, 0, x_40); +lean_ctor_set(x_345, 1, x_344); +x_346 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; lean_inc(x_40); -x_345 = l_Lean_Syntax_node2(x_40, x_338, x_324, x_344); -x_346 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_319); +x_347 = l_Lean_Syntax_node3(x_40, x_346, x_330, x_343, x_345); lean_inc(x_40); -x_347 = l_Lean_Syntax_node2(x_40, x_346, x_319, x_345); -x_348 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_265); +x_348 = l_Lean_Syntax_node1(x_40, x_267, x_347); lean_inc(x_40); -x_349 = l_Lean_Syntax_node2(x_40, x_348, x_265, x_347); -x_350 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +x_349 = l_Lean_Syntax_node2(x_40, x_342, x_328, x_348); +x_350 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_323); lean_inc(x_40); -x_351 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_351, 0, x_40); -lean_ctor_set(x_351, 1, x_350); -x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +x_351 = l_Lean_Syntax_node2(x_40, x_350, x_323, x_349); +x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_269); lean_inc(x_40); -x_353 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_353, 0, x_40); -lean_ctor_set(x_353, 1, x_352); -x_354 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +x_353 = l_Lean_Syntax_node2(x_40, x_352, x_269, x_351); +x_354 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_40); +x_355 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_355, 0, x_40); +lean_ctor_set(x_355, 1, x_354); +x_356 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_40); +x_357 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_357, 0, x_40); +lean_ctor_set(x_357, 1, x_356); +x_358 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; lean_inc(x_41); -lean_inc(x_262); -x_355 = l_Lean_addMacroScope(x_262, x_354, x_41); -x_356 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_357 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_266); +x_359 = l_Lean_addMacroScope(x_266, x_358, x_41); +x_360 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_361 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; lean_inc(x_40); -x_358 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_358, 0, x_40); -lean_ctor_set(x_358, 1, x_356); -lean_ctor_set(x_358, 2, x_355); -lean_ctor_set(x_358, 3, x_357); -x_359 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_265); +x_362 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_362, 0, x_40); +lean_ctor_set(x_362, 1, x_360); +lean_ctor_set(x_362, 2, x_359); +lean_ctor_set(x_362, 3, x_361); +x_363 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_269); lean_inc(x_40); -x_360 = l_Lean_Syntax_node2(x_40, x_359, x_358, x_265); -x_361 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; +x_364 = l_Lean_Syntax_node2(x_40, x_363, x_362, x_269); +x_365 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; lean_inc(x_41); -lean_inc(x_262); -x_362 = l_Lean_addMacroScope(x_262, x_361, x_41); -x_363 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_266); +x_366 = l_Lean_addMacroScope(x_266, x_365, x_41); +x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; lean_inc(x_40); -x_364 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_364, 0, x_40); -lean_ctor_set(x_364, 1, x_363); -lean_ctor_set(x_364, 2, x_362); -lean_ctor_set(x_364, 3, x_13); -x_365 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_364); -lean_inc(x_351); +x_368 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_368, 0, x_40); +lean_ctor_set(x_368, 1, x_367); +lean_ctor_set(x_368, 2, x_366); +lean_ctor_set(x_368, 3, x_13); +x_369 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_368); +lean_inc(x_355); lean_inc(x_40); -x_366 = l_Lean_Syntax_node3(x_40, x_365, x_360, x_351, x_364); -x_367 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +x_370 = l_Lean_Syntax_node3(x_40, x_369, x_364, x_355, x_368); +x_371 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; lean_inc(x_41); -lean_inc(x_262); -x_368 = l_Lean_addMacroScope(x_262, x_367, x_41); -x_369 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_370 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_266); +x_372 = l_Lean_addMacroScope(x_266, x_371, x_41); +x_373 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_374 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; lean_inc(x_40); -x_371 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_371, 0, x_40); -lean_ctor_set(x_371, 1, x_369); -lean_ctor_set(x_371, 2, x_368); -lean_ctor_set(x_371, 3, x_370); -lean_inc(x_265); +x_375 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_375, 0, x_40); +lean_ctor_set(x_375, 1, x_373); +lean_ctor_set(x_375, 2, x_372); +lean_ctor_set(x_375, 3, x_374); +lean_inc(x_269); lean_inc(x_40); -x_372 = l_Lean_Syntax_node2(x_40, x_359, x_371, x_265); -x_373 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; +x_376 = l_Lean_Syntax_node2(x_40, x_363, x_375, x_269); +x_377 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; lean_inc(x_41); -lean_inc(x_262); -x_374 = l_Lean_addMacroScope(x_262, x_373, x_41); -x_375 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_266); +x_378 = l_Lean_addMacroScope(x_266, x_377, x_41); +x_379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; lean_inc(x_40); -x_376 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_376, 0, x_40); -lean_ctor_set(x_376, 1, x_375); -lean_ctor_set(x_376, 2, x_374); -lean_ctor_set(x_376, 3, x_13); -lean_inc(x_376); -lean_inc(x_351); +x_380 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_380, 0, x_40); +lean_ctor_set(x_380, 1, x_379); +lean_ctor_set(x_380, 2, x_378); +lean_ctor_set(x_380, 3, x_13); +lean_inc(x_380); +lean_inc(x_355); lean_inc(x_40); -x_377 = l_Lean_Syntax_node3(x_40, x_365, x_372, x_351, x_376); +x_381 = l_Lean_Syntax_node3(x_40, x_369, x_376, x_355, x_380); lean_inc(x_40); -x_378 = l_Lean_Syntax_node3(x_40, x_263, x_366, x_18, x_377); -x_379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_265); +x_382 = l_Lean_Syntax_node3(x_40, x_267, x_370, x_18, x_381); +x_383 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; lean_inc(x_40); -x_380 = l_Lean_Syntax_node1(x_40, x_379, x_265); -x_381 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; +x_384 = l_Lean_Syntax_node1(x_40, x_383, x_382); +x_385 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_269); lean_inc(x_40); -x_382 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_382, 0, x_40); -lean_ctor_set(x_382, 1, x_381); -x_383 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc(x_382); -lean_inc(x_380); -lean_inc_n(x_265, 2); -lean_inc(x_353); +x_386 = l_Lean_Syntax_node1(x_40, x_385, x_269); +x_387 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_40); +x_388 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_388, 0, x_40); +lean_ctor_set(x_388, 1, x_387); +x_389 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_388); +lean_inc(x_386); +lean_inc_n(x_269, 2); +lean_inc(x_357); lean_inc(x_40); -x_384 = l_Lean_Syntax_node6(x_40, x_383, x_353, x_265, x_378, x_380, x_265, x_382); -x_385 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_265, 2); +x_390 = l_Lean_Syntax_node6(x_40, x_389, x_357, x_269, x_384, x_386, x_269, x_388); +x_391 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_269, 2); lean_inc(x_40); -x_386 = l_Lean_Syntax_node2(x_40, x_385, x_265, x_265); -x_387 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +x_392 = l_Lean_Syntax_node2(x_40, x_391, x_269, x_269); +x_393 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; lean_inc(x_41); -lean_inc(x_262); -x_388 = l_Lean_addMacroScope(x_262, x_387, x_41); -x_389 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +lean_inc(x_266); +x_394 = l_Lean_addMacroScope(x_266, x_393, x_41); +x_395 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; lean_inc(x_40); -x_390 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_390, 0, x_40); -lean_ctor_set(x_390, 1, x_389); -lean_ctor_set(x_390, 2, x_388); -lean_ctor_set(x_390, 3, x_13); -lean_inc(x_390); +x_396 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_396, 0, x_40); +lean_ctor_set(x_396, 1, x_395); +lean_ctor_set(x_396, 2, x_394); +lean_ctor_set(x_396, 3, x_13); +lean_inc(x_396); lean_inc(x_40); -x_391 = l_Lean_Syntax_node1(x_40, x_263, x_390); -x_392 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_397 = l_Lean_Syntax_node1(x_40, x_267, x_396); +x_398 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_40); -x_393 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_393, 0, x_40); -lean_ctor_set(x_393, 1, x_392); -x_394 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_399 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_399, 0, x_40); +lean_ctor_set(x_399, 1, x_398); +x_400 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_41); -lean_inc(x_262); -x_395 = l_Lean_addMacroScope(x_262, x_394, x_41); -x_396 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +lean_inc(x_266); +x_401 = l_Lean_addMacroScope(x_266, x_400, x_41); +x_402 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_403 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_40); -x_398 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_398, 0, x_40); -lean_ctor_set(x_398, 1, x_396); -lean_ctor_set(x_398, 2, x_395); -lean_ctor_set(x_398, 3, x_397); -x_399 = lean_array_size(x_29); -x_400 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_365, x_399, x_16, x_29); -x_401 = lean_array_size(x_400); -x_402 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_401, x_16, x_400); -x_403 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_404 = l_Lean_mkSepArray(x_402, x_403); -lean_dec(x_402); -x_405 = l_Array_append___rarg(x_264, x_404); -lean_dec(x_404); +x_404 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_404, 0, x_40); +lean_ctor_set(x_404, 1, x_402); +lean_ctor_set(x_404, 2, x_401); +lean_ctor_set(x_404, 3, x_403); +x_405 = lean_array_size(x_29); +x_406 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_369, x_405, x_16, x_29); +x_407 = lean_array_size(x_406); +x_408 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_407, x_16, x_406); +x_409 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_410 = l_Lean_mkSepArray(x_408, x_409); +lean_dec(x_408); +x_411 = l_Array_append___rarg(x_268, x_410); +lean_dec(x_410); lean_inc(x_40); -x_406 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_406, 0, x_40); -lean_ctor_set(x_406, 1, x_263); -lean_ctor_set(x_406, 2, x_405); -lean_inc(x_274); -lean_inc(x_319); +x_412 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_412, 0, x_40); +lean_ctor_set(x_412, 1, x_267); +lean_ctor_set(x_412, 2, x_411); lean_inc(x_40); -x_407 = l_Lean_Syntax_node2(x_40, x_263, x_319, x_274); -lean_inc(x_382); -lean_inc(x_380); -lean_inc(x_265); -lean_inc(x_353); +x_413 = l_Lean_Syntax_node1(x_40, x_383, x_412); +lean_inc(x_278); +lean_inc(x_323); lean_inc(x_40); -x_408 = l_Lean_Syntax_node6(x_40, x_383, x_353, x_265, x_406, x_380, x_407, x_382); +x_414 = l_Lean_Syntax_node2(x_40, x_267, x_323, x_278); +lean_inc(x_388); +lean_inc(x_386); +lean_inc(x_269); +lean_inc(x_357); lean_inc(x_40); -x_409 = l_Lean_Syntax_node1(x_40, x_263, x_408); +x_415 = l_Lean_Syntax_node6(x_40, x_389, x_357, x_269, x_413, x_386, x_414, x_388); lean_inc(x_40); -x_410 = l_Lean_Syntax_node2(x_40, x_338, x_398, x_409); +x_416 = l_Lean_Syntax_node1(x_40, x_267, x_415); lean_inc(x_40); -x_411 = l_Lean_Syntax_node1(x_40, x_263, x_410); -x_412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_393); +x_417 = l_Lean_Syntax_node2(x_40, x_342, x_404, x_416); lean_inc(x_40); -x_413 = l_Lean_Syntax_node2(x_40, x_412, x_393, x_411); -x_414 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_351); -lean_inc(x_265); +x_418 = l_Lean_Syntax_node1(x_40, x_267, x_417); +x_419 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_399); lean_inc(x_40); -x_415 = l_Lean_Syntax_node5(x_40, x_414, x_364, x_391, x_265, x_351, x_413); -x_416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +x_420 = l_Lean_Syntax_node2(x_40, x_419, x_399, x_418); +x_421 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_355); +lean_inc(x_269); lean_inc(x_40); -x_417 = l_Lean_Syntax_node1(x_40, x_416, x_415); -x_418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_386); -lean_inc_n(x_265, 2); +x_422 = l_Lean_Syntax_node5(x_40, x_421, x_368, x_397, x_269, x_355, x_420); +x_423 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; lean_inc(x_40); -x_419 = l_Lean_Syntax_node4(x_40, x_418, x_265, x_265, x_417, x_386); -x_420 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; +x_424 = l_Lean_Syntax_node1(x_40, x_423, x_422); +x_425 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_392); +lean_inc_n(x_269, 2); +lean_inc(x_40); +x_426 = l_Lean_Syntax_node4(x_40, x_425, x_269, x_269, x_424, x_392); +x_427 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_41); -lean_inc(x_262); -x_421 = l_Lean_addMacroScope(x_262, x_420, x_41); -x_422 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +lean_inc(x_266); +x_428 = l_Lean_addMacroScope(x_266, x_427, x_41); +x_429 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; lean_inc(x_40); -x_423 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_423, 0, x_40); -lean_ctor_set(x_423, 1, x_422); -lean_ctor_set(x_423, 2, x_421); -lean_ctor_set(x_423, 3, x_13); +x_430 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_430, 0, x_40); +lean_ctor_set(x_430, 1, x_429); +lean_ctor_set(x_430, 2, x_428); +lean_ctor_set(x_430, 3, x_13); lean_inc(x_40); -x_424 = l_Lean_Syntax_node1(x_40, x_263, x_423); -x_425 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +x_431 = l_Lean_Syntax_node1(x_40, x_267, x_430); +x_432 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; lean_inc(x_40); -x_426 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_426, 0, x_40); -lean_ctor_set(x_426, 1, x_425); -x_427 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +x_433 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_433, 0, x_40); +lean_ctor_set(x_433, 1, x_432); +x_434 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; lean_inc(x_40); -x_428 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_428, 0, x_40); -lean_ctor_set(x_428, 1, x_427); +x_435 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_435, 0, x_40); +lean_ctor_set(x_435, 1, x_434); lean_inc(x_40); -x_429 = l_Lean_Syntax_node2(x_40, x_346, x_319, x_274); +x_436 = l_Lean_Syntax_node2(x_40, x_350, x_323, x_278); lean_inc(x_40); -x_430 = l_Lean_Syntax_node1(x_40, x_263, x_429); -x_431 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_437 = l_Lean_Syntax_node1(x_40, x_267, x_436); +x_438 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_40); -x_432 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_432, 0, x_40); -lean_ctor_set(x_432, 1, x_431); -x_433 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_434 = l_Lean_addMacroScope(x_262, x_433, x_41); -x_435 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_436 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; +x_439 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_439, 0, x_40); +lean_ctor_set(x_439, 1, x_438); +x_440 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_441 = l_Lean_addMacroScope(x_266, x_440, x_41); +x_442 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; lean_inc(x_40); -x_437 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_437, 0, x_40); -lean_ctor_set(x_437, 1, x_435); -lean_ctor_set(x_437, 2, x_434); -lean_ctor_set(x_437, 3, x_436); -lean_inc(x_424); +x_444 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_444, 0, x_40); +lean_ctor_set(x_444, 1, x_442); +lean_ctor_set(x_444, 2, x_441); +lean_ctor_set(x_444, 3, x_443); +lean_inc(x_431); lean_inc(x_40); -x_438 = l_Lean_Syntax_node2(x_40, x_338, x_437, x_424); -x_439 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_445 = l_Lean_Syntax_node2(x_40, x_342, x_444, x_431); +x_446 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; lean_inc(x_40); -x_440 = l_Lean_Syntax_node1(x_40, x_439, x_438); -x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +x_447 = l_Lean_Syntax_node1(x_40, x_446, x_445); +x_448 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; lean_inc(x_40); -x_442 = l_Lean_Syntax_node4(x_40, x_441, x_390, x_430, x_432, x_440); -x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_265); +x_449 = l_Lean_Syntax_node4(x_40, x_448, x_396, x_437, x_439, x_447); +x_450 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_269); lean_inc(x_40); -x_444 = l_Lean_Syntax_node3(x_40, x_443, x_428, x_265, x_442); -x_445 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_265); +x_451 = l_Lean_Syntax_node3(x_40, x_450, x_435, x_269, x_449); +x_452 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_269); lean_inc(x_40); -x_446 = l_Lean_Syntax_node2(x_40, x_445, x_444, x_265); -x_447 = l_Lean_Syntax_SepArray_ofElems(x_292, x_26); +x_453 = l_Lean_Syntax_node2(x_40, x_452, x_451, x_269); +x_454 = l_Lean_Syntax_SepArray_ofElems(x_296, x_26); lean_dec(x_26); -x_448 = l_Array_append___rarg(x_264, x_447); -lean_dec(x_447); +x_455 = l_Array_append___rarg(x_268, x_454); +lean_dec(x_454); lean_inc(x_40); -x_449 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_449, 0, x_40); -lean_ctor_set(x_449, 1, x_263); -lean_ctor_set(x_449, 2, x_448); -lean_inc_n(x_265, 2); +x_456 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_456, 0, x_40); +lean_ctor_set(x_456, 1, x_267); +lean_ctor_set(x_456, 2, x_455); lean_inc(x_40); -x_450 = l_Lean_Syntax_node6(x_40, x_383, x_353, x_265, x_449, x_380, x_265, x_382); +x_457 = l_Lean_Syntax_node1(x_40, x_383, x_456); +lean_inc_n(x_269, 2); lean_inc(x_40); -x_451 = l_Lean_Syntax_node1(x_40, x_263, x_450); -x_452 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +x_458 = l_Lean_Syntax_node6(x_40, x_389, x_357, x_269, x_457, x_386, x_269, x_388); lean_inc(x_40); -x_453 = l_Lean_Syntax_node2(x_40, x_452, x_393, x_451); -lean_inc(x_265); +x_459 = l_Lean_Syntax_node1(x_40, x_267, x_458); +x_460 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; lean_inc(x_40); -x_454 = l_Lean_Syntax_node2(x_40, x_445, x_453, x_265); +x_461 = l_Lean_Syntax_node2(x_40, x_460, x_399, x_459); +lean_inc(x_269); lean_inc(x_40); -x_455 = l_Lean_Syntax_node2(x_40, x_263, x_446, x_454); -x_456 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_462 = l_Lean_Syntax_node2(x_40, x_452, x_461, x_269); lean_inc(x_40); -x_457 = l_Lean_Syntax_node1(x_40, x_456, x_455); -x_458 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_463 = l_Lean_Syntax_node2(x_40, x_267, x_453, x_462); +x_464 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_40); -x_459 = l_Lean_Syntax_node2(x_40, x_458, x_426, x_457); -lean_inc(x_351); -lean_inc(x_265); +x_465 = l_Lean_Syntax_node1(x_40, x_464, x_463); +x_466 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_40); -x_460 = l_Lean_Syntax_node5(x_40, x_414, x_376, x_424, x_265, x_351, x_459); +x_467 = l_Lean_Syntax_node2(x_40, x_466, x_433, x_465); +lean_inc(x_355); +lean_inc(x_269); lean_inc(x_40); -x_461 = l_Lean_Syntax_node1(x_40, x_416, x_460); -lean_inc(x_386); -lean_inc_n(x_265, 2); +x_468 = l_Lean_Syntax_node5(x_40, x_421, x_380, x_431, x_269, x_355, x_467); lean_inc(x_40); -x_462 = l_Lean_Syntax_node4(x_40, x_418, x_265, x_265, x_461, x_386); -lean_inc(x_265); +x_469 = l_Lean_Syntax_node1(x_40, x_423, x_468); +lean_inc(x_392); +lean_inc_n(x_269, 2); lean_inc(x_40); -x_463 = l_Lean_Syntax_node3(x_40, x_263, x_419, x_265, x_462); -x_464 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_470 = l_Lean_Syntax_node4(x_40, x_425, x_269, x_269, x_469, x_392); +lean_inc(x_269); lean_inc(x_40); -x_465 = l_Lean_Syntax_node2(x_40, x_464, x_20, x_463); +x_471 = l_Lean_Syntax_node3(x_40, x_267, x_426, x_269, x_470); +x_472 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_40); -x_466 = l_Lean_Syntax_node1(x_40, x_263, x_465); -x_467 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +x_473 = l_Lean_Syntax_node2(x_40, x_472, x_20, x_471); lean_inc(x_40); -x_468 = l_Lean_Syntax_node4(x_40, x_467, x_351, x_384, x_386, x_466); -x_469 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_265); +x_474 = l_Lean_Syntax_node1(x_40, x_267, x_473); +x_475 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; lean_inc(x_40); -x_470 = l_Lean_Syntax_node6(x_40, x_469, x_315, x_317, x_265, x_265, x_349, x_468); +x_476 = l_Lean_Syntax_node4(x_40, x_475, x_355, x_390, x_392, x_474); +x_477 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_269); lean_inc(x_40); -x_471 = l_Lean_Syntax_node2(x_40, x_304, x_267, x_470); -x_472 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +x_478 = l_Lean_Syntax_node6(x_40, x_477, x_319, x_321, x_269, x_269, x_353, x_476); lean_inc(x_40); -x_473 = l_Lean_Syntax_node3(x_40, x_472, x_311, x_313, x_471); -x_474 = l_Lean_Syntax_node2(x_40, x_263, x_305, x_473); -x_475 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_475, 0, x_474); -lean_ctor_set(x_475, 1, x_260); -return x_475; +x_479 = l_Lean_Syntax_node2(x_40, x_308, x_271, x_478); +x_480 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_40); +x_481 = l_Lean_Syntax_node3(x_40, x_480, x_315, x_317, x_479); +x_482 = l_Lean_Syntax_node2(x_40, x_267, x_309, x_481); +x_483 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_483, 0, x_482); +lean_ctor_set(x_483, 1, x_264); +return x_483; } } else { -uint8_t x_476; +uint8_t x_484; lean_free_object(x_21); lean_dec(x_33); lean_dec(x_32); @@ -5153,578 +5189,585 @@ lean_dec(x_26); lean_free_object(x_18); lean_dec(x_10); lean_dec(x_4); -x_476 = !lean_is_exclusive(x_35); -if (x_476 == 0) +x_484 = !lean_is_exclusive(x_35); +if (x_484 == 0) { return x_35; } else { -lean_object* x_477; lean_object* x_478; lean_object* x_479; -x_477 = lean_ctor_get(x_35, 0); -x_478 = lean_ctor_get(x_35, 1); -lean_inc(x_478); -lean_inc(x_477); +lean_object* x_485; lean_object* x_486; lean_object* x_487; +x_485 = lean_ctor_get(x_35, 0); +x_486 = lean_ctor_get(x_35, 1); +lean_inc(x_486); +lean_inc(x_485); lean_dec(x_35); -x_479 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_479, 0, x_477); -lean_ctor_set(x_479, 1, x_478); -return x_479; +x_487 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_487, 0, x_485); +lean_ctor_set(x_487, 1, x_486); +return x_487; } } } else { -lean_object* x_480; lean_object* x_481; size_t x_482; lean_object* x_483; -x_480 = lean_ctor_get(x_21, 0); -x_481 = lean_ctor_get(x_21, 1); -lean_inc(x_481); -lean_inc(x_480); -lean_dec(x_21); -x_482 = lean_array_size(x_2); -x_483 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_482, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -if (lean_obj_tag(x_483) == 0) -{ -lean_object* x_484; lean_object* x_485; lean_object* x_486; uint8_t x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; size_t x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; size_t x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; size_t x_633; lean_object* x_634; size_t x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; -x_484 = lean_ctor_get(x_483, 0); -lean_inc(x_484); -x_485 = lean_ctor_get(x_483, 1); -lean_inc(x_485); -lean_dec(x_483); -x_486 = lean_ctor_get(x_10, 5); -lean_inc(x_486); -x_487 = 0; -x_488 = l_Lean_SourceInfo_fromRef(x_486, x_487); -lean_dec(x_486); -x_489 = lean_ctor_get(x_10, 10); +lean_object* x_488; lean_object* x_489; size_t x_490; lean_object* x_491; +x_488 = lean_ctor_get(x_21, 0); +x_489 = lean_ctor_get(x_21, 1); lean_inc(x_489); -lean_dec(x_10); -x_490 = lean_st_ref_get(x_11, x_485); -x_491 = lean_ctor_get(x_490, 0); -lean_inc(x_491); -x_492 = lean_ctor_get(x_490, 1); -lean_inc(x_492); -if (lean_is_exclusive(x_490)) { - lean_ctor_release(x_490, 0); - lean_ctor_release(x_490, 1); - x_493 = x_490; -} else { - lean_dec_ref(x_490); - x_493 = lean_box(0); -} -x_494 = lean_ctor_get(x_491, 0); -lean_inc(x_494); -lean_dec(x_491); -x_495 = lean_environment_main_module(x_494); -x_496 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -x_497 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_488); -x_498 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_498, 0, x_488); -lean_ctor_set(x_498, 1, x_496); -lean_ctor_set(x_498, 2, x_497); -x_499 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_498, 6); -lean_inc(x_488); -x_500 = l_Lean_Syntax_node6(x_488, x_499, x_498, x_498, x_498, x_498, x_498, x_498); -x_501 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; -lean_inc(x_488); -x_502 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_502, 0, x_488); -lean_ctor_set(x_502, 1, x_501); -x_503 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; -lean_inc(x_488); -x_504 = l_Lean_Syntax_node1(x_488, x_503, x_502); -x_505 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_489); -lean_inc(x_495); -x_506 = l_Lean_addMacroScope(x_495, x_505, x_489); -x_507 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_488); -x_508 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_508, 0, x_488); -lean_ctor_set(x_508, 1, x_507); -lean_ctor_set(x_508, 2, x_506); -lean_ctor_set(x_508, 3, x_13); -x_509 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_498); -lean_inc(x_508); -lean_inc(x_488); -x_510 = l_Lean_Syntax_node2(x_488, x_509, x_508, x_498); -x_511 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_488); -lean_ctor_set_tag(x_20, 2); -lean_ctor_set(x_20, 1, x_511); -lean_ctor_set(x_20, 0, x_488); -x_512 = l_Array_zip___rarg(x_480, x_481); -lean_dec(x_481); -lean_dec(x_480); -x_513 = lean_array_size(x_512); +lean_dec(x_21); +x_490 = lean_array_size(x_2); +x_491 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_490, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_491) == 0) +{ +lean_object* x_492; lean_object* x_493; lean_object* x_494; uint8_t x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; size_t x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; size_t x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; size_t x_643; lean_object* x_644; size_t x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; +x_492 = lean_ctor_get(x_491, 0); +lean_inc(x_492); +x_493 = lean_ctor_get(x_491, 1); +lean_inc(x_493); +lean_dec(x_491); +x_494 = lean_ctor_get(x_10, 5); +lean_inc(x_494); +x_495 = 0; +x_496 = l_Lean_SourceInfo_fromRef(x_494, x_495); +lean_dec(x_494); +x_497 = lean_ctor_get(x_10, 10); +lean_inc(x_497); +lean_dec(x_10); +x_498 = lean_st_ref_get(x_11, x_493); +x_499 = lean_ctor_get(x_498, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_498, 1); lean_inc(x_500); -lean_inc(x_498); -lean_inc(x_488); -x_514 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_488, x_496, x_498, x_500, x_513, x_16, x_512); -x_515 = l_Array_append___rarg(x_497, x_514); -lean_dec(x_514); -lean_inc(x_488); -x_516 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_516, 0, x_488); -lean_ctor_set(x_516, 1, x_496); -lean_ctor_set(x_516, 2, x_515); -x_517 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_488); -x_518 = l_Lean_Syntax_node1(x_488, x_517, x_516); -lean_inc(x_498); +if (lean_is_exclusive(x_498)) { + lean_ctor_release(x_498, 0); + lean_ctor_release(x_498, 1); + x_501 = x_498; +} else { + lean_dec_ref(x_498); + x_501 = lean_box(0); +} +x_502 = lean_ctor_get(x_499, 0); +lean_inc(x_502); +lean_dec(x_499); +x_503 = lean_environment_main_module(x_502); +x_504 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +x_505 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +lean_inc(x_496); +x_506 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_506, 0, x_496); +lean_ctor_set(x_506, 1, x_504); +lean_ctor_set(x_506, 2, x_505); +x_507 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_506, 6); +lean_inc(x_496); +x_508 = l_Lean_Syntax_node6(x_496, x_507, x_506, x_506, x_506, x_506, x_506, x_506); +x_509 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; +lean_inc(x_496); +x_510 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_510, 0, x_496); +lean_ctor_set(x_510, 1, x_509); +x_511 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; +lean_inc(x_496); +x_512 = l_Lean_Syntax_node1(x_496, x_511, x_510); +x_513 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_497); +lean_inc(x_503); +x_514 = l_Lean_addMacroScope(x_503, x_513, x_497); +x_515 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_496); +x_516 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_516, 0, x_496); +lean_ctor_set(x_516, 1, x_515); +lean_ctor_set(x_516, 2, x_514); +lean_ctor_set(x_516, 3, x_13); +x_517 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_506); +lean_inc(x_516); +lean_inc(x_496); +x_518 = l_Lean_Syntax_node2(x_496, x_517, x_516, x_506); +x_519 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_496); +lean_ctor_set_tag(x_20, 2); +lean_ctor_set(x_20, 1, x_519); +lean_ctor_set(x_20, 0, x_496); +x_520 = l_Array_zip___rarg(x_488, x_489); +lean_dec(x_489); +lean_dec(x_488); +x_521 = lean_array_size(x_520); +lean_inc(x_508); +lean_inc(x_506); +lean_inc(x_496); +x_522 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_496, x_504, x_506, x_508, x_521, x_16, x_520); +x_523 = l_Array_append___rarg(x_505, x_522); +lean_dec(x_522); +lean_inc(x_496); +x_524 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_524, 0, x_496); +lean_ctor_set(x_524, 1, x_504); +lean_ctor_set(x_524, 2, x_523); +x_525 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; +lean_inc(x_496); +x_526 = l_Lean_Syntax_node1(x_496, x_525, x_524); +lean_inc(x_506); lean_inc(x_20); -lean_inc(x_488); -x_519 = l_Lean_Syntax_node3(x_488, x_496, x_20, x_498, x_518); -x_520 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_488); +lean_inc(x_496); +x_527 = l_Lean_Syntax_node3(x_496, x_504, x_20, x_506, x_526); +x_528 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_496); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_520); -lean_ctor_set(x_19, 0, x_488); -x_521 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_489); -lean_inc(x_495); -x_522 = l_Lean_addMacroScope(x_495, x_521, x_489); -x_523 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_524 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_488); -x_525 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_525, 0, x_488); -lean_ctor_set(x_525, 1, x_523); -lean_ctor_set(x_525, 2, x_522); -lean_ctor_set(x_525, 3, x_524); -x_526 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_488); +lean_ctor_set(x_19, 1, x_528); +lean_ctor_set(x_19, 0, x_496); +x_529 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_497); +lean_inc(x_503); +x_530 = l_Lean_addMacroScope(x_503, x_529, x_497); +x_531 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_532 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_496); +x_533 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_533, 0, x_496); +lean_ctor_set(x_533, 1, x_531); +lean_ctor_set(x_533, 2, x_530); +lean_ctor_set(x_533, 3, x_532); +x_534 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_496); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_526); -lean_ctor_set(x_18, 0, x_488); -x_527 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_489); -lean_inc(x_495); -x_528 = l_Lean_addMacroScope(x_495, x_527, x_489); -x_529 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_530 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_488); -x_531 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_531, 0, x_488); -lean_ctor_set(x_531, 1, x_529); -lean_ctor_set(x_531, 2, x_528); -lean_ctor_set(x_531, 3, x_530); +lean_ctor_set(x_18, 1, x_534); +lean_ctor_set(x_18, 0, x_496); +x_535 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_497); +lean_inc(x_503); +x_536 = l_Lean_addMacroScope(x_503, x_535, x_497); +x_537 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_538 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_496); +x_539 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_539, 0, x_496); +lean_ctor_set(x_539, 1, x_537); +lean_ctor_set(x_539, 2, x_536); +lean_ctor_set(x_539, 3, x_538); lean_inc(x_18); -lean_inc(x_488); -x_532 = l_Lean_Syntax_node3(x_488, x_496, x_525, x_18, x_531); -lean_inc(x_488); -x_533 = l_Lean_Syntax_node2(x_488, x_496, x_19, x_532); -x_534 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_488); -x_535 = l_Lean_Syntax_node1(x_488, x_534, x_533); -x_536 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; -lean_inc_n(x_498, 3); -lean_inc(x_488); -x_537 = l_Lean_Syntax_node7(x_488, x_536, x_504, x_510, x_498, x_498, x_498, x_519, x_535); -x_538 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_500); -lean_inc(x_488); -x_539 = l_Lean_Syntax_node2(x_488, x_538, x_500, x_537); -x_540 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_488); -x_541 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_541, 0, x_488); -lean_ctor_set(x_541, 1, x_540); -x_542 = l_Array_append___rarg(x_497, x_3); -lean_inc(x_488); -x_543 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_543, 0, x_488); -lean_ctor_set(x_543, 1, x_496); -lean_ctor_set(x_543, 2, x_542); -x_544 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_488); -x_545 = l_Lean_Syntax_node2(x_488, x_544, x_541, x_543); -x_546 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_488); -x_547 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_547, 0, x_488); -lean_ctor_set(x_547, 1, x_546); -x_548 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_498); -lean_inc(x_488); -x_549 = l_Lean_Syntax_node1(x_488, x_548, x_498); -x_550 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_488); -x_551 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_551, 0, x_488); -lean_ctor_set(x_551, 1, x_550); -x_552 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_488); -x_553 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_553, 0, x_488); -lean_ctor_set(x_553, 1, x_552); -x_554 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_489); -lean_inc(x_495); -x_555 = l_Lean_addMacroScope(x_495, x_554, x_489); -x_556 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_557 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_488); -x_558 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_558, 0, x_488); -lean_ctor_set(x_558, 1, x_556); -lean_ctor_set(x_558, 2, x_555); -lean_ctor_set(x_558, 3, x_557); -x_559 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; -lean_inc(x_488); -x_560 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_560, 0, x_488); -lean_ctor_set(x_560, 1, x_559); -x_561 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_488); -x_562 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_562, 0, x_488); -lean_ctor_set(x_562, 1, x_561); -x_563 = lean_box(0); -x_564 = l_Lean_mkCIdentFrom(x_563, x_4, x_487); -x_565 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_488); -x_566 = l_Lean_Syntax_node2(x_488, x_565, x_562, x_564); -x_567 = lean_array_size(x_484); -x_568 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; -x_569 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_568, x_567, x_16, x_484); -x_570 = l_Array_append___rarg(x_497, x_569); -lean_dec(x_569); -lean_inc(x_488); -x_571 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_571, 0, x_488); -lean_ctor_set(x_571, 1, x_496); -lean_ctor_set(x_571, 2, x_570); -x_572 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_488); -x_573 = l_Lean_Syntax_node2(x_488, x_572, x_566, x_571); -x_574 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; -lean_inc(x_488); -x_575 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_575, 0, x_488); -lean_ctor_set(x_575, 1, x_574); -x_576 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; -lean_inc(x_488); -x_577 = l_Lean_Syntax_node3(x_488, x_576, x_560, x_573, x_575); -lean_inc(x_488); -x_578 = l_Lean_Syntax_node1(x_488, x_496, x_577); -lean_inc(x_488); -x_579 = l_Lean_Syntax_node2(x_488, x_572, x_558, x_578); -x_580 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_553); -lean_inc(x_488); -x_581 = l_Lean_Syntax_node2(x_488, x_580, x_553, x_579); -x_582 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_498); -lean_inc(x_488); -x_583 = l_Lean_Syntax_node2(x_488, x_582, x_498, x_581); -x_584 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_488); -x_585 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_585, 0, x_488); -lean_ctor_set(x_585, 1, x_584); -x_586 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_488); -x_587 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_587, 0, x_488); -lean_ctor_set(x_587, 1, x_586); -x_588 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_489); -lean_inc(x_495); -x_589 = l_Lean_addMacroScope(x_495, x_588, x_489); -x_590 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_591 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_488); -x_592 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_592, 0, x_488); -lean_ctor_set(x_592, 1, x_590); -lean_ctor_set(x_592, 2, x_589); -lean_ctor_set(x_592, 3, x_591); -x_593 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_498); -lean_inc(x_488); -x_594 = l_Lean_Syntax_node2(x_488, x_593, x_592, x_498); -x_595 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_489); -lean_inc(x_495); -x_596 = l_Lean_addMacroScope(x_495, x_595, x_489); -x_597 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_488); -x_598 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_598, 0, x_488); -lean_ctor_set(x_598, 1, x_597); -lean_ctor_set(x_598, 2, x_596); -lean_ctor_set(x_598, 3, x_13); -x_599 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_598); -lean_inc(x_585); -lean_inc(x_488); -x_600 = l_Lean_Syntax_node3(x_488, x_599, x_594, x_585, x_598); -x_601 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_489); -lean_inc(x_495); -x_602 = l_Lean_addMacroScope(x_495, x_601, x_489); -x_603 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_604 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_488); -x_605 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_605, 0, x_488); -lean_ctor_set(x_605, 1, x_603); -lean_ctor_set(x_605, 2, x_602); -lean_ctor_set(x_605, 3, x_604); -lean_inc(x_498); -lean_inc(x_488); -x_606 = l_Lean_Syntax_node2(x_488, x_593, x_605, x_498); -x_607 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_489); -lean_inc(x_495); -x_608 = l_Lean_addMacroScope(x_495, x_607, x_489); -x_609 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_488); -x_610 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_610, 0, x_488); -lean_ctor_set(x_610, 1, x_609); -lean_ctor_set(x_610, 2, x_608); -lean_ctor_set(x_610, 3, x_13); -lean_inc(x_610); -lean_inc(x_585); -lean_inc(x_488); -x_611 = l_Lean_Syntax_node3(x_488, x_599, x_606, x_585, x_610); -lean_inc(x_488); -x_612 = l_Lean_Syntax_node3(x_488, x_496, x_600, x_18, x_611); -x_613 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_498); -lean_inc(x_488); -x_614 = l_Lean_Syntax_node1(x_488, x_613, x_498); -x_615 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_488); -x_616 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_616, 0, x_488); -lean_ctor_set(x_616, 1, x_615); -x_617 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc(x_616); -lean_inc(x_614); -lean_inc_n(x_498, 2); -lean_inc(x_587); -lean_inc(x_488); -x_618 = l_Lean_Syntax_node6(x_488, x_617, x_587, x_498, x_612, x_614, x_498, x_616); -x_619 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_498, 2); -lean_inc(x_488); -x_620 = l_Lean_Syntax_node2(x_488, x_619, x_498, x_498); -x_621 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; -lean_inc(x_489); -lean_inc(x_495); -x_622 = l_Lean_addMacroScope(x_495, x_621, x_489); -x_623 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; -lean_inc(x_488); -x_624 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_624, 0, x_488); -lean_ctor_set(x_624, 1, x_623); -lean_ctor_set(x_624, 2, x_622); -lean_ctor_set(x_624, 3, x_13); -lean_inc(x_624); -lean_inc(x_488); -x_625 = l_Lean_Syntax_node1(x_488, x_496, x_624); -x_626 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; -lean_inc(x_488); -x_627 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_627, 0, x_488); -lean_ctor_set(x_627, 1, x_626); -x_628 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; -lean_inc(x_489); -lean_inc(x_495); -x_629 = l_Lean_addMacroScope(x_495, x_628, x_489); -x_630 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_631 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; -lean_inc(x_488); -x_632 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_632, 0, x_488); -lean_ctor_set(x_632, 1, x_630); -lean_ctor_set(x_632, 2, x_629); -lean_ctor_set(x_632, 3, x_631); -x_633 = lean_array_size(x_29); -x_634 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_599, x_633, x_16, x_29); -x_635 = lean_array_size(x_634); -x_636 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_635, x_16, x_634); -x_637 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_638 = l_Lean_mkSepArray(x_636, x_637); -lean_dec(x_636); -x_639 = l_Array_append___rarg(x_497, x_638); -lean_dec(x_638); -lean_inc(x_488); -x_640 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_640, 0, x_488); -lean_ctor_set(x_640, 1, x_496); -lean_ctor_set(x_640, 2, x_639); +lean_inc(x_496); +x_540 = l_Lean_Syntax_node3(x_496, x_504, x_533, x_18, x_539); +lean_inc(x_496); +x_541 = l_Lean_Syntax_node2(x_496, x_504, x_19, x_540); +x_542 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_496); +x_543 = l_Lean_Syntax_node1(x_496, x_542, x_541); +x_544 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; +lean_inc_n(x_506, 3); +lean_inc(x_496); +x_545 = l_Lean_Syntax_node7(x_496, x_544, x_512, x_518, x_506, x_506, x_506, x_527, x_543); +x_546 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; lean_inc(x_508); -lean_inc(x_553); -lean_inc(x_488); -x_641 = l_Lean_Syntax_node2(x_488, x_496, x_553, x_508); -lean_inc(x_616); -lean_inc(x_614); -lean_inc(x_498); -lean_inc(x_587); -lean_inc(x_488); -x_642 = l_Lean_Syntax_node6(x_488, x_617, x_587, x_498, x_640, x_614, x_641, x_616); -lean_inc(x_488); -x_643 = l_Lean_Syntax_node1(x_488, x_496, x_642); -lean_inc(x_488); -x_644 = l_Lean_Syntax_node2(x_488, x_572, x_632, x_643); -lean_inc(x_488); -x_645 = l_Lean_Syntax_node1(x_488, x_496, x_644); -x_646 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_627); -lean_inc(x_488); -x_647 = l_Lean_Syntax_node2(x_488, x_646, x_627, x_645); -x_648 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_585); -lean_inc(x_498); -lean_inc(x_488); -x_649 = l_Lean_Syntax_node5(x_488, x_648, x_598, x_625, x_498, x_585, x_647); -x_650 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_488); -x_651 = l_Lean_Syntax_node1(x_488, x_650, x_649); -x_652 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_620); -lean_inc_n(x_498, 2); -lean_inc(x_488); -x_653 = l_Lean_Syntax_node4(x_488, x_652, x_498, x_498, x_651, x_620); -x_654 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_489); -lean_inc(x_495); -x_655 = l_Lean_addMacroScope(x_495, x_654, x_489); -x_656 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_488); -x_657 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_657, 0, x_488); -lean_ctor_set(x_657, 1, x_656); -lean_ctor_set(x_657, 2, x_655); -lean_ctor_set(x_657, 3, x_13); -lean_inc(x_488); -x_658 = l_Lean_Syntax_node1(x_488, x_496, x_657); -x_659 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_488); -x_660 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_660, 0, x_488); -lean_ctor_set(x_660, 1, x_659); -x_661 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_488); -x_662 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_662, 0, x_488); -lean_ctor_set(x_662, 1, x_661); -lean_inc(x_488); -x_663 = l_Lean_Syntax_node2(x_488, x_580, x_553, x_508); -lean_inc(x_488); -x_664 = l_Lean_Syntax_node1(x_488, x_496, x_663); -x_665 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_488); -x_666 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_666, 0, x_488); -lean_ctor_set(x_666, 1, x_665); -x_667 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_668 = l_Lean_addMacroScope(x_495, x_667, x_489); -x_669 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_670 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_488); -x_671 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_671, 0, x_488); -lean_ctor_set(x_671, 1, x_669); -lean_ctor_set(x_671, 2, x_668); -lean_ctor_set(x_671, 3, x_670); -lean_inc(x_658); -lean_inc(x_488); -x_672 = l_Lean_Syntax_node2(x_488, x_572, x_671, x_658); -x_673 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_488); -x_674 = l_Lean_Syntax_node1(x_488, x_673, x_672); -x_675 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_488); -x_676 = l_Lean_Syntax_node4(x_488, x_675, x_624, x_664, x_666, x_674); -x_677 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_498); -lean_inc(x_488); -x_678 = l_Lean_Syntax_node3(x_488, x_677, x_662, x_498, x_676); -x_679 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_498); -lean_inc(x_488); -x_680 = l_Lean_Syntax_node2(x_488, x_679, x_678, x_498); -x_681 = l_Lean_Syntax_SepArray_ofElems(x_526, x_26); +lean_inc(x_496); +x_547 = l_Lean_Syntax_node2(x_496, x_546, x_508, x_545); +x_548 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_496); +x_549 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_549, 0, x_496); +lean_ctor_set(x_549, 1, x_548); +x_550 = l_Array_append___rarg(x_505, x_3); +lean_inc(x_496); +x_551 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_551, 0, x_496); +lean_ctor_set(x_551, 1, x_504); +lean_ctor_set(x_551, 2, x_550); +x_552 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_496); +x_553 = l_Lean_Syntax_node2(x_496, x_552, x_549, x_551); +x_554 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_496); +x_555 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_555, 0, x_496); +lean_ctor_set(x_555, 1, x_554); +x_556 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_506); +lean_inc(x_496); +x_557 = l_Lean_Syntax_node1(x_496, x_556, x_506); +x_558 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_496); +x_559 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_559, 0, x_496); +lean_ctor_set(x_559, 1, x_558); +x_560 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_496); +x_561 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_561, 0, x_496); +lean_ctor_set(x_561, 1, x_560); +x_562 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_497); +lean_inc(x_503); +x_563 = l_Lean_addMacroScope(x_503, x_562, x_497); +x_564 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_565 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_496); +x_566 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_566, 0, x_496); +lean_ctor_set(x_566, 1, x_564); +lean_ctor_set(x_566, 2, x_563); +lean_ctor_set(x_566, 3, x_565); +x_567 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; +lean_inc(x_496); +x_568 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_568, 0, x_496); +lean_ctor_set(x_568, 1, x_567); +x_569 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_496); +x_570 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_570, 0, x_496); +lean_ctor_set(x_570, 1, x_569); +x_571 = lean_box(0); +x_572 = l_Lean_mkCIdentFrom(x_571, x_4, x_495); +x_573 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_496); +x_574 = l_Lean_Syntax_node2(x_496, x_573, x_570, x_572); +x_575 = lean_array_size(x_492); +x_576 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; +x_577 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_576, x_575, x_16, x_492); +x_578 = l_Array_append___rarg(x_505, x_577); +lean_dec(x_577); +lean_inc(x_496); +x_579 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_579, 0, x_496); +lean_ctor_set(x_579, 1, x_504); +lean_ctor_set(x_579, 2, x_578); +x_580 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_496); +x_581 = l_Lean_Syntax_node2(x_496, x_580, x_574, x_579); +x_582 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; +lean_inc(x_496); +x_583 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_583, 0, x_496); +lean_ctor_set(x_583, 1, x_582); +x_584 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; +lean_inc(x_496); +x_585 = l_Lean_Syntax_node3(x_496, x_584, x_568, x_581, x_583); +lean_inc(x_496); +x_586 = l_Lean_Syntax_node1(x_496, x_504, x_585); +lean_inc(x_496); +x_587 = l_Lean_Syntax_node2(x_496, x_580, x_566, x_586); +x_588 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_561); +lean_inc(x_496); +x_589 = l_Lean_Syntax_node2(x_496, x_588, x_561, x_587); +x_590 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_506); +lean_inc(x_496); +x_591 = l_Lean_Syntax_node2(x_496, x_590, x_506, x_589); +x_592 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_496); +x_593 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_593, 0, x_496); +lean_ctor_set(x_593, 1, x_592); +x_594 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_496); +x_595 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_595, 0, x_496); +lean_ctor_set(x_595, 1, x_594); +x_596 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_497); +lean_inc(x_503); +x_597 = l_Lean_addMacroScope(x_503, x_596, x_497); +x_598 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_599 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_496); +x_600 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_600, 0, x_496); +lean_ctor_set(x_600, 1, x_598); +lean_ctor_set(x_600, 2, x_597); +lean_ctor_set(x_600, 3, x_599); +x_601 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_506); +lean_inc(x_496); +x_602 = l_Lean_Syntax_node2(x_496, x_601, x_600, x_506); +x_603 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_497); +lean_inc(x_503); +x_604 = l_Lean_addMacroScope(x_503, x_603, x_497); +x_605 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_496); +x_606 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_606, 0, x_496); +lean_ctor_set(x_606, 1, x_605); +lean_ctor_set(x_606, 2, x_604); +lean_ctor_set(x_606, 3, x_13); +x_607 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_606); +lean_inc(x_593); +lean_inc(x_496); +x_608 = l_Lean_Syntax_node3(x_496, x_607, x_602, x_593, x_606); +x_609 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_497); +lean_inc(x_503); +x_610 = l_Lean_addMacroScope(x_503, x_609, x_497); +x_611 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_612 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_496); +x_613 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_613, 0, x_496); +lean_ctor_set(x_613, 1, x_611); +lean_ctor_set(x_613, 2, x_610); +lean_ctor_set(x_613, 3, x_612); +lean_inc(x_506); +lean_inc(x_496); +x_614 = l_Lean_Syntax_node2(x_496, x_601, x_613, x_506); +x_615 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_497); +lean_inc(x_503); +x_616 = l_Lean_addMacroScope(x_503, x_615, x_497); +x_617 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_496); +x_618 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_618, 0, x_496); +lean_ctor_set(x_618, 1, x_617); +lean_ctor_set(x_618, 2, x_616); +lean_ctor_set(x_618, 3, x_13); +lean_inc(x_618); +lean_inc(x_593); +lean_inc(x_496); +x_619 = l_Lean_Syntax_node3(x_496, x_607, x_614, x_593, x_618); +lean_inc(x_496); +x_620 = l_Lean_Syntax_node3(x_496, x_504, x_608, x_18, x_619); +x_621 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_496); +x_622 = l_Lean_Syntax_node1(x_496, x_621, x_620); +x_623 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_506); +lean_inc(x_496); +x_624 = l_Lean_Syntax_node1(x_496, x_623, x_506); +x_625 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_496); +x_626 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_626, 0, x_496); +lean_ctor_set(x_626, 1, x_625); +x_627 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_626); +lean_inc(x_624); +lean_inc_n(x_506, 2); +lean_inc(x_595); +lean_inc(x_496); +x_628 = l_Lean_Syntax_node6(x_496, x_627, x_595, x_506, x_622, x_624, x_506, x_626); +x_629 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_506, 2); +lean_inc(x_496); +x_630 = l_Lean_Syntax_node2(x_496, x_629, x_506, x_506); +x_631 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +lean_inc(x_497); +lean_inc(x_503); +x_632 = l_Lean_addMacroScope(x_503, x_631, x_497); +x_633 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +lean_inc(x_496); +x_634 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_634, 0, x_496); +lean_ctor_set(x_634, 1, x_633); +lean_ctor_set(x_634, 2, x_632); +lean_ctor_set(x_634, 3, x_13); +lean_inc(x_634); +lean_inc(x_496); +x_635 = l_Lean_Syntax_node1(x_496, x_504, x_634); +x_636 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; +lean_inc(x_496); +x_637 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_637, 0, x_496); +lean_ctor_set(x_637, 1, x_636); +x_638 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; +lean_inc(x_497); +lean_inc(x_503); +x_639 = l_Lean_addMacroScope(x_503, x_638, x_497); +x_640 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_641 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +lean_inc(x_496); +x_642 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_642, 0, x_496); +lean_ctor_set(x_642, 1, x_640); +lean_ctor_set(x_642, 2, x_639); +lean_ctor_set(x_642, 3, x_641); +x_643 = lean_array_size(x_29); +x_644 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_607, x_643, x_16, x_29); +x_645 = lean_array_size(x_644); +x_646 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_645, x_16, x_644); +x_647 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_648 = l_Lean_mkSepArray(x_646, x_647); +lean_dec(x_646); +x_649 = l_Array_append___rarg(x_505, x_648); +lean_dec(x_648); +lean_inc(x_496); +x_650 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_650, 0, x_496); +lean_ctor_set(x_650, 1, x_504); +lean_ctor_set(x_650, 2, x_649); +lean_inc(x_496); +x_651 = l_Lean_Syntax_node1(x_496, x_621, x_650); +lean_inc(x_516); +lean_inc(x_561); +lean_inc(x_496); +x_652 = l_Lean_Syntax_node2(x_496, x_504, x_561, x_516); +lean_inc(x_626); +lean_inc(x_624); +lean_inc(x_506); +lean_inc(x_595); +lean_inc(x_496); +x_653 = l_Lean_Syntax_node6(x_496, x_627, x_595, x_506, x_651, x_624, x_652, x_626); +lean_inc(x_496); +x_654 = l_Lean_Syntax_node1(x_496, x_504, x_653); +lean_inc(x_496); +x_655 = l_Lean_Syntax_node2(x_496, x_580, x_642, x_654); +lean_inc(x_496); +x_656 = l_Lean_Syntax_node1(x_496, x_504, x_655); +x_657 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_637); +lean_inc(x_496); +x_658 = l_Lean_Syntax_node2(x_496, x_657, x_637, x_656); +x_659 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_593); +lean_inc(x_506); +lean_inc(x_496); +x_660 = l_Lean_Syntax_node5(x_496, x_659, x_606, x_635, x_506, x_593, x_658); +x_661 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_496); +x_662 = l_Lean_Syntax_node1(x_496, x_661, x_660); +x_663 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_630); +lean_inc_n(x_506, 2); +lean_inc(x_496); +x_664 = l_Lean_Syntax_node4(x_496, x_663, x_506, x_506, x_662, x_630); +x_665 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_497); +lean_inc(x_503); +x_666 = l_Lean_addMacroScope(x_503, x_665, x_497); +x_667 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_496); +x_668 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_668, 0, x_496); +lean_ctor_set(x_668, 1, x_667); +lean_ctor_set(x_668, 2, x_666); +lean_ctor_set(x_668, 3, x_13); +lean_inc(x_496); +x_669 = l_Lean_Syntax_node1(x_496, x_504, x_668); +x_670 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_496); +x_671 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_671, 0, x_496); +lean_ctor_set(x_671, 1, x_670); +x_672 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_496); +x_673 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_673, 0, x_496); +lean_ctor_set(x_673, 1, x_672); +lean_inc(x_496); +x_674 = l_Lean_Syntax_node2(x_496, x_588, x_561, x_516); +lean_inc(x_496); +x_675 = l_Lean_Syntax_node1(x_496, x_504, x_674); +x_676 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_496); +x_677 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_677, 0, x_496); +lean_ctor_set(x_677, 1, x_676); +x_678 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_679 = l_Lean_addMacroScope(x_503, x_678, x_497); +x_680 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_681 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_496); +x_682 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_682, 0, x_496); +lean_ctor_set(x_682, 1, x_680); +lean_ctor_set(x_682, 2, x_679); +lean_ctor_set(x_682, 3, x_681); +lean_inc(x_669); +lean_inc(x_496); +x_683 = l_Lean_Syntax_node2(x_496, x_580, x_682, x_669); +x_684 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_496); +x_685 = l_Lean_Syntax_node1(x_496, x_684, x_683); +x_686 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_496); +x_687 = l_Lean_Syntax_node4(x_496, x_686, x_634, x_675, x_677, x_685); +x_688 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_506); +lean_inc(x_496); +x_689 = l_Lean_Syntax_node3(x_496, x_688, x_673, x_506, x_687); +x_690 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_506); +lean_inc(x_496); +x_691 = l_Lean_Syntax_node2(x_496, x_690, x_689, x_506); +x_692 = l_Lean_Syntax_SepArray_ofElems(x_534, x_26); lean_dec(x_26); -x_682 = l_Array_append___rarg(x_497, x_681); -lean_dec(x_681); -lean_inc(x_488); -x_683 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_683, 0, x_488); -lean_ctor_set(x_683, 1, x_496); -lean_ctor_set(x_683, 2, x_682); -lean_inc_n(x_498, 2); -lean_inc(x_488); -x_684 = l_Lean_Syntax_node6(x_488, x_617, x_587, x_498, x_683, x_614, x_498, x_616); -lean_inc(x_488); -x_685 = l_Lean_Syntax_node1(x_488, x_496, x_684); -x_686 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; -lean_inc(x_488); -x_687 = l_Lean_Syntax_node2(x_488, x_686, x_627, x_685); -lean_inc(x_498); -lean_inc(x_488); -x_688 = l_Lean_Syntax_node2(x_488, x_679, x_687, x_498); -lean_inc(x_488); -x_689 = l_Lean_Syntax_node2(x_488, x_496, x_680, x_688); -x_690 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_488); -x_691 = l_Lean_Syntax_node1(x_488, x_690, x_689); -x_692 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_488); -x_693 = l_Lean_Syntax_node2(x_488, x_692, x_660, x_691); -lean_inc(x_585); -lean_inc(x_498); -lean_inc(x_488); -x_694 = l_Lean_Syntax_node5(x_488, x_648, x_610, x_658, x_498, x_585, x_693); -lean_inc(x_488); -x_695 = l_Lean_Syntax_node1(x_488, x_650, x_694); -lean_inc(x_620); -lean_inc_n(x_498, 2); -lean_inc(x_488); -x_696 = l_Lean_Syntax_node4(x_488, x_652, x_498, x_498, x_695, x_620); -lean_inc(x_498); -lean_inc(x_488); -x_697 = l_Lean_Syntax_node3(x_488, x_496, x_653, x_498, x_696); -x_698 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_488); -x_699 = l_Lean_Syntax_node2(x_488, x_698, x_20, x_697); -lean_inc(x_488); -x_700 = l_Lean_Syntax_node1(x_488, x_496, x_699); -x_701 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_488); -x_702 = l_Lean_Syntax_node4(x_488, x_701, x_585, x_618, x_620, x_700); -x_703 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_498); -lean_inc(x_488); -x_704 = l_Lean_Syntax_node6(x_488, x_703, x_549, x_551, x_498, x_498, x_583, x_702); -lean_inc(x_488); -x_705 = l_Lean_Syntax_node2(x_488, x_538, x_500, x_704); -x_706 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_488); -x_707 = l_Lean_Syntax_node3(x_488, x_706, x_545, x_547, x_705); -x_708 = l_Lean_Syntax_node2(x_488, x_496, x_539, x_707); -if (lean_is_scalar(x_493)) { - x_709 = lean_alloc_ctor(0, 2, 0); +x_693 = l_Array_append___rarg(x_505, x_692); +lean_dec(x_692); +lean_inc(x_496); +x_694 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_694, 0, x_496); +lean_ctor_set(x_694, 1, x_504); +lean_ctor_set(x_694, 2, x_693); +lean_inc(x_496); +x_695 = l_Lean_Syntax_node1(x_496, x_621, x_694); +lean_inc_n(x_506, 2); +lean_inc(x_496); +x_696 = l_Lean_Syntax_node6(x_496, x_627, x_595, x_506, x_695, x_624, x_506, x_626); +lean_inc(x_496); +x_697 = l_Lean_Syntax_node1(x_496, x_504, x_696); +x_698 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; +lean_inc(x_496); +x_699 = l_Lean_Syntax_node2(x_496, x_698, x_637, x_697); +lean_inc(x_506); +lean_inc(x_496); +x_700 = l_Lean_Syntax_node2(x_496, x_690, x_699, x_506); +lean_inc(x_496); +x_701 = l_Lean_Syntax_node2(x_496, x_504, x_691, x_700); +x_702 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_496); +x_703 = l_Lean_Syntax_node1(x_496, x_702, x_701); +x_704 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_496); +x_705 = l_Lean_Syntax_node2(x_496, x_704, x_671, x_703); +lean_inc(x_593); +lean_inc(x_506); +lean_inc(x_496); +x_706 = l_Lean_Syntax_node5(x_496, x_659, x_618, x_669, x_506, x_593, x_705); +lean_inc(x_496); +x_707 = l_Lean_Syntax_node1(x_496, x_661, x_706); +lean_inc(x_630); +lean_inc_n(x_506, 2); +lean_inc(x_496); +x_708 = l_Lean_Syntax_node4(x_496, x_663, x_506, x_506, x_707, x_630); +lean_inc(x_506); +lean_inc(x_496); +x_709 = l_Lean_Syntax_node3(x_496, x_504, x_664, x_506, x_708); +x_710 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_496); +x_711 = l_Lean_Syntax_node2(x_496, x_710, x_20, x_709); +lean_inc(x_496); +x_712 = l_Lean_Syntax_node1(x_496, x_504, x_711); +x_713 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_496); +x_714 = l_Lean_Syntax_node4(x_496, x_713, x_593, x_628, x_630, x_712); +x_715 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_506); +lean_inc(x_496); +x_716 = l_Lean_Syntax_node6(x_496, x_715, x_557, x_559, x_506, x_506, x_591, x_714); +lean_inc(x_496); +x_717 = l_Lean_Syntax_node2(x_496, x_546, x_508, x_716); +x_718 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_496); +x_719 = l_Lean_Syntax_node3(x_496, x_718, x_553, x_555, x_717); +x_720 = l_Lean_Syntax_node2(x_496, x_504, x_547, x_719); +if (lean_is_scalar(x_501)) { + x_721 = lean_alloc_ctor(0, 2, 0); } else { - x_709 = x_493; + x_721 = x_501; } -lean_ctor_set(x_709, 0, x_708); -lean_ctor_set(x_709, 1, x_492); -return x_709; +lean_ctor_set(x_721, 0, x_720); +lean_ctor_set(x_721, 1, x_500); +return x_721; } else { -lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; -lean_dec(x_481); -lean_dec(x_480); +lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; +lean_dec(x_489); +lean_dec(x_488); lean_free_object(x_20); lean_dec(x_29); lean_free_object(x_19); @@ -5732,1862 +5775,1883 @@ lean_dec(x_26); lean_free_object(x_18); lean_dec(x_10); lean_dec(x_4); -x_710 = lean_ctor_get(x_483, 0); -lean_inc(x_710); -x_711 = lean_ctor_get(x_483, 1); -lean_inc(x_711); -if (lean_is_exclusive(x_483)) { - lean_ctor_release(x_483, 0); - lean_ctor_release(x_483, 1); - x_712 = x_483; +x_722 = lean_ctor_get(x_491, 0); +lean_inc(x_722); +x_723 = lean_ctor_get(x_491, 1); +lean_inc(x_723); +if (lean_is_exclusive(x_491)) { + lean_ctor_release(x_491, 0); + lean_ctor_release(x_491, 1); + x_724 = x_491; } else { - lean_dec_ref(x_483); - x_712 = lean_box(0); + lean_dec_ref(x_491); + x_724 = lean_box(0); } -if (lean_is_scalar(x_712)) { - x_713 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_724)) { + x_725 = lean_alloc_ctor(1, 2, 0); } else { - x_713 = x_712; + x_725 = x_724; } -lean_ctor_set(x_713, 0, x_710); -lean_ctor_set(x_713, 1, x_711); -return x_713; +lean_ctor_set(x_725, 0, x_722); +lean_ctor_set(x_725, 1, x_723); +return x_725; } } } else { -lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; size_t x_718; lean_object* x_719; -x_714 = lean_ctor_get(x_20, 0); -lean_inc(x_714); +lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; size_t x_730; lean_object* x_731; +x_726 = lean_ctor_get(x_20, 0); +lean_inc(x_726); lean_dec(x_20); -x_715 = lean_ctor_get(x_21, 0); -lean_inc(x_715); -x_716 = lean_ctor_get(x_21, 1); -lean_inc(x_716); +x_727 = lean_ctor_get(x_21, 0); +lean_inc(x_727); +x_728 = lean_ctor_get(x_21, 1); +lean_inc(x_728); if (lean_is_exclusive(x_21)) { lean_ctor_release(x_21, 0); lean_ctor_release(x_21, 1); - x_717 = x_21; + x_729 = x_21; } else { lean_dec_ref(x_21); - x_717 = lean_box(0); -} -x_718 = lean_array_size(x_2); -x_719 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_718, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -if (lean_obj_tag(x_719) == 0) -{ -lean_object* x_720; lean_object* x_721; lean_object* x_722; uint8_t x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; size_t x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; size_t x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; size_t x_870; lean_object* x_871; size_t x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; -x_720 = lean_ctor_get(x_719, 0); -lean_inc(x_720); -x_721 = lean_ctor_get(x_719, 1); -lean_inc(x_721); -lean_dec(x_719); -x_722 = lean_ctor_get(x_10, 5); -lean_inc(x_722); -x_723 = 0; -x_724 = l_Lean_SourceInfo_fromRef(x_722, x_723); -lean_dec(x_722); -x_725 = lean_ctor_get(x_10, 10); -lean_inc(x_725); + x_729 = lean_box(0); +} +x_730 = lean_array_size(x_2); +x_731 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_730, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_731) == 0) +{ +lean_object* x_732; lean_object* x_733; lean_object* x_734; uint8_t x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; size_t x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; size_t x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; size_t x_884; lean_object* x_885; size_t x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; +x_732 = lean_ctor_get(x_731, 0); +lean_inc(x_732); +x_733 = lean_ctor_get(x_731, 1); +lean_inc(x_733); +lean_dec(x_731); +x_734 = lean_ctor_get(x_10, 5); +lean_inc(x_734); +x_735 = 0; +x_736 = l_Lean_SourceInfo_fromRef(x_734, x_735); +lean_dec(x_734); +x_737 = lean_ctor_get(x_10, 10); +lean_inc(x_737); lean_dec(x_10); -x_726 = lean_st_ref_get(x_11, x_721); -x_727 = lean_ctor_get(x_726, 0); -lean_inc(x_727); -x_728 = lean_ctor_get(x_726, 1); -lean_inc(x_728); -if (lean_is_exclusive(x_726)) { - lean_ctor_release(x_726, 0); - lean_ctor_release(x_726, 1); - x_729 = x_726; +x_738 = lean_st_ref_get(x_11, x_733); +x_739 = lean_ctor_get(x_738, 0); +lean_inc(x_739); +x_740 = lean_ctor_get(x_738, 1); +lean_inc(x_740); +if (lean_is_exclusive(x_738)) { + lean_ctor_release(x_738, 0); + lean_ctor_release(x_738, 1); + x_741 = x_738; } else { - lean_dec_ref(x_726); - x_729 = lean_box(0); + lean_dec_ref(x_738); + x_741 = lean_box(0); } -x_730 = lean_ctor_get(x_727, 0); -lean_inc(x_730); -lean_dec(x_727); -x_731 = lean_environment_main_module(x_730); -x_732 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -x_733 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -lean_inc(x_724); -x_734 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_734, 0, x_724); -lean_ctor_set(x_734, 1, x_732); -lean_ctor_set(x_734, 2, x_733); -x_735 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_734, 6); -lean_inc(x_724); -x_736 = l_Lean_Syntax_node6(x_724, x_735, x_734, x_734, x_734, x_734, x_734, x_734); -x_737 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; -lean_inc(x_724); -if (lean_is_scalar(x_717)) { - x_738 = lean_alloc_ctor(2, 2, 0); +x_742 = lean_ctor_get(x_739, 0); +lean_inc(x_742); +lean_dec(x_739); +x_743 = lean_environment_main_module(x_742); +x_744 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +x_745 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +lean_inc(x_736); +x_746 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_746, 0, x_736); +lean_ctor_set(x_746, 1, x_744); +lean_ctor_set(x_746, 2, x_745); +x_747 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_746, 6); +lean_inc(x_736); +x_748 = l_Lean_Syntax_node6(x_736, x_747, x_746, x_746, x_746, x_746, x_746, x_746); +x_749 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; +lean_inc(x_736); +if (lean_is_scalar(x_729)) { + x_750 = lean_alloc_ctor(2, 2, 0); } else { - x_738 = x_717; - lean_ctor_set_tag(x_738, 2); -} -lean_ctor_set(x_738, 0, x_724); -lean_ctor_set(x_738, 1, x_737); -x_739 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; -lean_inc(x_724); -x_740 = l_Lean_Syntax_node1(x_724, x_739, x_738); -x_741 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_725); -lean_inc(x_731); -x_742 = l_Lean_addMacroScope(x_731, x_741, x_725); -x_743 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_724); -x_744 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_744, 0, x_724); -lean_ctor_set(x_744, 1, x_743); -lean_ctor_set(x_744, 2, x_742); -lean_ctor_set(x_744, 3, x_13); -x_745 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_734); -lean_inc(x_744); -lean_inc(x_724); -x_746 = l_Lean_Syntax_node2(x_724, x_745, x_744, x_734); -x_747 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_724); -x_748 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_748, 0, x_724); -lean_ctor_set(x_748, 1, x_747); -x_749 = l_Array_zip___rarg(x_715, x_716); -lean_dec(x_716); -lean_dec(x_715); -x_750 = lean_array_size(x_749); + x_750 = x_729; + lean_ctor_set_tag(x_750, 2); +} +lean_ctor_set(x_750, 0, x_736); +lean_ctor_set(x_750, 1, x_749); +x_751 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; lean_inc(x_736); -lean_inc(x_734); -lean_inc(x_724); -x_751 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_724, x_732, x_734, x_736, x_750, x_16, x_749); -x_752 = l_Array_append___rarg(x_733, x_751); -lean_dec(x_751); -lean_inc(x_724); -x_753 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_753, 0, x_724); -lean_ctor_set(x_753, 1, x_732); -lean_ctor_set(x_753, 2, x_752); -x_754 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_724); -x_755 = l_Lean_Syntax_node1(x_724, x_754, x_753); -lean_inc(x_734); +x_752 = l_Lean_Syntax_node1(x_736, x_751, x_750); +x_753 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_737); +lean_inc(x_743); +x_754 = l_Lean_addMacroScope(x_743, x_753, x_737); +x_755 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_736); +x_756 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_756, 0, x_736); +lean_ctor_set(x_756, 1, x_755); +lean_ctor_set(x_756, 2, x_754); +lean_ctor_set(x_756, 3, x_13); +x_757 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_746); +lean_inc(x_756); +lean_inc(x_736); +x_758 = l_Lean_Syntax_node2(x_736, x_757, x_756, x_746); +x_759 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_736); +x_760 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_760, 0, x_736); +lean_ctor_set(x_760, 1, x_759); +x_761 = l_Array_zip___rarg(x_727, x_728); +lean_dec(x_728); +lean_dec(x_727); +x_762 = lean_array_size(x_761); lean_inc(x_748); -lean_inc(x_724); -x_756 = l_Lean_Syntax_node3(x_724, x_732, x_748, x_734, x_755); -x_757 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_724); +lean_inc(x_746); +lean_inc(x_736); +x_763 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_736, x_744, x_746, x_748, x_762, x_16, x_761); +x_764 = l_Array_append___rarg(x_745, x_763); +lean_dec(x_763); +lean_inc(x_736); +x_765 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_765, 0, x_736); +lean_ctor_set(x_765, 1, x_744); +lean_ctor_set(x_765, 2, x_764); +x_766 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; +lean_inc(x_736); +x_767 = l_Lean_Syntax_node1(x_736, x_766, x_765); +lean_inc(x_746); +lean_inc(x_760); +lean_inc(x_736); +x_768 = l_Lean_Syntax_node3(x_736, x_744, x_760, x_746, x_767); +x_769 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_736); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_757); -lean_ctor_set(x_19, 0, x_724); -x_758 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_725); -lean_inc(x_731); -x_759 = l_Lean_addMacroScope(x_731, x_758, x_725); -x_760 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_761 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_724); -x_762 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_762, 0, x_724); -lean_ctor_set(x_762, 1, x_760); -lean_ctor_set(x_762, 2, x_759); -lean_ctor_set(x_762, 3, x_761); -x_763 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_724); +lean_ctor_set(x_19, 1, x_769); +lean_ctor_set(x_19, 0, x_736); +x_770 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_737); +lean_inc(x_743); +x_771 = l_Lean_addMacroScope(x_743, x_770, x_737); +x_772 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_773 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_736); +x_774 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_774, 0, x_736); +lean_ctor_set(x_774, 1, x_772); +lean_ctor_set(x_774, 2, x_771); +lean_ctor_set(x_774, 3, x_773); +x_775 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_736); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_763); -lean_ctor_set(x_18, 0, x_724); -x_764 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_725); -lean_inc(x_731); -x_765 = l_Lean_addMacroScope(x_731, x_764, x_725); -x_766 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_767 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_724); -x_768 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_768, 0, x_724); -lean_ctor_set(x_768, 1, x_766); -lean_ctor_set(x_768, 2, x_765); -lean_ctor_set(x_768, 3, x_767); +lean_ctor_set(x_18, 1, x_775); +lean_ctor_set(x_18, 0, x_736); +x_776 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_737); +lean_inc(x_743); +x_777 = l_Lean_addMacroScope(x_743, x_776, x_737); +x_778 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_779 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_736); +x_780 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_780, 0, x_736); +lean_ctor_set(x_780, 1, x_778); +lean_ctor_set(x_780, 2, x_777); +lean_ctor_set(x_780, 3, x_779); lean_inc(x_18); -lean_inc(x_724); -x_769 = l_Lean_Syntax_node3(x_724, x_732, x_762, x_18, x_768); -lean_inc(x_724); -x_770 = l_Lean_Syntax_node2(x_724, x_732, x_19, x_769); -x_771 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_724); -x_772 = l_Lean_Syntax_node1(x_724, x_771, x_770); -x_773 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; -lean_inc_n(x_734, 3); -lean_inc(x_724); -x_774 = l_Lean_Syntax_node7(x_724, x_773, x_740, x_746, x_734, x_734, x_734, x_756, x_772); -x_775 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; lean_inc(x_736); -lean_inc(x_724); -x_776 = l_Lean_Syntax_node2(x_724, x_775, x_736, x_774); -x_777 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_724); -x_778 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_778, 0, x_724); -lean_ctor_set(x_778, 1, x_777); -x_779 = l_Array_append___rarg(x_733, x_3); -lean_inc(x_724); -x_780 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_780, 0, x_724); -lean_ctor_set(x_780, 1, x_732); -lean_ctor_set(x_780, 2, x_779); -x_781 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_724); -x_782 = l_Lean_Syntax_node2(x_724, x_781, x_778, x_780); -x_783 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_724); -x_784 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_784, 0, x_724); -lean_ctor_set(x_784, 1, x_783); -x_785 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_734); -lean_inc(x_724); -x_786 = l_Lean_Syntax_node1(x_724, x_785, x_734); -x_787 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_724); -x_788 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_788, 0, x_724); -lean_ctor_set(x_788, 1, x_787); -x_789 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_724); +x_781 = l_Lean_Syntax_node3(x_736, x_744, x_774, x_18, x_780); +lean_inc(x_736); +x_782 = l_Lean_Syntax_node2(x_736, x_744, x_19, x_781); +x_783 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_736); +x_784 = l_Lean_Syntax_node1(x_736, x_783, x_782); +x_785 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; +lean_inc_n(x_746, 3); +lean_inc(x_736); +x_786 = l_Lean_Syntax_node7(x_736, x_785, x_752, x_758, x_746, x_746, x_746, x_768, x_784); +x_787 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_748); +lean_inc(x_736); +x_788 = l_Lean_Syntax_node2(x_736, x_787, x_748, x_786); +x_789 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_736); x_790 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_790, 0, x_724); +lean_ctor_set(x_790, 0, x_736); lean_ctor_set(x_790, 1, x_789); -x_791 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_725); -lean_inc(x_731); -x_792 = l_Lean_addMacroScope(x_731, x_791, x_725); -x_793 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_794 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_724); -x_795 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_795, 0, x_724); -lean_ctor_set(x_795, 1, x_793); -lean_ctor_set(x_795, 2, x_792); -lean_ctor_set(x_795, 3, x_794); -x_796 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; -lean_inc(x_724); -x_797 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_797, 0, x_724); -lean_ctor_set(x_797, 1, x_796); -x_798 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_724); -x_799 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_799, 0, x_724); -lean_ctor_set(x_799, 1, x_798); -x_800 = lean_box(0); -x_801 = l_Lean_mkCIdentFrom(x_800, x_4, x_723); -x_802 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_724); -x_803 = l_Lean_Syntax_node2(x_724, x_802, x_799, x_801); -x_804 = lean_array_size(x_720); -x_805 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; -x_806 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_805, x_804, x_16, x_720); -x_807 = l_Array_append___rarg(x_733, x_806); -lean_dec(x_806); -lean_inc(x_724); -x_808 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_808, 0, x_724); -lean_ctor_set(x_808, 1, x_732); -lean_ctor_set(x_808, 2, x_807); -x_809 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_724); -x_810 = l_Lean_Syntax_node2(x_724, x_809, x_803, x_808); -x_811 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; -lean_inc(x_724); -x_812 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_812, 0, x_724); -lean_ctor_set(x_812, 1, x_811); -x_813 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; -lean_inc(x_724); -x_814 = l_Lean_Syntax_node3(x_724, x_813, x_797, x_810, x_812); -lean_inc(x_724); -x_815 = l_Lean_Syntax_node1(x_724, x_732, x_814); -lean_inc(x_724); -x_816 = l_Lean_Syntax_node2(x_724, x_809, x_795, x_815); -x_817 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_790); -lean_inc(x_724); -x_818 = l_Lean_Syntax_node2(x_724, x_817, x_790, x_816); -x_819 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_734); -lean_inc(x_724); -x_820 = l_Lean_Syntax_node2(x_724, x_819, x_734, x_818); -x_821 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_724); -x_822 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_822, 0, x_724); -lean_ctor_set(x_822, 1, x_821); -x_823 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_724); +x_791 = l_Array_append___rarg(x_745, x_3); +lean_inc(x_736); +x_792 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_792, 0, x_736); +lean_ctor_set(x_792, 1, x_744); +lean_ctor_set(x_792, 2, x_791); +x_793 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_736); +x_794 = l_Lean_Syntax_node2(x_736, x_793, x_790, x_792); +x_795 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_736); +x_796 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_796, 0, x_736); +lean_ctor_set(x_796, 1, x_795); +x_797 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_746); +lean_inc(x_736); +x_798 = l_Lean_Syntax_node1(x_736, x_797, x_746); +x_799 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_736); +x_800 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_800, 0, x_736); +lean_ctor_set(x_800, 1, x_799); +x_801 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_736); +x_802 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_802, 0, x_736); +lean_ctor_set(x_802, 1, x_801); +x_803 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_737); +lean_inc(x_743); +x_804 = l_Lean_addMacroScope(x_743, x_803, x_737); +x_805 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_806 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_736); +x_807 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_807, 0, x_736); +lean_ctor_set(x_807, 1, x_805); +lean_ctor_set(x_807, 2, x_804); +lean_ctor_set(x_807, 3, x_806); +x_808 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; +lean_inc(x_736); +x_809 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_809, 0, x_736); +lean_ctor_set(x_809, 1, x_808); +x_810 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_736); +x_811 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_811, 0, x_736); +lean_ctor_set(x_811, 1, x_810); +x_812 = lean_box(0); +x_813 = l_Lean_mkCIdentFrom(x_812, x_4, x_735); +x_814 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_736); +x_815 = l_Lean_Syntax_node2(x_736, x_814, x_811, x_813); +x_816 = lean_array_size(x_732); +x_817 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; +x_818 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_817, x_816, x_16, x_732); +x_819 = l_Array_append___rarg(x_745, x_818); +lean_dec(x_818); +lean_inc(x_736); +x_820 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_820, 0, x_736); +lean_ctor_set(x_820, 1, x_744); +lean_ctor_set(x_820, 2, x_819); +x_821 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_736); +x_822 = l_Lean_Syntax_node2(x_736, x_821, x_815, x_820); +x_823 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; +lean_inc(x_736); x_824 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_824, 0, x_724); +lean_ctor_set(x_824, 0, x_736); lean_ctor_set(x_824, 1, x_823); -x_825 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_725); -lean_inc(x_731); -x_826 = l_Lean_addMacroScope(x_731, x_825, x_725); -x_827 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_828 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_724); -x_829 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_829, 0, x_724); -lean_ctor_set(x_829, 1, x_827); -lean_ctor_set(x_829, 2, x_826); -lean_ctor_set(x_829, 3, x_828); -x_830 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_734); -lean_inc(x_724); -x_831 = l_Lean_Syntax_node2(x_724, x_830, x_829, x_734); -x_832 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_725); -lean_inc(x_731); -x_833 = l_Lean_addMacroScope(x_731, x_832, x_725); -x_834 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_724); -x_835 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_835, 0, x_724); -lean_ctor_set(x_835, 1, x_834); -lean_ctor_set(x_835, 2, x_833); -lean_ctor_set(x_835, 3, x_13); -x_836 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_835); -lean_inc(x_822); -lean_inc(x_724); -x_837 = l_Lean_Syntax_node3(x_724, x_836, x_831, x_822, x_835); -x_838 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_725); -lean_inc(x_731); -x_839 = l_Lean_addMacroScope(x_731, x_838, x_725); -x_840 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_841 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_724); -x_842 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_842, 0, x_724); -lean_ctor_set(x_842, 1, x_840); -lean_ctor_set(x_842, 2, x_839); -lean_ctor_set(x_842, 3, x_841); -lean_inc(x_734); -lean_inc(x_724); -x_843 = l_Lean_Syntax_node2(x_724, x_830, x_842, x_734); -x_844 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_725); -lean_inc(x_731); -x_845 = l_Lean_addMacroScope(x_731, x_844, x_725); -x_846 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_724); +x_825 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; +lean_inc(x_736); +x_826 = l_Lean_Syntax_node3(x_736, x_825, x_809, x_822, x_824); +lean_inc(x_736); +x_827 = l_Lean_Syntax_node1(x_736, x_744, x_826); +lean_inc(x_736); +x_828 = l_Lean_Syntax_node2(x_736, x_821, x_807, x_827); +x_829 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_802); +lean_inc(x_736); +x_830 = l_Lean_Syntax_node2(x_736, x_829, x_802, x_828); +x_831 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_746); +lean_inc(x_736); +x_832 = l_Lean_Syntax_node2(x_736, x_831, x_746, x_830); +x_833 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_736); +x_834 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_834, 0, x_736); +lean_ctor_set(x_834, 1, x_833); +x_835 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_736); +x_836 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_836, 0, x_736); +lean_ctor_set(x_836, 1, x_835); +x_837 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_737); +lean_inc(x_743); +x_838 = l_Lean_addMacroScope(x_743, x_837, x_737); +x_839 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_840 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_736); +x_841 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_841, 0, x_736); +lean_ctor_set(x_841, 1, x_839); +lean_ctor_set(x_841, 2, x_838); +lean_ctor_set(x_841, 3, x_840); +x_842 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_746); +lean_inc(x_736); +x_843 = l_Lean_Syntax_node2(x_736, x_842, x_841, x_746); +x_844 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_737); +lean_inc(x_743); +x_845 = l_Lean_addMacroScope(x_743, x_844, x_737); +x_846 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_736); x_847 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_847, 0, x_724); +lean_ctor_set(x_847, 0, x_736); lean_ctor_set(x_847, 1, x_846); lean_ctor_set(x_847, 2, x_845); lean_ctor_set(x_847, 3, x_13); +x_848 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; lean_inc(x_847); -lean_inc(x_822); -lean_inc(x_724); -x_848 = l_Lean_Syntax_node3(x_724, x_836, x_843, x_822, x_847); -lean_inc(x_724); -x_849 = l_Lean_Syntax_node3(x_724, x_732, x_837, x_18, x_848); -x_850 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_734); -lean_inc(x_724); -x_851 = l_Lean_Syntax_node1(x_724, x_850, x_734); -x_852 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_724); -x_853 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_853, 0, x_724); -lean_ctor_set(x_853, 1, x_852); -x_854 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc(x_853); -lean_inc(x_851); -lean_inc_n(x_734, 2); -lean_inc(x_824); -lean_inc(x_724); -x_855 = l_Lean_Syntax_node6(x_724, x_854, x_824, x_734, x_849, x_851, x_734, x_853); -x_856 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_734, 2); -lean_inc(x_724); -x_857 = l_Lean_Syntax_node2(x_724, x_856, x_734, x_734); -x_858 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; -lean_inc(x_725); -lean_inc(x_731); -x_859 = l_Lean_addMacroScope(x_731, x_858, x_725); -x_860 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; -lean_inc(x_724); -x_861 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_861, 0, x_724); -lean_ctor_set(x_861, 1, x_860); -lean_ctor_set(x_861, 2, x_859); -lean_ctor_set(x_861, 3, x_13); -lean_inc(x_861); -lean_inc(x_724); -x_862 = l_Lean_Syntax_node1(x_724, x_732, x_861); -x_863 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; -lean_inc(x_724); -x_864 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_864, 0, x_724); -lean_ctor_set(x_864, 1, x_863); -x_865 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; -lean_inc(x_725); -lean_inc(x_731); -x_866 = l_Lean_addMacroScope(x_731, x_865, x_725); -x_867 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_868 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; -lean_inc(x_724); -x_869 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_869, 0, x_724); -lean_ctor_set(x_869, 1, x_867); -lean_ctor_set(x_869, 2, x_866); -lean_ctor_set(x_869, 3, x_868); -x_870 = lean_array_size(x_714); -x_871 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_836, x_870, x_16, x_714); -x_872 = lean_array_size(x_871); -x_873 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_872, x_16, x_871); -x_874 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_875 = l_Lean_mkSepArray(x_873, x_874); -lean_dec(x_873); -x_876 = l_Array_append___rarg(x_733, x_875); -lean_dec(x_875); -lean_inc(x_724); -x_877 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_877, 0, x_724); -lean_ctor_set(x_877, 1, x_732); -lean_ctor_set(x_877, 2, x_876); -lean_inc(x_744); -lean_inc(x_790); -lean_inc(x_724); -x_878 = l_Lean_Syntax_node2(x_724, x_732, x_790, x_744); -lean_inc(x_853); -lean_inc(x_851); -lean_inc(x_734); -lean_inc(x_824); -lean_inc(x_724); -x_879 = l_Lean_Syntax_node6(x_724, x_854, x_824, x_734, x_877, x_851, x_878, x_853); -lean_inc(x_724); -x_880 = l_Lean_Syntax_node1(x_724, x_732, x_879); -lean_inc(x_724); -x_881 = l_Lean_Syntax_node2(x_724, x_809, x_869, x_880); -lean_inc(x_724); -x_882 = l_Lean_Syntax_node1(x_724, x_732, x_881); -x_883 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_864); -lean_inc(x_724); -x_884 = l_Lean_Syntax_node2(x_724, x_883, x_864, x_882); -x_885 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_822); -lean_inc(x_734); -lean_inc(x_724); -x_886 = l_Lean_Syntax_node5(x_724, x_885, x_835, x_862, x_734, x_822, x_884); -x_887 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_724); -x_888 = l_Lean_Syntax_node1(x_724, x_887, x_886); -x_889 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_857); -lean_inc_n(x_734, 2); -lean_inc(x_724); -x_890 = l_Lean_Syntax_node4(x_724, x_889, x_734, x_734, x_888, x_857); -x_891 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_725); -lean_inc(x_731); -x_892 = l_Lean_addMacroScope(x_731, x_891, x_725); -x_893 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_724); -x_894 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_894, 0, x_724); -lean_ctor_set(x_894, 1, x_893); -lean_ctor_set(x_894, 2, x_892); -lean_ctor_set(x_894, 3, x_13); -lean_inc(x_724); -x_895 = l_Lean_Syntax_node1(x_724, x_732, x_894); -x_896 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_724); -x_897 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_897, 0, x_724); -lean_ctor_set(x_897, 1, x_896); -x_898 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_724); -x_899 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_899, 0, x_724); -lean_ctor_set(x_899, 1, x_898); -lean_inc(x_724); -x_900 = l_Lean_Syntax_node2(x_724, x_817, x_790, x_744); -lean_inc(x_724); -x_901 = l_Lean_Syntax_node1(x_724, x_732, x_900); -x_902 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_724); -x_903 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_903, 0, x_724); -lean_ctor_set(x_903, 1, x_902); -x_904 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_905 = l_Lean_addMacroScope(x_731, x_904, x_725); -x_906 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_907 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_724); -x_908 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_908, 0, x_724); -lean_ctor_set(x_908, 1, x_906); -lean_ctor_set(x_908, 2, x_905); -lean_ctor_set(x_908, 3, x_907); -lean_inc(x_895); -lean_inc(x_724); -x_909 = l_Lean_Syntax_node2(x_724, x_809, x_908, x_895); -x_910 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_724); -x_911 = l_Lean_Syntax_node1(x_724, x_910, x_909); -x_912 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_724); -x_913 = l_Lean_Syntax_node4(x_724, x_912, x_861, x_901, x_903, x_911); -x_914 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_734); -lean_inc(x_724); -x_915 = l_Lean_Syntax_node3(x_724, x_914, x_899, x_734, x_913); -x_916 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_734); -lean_inc(x_724); -x_917 = l_Lean_Syntax_node2(x_724, x_916, x_915, x_734); -x_918 = l_Lean_Syntax_SepArray_ofElems(x_763, x_26); -lean_dec(x_26); -x_919 = l_Array_append___rarg(x_733, x_918); -lean_dec(x_918); -lean_inc(x_724); -x_920 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_920, 0, x_724); -lean_ctor_set(x_920, 1, x_732); -lean_ctor_set(x_920, 2, x_919); -lean_inc_n(x_734, 2); -lean_inc(x_724); -x_921 = l_Lean_Syntax_node6(x_724, x_854, x_824, x_734, x_920, x_851, x_734, x_853); -lean_inc(x_724); -x_922 = l_Lean_Syntax_node1(x_724, x_732, x_921); -x_923 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; -lean_inc(x_724); -x_924 = l_Lean_Syntax_node2(x_724, x_923, x_864, x_922); -lean_inc(x_734); -lean_inc(x_724); -x_925 = l_Lean_Syntax_node2(x_724, x_916, x_924, x_734); -lean_inc(x_724); -x_926 = l_Lean_Syntax_node2(x_724, x_732, x_917, x_925); -x_927 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_724); -x_928 = l_Lean_Syntax_node1(x_724, x_927, x_926); -x_929 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_724); -x_930 = l_Lean_Syntax_node2(x_724, x_929, x_897, x_928); -lean_inc(x_822); -lean_inc(x_734); -lean_inc(x_724); -x_931 = l_Lean_Syntax_node5(x_724, x_885, x_847, x_895, x_734, x_822, x_930); -lean_inc(x_724); -x_932 = l_Lean_Syntax_node1(x_724, x_887, x_931); -lean_inc(x_857); -lean_inc_n(x_734, 2); -lean_inc(x_724); -x_933 = l_Lean_Syntax_node4(x_724, x_889, x_734, x_734, x_932, x_857); -lean_inc(x_734); -lean_inc(x_724); -x_934 = l_Lean_Syntax_node3(x_724, x_732, x_890, x_734, x_933); -x_935 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_724); -x_936 = l_Lean_Syntax_node2(x_724, x_935, x_748, x_934); -lean_inc(x_724); -x_937 = l_Lean_Syntax_node1(x_724, x_732, x_936); -x_938 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_724); -x_939 = l_Lean_Syntax_node4(x_724, x_938, x_822, x_855, x_857, x_937); -x_940 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_734); -lean_inc(x_724); -x_941 = l_Lean_Syntax_node6(x_724, x_940, x_786, x_788, x_734, x_734, x_820, x_939); -lean_inc(x_724); -x_942 = l_Lean_Syntax_node2(x_724, x_775, x_736, x_941); -x_943 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_724); -x_944 = l_Lean_Syntax_node3(x_724, x_943, x_782, x_784, x_942); -x_945 = l_Lean_Syntax_node2(x_724, x_732, x_776, x_944); -if (lean_is_scalar(x_729)) { - x_946 = lean_alloc_ctor(0, 2, 0); -} else { - x_946 = x_729; -} -lean_ctor_set(x_946, 0, x_945); -lean_ctor_set(x_946, 1, x_728); -return x_946; -} -else -{ -lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; -lean_dec(x_717); -lean_dec(x_716); -lean_dec(x_715); -lean_dec(x_714); -lean_free_object(x_19); -lean_dec(x_26); -lean_free_object(x_18); -lean_dec(x_10); -lean_dec(x_4); -x_947 = lean_ctor_get(x_719, 0); -lean_inc(x_947); -x_948 = lean_ctor_get(x_719, 1); -lean_inc(x_948); -if (lean_is_exclusive(x_719)) { - lean_ctor_release(x_719, 0); - lean_ctor_release(x_719, 1); - x_949 = x_719; -} else { - lean_dec_ref(x_719); - x_949 = lean_box(0); -} -if (lean_is_scalar(x_949)) { - x_950 = lean_alloc_ctor(1, 2, 0); -} else { - x_950 = x_949; -} -lean_ctor_set(x_950, 0, x_947); -lean_ctor_set(x_950, 1, x_948); -return x_950; -} -} +lean_inc(x_834); +lean_inc(x_736); +x_849 = l_Lean_Syntax_node3(x_736, x_848, x_843, x_834, x_847); +x_850 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_737); +lean_inc(x_743); +x_851 = l_Lean_addMacroScope(x_743, x_850, x_737); +x_852 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_853 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_736); +x_854 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_854, 0, x_736); +lean_ctor_set(x_854, 1, x_852); +lean_ctor_set(x_854, 2, x_851); +lean_ctor_set(x_854, 3, x_853); +lean_inc(x_746); +lean_inc(x_736); +x_855 = l_Lean_Syntax_node2(x_736, x_842, x_854, x_746); +x_856 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_737); +lean_inc(x_743); +x_857 = l_Lean_addMacroScope(x_743, x_856, x_737); +x_858 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_736); +x_859 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_859, 0, x_736); +lean_ctor_set(x_859, 1, x_858); +lean_ctor_set(x_859, 2, x_857); +lean_ctor_set(x_859, 3, x_13); +lean_inc(x_859); +lean_inc(x_834); +lean_inc(x_736); +x_860 = l_Lean_Syntax_node3(x_736, x_848, x_855, x_834, x_859); +lean_inc(x_736); +x_861 = l_Lean_Syntax_node3(x_736, x_744, x_849, x_18, x_860); +x_862 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_736); +x_863 = l_Lean_Syntax_node1(x_736, x_862, x_861); +x_864 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_746); +lean_inc(x_736); +x_865 = l_Lean_Syntax_node1(x_736, x_864, x_746); +x_866 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_736); +x_867 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_867, 0, x_736); +lean_ctor_set(x_867, 1, x_866); +x_868 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_867); +lean_inc(x_865); +lean_inc_n(x_746, 2); +lean_inc(x_836); +lean_inc(x_736); +x_869 = l_Lean_Syntax_node6(x_736, x_868, x_836, x_746, x_863, x_865, x_746, x_867); +x_870 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_746, 2); +lean_inc(x_736); +x_871 = l_Lean_Syntax_node2(x_736, x_870, x_746, x_746); +x_872 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +lean_inc(x_737); +lean_inc(x_743); +x_873 = l_Lean_addMacroScope(x_743, x_872, x_737); +x_874 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +lean_inc(x_736); +x_875 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_875, 0, x_736); +lean_ctor_set(x_875, 1, x_874); +lean_ctor_set(x_875, 2, x_873); +lean_ctor_set(x_875, 3, x_13); +lean_inc(x_875); +lean_inc(x_736); +x_876 = l_Lean_Syntax_node1(x_736, x_744, x_875); +x_877 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; +lean_inc(x_736); +x_878 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_878, 0, x_736); +lean_ctor_set(x_878, 1, x_877); +x_879 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; +lean_inc(x_737); +lean_inc(x_743); +x_880 = l_Lean_addMacroScope(x_743, x_879, x_737); +x_881 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_882 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +lean_inc(x_736); +x_883 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_883, 0, x_736); +lean_ctor_set(x_883, 1, x_881); +lean_ctor_set(x_883, 2, x_880); +lean_ctor_set(x_883, 3, x_882); +x_884 = lean_array_size(x_726); +x_885 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_848, x_884, x_16, x_726); +x_886 = lean_array_size(x_885); +x_887 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_886, x_16, x_885); +x_888 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_889 = l_Lean_mkSepArray(x_887, x_888); +lean_dec(x_887); +x_890 = l_Array_append___rarg(x_745, x_889); +lean_dec(x_889); +lean_inc(x_736); +x_891 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_891, 0, x_736); +lean_ctor_set(x_891, 1, x_744); +lean_ctor_set(x_891, 2, x_890); +lean_inc(x_736); +x_892 = l_Lean_Syntax_node1(x_736, x_862, x_891); +lean_inc(x_756); +lean_inc(x_802); +lean_inc(x_736); +x_893 = l_Lean_Syntax_node2(x_736, x_744, x_802, x_756); +lean_inc(x_867); +lean_inc(x_865); +lean_inc(x_746); +lean_inc(x_836); +lean_inc(x_736); +x_894 = l_Lean_Syntax_node6(x_736, x_868, x_836, x_746, x_892, x_865, x_893, x_867); +lean_inc(x_736); +x_895 = l_Lean_Syntax_node1(x_736, x_744, x_894); +lean_inc(x_736); +x_896 = l_Lean_Syntax_node2(x_736, x_821, x_883, x_895); +lean_inc(x_736); +x_897 = l_Lean_Syntax_node1(x_736, x_744, x_896); +x_898 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_878); +lean_inc(x_736); +x_899 = l_Lean_Syntax_node2(x_736, x_898, x_878, x_897); +x_900 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_834); +lean_inc(x_746); +lean_inc(x_736); +x_901 = l_Lean_Syntax_node5(x_736, x_900, x_847, x_876, x_746, x_834, x_899); +x_902 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_736); +x_903 = l_Lean_Syntax_node1(x_736, x_902, x_901); +x_904 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_871); +lean_inc_n(x_746, 2); +lean_inc(x_736); +x_905 = l_Lean_Syntax_node4(x_736, x_904, x_746, x_746, x_903, x_871); +x_906 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_737); +lean_inc(x_743); +x_907 = l_Lean_addMacroScope(x_743, x_906, x_737); +x_908 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_736); +x_909 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_909, 0, x_736); +lean_ctor_set(x_909, 1, x_908); +lean_ctor_set(x_909, 2, x_907); +lean_ctor_set(x_909, 3, x_13); +lean_inc(x_736); +x_910 = l_Lean_Syntax_node1(x_736, x_744, x_909); +x_911 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_736); +x_912 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_912, 0, x_736); +lean_ctor_set(x_912, 1, x_911); +x_913 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_736); +x_914 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_914, 0, x_736); +lean_ctor_set(x_914, 1, x_913); +lean_inc(x_736); +x_915 = l_Lean_Syntax_node2(x_736, x_829, x_802, x_756); +lean_inc(x_736); +x_916 = l_Lean_Syntax_node1(x_736, x_744, x_915); +x_917 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_736); +x_918 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_918, 0, x_736); +lean_ctor_set(x_918, 1, x_917); +x_919 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_920 = l_Lean_addMacroScope(x_743, x_919, x_737); +x_921 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_922 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_736); +x_923 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_923, 0, x_736); +lean_ctor_set(x_923, 1, x_921); +lean_ctor_set(x_923, 2, x_920); +lean_ctor_set(x_923, 3, x_922); +lean_inc(x_910); +lean_inc(x_736); +x_924 = l_Lean_Syntax_node2(x_736, x_821, x_923, x_910); +x_925 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_736); +x_926 = l_Lean_Syntax_node1(x_736, x_925, x_924); +x_927 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_736); +x_928 = l_Lean_Syntax_node4(x_736, x_927, x_875, x_916, x_918, x_926); +x_929 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_746); +lean_inc(x_736); +x_930 = l_Lean_Syntax_node3(x_736, x_929, x_914, x_746, x_928); +x_931 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_746); +lean_inc(x_736); +x_932 = l_Lean_Syntax_node2(x_736, x_931, x_930, x_746); +x_933 = l_Lean_Syntax_SepArray_ofElems(x_775, x_26); +lean_dec(x_26); +x_934 = l_Array_append___rarg(x_745, x_933); +lean_dec(x_933); +lean_inc(x_736); +x_935 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_935, 0, x_736); +lean_ctor_set(x_935, 1, x_744); +lean_ctor_set(x_935, 2, x_934); +lean_inc(x_736); +x_936 = l_Lean_Syntax_node1(x_736, x_862, x_935); +lean_inc_n(x_746, 2); +lean_inc(x_736); +x_937 = l_Lean_Syntax_node6(x_736, x_868, x_836, x_746, x_936, x_865, x_746, x_867); +lean_inc(x_736); +x_938 = l_Lean_Syntax_node1(x_736, x_744, x_937); +x_939 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; +lean_inc(x_736); +x_940 = l_Lean_Syntax_node2(x_736, x_939, x_878, x_938); +lean_inc(x_746); +lean_inc(x_736); +x_941 = l_Lean_Syntax_node2(x_736, x_931, x_940, x_746); +lean_inc(x_736); +x_942 = l_Lean_Syntax_node2(x_736, x_744, x_932, x_941); +x_943 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_736); +x_944 = l_Lean_Syntax_node1(x_736, x_943, x_942); +x_945 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_736); +x_946 = l_Lean_Syntax_node2(x_736, x_945, x_912, x_944); +lean_inc(x_834); +lean_inc(x_746); +lean_inc(x_736); +x_947 = l_Lean_Syntax_node5(x_736, x_900, x_859, x_910, x_746, x_834, x_946); +lean_inc(x_736); +x_948 = l_Lean_Syntax_node1(x_736, x_902, x_947); +lean_inc(x_871); +lean_inc_n(x_746, 2); +lean_inc(x_736); +x_949 = l_Lean_Syntax_node4(x_736, x_904, x_746, x_746, x_948, x_871); +lean_inc(x_746); +lean_inc(x_736); +x_950 = l_Lean_Syntax_node3(x_736, x_744, x_905, x_746, x_949); +x_951 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_736); +x_952 = l_Lean_Syntax_node2(x_736, x_951, x_760, x_950); +lean_inc(x_736); +x_953 = l_Lean_Syntax_node1(x_736, x_744, x_952); +x_954 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_736); +x_955 = l_Lean_Syntax_node4(x_736, x_954, x_834, x_869, x_871, x_953); +x_956 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_746); +lean_inc(x_736); +x_957 = l_Lean_Syntax_node6(x_736, x_956, x_798, x_800, x_746, x_746, x_832, x_955); +lean_inc(x_736); +x_958 = l_Lean_Syntax_node2(x_736, x_787, x_748, x_957); +x_959 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_736); +x_960 = l_Lean_Syntax_node3(x_736, x_959, x_794, x_796, x_958); +x_961 = l_Lean_Syntax_node2(x_736, x_744, x_788, x_960); +if (lean_is_scalar(x_741)) { + x_962 = lean_alloc_ctor(0, 2, 0); +} else { + x_962 = x_741; +} +lean_ctor_set(x_962, 0, x_961); +lean_ctor_set(x_962, 1, x_740); +return x_962; +} +else +{ +lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; +lean_dec(x_729); +lean_dec(x_728); +lean_dec(x_727); +lean_dec(x_726); +lean_free_object(x_19); +lean_dec(x_26); +lean_free_object(x_18); +lean_dec(x_10); +lean_dec(x_4); +x_963 = lean_ctor_get(x_731, 0); +lean_inc(x_963); +x_964 = lean_ctor_get(x_731, 1); +lean_inc(x_964); +if (lean_is_exclusive(x_731)) { + lean_ctor_release(x_731, 0); + lean_ctor_release(x_731, 1); + x_965 = x_731; +} else { + lean_dec_ref(x_731); + x_965 = lean_box(0); +} +if (lean_is_scalar(x_965)) { + x_966 = lean_alloc_ctor(1, 2, 0); +} else { + x_966 = x_965; +} +lean_ctor_set(x_966, 0, x_963); +lean_ctor_set(x_966, 1, x_964); +return x_966; +} +} } else { -lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; size_t x_957; lean_object* x_958; -x_951 = lean_ctor_get(x_19, 0); -lean_inc(x_951); +lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; size_t x_973; lean_object* x_974; +x_967 = lean_ctor_get(x_19, 0); +lean_inc(x_967); lean_dec(x_19); -x_952 = lean_ctor_get(x_20, 0); -lean_inc(x_952); +x_968 = lean_ctor_get(x_20, 0); +lean_inc(x_968); if (lean_is_exclusive(x_20)) { lean_ctor_release(x_20, 0); lean_ctor_release(x_20, 1); - x_953 = x_20; + x_969 = x_20; } else { lean_dec_ref(x_20); - x_953 = lean_box(0); + x_969 = lean_box(0); } -x_954 = lean_ctor_get(x_21, 0); -lean_inc(x_954); -x_955 = lean_ctor_get(x_21, 1); -lean_inc(x_955); +x_970 = lean_ctor_get(x_21, 0); +lean_inc(x_970); +x_971 = lean_ctor_get(x_21, 1); +lean_inc(x_971); if (lean_is_exclusive(x_21)) { lean_ctor_release(x_21, 0); lean_ctor_release(x_21, 1); - x_956 = x_21; + x_972 = x_21; } else { lean_dec_ref(x_21); - x_956 = lean_box(0); -} -x_957 = lean_array_size(x_2); -x_958 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_957, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); -if (lean_obj_tag(x_958) == 0) -{ -lean_object* x_959; lean_object* x_960; lean_object* x_961; uint8_t x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; size_t x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; size_t x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; size_t x_1110; lean_object* x_1111; size_t x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; -x_959 = lean_ctor_get(x_958, 0); -lean_inc(x_959); -x_960 = lean_ctor_get(x_958, 1); -lean_inc(x_960); -lean_dec(x_958); -x_961 = lean_ctor_get(x_10, 5); -lean_inc(x_961); -x_962 = 0; -x_963 = l_Lean_SourceInfo_fromRef(x_961, x_962); -lean_dec(x_961); -x_964 = lean_ctor_get(x_10, 10); -lean_inc(x_964); + x_972 = lean_box(0); +} +x_973 = lean_array_size(x_2); +x_974 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_973, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_974) == 0) +{ +lean_object* x_975; lean_object* x_976; lean_object* x_977; uint8_t x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; size_t x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; size_t x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; size_t x_1128; lean_object* x_1129; size_t x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; +x_975 = lean_ctor_get(x_974, 0); +lean_inc(x_975); +x_976 = lean_ctor_get(x_974, 1); +lean_inc(x_976); +lean_dec(x_974); +x_977 = lean_ctor_get(x_10, 5); +lean_inc(x_977); +x_978 = 0; +x_979 = l_Lean_SourceInfo_fromRef(x_977, x_978); +lean_dec(x_977); +x_980 = lean_ctor_get(x_10, 10); +lean_inc(x_980); lean_dec(x_10); -x_965 = lean_st_ref_get(x_11, x_960); -x_966 = lean_ctor_get(x_965, 0); -lean_inc(x_966); -x_967 = lean_ctor_get(x_965, 1); -lean_inc(x_967); -if (lean_is_exclusive(x_965)) { - lean_ctor_release(x_965, 0); - lean_ctor_release(x_965, 1); - x_968 = x_965; +x_981 = lean_st_ref_get(x_11, x_976); +x_982 = lean_ctor_get(x_981, 0); +lean_inc(x_982); +x_983 = lean_ctor_get(x_981, 1); +lean_inc(x_983); +if (lean_is_exclusive(x_981)) { + lean_ctor_release(x_981, 0); + lean_ctor_release(x_981, 1); + x_984 = x_981; } else { - lean_dec_ref(x_965); - x_968 = lean_box(0); -} -x_969 = lean_ctor_get(x_966, 0); -lean_inc(x_969); -lean_dec(x_966); -x_970 = lean_environment_main_module(x_969); -x_971 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -x_972 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -lean_inc(x_963); -x_973 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_973, 0, x_963); -lean_ctor_set(x_973, 1, x_971); -lean_ctor_set(x_973, 2, x_972); -x_974 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_973, 6); -lean_inc(x_963); -x_975 = l_Lean_Syntax_node6(x_963, x_974, x_973, x_973, x_973, x_973, x_973, x_973); -x_976 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; -lean_inc(x_963); -if (lean_is_scalar(x_956)) { - x_977 = lean_alloc_ctor(2, 2, 0); + lean_dec_ref(x_981); + x_984 = lean_box(0); +} +x_985 = lean_ctor_get(x_982, 0); +lean_inc(x_985); +lean_dec(x_982); +x_986 = lean_environment_main_module(x_985); +x_987 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +x_988 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +lean_inc(x_979); +x_989 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_989, 0, x_979); +lean_ctor_set(x_989, 1, x_987); +lean_ctor_set(x_989, 2, x_988); +x_990 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_989, 6); +lean_inc(x_979); +x_991 = l_Lean_Syntax_node6(x_979, x_990, x_989, x_989, x_989, x_989, x_989, x_989); +x_992 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; +lean_inc(x_979); +if (lean_is_scalar(x_972)) { + x_993 = lean_alloc_ctor(2, 2, 0); } else { - x_977 = x_956; - lean_ctor_set_tag(x_977, 2); -} -lean_ctor_set(x_977, 0, x_963); -lean_ctor_set(x_977, 1, x_976); -x_978 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; -lean_inc(x_963); -x_979 = l_Lean_Syntax_node1(x_963, x_978, x_977); -x_980 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_964); -lean_inc(x_970); -x_981 = l_Lean_addMacroScope(x_970, x_980, x_964); -x_982 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_963); -x_983 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_983, 0, x_963); -lean_ctor_set(x_983, 1, x_982); -lean_ctor_set(x_983, 2, x_981); -lean_ctor_set(x_983, 3, x_13); -x_984 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_973); -lean_inc(x_983); -lean_inc(x_963); -x_985 = l_Lean_Syntax_node2(x_963, x_984, x_983, x_973); -x_986 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_963); -if (lean_is_scalar(x_953)) { - x_987 = lean_alloc_ctor(2, 2, 0); + x_993 = x_972; + lean_ctor_set_tag(x_993, 2); +} +lean_ctor_set(x_993, 0, x_979); +lean_ctor_set(x_993, 1, x_992); +x_994 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; +lean_inc(x_979); +x_995 = l_Lean_Syntax_node1(x_979, x_994, x_993); +x_996 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_980); +lean_inc(x_986); +x_997 = l_Lean_addMacroScope(x_986, x_996, x_980); +x_998 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_979); +x_999 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_999, 0, x_979); +lean_ctor_set(x_999, 1, x_998); +lean_ctor_set(x_999, 2, x_997); +lean_ctor_set(x_999, 3, x_13); +x_1000 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_989); +lean_inc(x_999); +lean_inc(x_979); +x_1001 = l_Lean_Syntax_node2(x_979, x_1000, x_999, x_989); +x_1002 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_979); +if (lean_is_scalar(x_969)) { + x_1003 = lean_alloc_ctor(2, 2, 0); } else { - x_987 = x_953; - lean_ctor_set_tag(x_987, 2); -} -lean_ctor_set(x_987, 0, x_963); -lean_ctor_set(x_987, 1, x_986); -x_988 = l_Array_zip___rarg(x_954, x_955); -lean_dec(x_955); -lean_dec(x_954); -x_989 = lean_array_size(x_988); -lean_inc(x_975); -lean_inc(x_973); -lean_inc(x_963); -x_990 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_963, x_971, x_973, x_975, x_989, x_16, x_988); -x_991 = l_Array_append___rarg(x_972, x_990); -lean_dec(x_990); -lean_inc(x_963); -x_992 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_992, 0, x_963); -lean_ctor_set(x_992, 1, x_971); -lean_ctor_set(x_992, 2, x_991); -x_993 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_963); -x_994 = l_Lean_Syntax_node1(x_963, x_993, x_992); -lean_inc(x_973); -lean_inc(x_987); -lean_inc(x_963); -x_995 = l_Lean_Syntax_node3(x_963, x_971, x_987, x_973, x_994); -x_996 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_963); -x_997 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_997, 0, x_963); -lean_ctor_set(x_997, 1, x_996); -x_998 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_964); -lean_inc(x_970); -x_999 = l_Lean_addMacroScope(x_970, x_998, x_964); -x_1000 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_1001 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_963); -x_1002 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1002, 0, x_963); -lean_ctor_set(x_1002, 1, x_1000); -lean_ctor_set(x_1002, 2, x_999); -lean_ctor_set(x_1002, 3, x_1001); -x_1003 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_963); + x_1003 = x_969; + lean_ctor_set_tag(x_1003, 2); +} +lean_ctor_set(x_1003, 0, x_979); +lean_ctor_set(x_1003, 1, x_1002); +x_1004 = l_Array_zip___rarg(x_970, x_971); +lean_dec(x_971); +lean_dec(x_970); +x_1005 = lean_array_size(x_1004); +lean_inc(x_991); +lean_inc(x_989); +lean_inc(x_979); +x_1006 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_979, x_987, x_989, x_991, x_1005, x_16, x_1004); +x_1007 = l_Array_append___rarg(x_988, x_1006); +lean_dec(x_1006); +lean_inc(x_979); +x_1008 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1008, 0, x_979); +lean_ctor_set(x_1008, 1, x_987); +lean_ctor_set(x_1008, 2, x_1007); +x_1009 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; +lean_inc(x_979); +x_1010 = l_Lean_Syntax_node1(x_979, x_1009, x_1008); +lean_inc(x_989); +lean_inc(x_1003); +lean_inc(x_979); +x_1011 = l_Lean_Syntax_node3(x_979, x_987, x_1003, x_989, x_1010); +x_1012 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_979); +x_1013 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1013, 0, x_979); +lean_ctor_set(x_1013, 1, x_1012); +x_1014 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_980); +lean_inc(x_986); +x_1015 = l_Lean_addMacroScope(x_986, x_1014, x_980); +x_1016 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_1017 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_979); +x_1018 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1018, 0, x_979); +lean_ctor_set(x_1018, 1, x_1016); +lean_ctor_set(x_1018, 2, x_1015); +lean_ctor_set(x_1018, 3, x_1017); +x_1019 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_979); lean_ctor_set_tag(x_18, 2); -lean_ctor_set(x_18, 1, x_1003); -lean_ctor_set(x_18, 0, x_963); -x_1004 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_964); -lean_inc(x_970); -x_1005 = l_Lean_addMacroScope(x_970, x_1004, x_964); -x_1006 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_1007 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_963); -x_1008 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1008, 0, x_963); -lean_ctor_set(x_1008, 1, x_1006); -lean_ctor_set(x_1008, 2, x_1005); -lean_ctor_set(x_1008, 3, x_1007); +lean_ctor_set(x_18, 1, x_1019); +lean_ctor_set(x_18, 0, x_979); +x_1020 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_980); +lean_inc(x_986); +x_1021 = l_Lean_addMacroScope(x_986, x_1020, x_980); +x_1022 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_1023 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_979); +x_1024 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1024, 0, x_979); +lean_ctor_set(x_1024, 1, x_1022); +lean_ctor_set(x_1024, 2, x_1021); +lean_ctor_set(x_1024, 3, x_1023); lean_inc(x_18); -lean_inc(x_963); -x_1009 = l_Lean_Syntax_node3(x_963, x_971, x_1002, x_18, x_1008); -lean_inc(x_963); -x_1010 = l_Lean_Syntax_node2(x_963, x_971, x_997, x_1009); -x_1011 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_963); -x_1012 = l_Lean_Syntax_node1(x_963, x_1011, x_1010); -x_1013 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; -lean_inc_n(x_973, 3); -lean_inc(x_963); -x_1014 = l_Lean_Syntax_node7(x_963, x_1013, x_979, x_985, x_973, x_973, x_973, x_995, x_1012); -x_1015 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_975); -lean_inc(x_963); -x_1016 = l_Lean_Syntax_node2(x_963, x_1015, x_975, x_1014); -x_1017 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_963); -x_1018 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1018, 0, x_963); -lean_ctor_set(x_1018, 1, x_1017); -x_1019 = l_Array_append___rarg(x_972, x_3); -lean_inc(x_963); -x_1020 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1020, 0, x_963); -lean_ctor_set(x_1020, 1, x_971); -lean_ctor_set(x_1020, 2, x_1019); -x_1021 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_963); -x_1022 = l_Lean_Syntax_node2(x_963, x_1021, x_1018, x_1020); -x_1023 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_963); -x_1024 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1024, 0, x_963); -lean_ctor_set(x_1024, 1, x_1023); -x_1025 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_973); -lean_inc(x_963); -x_1026 = l_Lean_Syntax_node1(x_963, x_1025, x_973); -x_1027 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_963); -x_1028 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1028, 0, x_963); -lean_ctor_set(x_1028, 1, x_1027); -x_1029 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_963); -x_1030 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1030, 0, x_963); -lean_ctor_set(x_1030, 1, x_1029); -x_1031 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_964); -lean_inc(x_970); -x_1032 = l_Lean_addMacroScope(x_970, x_1031, x_964); -x_1033 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_1034 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_963); -x_1035 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1035, 0, x_963); -lean_ctor_set(x_1035, 1, x_1033); -lean_ctor_set(x_1035, 2, x_1032); -lean_ctor_set(x_1035, 3, x_1034); -x_1036 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; -lean_inc(x_963); -x_1037 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1037, 0, x_963); -lean_ctor_set(x_1037, 1, x_1036); -x_1038 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_963); -x_1039 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1039, 0, x_963); -lean_ctor_set(x_1039, 1, x_1038); -x_1040 = lean_box(0); -x_1041 = l_Lean_mkCIdentFrom(x_1040, x_4, x_962); -x_1042 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_963); -x_1043 = l_Lean_Syntax_node2(x_963, x_1042, x_1039, x_1041); -x_1044 = lean_array_size(x_959); -x_1045 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; -x_1046 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_1045, x_1044, x_16, x_959); -x_1047 = l_Array_append___rarg(x_972, x_1046); -lean_dec(x_1046); -lean_inc(x_963); -x_1048 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1048, 0, x_963); -lean_ctor_set(x_1048, 1, x_971); -lean_ctor_set(x_1048, 2, x_1047); -x_1049 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_963); -x_1050 = l_Lean_Syntax_node2(x_963, x_1049, x_1043, x_1048); -x_1051 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; -lean_inc(x_963); -x_1052 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1052, 0, x_963); -lean_ctor_set(x_1052, 1, x_1051); -x_1053 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; -lean_inc(x_963); -x_1054 = l_Lean_Syntax_node3(x_963, x_1053, x_1037, x_1050, x_1052); -lean_inc(x_963); -x_1055 = l_Lean_Syntax_node1(x_963, x_971, x_1054); -lean_inc(x_963); -x_1056 = l_Lean_Syntax_node2(x_963, x_1049, x_1035, x_1055); -x_1057 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_1030); -lean_inc(x_963); -x_1058 = l_Lean_Syntax_node2(x_963, x_1057, x_1030, x_1056); -x_1059 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_973); -lean_inc(x_963); -x_1060 = l_Lean_Syntax_node2(x_963, x_1059, x_973, x_1058); -x_1061 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_963); -x_1062 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1062, 0, x_963); -lean_ctor_set(x_1062, 1, x_1061); -x_1063 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_963); -x_1064 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1064, 0, x_963); -lean_ctor_set(x_1064, 1, x_1063); -x_1065 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_964); -lean_inc(x_970); -x_1066 = l_Lean_addMacroScope(x_970, x_1065, x_964); -x_1067 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_1068 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_963); -x_1069 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1069, 0, x_963); -lean_ctor_set(x_1069, 1, x_1067); -lean_ctor_set(x_1069, 2, x_1066); -lean_ctor_set(x_1069, 3, x_1068); -x_1070 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_973); -lean_inc(x_963); -x_1071 = l_Lean_Syntax_node2(x_963, x_1070, x_1069, x_973); -x_1072 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_964); -lean_inc(x_970); -x_1073 = l_Lean_addMacroScope(x_970, x_1072, x_964); -x_1074 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_963); -x_1075 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1075, 0, x_963); -lean_ctor_set(x_1075, 1, x_1074); -lean_ctor_set(x_1075, 2, x_1073); -lean_ctor_set(x_1075, 3, x_13); -x_1076 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_1075); -lean_inc(x_1062); -lean_inc(x_963); -x_1077 = l_Lean_Syntax_node3(x_963, x_1076, x_1071, x_1062, x_1075); -x_1078 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_964); -lean_inc(x_970); -x_1079 = l_Lean_addMacroScope(x_970, x_1078, x_964); -x_1080 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_1081 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_963); -x_1082 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1082, 0, x_963); -lean_ctor_set(x_1082, 1, x_1080); -lean_ctor_set(x_1082, 2, x_1079); -lean_ctor_set(x_1082, 3, x_1081); -lean_inc(x_973); -lean_inc(x_963); -x_1083 = l_Lean_Syntax_node2(x_963, x_1070, x_1082, x_973); -x_1084 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_964); -lean_inc(x_970); -x_1085 = l_Lean_addMacroScope(x_970, x_1084, x_964); -x_1086 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_963); -x_1087 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1087, 0, x_963); -lean_ctor_set(x_1087, 1, x_1086); -lean_ctor_set(x_1087, 2, x_1085); -lean_ctor_set(x_1087, 3, x_13); -lean_inc(x_1087); -lean_inc(x_1062); -lean_inc(x_963); -x_1088 = l_Lean_Syntax_node3(x_963, x_1076, x_1083, x_1062, x_1087); -lean_inc(x_963); -x_1089 = l_Lean_Syntax_node3(x_963, x_971, x_1077, x_18, x_1088); -x_1090 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_973); -lean_inc(x_963); -x_1091 = l_Lean_Syntax_node1(x_963, x_1090, x_973); -x_1092 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_963); -x_1093 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1093, 0, x_963); -lean_ctor_set(x_1093, 1, x_1092); -x_1094 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc(x_1093); -lean_inc(x_1091); -lean_inc_n(x_973, 2); -lean_inc(x_1064); -lean_inc(x_963); -x_1095 = l_Lean_Syntax_node6(x_963, x_1094, x_1064, x_973, x_1089, x_1091, x_973, x_1093); -x_1096 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_973, 2); -lean_inc(x_963); -x_1097 = l_Lean_Syntax_node2(x_963, x_1096, x_973, x_973); -x_1098 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; -lean_inc(x_964); -lean_inc(x_970); -x_1099 = l_Lean_addMacroScope(x_970, x_1098, x_964); -x_1100 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; -lean_inc(x_963); -x_1101 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1101, 0, x_963); -lean_ctor_set(x_1101, 1, x_1100); -lean_ctor_set(x_1101, 2, x_1099); -lean_ctor_set(x_1101, 3, x_13); -lean_inc(x_1101); -lean_inc(x_963); -x_1102 = l_Lean_Syntax_node1(x_963, x_971, x_1101); -x_1103 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; -lean_inc(x_963); -x_1104 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1104, 0, x_963); -lean_ctor_set(x_1104, 1, x_1103); -x_1105 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; -lean_inc(x_964); -lean_inc(x_970); -x_1106 = l_Lean_addMacroScope(x_970, x_1105, x_964); -x_1107 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_1108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; -lean_inc(x_963); -x_1109 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1109, 0, x_963); -lean_ctor_set(x_1109, 1, x_1107); -lean_ctor_set(x_1109, 2, x_1106); -lean_ctor_set(x_1109, 3, x_1108); -x_1110 = lean_array_size(x_952); -x_1111 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_1076, x_1110, x_16, x_952); -x_1112 = lean_array_size(x_1111); -x_1113 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_1112, x_16, x_1111); -x_1114 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_1115 = l_Lean_mkSepArray(x_1113, x_1114); -lean_dec(x_1113); -x_1116 = l_Array_append___rarg(x_972, x_1115); -lean_dec(x_1115); -lean_inc(x_963); -x_1117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1117, 0, x_963); -lean_ctor_set(x_1117, 1, x_971); -lean_ctor_set(x_1117, 2, x_1116); -lean_inc(x_983); -lean_inc(x_1030); -lean_inc(x_963); -x_1118 = l_Lean_Syntax_node2(x_963, x_971, x_1030, x_983); -lean_inc(x_1093); +lean_inc(x_979); +x_1025 = l_Lean_Syntax_node3(x_979, x_987, x_1018, x_18, x_1024); +lean_inc(x_979); +x_1026 = l_Lean_Syntax_node2(x_979, x_987, x_1013, x_1025); +x_1027 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_979); +x_1028 = l_Lean_Syntax_node1(x_979, x_1027, x_1026); +x_1029 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; +lean_inc_n(x_989, 3); +lean_inc(x_979); +x_1030 = l_Lean_Syntax_node7(x_979, x_1029, x_995, x_1001, x_989, x_989, x_989, x_1011, x_1028); +x_1031 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_991); +lean_inc(x_979); +x_1032 = l_Lean_Syntax_node2(x_979, x_1031, x_991, x_1030); +x_1033 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_979); +x_1034 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1034, 0, x_979); +lean_ctor_set(x_1034, 1, x_1033); +x_1035 = l_Array_append___rarg(x_988, x_3); +lean_inc(x_979); +x_1036 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1036, 0, x_979); +lean_ctor_set(x_1036, 1, x_987); +lean_ctor_set(x_1036, 2, x_1035); +x_1037 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_979); +x_1038 = l_Lean_Syntax_node2(x_979, x_1037, x_1034, x_1036); +x_1039 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_979); +x_1040 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1040, 0, x_979); +lean_ctor_set(x_1040, 1, x_1039); +x_1041 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_989); +lean_inc(x_979); +x_1042 = l_Lean_Syntax_node1(x_979, x_1041, x_989); +x_1043 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_979); +x_1044 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1044, 0, x_979); +lean_ctor_set(x_1044, 1, x_1043); +x_1045 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_979); +x_1046 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1046, 0, x_979); +lean_ctor_set(x_1046, 1, x_1045); +x_1047 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_980); +lean_inc(x_986); +x_1048 = l_Lean_addMacroScope(x_986, x_1047, x_980); +x_1049 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_1050 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_979); +x_1051 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1051, 0, x_979); +lean_ctor_set(x_1051, 1, x_1049); +lean_ctor_set(x_1051, 2, x_1048); +lean_ctor_set(x_1051, 3, x_1050); +x_1052 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; +lean_inc(x_979); +x_1053 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1053, 0, x_979); +lean_ctor_set(x_1053, 1, x_1052); +x_1054 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_979); +x_1055 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1055, 0, x_979); +lean_ctor_set(x_1055, 1, x_1054); +x_1056 = lean_box(0); +x_1057 = l_Lean_mkCIdentFrom(x_1056, x_4, x_978); +x_1058 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_979); +x_1059 = l_Lean_Syntax_node2(x_979, x_1058, x_1055, x_1057); +x_1060 = lean_array_size(x_975); +x_1061 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; +x_1062 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_1061, x_1060, x_16, x_975); +x_1063 = l_Array_append___rarg(x_988, x_1062); +lean_dec(x_1062); +lean_inc(x_979); +x_1064 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1064, 0, x_979); +lean_ctor_set(x_1064, 1, x_987); +lean_ctor_set(x_1064, 2, x_1063); +x_1065 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_979); +x_1066 = l_Lean_Syntax_node2(x_979, x_1065, x_1059, x_1064); +x_1067 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; +lean_inc(x_979); +x_1068 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1068, 0, x_979); +lean_ctor_set(x_1068, 1, x_1067); +x_1069 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; +lean_inc(x_979); +x_1070 = l_Lean_Syntax_node3(x_979, x_1069, x_1053, x_1066, x_1068); +lean_inc(x_979); +x_1071 = l_Lean_Syntax_node1(x_979, x_987, x_1070); +lean_inc(x_979); +x_1072 = l_Lean_Syntax_node2(x_979, x_1065, x_1051, x_1071); +x_1073 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_1046); +lean_inc(x_979); +x_1074 = l_Lean_Syntax_node2(x_979, x_1073, x_1046, x_1072); +x_1075 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_989); +lean_inc(x_979); +x_1076 = l_Lean_Syntax_node2(x_979, x_1075, x_989, x_1074); +x_1077 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_979); +x_1078 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1078, 0, x_979); +lean_ctor_set(x_1078, 1, x_1077); +x_1079 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_979); +x_1080 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1080, 0, x_979); +lean_ctor_set(x_1080, 1, x_1079); +x_1081 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_980); +lean_inc(x_986); +x_1082 = l_Lean_addMacroScope(x_986, x_1081, x_980); +x_1083 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_1084 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_979); +x_1085 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1085, 0, x_979); +lean_ctor_set(x_1085, 1, x_1083); +lean_ctor_set(x_1085, 2, x_1082); +lean_ctor_set(x_1085, 3, x_1084); +x_1086 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_989); +lean_inc(x_979); +x_1087 = l_Lean_Syntax_node2(x_979, x_1086, x_1085, x_989); +x_1088 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_980); +lean_inc(x_986); +x_1089 = l_Lean_addMacroScope(x_986, x_1088, x_980); +x_1090 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_979); +x_1091 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1091, 0, x_979); +lean_ctor_set(x_1091, 1, x_1090); +lean_ctor_set(x_1091, 2, x_1089); +lean_ctor_set(x_1091, 3, x_13); +x_1092 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; lean_inc(x_1091); -lean_inc(x_973); -lean_inc(x_1064); -lean_inc(x_963); -x_1119 = l_Lean_Syntax_node6(x_963, x_1094, x_1064, x_973, x_1117, x_1091, x_1118, x_1093); -lean_inc(x_963); -x_1120 = l_Lean_Syntax_node1(x_963, x_971, x_1119); -lean_inc(x_963); -x_1121 = l_Lean_Syntax_node2(x_963, x_1049, x_1109, x_1120); -lean_inc(x_963); -x_1122 = l_Lean_Syntax_node1(x_963, x_971, x_1121); -x_1123 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_1104); -lean_inc(x_963); -x_1124 = l_Lean_Syntax_node2(x_963, x_1123, x_1104, x_1122); -x_1125 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_1062); -lean_inc(x_973); -lean_inc(x_963); -x_1126 = l_Lean_Syntax_node5(x_963, x_1125, x_1075, x_1102, x_973, x_1062, x_1124); -x_1127 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_963); -x_1128 = l_Lean_Syntax_node1(x_963, x_1127, x_1126); -x_1129 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_1097); -lean_inc_n(x_973, 2); -lean_inc(x_963); -x_1130 = l_Lean_Syntax_node4(x_963, x_1129, x_973, x_973, x_1128, x_1097); -x_1131 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_964); -lean_inc(x_970); -x_1132 = l_Lean_addMacroScope(x_970, x_1131, x_964); -x_1133 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_963); -x_1134 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1134, 0, x_963); -lean_ctor_set(x_1134, 1, x_1133); -lean_ctor_set(x_1134, 2, x_1132); -lean_ctor_set(x_1134, 3, x_13); -lean_inc(x_963); -x_1135 = l_Lean_Syntax_node1(x_963, x_971, x_1134); -x_1136 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_963); -x_1137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1137, 0, x_963); -lean_ctor_set(x_1137, 1, x_1136); -x_1138 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_963); -x_1139 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1139, 0, x_963); -lean_ctor_set(x_1139, 1, x_1138); -lean_inc(x_963); -x_1140 = l_Lean_Syntax_node2(x_963, x_1057, x_1030, x_983); -lean_inc(x_963); -x_1141 = l_Lean_Syntax_node1(x_963, x_971, x_1140); -x_1142 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_963); -x_1143 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1143, 0, x_963); -lean_ctor_set(x_1143, 1, x_1142); -x_1144 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_1145 = l_Lean_addMacroScope(x_970, x_1144, x_964); -x_1146 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_1147 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_963); -x_1148 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1148, 0, x_963); -lean_ctor_set(x_1148, 1, x_1146); -lean_ctor_set(x_1148, 2, x_1145); -lean_ctor_set(x_1148, 3, x_1147); -lean_inc(x_1135); -lean_inc(x_963); -x_1149 = l_Lean_Syntax_node2(x_963, x_1049, x_1148, x_1135); -x_1150 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_963); -x_1151 = l_Lean_Syntax_node1(x_963, x_1150, x_1149); -x_1152 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_963); -x_1153 = l_Lean_Syntax_node4(x_963, x_1152, x_1101, x_1141, x_1143, x_1151); -x_1154 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_973); -lean_inc(x_963); -x_1155 = l_Lean_Syntax_node3(x_963, x_1154, x_1139, x_973, x_1153); -x_1156 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_973); -lean_inc(x_963); -x_1157 = l_Lean_Syntax_node2(x_963, x_1156, x_1155, x_973); -x_1158 = l_Lean_Syntax_SepArray_ofElems(x_1003, x_951); -lean_dec(x_951); -x_1159 = l_Array_append___rarg(x_972, x_1158); -lean_dec(x_1158); -lean_inc(x_963); -x_1160 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1160, 0, x_963); -lean_ctor_set(x_1160, 1, x_971); -lean_ctor_set(x_1160, 2, x_1159); -lean_inc_n(x_973, 2); -lean_inc(x_963); -x_1161 = l_Lean_Syntax_node6(x_963, x_1094, x_1064, x_973, x_1160, x_1091, x_973, x_1093); -lean_inc(x_963); -x_1162 = l_Lean_Syntax_node1(x_963, x_971, x_1161); -x_1163 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; -lean_inc(x_963); -x_1164 = l_Lean_Syntax_node2(x_963, x_1163, x_1104, x_1162); -lean_inc(x_973); -lean_inc(x_963); -x_1165 = l_Lean_Syntax_node2(x_963, x_1156, x_1164, x_973); -lean_inc(x_963); -x_1166 = l_Lean_Syntax_node2(x_963, x_971, x_1157, x_1165); -x_1167 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_963); -x_1168 = l_Lean_Syntax_node1(x_963, x_1167, x_1166); -x_1169 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_963); -x_1170 = l_Lean_Syntax_node2(x_963, x_1169, x_1137, x_1168); -lean_inc(x_1062); -lean_inc(x_973); -lean_inc(x_963); -x_1171 = l_Lean_Syntax_node5(x_963, x_1125, x_1087, x_1135, x_973, x_1062, x_1170); -lean_inc(x_963); -x_1172 = l_Lean_Syntax_node1(x_963, x_1127, x_1171); -lean_inc(x_1097); -lean_inc_n(x_973, 2); -lean_inc(x_963); -x_1173 = l_Lean_Syntax_node4(x_963, x_1129, x_973, x_973, x_1172, x_1097); -lean_inc(x_973); -lean_inc(x_963); -x_1174 = l_Lean_Syntax_node3(x_963, x_971, x_1130, x_973, x_1173); -x_1175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_963); -x_1176 = l_Lean_Syntax_node2(x_963, x_1175, x_987, x_1174); -lean_inc(x_963); -x_1177 = l_Lean_Syntax_node1(x_963, x_971, x_1176); -x_1178 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_963); -x_1179 = l_Lean_Syntax_node4(x_963, x_1178, x_1062, x_1095, x_1097, x_1177); -x_1180 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_973); -lean_inc(x_963); -x_1181 = l_Lean_Syntax_node6(x_963, x_1180, x_1026, x_1028, x_973, x_973, x_1060, x_1179); -lean_inc(x_963); -x_1182 = l_Lean_Syntax_node2(x_963, x_1015, x_975, x_1181); -x_1183 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_963); -x_1184 = l_Lean_Syntax_node3(x_963, x_1183, x_1022, x_1024, x_1182); -x_1185 = l_Lean_Syntax_node2(x_963, x_971, x_1016, x_1184); -if (lean_is_scalar(x_968)) { - x_1186 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_1078); +lean_inc(x_979); +x_1093 = l_Lean_Syntax_node3(x_979, x_1092, x_1087, x_1078, x_1091); +x_1094 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_980); +lean_inc(x_986); +x_1095 = l_Lean_addMacroScope(x_986, x_1094, x_980); +x_1096 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_1097 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_979); +x_1098 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1098, 0, x_979); +lean_ctor_set(x_1098, 1, x_1096); +lean_ctor_set(x_1098, 2, x_1095); +lean_ctor_set(x_1098, 3, x_1097); +lean_inc(x_989); +lean_inc(x_979); +x_1099 = l_Lean_Syntax_node2(x_979, x_1086, x_1098, x_989); +x_1100 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_980); +lean_inc(x_986); +x_1101 = l_Lean_addMacroScope(x_986, x_1100, x_980); +x_1102 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_979); +x_1103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1103, 0, x_979); +lean_ctor_set(x_1103, 1, x_1102); +lean_ctor_set(x_1103, 2, x_1101); +lean_ctor_set(x_1103, 3, x_13); +lean_inc(x_1103); +lean_inc(x_1078); +lean_inc(x_979); +x_1104 = l_Lean_Syntax_node3(x_979, x_1092, x_1099, x_1078, x_1103); +lean_inc(x_979); +x_1105 = l_Lean_Syntax_node3(x_979, x_987, x_1093, x_18, x_1104); +x_1106 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_979); +x_1107 = l_Lean_Syntax_node1(x_979, x_1106, x_1105); +x_1108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_989); +lean_inc(x_979); +x_1109 = l_Lean_Syntax_node1(x_979, x_1108, x_989); +x_1110 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_979); +x_1111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1111, 0, x_979); +lean_ctor_set(x_1111, 1, x_1110); +x_1112 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_1111); +lean_inc(x_1109); +lean_inc_n(x_989, 2); +lean_inc(x_1080); +lean_inc(x_979); +x_1113 = l_Lean_Syntax_node6(x_979, x_1112, x_1080, x_989, x_1107, x_1109, x_989, x_1111); +x_1114 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_989, 2); +lean_inc(x_979); +x_1115 = l_Lean_Syntax_node2(x_979, x_1114, x_989, x_989); +x_1116 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +lean_inc(x_980); +lean_inc(x_986); +x_1117 = l_Lean_addMacroScope(x_986, x_1116, x_980); +x_1118 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +lean_inc(x_979); +x_1119 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1119, 0, x_979); +lean_ctor_set(x_1119, 1, x_1118); +lean_ctor_set(x_1119, 2, x_1117); +lean_ctor_set(x_1119, 3, x_13); +lean_inc(x_1119); +lean_inc(x_979); +x_1120 = l_Lean_Syntax_node1(x_979, x_987, x_1119); +x_1121 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; +lean_inc(x_979); +x_1122 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1122, 0, x_979); +lean_ctor_set(x_1122, 1, x_1121); +x_1123 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; +lean_inc(x_980); +lean_inc(x_986); +x_1124 = l_Lean_addMacroScope(x_986, x_1123, x_980); +x_1125 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_1126 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +lean_inc(x_979); +x_1127 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1127, 0, x_979); +lean_ctor_set(x_1127, 1, x_1125); +lean_ctor_set(x_1127, 2, x_1124); +lean_ctor_set(x_1127, 3, x_1126); +x_1128 = lean_array_size(x_968); +x_1129 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_1092, x_1128, x_16, x_968); +x_1130 = lean_array_size(x_1129); +x_1131 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_1130, x_16, x_1129); +x_1132 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_1133 = l_Lean_mkSepArray(x_1131, x_1132); +lean_dec(x_1131); +x_1134 = l_Array_append___rarg(x_988, x_1133); +lean_dec(x_1133); +lean_inc(x_979); +x_1135 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1135, 0, x_979); +lean_ctor_set(x_1135, 1, x_987); +lean_ctor_set(x_1135, 2, x_1134); +lean_inc(x_979); +x_1136 = l_Lean_Syntax_node1(x_979, x_1106, x_1135); +lean_inc(x_999); +lean_inc(x_1046); +lean_inc(x_979); +x_1137 = l_Lean_Syntax_node2(x_979, x_987, x_1046, x_999); +lean_inc(x_1111); +lean_inc(x_1109); +lean_inc(x_989); +lean_inc(x_1080); +lean_inc(x_979); +x_1138 = l_Lean_Syntax_node6(x_979, x_1112, x_1080, x_989, x_1136, x_1109, x_1137, x_1111); +lean_inc(x_979); +x_1139 = l_Lean_Syntax_node1(x_979, x_987, x_1138); +lean_inc(x_979); +x_1140 = l_Lean_Syntax_node2(x_979, x_1065, x_1127, x_1139); +lean_inc(x_979); +x_1141 = l_Lean_Syntax_node1(x_979, x_987, x_1140); +x_1142 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_1122); +lean_inc(x_979); +x_1143 = l_Lean_Syntax_node2(x_979, x_1142, x_1122, x_1141); +x_1144 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_1078); +lean_inc(x_989); +lean_inc(x_979); +x_1145 = l_Lean_Syntax_node5(x_979, x_1144, x_1091, x_1120, x_989, x_1078, x_1143); +x_1146 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_979); +x_1147 = l_Lean_Syntax_node1(x_979, x_1146, x_1145); +x_1148 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_1115); +lean_inc_n(x_989, 2); +lean_inc(x_979); +x_1149 = l_Lean_Syntax_node4(x_979, x_1148, x_989, x_989, x_1147, x_1115); +x_1150 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_980); +lean_inc(x_986); +x_1151 = l_Lean_addMacroScope(x_986, x_1150, x_980); +x_1152 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_979); +x_1153 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1153, 0, x_979); +lean_ctor_set(x_1153, 1, x_1152); +lean_ctor_set(x_1153, 2, x_1151); +lean_ctor_set(x_1153, 3, x_13); +lean_inc(x_979); +x_1154 = l_Lean_Syntax_node1(x_979, x_987, x_1153); +x_1155 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_979); +x_1156 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1156, 0, x_979); +lean_ctor_set(x_1156, 1, x_1155); +x_1157 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_979); +x_1158 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1158, 0, x_979); +lean_ctor_set(x_1158, 1, x_1157); +lean_inc(x_979); +x_1159 = l_Lean_Syntax_node2(x_979, x_1073, x_1046, x_999); +lean_inc(x_979); +x_1160 = l_Lean_Syntax_node1(x_979, x_987, x_1159); +x_1161 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_979); +x_1162 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1162, 0, x_979); +lean_ctor_set(x_1162, 1, x_1161); +x_1163 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_1164 = l_Lean_addMacroScope(x_986, x_1163, x_980); +x_1165 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_1166 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_979); +x_1167 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1167, 0, x_979); +lean_ctor_set(x_1167, 1, x_1165); +lean_ctor_set(x_1167, 2, x_1164); +lean_ctor_set(x_1167, 3, x_1166); +lean_inc(x_1154); +lean_inc(x_979); +x_1168 = l_Lean_Syntax_node2(x_979, x_1065, x_1167, x_1154); +x_1169 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_979); +x_1170 = l_Lean_Syntax_node1(x_979, x_1169, x_1168); +x_1171 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_979); +x_1172 = l_Lean_Syntax_node4(x_979, x_1171, x_1119, x_1160, x_1162, x_1170); +x_1173 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_989); +lean_inc(x_979); +x_1174 = l_Lean_Syntax_node3(x_979, x_1173, x_1158, x_989, x_1172); +x_1175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_989); +lean_inc(x_979); +x_1176 = l_Lean_Syntax_node2(x_979, x_1175, x_1174, x_989); +x_1177 = l_Lean_Syntax_SepArray_ofElems(x_1019, x_967); +lean_dec(x_967); +x_1178 = l_Array_append___rarg(x_988, x_1177); +lean_dec(x_1177); +lean_inc(x_979); +x_1179 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1179, 0, x_979); +lean_ctor_set(x_1179, 1, x_987); +lean_ctor_set(x_1179, 2, x_1178); +lean_inc(x_979); +x_1180 = l_Lean_Syntax_node1(x_979, x_1106, x_1179); +lean_inc_n(x_989, 2); +lean_inc(x_979); +x_1181 = l_Lean_Syntax_node6(x_979, x_1112, x_1080, x_989, x_1180, x_1109, x_989, x_1111); +lean_inc(x_979); +x_1182 = l_Lean_Syntax_node1(x_979, x_987, x_1181); +x_1183 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; +lean_inc(x_979); +x_1184 = l_Lean_Syntax_node2(x_979, x_1183, x_1122, x_1182); +lean_inc(x_989); +lean_inc(x_979); +x_1185 = l_Lean_Syntax_node2(x_979, x_1175, x_1184, x_989); +lean_inc(x_979); +x_1186 = l_Lean_Syntax_node2(x_979, x_987, x_1176, x_1185); +x_1187 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_979); +x_1188 = l_Lean_Syntax_node1(x_979, x_1187, x_1186); +x_1189 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_979); +x_1190 = l_Lean_Syntax_node2(x_979, x_1189, x_1156, x_1188); +lean_inc(x_1078); +lean_inc(x_989); +lean_inc(x_979); +x_1191 = l_Lean_Syntax_node5(x_979, x_1144, x_1103, x_1154, x_989, x_1078, x_1190); +lean_inc(x_979); +x_1192 = l_Lean_Syntax_node1(x_979, x_1146, x_1191); +lean_inc(x_1115); +lean_inc_n(x_989, 2); +lean_inc(x_979); +x_1193 = l_Lean_Syntax_node4(x_979, x_1148, x_989, x_989, x_1192, x_1115); +lean_inc(x_989); +lean_inc(x_979); +x_1194 = l_Lean_Syntax_node3(x_979, x_987, x_1149, x_989, x_1193); +x_1195 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_979); +x_1196 = l_Lean_Syntax_node2(x_979, x_1195, x_1003, x_1194); +lean_inc(x_979); +x_1197 = l_Lean_Syntax_node1(x_979, x_987, x_1196); +x_1198 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_979); +x_1199 = l_Lean_Syntax_node4(x_979, x_1198, x_1078, x_1113, x_1115, x_1197); +x_1200 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_989); +lean_inc(x_979); +x_1201 = l_Lean_Syntax_node6(x_979, x_1200, x_1042, x_1044, x_989, x_989, x_1076, x_1199); +lean_inc(x_979); +x_1202 = l_Lean_Syntax_node2(x_979, x_1031, x_991, x_1201); +x_1203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_979); +x_1204 = l_Lean_Syntax_node3(x_979, x_1203, x_1038, x_1040, x_1202); +x_1205 = l_Lean_Syntax_node2(x_979, x_987, x_1032, x_1204); +if (lean_is_scalar(x_984)) { + x_1206 = lean_alloc_ctor(0, 2, 0); } else { - x_1186 = x_968; + x_1206 = x_984; } -lean_ctor_set(x_1186, 0, x_1185); -lean_ctor_set(x_1186, 1, x_967); -return x_1186; +lean_ctor_set(x_1206, 0, x_1205); +lean_ctor_set(x_1206, 1, x_983); +return x_1206; } else { -lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; -lean_dec(x_956); -lean_dec(x_955); -lean_dec(x_954); -lean_dec(x_953); -lean_dec(x_952); -lean_dec(x_951); +lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; +lean_dec(x_972); +lean_dec(x_971); +lean_dec(x_970); +lean_dec(x_969); +lean_dec(x_968); +lean_dec(x_967); lean_free_object(x_18); lean_dec(x_10); lean_dec(x_4); -x_1187 = lean_ctor_get(x_958, 0); -lean_inc(x_1187); -x_1188 = lean_ctor_get(x_958, 1); -lean_inc(x_1188); -if (lean_is_exclusive(x_958)) { - lean_ctor_release(x_958, 0); - lean_ctor_release(x_958, 1); - x_1189 = x_958; +x_1207 = lean_ctor_get(x_974, 0); +lean_inc(x_1207); +x_1208 = lean_ctor_get(x_974, 1); +lean_inc(x_1208); +if (lean_is_exclusive(x_974)) { + lean_ctor_release(x_974, 0); + lean_ctor_release(x_974, 1); + x_1209 = x_974; } else { - lean_dec_ref(x_958); - x_1189 = lean_box(0); + lean_dec_ref(x_974); + x_1209 = lean_box(0); } -if (lean_is_scalar(x_1189)) { - x_1190 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1209)) { + x_1210 = lean_alloc_ctor(1, 2, 0); } else { - x_1190 = x_1189; + x_1210 = x_1209; } -lean_ctor_set(x_1190, 0, x_1187); -lean_ctor_set(x_1190, 1, x_1188); -return x_1190; +lean_ctor_set(x_1210, 0, x_1207); +lean_ctor_set(x_1210, 1, x_1208); +return x_1210; } } } else { -lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; size_t x_1199; lean_object* x_1200; -x_1191 = lean_ctor_get(x_18, 1); -lean_inc(x_1191); +lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; size_t x_1219; lean_object* x_1220; +x_1211 = lean_ctor_get(x_18, 1); +lean_inc(x_1211); lean_dec(x_18); -x_1192 = lean_ctor_get(x_19, 0); -lean_inc(x_1192); +x_1212 = lean_ctor_get(x_19, 0); +lean_inc(x_1212); if (lean_is_exclusive(x_19)) { lean_ctor_release(x_19, 0); lean_ctor_release(x_19, 1); - x_1193 = x_19; + x_1213 = x_19; } else { lean_dec_ref(x_19); - x_1193 = lean_box(0); + x_1213 = lean_box(0); } -x_1194 = lean_ctor_get(x_20, 0); -lean_inc(x_1194); +x_1214 = lean_ctor_get(x_20, 0); +lean_inc(x_1214); if (lean_is_exclusive(x_20)) { lean_ctor_release(x_20, 0); lean_ctor_release(x_20, 1); - x_1195 = x_20; + x_1215 = x_20; } else { lean_dec_ref(x_20); - x_1195 = lean_box(0); + x_1215 = lean_box(0); } -x_1196 = lean_ctor_get(x_21, 0); -lean_inc(x_1196); -x_1197 = lean_ctor_get(x_21, 1); -lean_inc(x_1197); +x_1216 = lean_ctor_get(x_21, 0); +lean_inc(x_1216); +x_1217 = lean_ctor_get(x_21, 1); +lean_inc(x_1217); if (lean_is_exclusive(x_21)) { lean_ctor_release(x_21, 0); lean_ctor_release(x_21, 1); - x_1198 = x_21; + x_1218 = x_21; } else { lean_dec_ref(x_21); - x_1198 = lean_box(0); -} -x_1199 = lean_array_size(x_2); -x_1200 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_1199, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_1191); -if (lean_obj_tag(x_1200) == 0) -{ -lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; uint8_t x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; size_t x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; size_t x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; size_t x_1353; lean_object* x_1354; size_t x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; -x_1201 = lean_ctor_get(x_1200, 0); -lean_inc(x_1201); -x_1202 = lean_ctor_get(x_1200, 1); -lean_inc(x_1202); -lean_dec(x_1200); -x_1203 = lean_ctor_get(x_10, 5); -lean_inc(x_1203); -x_1204 = 0; -x_1205 = l_Lean_SourceInfo_fromRef(x_1203, x_1204); -lean_dec(x_1203); -x_1206 = lean_ctor_get(x_10, 10); -lean_inc(x_1206); + x_1218 = lean_box(0); +} +x_1219 = lean_array_size(x_2); +x_1220 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_1219, x_16, x_2, x_6, x_7, x_8, x_9, x_10, x_11, x_1211); +if (lean_obj_tag(x_1220) == 0) +{ +lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; uint8_t x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; size_t x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; size_t x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; size_t x_1375; lean_object* x_1376; size_t x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; lean_object* x_1433; lean_object* x_1434; lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; lean_object* x_1447; lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; lean_object* x_1451; lean_object* x_1452; lean_object* x_1453; +x_1221 = lean_ctor_get(x_1220, 0); +lean_inc(x_1221); +x_1222 = lean_ctor_get(x_1220, 1); +lean_inc(x_1222); +lean_dec(x_1220); +x_1223 = lean_ctor_get(x_10, 5); +lean_inc(x_1223); +x_1224 = 0; +x_1225 = l_Lean_SourceInfo_fromRef(x_1223, x_1224); +lean_dec(x_1223); +x_1226 = lean_ctor_get(x_10, 10); +lean_inc(x_1226); lean_dec(x_10); -x_1207 = lean_st_ref_get(x_11, x_1202); -x_1208 = lean_ctor_get(x_1207, 0); -lean_inc(x_1208); -x_1209 = lean_ctor_get(x_1207, 1); -lean_inc(x_1209); -if (lean_is_exclusive(x_1207)) { - lean_ctor_release(x_1207, 0); - lean_ctor_release(x_1207, 1); - x_1210 = x_1207; +x_1227 = lean_st_ref_get(x_11, x_1222); +x_1228 = lean_ctor_get(x_1227, 0); +lean_inc(x_1228); +x_1229 = lean_ctor_get(x_1227, 1); +lean_inc(x_1229); +if (lean_is_exclusive(x_1227)) { + lean_ctor_release(x_1227, 0); + lean_ctor_release(x_1227, 1); + x_1230 = x_1227; } else { - lean_dec_ref(x_1207); - x_1210 = lean_box(0); -} -x_1211 = lean_ctor_get(x_1208, 0); -lean_inc(x_1211); -lean_dec(x_1208); -x_1212 = lean_environment_main_module(x_1211); -x_1213 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -x_1214 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -lean_inc(x_1205); -x_1215 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1215, 0, x_1205); -lean_ctor_set(x_1215, 1, x_1213); -lean_ctor_set(x_1215, 2, x_1214); -x_1216 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_1215, 6); -lean_inc(x_1205); -x_1217 = l_Lean_Syntax_node6(x_1205, x_1216, x_1215, x_1215, x_1215, x_1215, x_1215, x_1215); -x_1218 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; -lean_inc(x_1205); -if (lean_is_scalar(x_1198)) { - x_1219 = lean_alloc_ctor(2, 2, 0); -} else { - x_1219 = x_1198; - lean_ctor_set_tag(x_1219, 2); -} -lean_ctor_set(x_1219, 0, x_1205); -lean_ctor_set(x_1219, 1, x_1218); -x_1220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; -lean_inc(x_1205); -x_1221 = l_Lean_Syntax_node1(x_1205, x_1220, x_1219); -x_1222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_1206); -lean_inc(x_1212); -x_1223 = l_Lean_addMacroScope(x_1212, x_1222, x_1206); -x_1224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_1205); -x_1225 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1225, 0, x_1205); -lean_ctor_set(x_1225, 1, x_1224); -lean_ctor_set(x_1225, 2, x_1223); -lean_ctor_set(x_1225, 3, x_13); -x_1226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_1215); + lean_dec_ref(x_1227); + x_1230 = lean_box(0); +} +x_1231 = lean_ctor_get(x_1228, 0); +lean_inc(x_1231); +lean_dec(x_1228); +x_1232 = lean_environment_main_module(x_1231); +x_1233 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +x_1234 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; lean_inc(x_1225); -lean_inc(x_1205); -x_1227 = l_Lean_Syntax_node2(x_1205, x_1226, x_1225, x_1215); -x_1228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_1205); -if (lean_is_scalar(x_1195)) { - x_1229 = lean_alloc_ctor(2, 2, 0); -} else { - x_1229 = x_1195; - lean_ctor_set_tag(x_1229, 2); -} -lean_ctor_set(x_1229, 0, x_1205); -lean_ctor_set(x_1229, 1, x_1228); -x_1230 = l_Array_zip___rarg(x_1196, x_1197); -lean_dec(x_1197); -lean_dec(x_1196); -x_1231 = lean_array_size(x_1230); -lean_inc(x_1217); -lean_inc(x_1215); -lean_inc(x_1205); -x_1232 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_1205, x_1213, x_1215, x_1217, x_1231, x_16, x_1230); -x_1233 = l_Array_append___rarg(x_1214, x_1232); -lean_dec(x_1232); -lean_inc(x_1205); -x_1234 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1234, 0, x_1205); -lean_ctor_set(x_1234, 1, x_1213); -lean_ctor_set(x_1234, 2, x_1233); -x_1235 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; -lean_inc(x_1205); -x_1236 = l_Lean_Syntax_node1(x_1205, x_1235, x_1234); -lean_inc(x_1215); -lean_inc(x_1229); -lean_inc(x_1205); -x_1237 = l_Lean_Syntax_node3(x_1205, x_1213, x_1229, x_1215, x_1236); -x_1238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_1205); -if (lean_is_scalar(x_1193)) { +x_1235 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1235, 0, x_1225); +lean_ctor_set(x_1235, 1, x_1233); +lean_ctor_set(x_1235, 2, x_1234); +x_1236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_1235, 6); +lean_inc(x_1225); +x_1237 = l_Lean_Syntax_node6(x_1225, x_1236, x_1235, x_1235, x_1235, x_1235, x_1235, x_1235); +x_1238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; +lean_inc(x_1225); +if (lean_is_scalar(x_1218)) { x_1239 = lean_alloc_ctor(2, 2, 0); } else { - x_1239 = x_1193; + x_1239 = x_1218; lean_ctor_set_tag(x_1239, 2); } -lean_ctor_set(x_1239, 0, x_1205); +lean_ctor_set(x_1239, 0, x_1225); lean_ctor_set(x_1239, 1, x_1238); -x_1240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_1206); -lean_inc(x_1212); -x_1241 = l_Lean_addMacroScope(x_1212, x_1240, x_1206); -x_1242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_1243 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_1205); -x_1244 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1244, 0, x_1205); -lean_ctor_set(x_1244, 1, x_1242); -lean_ctor_set(x_1244, 2, x_1241); -lean_ctor_set(x_1244, 3, x_1243); -x_1245 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_1205); -x_1246 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1246, 0, x_1205); -lean_ctor_set(x_1246, 1, x_1245); -x_1247 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_1206); -lean_inc(x_1212); -x_1248 = l_Lean_addMacroScope(x_1212, x_1247, x_1206); -x_1249 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_1250 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_1205); -x_1251 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1251, 0, x_1205); -lean_ctor_set(x_1251, 1, x_1249); -lean_ctor_set(x_1251, 2, x_1248); -lean_ctor_set(x_1251, 3, x_1250); -lean_inc(x_1246); -lean_inc(x_1205); -x_1252 = l_Lean_Syntax_node3(x_1205, x_1213, x_1244, x_1246, x_1251); -lean_inc(x_1205); -x_1253 = l_Lean_Syntax_node2(x_1205, x_1213, x_1239, x_1252); -x_1254 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_1205); -x_1255 = l_Lean_Syntax_node1(x_1205, x_1254, x_1253); -x_1256 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; -lean_inc_n(x_1215, 3); -lean_inc(x_1205); -x_1257 = l_Lean_Syntax_node7(x_1205, x_1256, x_1221, x_1227, x_1215, x_1215, x_1215, x_1237, x_1255); -x_1258 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_1217); -lean_inc(x_1205); -x_1259 = l_Lean_Syntax_node2(x_1205, x_1258, x_1217, x_1257); -x_1260 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_1205); -x_1261 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1261, 0, x_1205); -lean_ctor_set(x_1261, 1, x_1260); -x_1262 = l_Array_append___rarg(x_1214, x_3); -lean_inc(x_1205); -x_1263 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1263, 0, x_1205); -lean_ctor_set(x_1263, 1, x_1213); -lean_ctor_set(x_1263, 2, x_1262); -x_1264 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_1205); -x_1265 = l_Lean_Syntax_node2(x_1205, x_1264, x_1261, x_1263); -x_1266 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_1205); -x_1267 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1267, 0, x_1205); -lean_ctor_set(x_1267, 1, x_1266); -x_1268 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_1215); -lean_inc(x_1205); -x_1269 = l_Lean_Syntax_node1(x_1205, x_1268, x_1215); -x_1270 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_1205); -x_1271 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1271, 0, x_1205); -lean_ctor_set(x_1271, 1, x_1270); -x_1272 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_1205); -x_1273 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1273, 0, x_1205); -lean_ctor_set(x_1273, 1, x_1272); -x_1274 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_1206); -lean_inc(x_1212); -x_1275 = l_Lean_addMacroScope(x_1212, x_1274, x_1206); -x_1276 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_1277 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_1205); -x_1278 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1278, 0, x_1205); -lean_ctor_set(x_1278, 1, x_1276); -lean_ctor_set(x_1278, 2, x_1275); -lean_ctor_set(x_1278, 3, x_1277); -x_1279 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; -lean_inc(x_1205); -x_1280 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1280, 0, x_1205); -lean_ctor_set(x_1280, 1, x_1279); -x_1281 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_1205); -x_1282 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1282, 0, x_1205); -lean_ctor_set(x_1282, 1, x_1281); -x_1283 = lean_box(0); -x_1284 = l_Lean_mkCIdentFrom(x_1283, x_4, x_1204); -x_1285 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_1205); -x_1286 = l_Lean_Syntax_node2(x_1205, x_1285, x_1282, x_1284); -x_1287 = lean_array_size(x_1201); -x_1288 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; -x_1289 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_1288, x_1287, x_16, x_1201); -x_1290 = l_Array_append___rarg(x_1214, x_1289); -lean_dec(x_1289); -lean_inc(x_1205); -x_1291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1291, 0, x_1205); -lean_ctor_set(x_1291, 1, x_1213); -lean_ctor_set(x_1291, 2, x_1290); -x_1292 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_1205); -x_1293 = l_Lean_Syntax_node2(x_1205, x_1292, x_1286, x_1291); -x_1294 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; -lean_inc(x_1205); -x_1295 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1295, 0, x_1205); -lean_ctor_set(x_1295, 1, x_1294); -x_1296 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; -lean_inc(x_1205); -x_1297 = l_Lean_Syntax_node3(x_1205, x_1296, x_1280, x_1293, x_1295); -lean_inc(x_1205); -x_1298 = l_Lean_Syntax_node1(x_1205, x_1213, x_1297); -lean_inc(x_1205); -x_1299 = l_Lean_Syntax_node2(x_1205, x_1292, x_1278, x_1298); -x_1300 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_1273); -lean_inc(x_1205); -x_1301 = l_Lean_Syntax_node2(x_1205, x_1300, x_1273, x_1299); -x_1302 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_1215); -lean_inc(x_1205); -x_1303 = l_Lean_Syntax_node2(x_1205, x_1302, x_1215, x_1301); -x_1304 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_1205); -x_1305 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1305, 0, x_1205); -lean_ctor_set(x_1305, 1, x_1304); -x_1306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_1205); -x_1307 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1307, 0, x_1205); -lean_ctor_set(x_1307, 1, x_1306); -x_1308 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_1206); -lean_inc(x_1212); -x_1309 = l_Lean_addMacroScope(x_1212, x_1308, x_1206); -x_1310 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_1311 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_1205); -x_1312 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1312, 0, x_1205); -lean_ctor_set(x_1312, 1, x_1310); -lean_ctor_set(x_1312, 2, x_1309); -lean_ctor_set(x_1312, 3, x_1311); -x_1313 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_1215); -lean_inc(x_1205); -x_1314 = l_Lean_Syntax_node2(x_1205, x_1313, x_1312, x_1215); -x_1315 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_1206); -lean_inc(x_1212); -x_1316 = l_Lean_addMacroScope(x_1212, x_1315, x_1206); -x_1317 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_1205); -x_1318 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1318, 0, x_1205); -lean_ctor_set(x_1318, 1, x_1317); -lean_ctor_set(x_1318, 2, x_1316); -lean_ctor_set(x_1318, 3, x_13); -x_1319 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_1318); -lean_inc(x_1305); -lean_inc(x_1205); -x_1320 = l_Lean_Syntax_node3(x_1205, x_1319, x_1314, x_1305, x_1318); -x_1321 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_1206); -lean_inc(x_1212); -x_1322 = l_Lean_addMacroScope(x_1212, x_1321, x_1206); -x_1323 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_1324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_1205); -x_1325 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1325, 0, x_1205); -lean_ctor_set(x_1325, 1, x_1323); -lean_ctor_set(x_1325, 2, x_1322); -lean_ctor_set(x_1325, 3, x_1324); -lean_inc(x_1215); -lean_inc(x_1205); -x_1326 = l_Lean_Syntax_node2(x_1205, x_1313, x_1325, x_1215); -x_1327 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_1206); -lean_inc(x_1212); -x_1328 = l_Lean_addMacroScope(x_1212, x_1327, x_1206); -x_1329 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_1205); -x_1330 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1330, 0, x_1205); -lean_ctor_set(x_1330, 1, x_1329); -lean_ctor_set(x_1330, 2, x_1328); -lean_ctor_set(x_1330, 3, x_13); -lean_inc(x_1330); -lean_inc(x_1305); -lean_inc(x_1205); -x_1331 = l_Lean_Syntax_node3(x_1205, x_1319, x_1326, x_1305, x_1330); -lean_inc(x_1205); -x_1332 = l_Lean_Syntax_node3(x_1205, x_1213, x_1320, x_1246, x_1331); -x_1333 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_1215); -lean_inc(x_1205); -x_1334 = l_Lean_Syntax_node1(x_1205, x_1333, x_1215); -x_1335 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_1205); -x_1336 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1336, 0, x_1205); -lean_ctor_set(x_1336, 1, x_1335); -x_1337 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc(x_1336); -lean_inc(x_1334); -lean_inc_n(x_1215, 2); -lean_inc(x_1307); -lean_inc(x_1205); -x_1338 = l_Lean_Syntax_node6(x_1205, x_1337, x_1307, x_1215, x_1332, x_1334, x_1215, x_1336); -x_1339 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_1215, 2); -lean_inc(x_1205); -x_1340 = l_Lean_Syntax_node2(x_1205, x_1339, x_1215, x_1215); -x_1341 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; -lean_inc(x_1206); -lean_inc(x_1212); -x_1342 = l_Lean_addMacroScope(x_1212, x_1341, x_1206); -x_1343 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; -lean_inc(x_1205); -x_1344 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1344, 0, x_1205); -lean_ctor_set(x_1344, 1, x_1343); -lean_ctor_set(x_1344, 2, x_1342); -lean_ctor_set(x_1344, 3, x_13); -lean_inc(x_1344); -lean_inc(x_1205); -x_1345 = l_Lean_Syntax_node1(x_1205, x_1213, x_1344); -x_1346 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; -lean_inc(x_1205); -x_1347 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1347, 0, x_1205); -lean_ctor_set(x_1347, 1, x_1346); -x_1348 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; -lean_inc(x_1206); -lean_inc(x_1212); -x_1349 = l_Lean_addMacroScope(x_1212, x_1348, x_1206); -x_1350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_1351 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; -lean_inc(x_1205); -x_1352 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1352, 0, x_1205); -lean_ctor_set(x_1352, 1, x_1350); -lean_ctor_set(x_1352, 2, x_1349); -lean_ctor_set(x_1352, 3, x_1351); -x_1353 = lean_array_size(x_1194); -x_1354 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_1319, x_1353, x_16, x_1194); -x_1355 = lean_array_size(x_1354); -x_1356 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_1355, x_16, x_1354); -x_1357 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; -x_1358 = l_Lean_mkSepArray(x_1356, x_1357); -lean_dec(x_1356); -x_1359 = l_Array_append___rarg(x_1214, x_1358); -lean_dec(x_1358); -lean_inc(x_1205); -x_1360 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1360, 0, x_1205); -lean_ctor_set(x_1360, 1, x_1213); -lean_ctor_set(x_1360, 2, x_1359); +x_1240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__12; lean_inc(x_1225); -lean_inc(x_1273); -lean_inc(x_1205); -x_1361 = l_Lean_Syntax_node2(x_1205, x_1213, x_1273, x_1225); -lean_inc(x_1336); -lean_inc(x_1334); -lean_inc(x_1215); -lean_inc(x_1307); -lean_inc(x_1205); -x_1362 = l_Lean_Syntax_node6(x_1205, x_1337, x_1307, x_1215, x_1360, x_1334, x_1361, x_1336); -lean_inc(x_1205); -x_1363 = l_Lean_Syntax_node1(x_1205, x_1213, x_1362); -lean_inc(x_1205); -x_1364 = l_Lean_Syntax_node2(x_1205, x_1292, x_1352, x_1363); -lean_inc(x_1205); -x_1365 = l_Lean_Syntax_node1(x_1205, x_1213, x_1364); -x_1366 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; -lean_inc(x_1347); -lean_inc(x_1205); -x_1367 = l_Lean_Syntax_node2(x_1205, x_1366, x_1347, x_1365); -x_1368 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_1305); -lean_inc(x_1215); -lean_inc(x_1205); -x_1369 = l_Lean_Syntax_node5(x_1205, x_1368, x_1318, x_1345, x_1215, x_1305, x_1367); -x_1370 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_1205); -x_1371 = l_Lean_Syntax_node1(x_1205, x_1370, x_1369); -x_1372 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_1340); -lean_inc_n(x_1215, 2); -lean_inc(x_1205); -x_1373 = l_Lean_Syntax_node4(x_1205, x_1372, x_1215, x_1215, x_1371, x_1340); -x_1374 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_1206); -lean_inc(x_1212); -x_1375 = l_Lean_addMacroScope(x_1212, x_1374, x_1206); -x_1376 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_1205); -x_1377 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1377, 0, x_1205); -lean_ctor_set(x_1377, 1, x_1376); -lean_ctor_set(x_1377, 2, x_1375); -lean_ctor_set(x_1377, 3, x_13); -lean_inc(x_1205); -x_1378 = l_Lean_Syntax_node1(x_1205, x_1213, x_1377); -x_1379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_1205); -x_1380 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1380, 0, x_1205); -lean_ctor_set(x_1380, 1, x_1379); -x_1381 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_1205); -x_1382 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1382, 0, x_1205); -lean_ctor_set(x_1382, 1, x_1381); -lean_inc(x_1205); -x_1383 = l_Lean_Syntax_node2(x_1205, x_1300, x_1273, x_1225); -lean_inc(x_1205); -x_1384 = l_Lean_Syntax_node1(x_1205, x_1213, x_1383); -x_1385 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_1205); -x_1386 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1386, 0, x_1205); -lean_ctor_set(x_1386, 1, x_1385); -x_1387 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -x_1388 = l_Lean_addMacroScope(x_1212, x_1387, x_1206); -x_1389 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_1390 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_1205); -x_1391 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1391, 0, x_1205); -lean_ctor_set(x_1391, 1, x_1389); -lean_ctor_set(x_1391, 2, x_1388); -lean_ctor_set(x_1391, 3, x_1390); -lean_inc(x_1378); -lean_inc(x_1205); -x_1392 = l_Lean_Syntax_node2(x_1205, x_1292, x_1391, x_1378); -x_1393 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_1205); -x_1394 = l_Lean_Syntax_node1(x_1205, x_1393, x_1392); -x_1395 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_1205); -x_1396 = l_Lean_Syntax_node4(x_1205, x_1395, x_1344, x_1384, x_1386, x_1394); -x_1397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_1215); -lean_inc(x_1205); -x_1398 = l_Lean_Syntax_node3(x_1205, x_1397, x_1382, x_1215, x_1396); -x_1399 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_1215); -lean_inc(x_1205); -x_1400 = l_Lean_Syntax_node2(x_1205, x_1399, x_1398, x_1215); -x_1401 = l_Lean_Syntax_SepArray_ofElems(x_1245, x_1192); -lean_dec(x_1192); -x_1402 = l_Array_append___rarg(x_1214, x_1401); -lean_dec(x_1401); -lean_inc(x_1205); -x_1403 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1403, 0, x_1205); -lean_ctor_set(x_1403, 1, x_1213); -lean_ctor_set(x_1403, 2, x_1402); -lean_inc_n(x_1215, 2); -lean_inc(x_1205); -x_1404 = l_Lean_Syntax_node6(x_1205, x_1337, x_1307, x_1215, x_1403, x_1334, x_1215, x_1336); -lean_inc(x_1205); -x_1405 = l_Lean_Syntax_node1(x_1205, x_1213, x_1404); -x_1406 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; -lean_inc(x_1205); -x_1407 = l_Lean_Syntax_node2(x_1205, x_1406, x_1347, x_1405); -lean_inc(x_1215); -lean_inc(x_1205); -x_1408 = l_Lean_Syntax_node2(x_1205, x_1399, x_1407, x_1215); -lean_inc(x_1205); -x_1409 = l_Lean_Syntax_node2(x_1205, x_1213, x_1400, x_1408); -x_1410 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_1205); -x_1411 = l_Lean_Syntax_node1(x_1205, x_1410, x_1409); -x_1412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_1205); -x_1413 = l_Lean_Syntax_node2(x_1205, x_1412, x_1380, x_1411); -lean_inc(x_1305); -lean_inc(x_1215); -lean_inc(x_1205); -x_1414 = l_Lean_Syntax_node5(x_1205, x_1368, x_1330, x_1378, x_1215, x_1305, x_1413); -lean_inc(x_1205); -x_1415 = l_Lean_Syntax_node1(x_1205, x_1370, x_1414); -lean_inc(x_1340); -lean_inc_n(x_1215, 2); -lean_inc(x_1205); -x_1416 = l_Lean_Syntax_node4(x_1205, x_1372, x_1215, x_1215, x_1415, x_1340); -lean_inc(x_1215); -lean_inc(x_1205); -x_1417 = l_Lean_Syntax_node3(x_1205, x_1213, x_1373, x_1215, x_1416); -x_1418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_1205); -x_1419 = l_Lean_Syntax_node2(x_1205, x_1418, x_1229, x_1417); -lean_inc(x_1205); -x_1420 = l_Lean_Syntax_node1(x_1205, x_1213, x_1419); -x_1421 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_1205); -x_1422 = l_Lean_Syntax_node4(x_1205, x_1421, x_1305, x_1338, x_1340, x_1420); -x_1423 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_1215); -lean_inc(x_1205); -x_1424 = l_Lean_Syntax_node6(x_1205, x_1423, x_1269, x_1271, x_1215, x_1215, x_1303, x_1422); -lean_inc(x_1205); -x_1425 = l_Lean_Syntax_node2(x_1205, x_1258, x_1217, x_1424); -x_1426 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_1205); -x_1427 = l_Lean_Syntax_node3(x_1205, x_1426, x_1265, x_1267, x_1425); -x_1428 = l_Lean_Syntax_node2(x_1205, x_1213, x_1259, x_1427); -if (lean_is_scalar(x_1210)) { - x_1429 = lean_alloc_ctor(0, 2, 0); +x_1241 = l_Lean_Syntax_node1(x_1225, x_1240, x_1239); +x_1242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_1226); +lean_inc(x_1232); +x_1243 = l_Lean_addMacroScope(x_1232, x_1242, x_1226); +x_1244 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_1225); +x_1245 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1245, 0, x_1225); +lean_ctor_set(x_1245, 1, x_1244); +lean_ctor_set(x_1245, 2, x_1243); +lean_ctor_set(x_1245, 3, x_13); +x_1246 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_1235); +lean_inc(x_1245); +lean_inc(x_1225); +x_1247 = l_Lean_Syntax_node2(x_1225, x_1246, x_1245, x_1235); +x_1248 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_1225); +if (lean_is_scalar(x_1215)) { + x_1249 = lean_alloc_ctor(2, 2, 0); +} else { + x_1249 = x_1215; + lean_ctor_set_tag(x_1249, 2); +} +lean_ctor_set(x_1249, 0, x_1225); +lean_ctor_set(x_1249, 1, x_1248); +x_1250 = l_Array_zip___rarg(x_1216, x_1217); +lean_dec(x_1217); +lean_dec(x_1216); +x_1251 = lean_array_size(x_1250); +lean_inc(x_1237); +lean_inc(x_1235); +lean_inc(x_1225); +x_1252 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3(x_1225, x_1233, x_1235, x_1237, x_1251, x_16, x_1250); +x_1253 = l_Array_append___rarg(x_1234, x_1252); +lean_dec(x_1252); +lean_inc(x_1225); +x_1254 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1254, 0, x_1225); +lean_ctor_set(x_1254, 1, x_1233); +lean_ctor_set(x_1254, 2, x_1253); +x_1255 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; +lean_inc(x_1225); +x_1256 = l_Lean_Syntax_node1(x_1225, x_1255, x_1254); +lean_inc(x_1235); +lean_inc(x_1249); +lean_inc(x_1225); +x_1257 = l_Lean_Syntax_node3(x_1225, x_1233, x_1249, x_1235, x_1256); +x_1258 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_1225); +if (lean_is_scalar(x_1213)) { + x_1259 = lean_alloc_ctor(2, 2, 0); +} else { + x_1259 = x_1213; + lean_ctor_set_tag(x_1259, 2); +} +lean_ctor_set(x_1259, 0, x_1225); +lean_ctor_set(x_1259, 1, x_1258); +x_1260 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_1226); +lean_inc(x_1232); +x_1261 = l_Lean_addMacroScope(x_1232, x_1260, x_1226); +x_1262 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_1263 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_1225); +x_1264 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1264, 0, x_1225); +lean_ctor_set(x_1264, 1, x_1262); +lean_ctor_set(x_1264, 2, x_1261); +lean_ctor_set(x_1264, 3, x_1263); +x_1265 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_1225); +x_1266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1266, 0, x_1225); +lean_ctor_set(x_1266, 1, x_1265); +x_1267 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_1226); +lean_inc(x_1232); +x_1268 = l_Lean_addMacroScope(x_1232, x_1267, x_1226); +x_1269 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_1270 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_1225); +x_1271 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1271, 0, x_1225); +lean_ctor_set(x_1271, 1, x_1269); +lean_ctor_set(x_1271, 2, x_1268); +lean_ctor_set(x_1271, 3, x_1270); +lean_inc(x_1266); +lean_inc(x_1225); +x_1272 = l_Lean_Syntax_node3(x_1225, x_1233, x_1264, x_1266, x_1271); +lean_inc(x_1225); +x_1273 = l_Lean_Syntax_node2(x_1225, x_1233, x_1259, x_1272); +x_1274 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_1225); +x_1275 = l_Lean_Syntax_node1(x_1225, x_1274, x_1273); +x_1276 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__10; +lean_inc_n(x_1235, 3); +lean_inc(x_1225); +x_1277 = l_Lean_Syntax_node7(x_1225, x_1276, x_1241, x_1247, x_1235, x_1235, x_1235, x_1257, x_1275); +x_1278 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_1237); +lean_inc(x_1225); +x_1279 = l_Lean_Syntax_node2(x_1225, x_1278, x_1237, x_1277); +x_1280 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_1225); +x_1281 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1281, 0, x_1225); +lean_ctor_set(x_1281, 1, x_1280); +x_1282 = l_Array_append___rarg(x_1234, x_3); +lean_inc(x_1225); +x_1283 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1283, 0, x_1225); +lean_ctor_set(x_1283, 1, x_1233); +lean_ctor_set(x_1283, 2, x_1282); +x_1284 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_1225); +x_1285 = l_Lean_Syntax_node2(x_1225, x_1284, x_1281, x_1283); +x_1286 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_1225); +x_1287 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1287, 0, x_1225); +lean_ctor_set(x_1287, 1, x_1286); +x_1288 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_1235); +lean_inc(x_1225); +x_1289 = l_Lean_Syntax_node1(x_1225, x_1288, x_1235); +x_1290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_1225); +x_1291 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1291, 0, x_1225); +lean_ctor_set(x_1291, 1, x_1290); +x_1292 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_1225); +x_1293 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1293, 0, x_1225); +lean_ctor_set(x_1293, 1, x_1292); +x_1294 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_1226); +lean_inc(x_1232); +x_1295 = l_Lean_addMacroScope(x_1232, x_1294, x_1226); +x_1296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_1297 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_1225); +x_1298 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1298, 0, x_1225); +lean_ctor_set(x_1298, 1, x_1296); +lean_ctor_set(x_1298, 2, x_1295); +lean_ctor_set(x_1298, 3, x_1297); +x_1299 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__57; +lean_inc(x_1225); +x_1300 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1300, 0, x_1225); +lean_ctor_set(x_1300, 1, x_1299); +x_1301 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_1225); +x_1302 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1302, 0, x_1225); +lean_ctor_set(x_1302, 1, x_1301); +x_1303 = lean_box(0); +x_1304 = l_Lean_mkCIdentFrom(x_1303, x_4, x_1224); +x_1305 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_1225); +x_1306 = l_Lean_Syntax_node2(x_1225, x_1305, x_1302, x_1304); +x_1307 = lean_array_size(x_1221); +x_1308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; +x_1309 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(x_1308, x_1307, x_16, x_1221); +x_1310 = l_Array_append___rarg(x_1234, x_1309); +lean_dec(x_1309); +lean_inc(x_1225); +x_1311 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1311, 0, x_1225); +lean_ctor_set(x_1311, 1, x_1233); +lean_ctor_set(x_1311, 2, x_1310); +x_1312 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_1225); +x_1313 = l_Lean_Syntax_node2(x_1225, x_1312, x_1306, x_1311); +x_1314 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; +lean_inc(x_1225); +x_1315 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1315, 0, x_1225); +lean_ctor_set(x_1315, 1, x_1314); +x_1316 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; +lean_inc(x_1225); +x_1317 = l_Lean_Syntax_node3(x_1225, x_1316, x_1300, x_1313, x_1315); +lean_inc(x_1225); +x_1318 = l_Lean_Syntax_node1(x_1225, x_1233, x_1317); +lean_inc(x_1225); +x_1319 = l_Lean_Syntax_node2(x_1225, x_1312, x_1298, x_1318); +x_1320 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_1293); +lean_inc(x_1225); +x_1321 = l_Lean_Syntax_node2(x_1225, x_1320, x_1293, x_1319); +x_1322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_1235); +lean_inc(x_1225); +x_1323 = l_Lean_Syntax_node2(x_1225, x_1322, x_1235, x_1321); +x_1324 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_1225); +x_1325 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1325, 0, x_1225); +lean_ctor_set(x_1325, 1, x_1324); +x_1326 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_1225); +x_1327 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1327, 0, x_1225); +lean_ctor_set(x_1327, 1, x_1326); +x_1328 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_1226); +lean_inc(x_1232); +x_1329 = l_Lean_addMacroScope(x_1232, x_1328, x_1226); +x_1330 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_1331 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_1225); +x_1332 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1332, 0, x_1225); +lean_ctor_set(x_1332, 1, x_1330); +lean_ctor_set(x_1332, 2, x_1329); +lean_ctor_set(x_1332, 3, x_1331); +x_1333 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_1235); +lean_inc(x_1225); +x_1334 = l_Lean_Syntax_node2(x_1225, x_1333, x_1332, x_1235); +x_1335 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_1226); +lean_inc(x_1232); +x_1336 = l_Lean_addMacroScope(x_1232, x_1335, x_1226); +x_1337 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_1225); +x_1338 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1338, 0, x_1225); +lean_ctor_set(x_1338, 1, x_1337); +lean_ctor_set(x_1338, 2, x_1336); +lean_ctor_set(x_1338, 3, x_13); +x_1339 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_1338); +lean_inc(x_1325); +lean_inc(x_1225); +x_1340 = l_Lean_Syntax_node3(x_1225, x_1339, x_1334, x_1325, x_1338); +x_1341 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_1226); +lean_inc(x_1232); +x_1342 = l_Lean_addMacroScope(x_1232, x_1341, x_1226); +x_1343 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_1344 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_1225); +x_1345 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1345, 0, x_1225); +lean_ctor_set(x_1345, 1, x_1343); +lean_ctor_set(x_1345, 2, x_1342); +lean_ctor_set(x_1345, 3, x_1344); +lean_inc(x_1235); +lean_inc(x_1225); +x_1346 = l_Lean_Syntax_node2(x_1225, x_1333, x_1345, x_1235); +x_1347 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_1226); +lean_inc(x_1232); +x_1348 = l_Lean_addMacroScope(x_1232, x_1347, x_1226); +x_1349 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_1225); +x_1350 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1350, 0, x_1225); +lean_ctor_set(x_1350, 1, x_1349); +lean_ctor_set(x_1350, 2, x_1348); +lean_ctor_set(x_1350, 3, x_13); +lean_inc(x_1350); +lean_inc(x_1325); +lean_inc(x_1225); +x_1351 = l_Lean_Syntax_node3(x_1225, x_1339, x_1346, x_1325, x_1350); +lean_inc(x_1225); +x_1352 = l_Lean_Syntax_node3(x_1225, x_1233, x_1340, x_1266, x_1351); +x_1353 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_1225); +x_1354 = l_Lean_Syntax_node1(x_1225, x_1353, x_1352); +x_1355 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_1235); +lean_inc(x_1225); +x_1356 = l_Lean_Syntax_node1(x_1225, x_1355, x_1235); +x_1357 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_1225); +x_1358 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1358, 0, x_1225); +lean_ctor_set(x_1358, 1, x_1357); +x_1359 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc(x_1358); +lean_inc(x_1356); +lean_inc_n(x_1235, 2); +lean_inc(x_1327); +lean_inc(x_1225); +x_1360 = l_Lean_Syntax_node6(x_1225, x_1359, x_1327, x_1235, x_1354, x_1356, x_1235, x_1358); +x_1361 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_1235, 2); +lean_inc(x_1225); +x_1362 = l_Lean_Syntax_node2(x_1225, x_1361, x_1235, x_1235); +x_1363 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__37; +lean_inc(x_1226); +lean_inc(x_1232); +x_1364 = l_Lean_addMacroScope(x_1232, x_1363, x_1226); +x_1365 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; +lean_inc(x_1225); +x_1366 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1366, 0, x_1225); +lean_ctor_set(x_1366, 1, x_1365); +lean_ctor_set(x_1366, 2, x_1364); +lean_ctor_set(x_1366, 3, x_13); +lean_inc(x_1366); +lean_inc(x_1225); +x_1367 = l_Lean_Syntax_node1(x_1225, x_1233, x_1366); +x_1368 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; +lean_inc(x_1225); +x_1369 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1369, 0, x_1225); +lean_ctor_set(x_1369, 1, x_1368); +x_1370 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; +lean_inc(x_1226); +lean_inc(x_1232); +x_1371 = l_Lean_addMacroScope(x_1232, x_1370, x_1226); +x_1372 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_1373 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +lean_inc(x_1225); +x_1374 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1374, 0, x_1225); +lean_ctor_set(x_1374, 1, x_1372); +lean_ctor_set(x_1374, 2, x_1371); +lean_ctor_set(x_1374, 3, x_1373); +x_1375 = lean_array_size(x_1214); +x_1376 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__5(x_13, x_1339, x_1375, x_16, x_1214); +x_1377 = lean_array_size(x_1376); +x_1378 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__6(x_1377, x_16, x_1376); +x_1379 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_1380 = l_Lean_mkSepArray(x_1378, x_1379); +lean_dec(x_1378); +x_1381 = l_Array_append___rarg(x_1234, x_1380); +lean_dec(x_1380); +lean_inc(x_1225); +x_1382 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1382, 0, x_1225); +lean_ctor_set(x_1382, 1, x_1233); +lean_ctor_set(x_1382, 2, x_1381); +lean_inc(x_1225); +x_1383 = l_Lean_Syntax_node1(x_1225, x_1353, x_1382); +lean_inc(x_1245); +lean_inc(x_1293); +lean_inc(x_1225); +x_1384 = l_Lean_Syntax_node2(x_1225, x_1233, x_1293, x_1245); +lean_inc(x_1358); +lean_inc(x_1356); +lean_inc(x_1235); +lean_inc(x_1327); +lean_inc(x_1225); +x_1385 = l_Lean_Syntax_node6(x_1225, x_1359, x_1327, x_1235, x_1383, x_1356, x_1384, x_1358); +lean_inc(x_1225); +x_1386 = l_Lean_Syntax_node1(x_1225, x_1233, x_1385); +lean_inc(x_1225); +x_1387 = l_Lean_Syntax_node2(x_1225, x_1312, x_1374, x_1386); +lean_inc(x_1225); +x_1388 = l_Lean_Syntax_node1(x_1225, x_1233, x_1387); +x_1389 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; +lean_inc(x_1369); +lean_inc(x_1225); +x_1390 = l_Lean_Syntax_node2(x_1225, x_1389, x_1369, x_1388); +x_1391 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_1325); +lean_inc(x_1235); +lean_inc(x_1225); +x_1392 = l_Lean_Syntax_node5(x_1225, x_1391, x_1338, x_1367, x_1235, x_1325, x_1390); +x_1393 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_1225); +x_1394 = l_Lean_Syntax_node1(x_1225, x_1393, x_1392); +x_1395 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_1362); +lean_inc_n(x_1235, 2); +lean_inc(x_1225); +x_1396 = l_Lean_Syntax_node4(x_1225, x_1395, x_1235, x_1235, x_1394, x_1362); +x_1397 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_1226); +lean_inc(x_1232); +x_1398 = l_Lean_addMacroScope(x_1232, x_1397, x_1226); +x_1399 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_1225); +x_1400 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1400, 0, x_1225); +lean_ctor_set(x_1400, 1, x_1399); +lean_ctor_set(x_1400, 2, x_1398); +lean_ctor_set(x_1400, 3, x_13); +lean_inc(x_1225); +x_1401 = l_Lean_Syntax_node1(x_1225, x_1233, x_1400); +x_1402 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_1225); +x_1403 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1403, 0, x_1225); +lean_ctor_set(x_1403, 1, x_1402); +x_1404 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_1225); +x_1405 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1405, 0, x_1225); +lean_ctor_set(x_1405, 1, x_1404); +lean_inc(x_1225); +x_1406 = l_Lean_Syntax_node2(x_1225, x_1320, x_1293, x_1245); +lean_inc(x_1225); +x_1407 = l_Lean_Syntax_node1(x_1225, x_1233, x_1406); +x_1408 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_1225); +x_1409 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1409, 0, x_1225); +lean_ctor_set(x_1409, 1, x_1408); +x_1410 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +x_1411 = l_Lean_addMacroScope(x_1232, x_1410, x_1226); +x_1412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_1413 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_1225); +x_1414 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1414, 0, x_1225); +lean_ctor_set(x_1414, 1, x_1412); +lean_ctor_set(x_1414, 2, x_1411); +lean_ctor_set(x_1414, 3, x_1413); +lean_inc(x_1401); +lean_inc(x_1225); +x_1415 = l_Lean_Syntax_node2(x_1225, x_1312, x_1414, x_1401); +x_1416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_1225); +x_1417 = l_Lean_Syntax_node1(x_1225, x_1416, x_1415); +x_1418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_1225); +x_1419 = l_Lean_Syntax_node4(x_1225, x_1418, x_1366, x_1407, x_1409, x_1417); +x_1420 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_1235); +lean_inc(x_1225); +x_1421 = l_Lean_Syntax_node3(x_1225, x_1420, x_1405, x_1235, x_1419); +x_1422 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_1235); +lean_inc(x_1225); +x_1423 = l_Lean_Syntax_node2(x_1225, x_1422, x_1421, x_1235); +x_1424 = l_Lean_Syntax_SepArray_ofElems(x_1265, x_1212); +lean_dec(x_1212); +x_1425 = l_Array_append___rarg(x_1234, x_1424); +lean_dec(x_1424); +lean_inc(x_1225); +x_1426 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1426, 0, x_1225); +lean_ctor_set(x_1426, 1, x_1233); +lean_ctor_set(x_1426, 2, x_1425); +lean_inc(x_1225); +x_1427 = l_Lean_Syntax_node1(x_1225, x_1353, x_1426); +lean_inc_n(x_1235, 2); +lean_inc(x_1225); +x_1428 = l_Lean_Syntax_node6(x_1225, x_1359, x_1327, x_1235, x_1427, x_1356, x_1235, x_1358); +lean_inc(x_1225); +x_1429 = l_Lean_Syntax_node1(x_1225, x_1233, x_1428); +x_1430 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124; +lean_inc(x_1225); +x_1431 = l_Lean_Syntax_node2(x_1225, x_1430, x_1369, x_1429); +lean_inc(x_1235); +lean_inc(x_1225); +x_1432 = l_Lean_Syntax_node2(x_1225, x_1422, x_1431, x_1235); +lean_inc(x_1225); +x_1433 = l_Lean_Syntax_node2(x_1225, x_1233, x_1423, x_1432); +x_1434 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_1225); +x_1435 = l_Lean_Syntax_node1(x_1225, x_1434, x_1433); +x_1436 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_1225); +x_1437 = l_Lean_Syntax_node2(x_1225, x_1436, x_1403, x_1435); +lean_inc(x_1325); +lean_inc(x_1235); +lean_inc(x_1225); +x_1438 = l_Lean_Syntax_node5(x_1225, x_1391, x_1350, x_1401, x_1235, x_1325, x_1437); +lean_inc(x_1225); +x_1439 = l_Lean_Syntax_node1(x_1225, x_1393, x_1438); +lean_inc(x_1362); +lean_inc_n(x_1235, 2); +lean_inc(x_1225); +x_1440 = l_Lean_Syntax_node4(x_1225, x_1395, x_1235, x_1235, x_1439, x_1362); +lean_inc(x_1235); +lean_inc(x_1225); +x_1441 = l_Lean_Syntax_node3(x_1225, x_1233, x_1396, x_1235, x_1440); +x_1442 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_1225); +x_1443 = l_Lean_Syntax_node2(x_1225, x_1442, x_1249, x_1441); +lean_inc(x_1225); +x_1444 = l_Lean_Syntax_node1(x_1225, x_1233, x_1443); +x_1445 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_1225); +x_1446 = l_Lean_Syntax_node4(x_1225, x_1445, x_1325, x_1360, x_1362, x_1444); +x_1447 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_1235); +lean_inc(x_1225); +x_1448 = l_Lean_Syntax_node6(x_1225, x_1447, x_1289, x_1291, x_1235, x_1235, x_1323, x_1446); +lean_inc(x_1225); +x_1449 = l_Lean_Syntax_node2(x_1225, x_1278, x_1237, x_1448); +x_1450 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_1225); +x_1451 = l_Lean_Syntax_node3(x_1225, x_1450, x_1285, x_1287, x_1449); +x_1452 = l_Lean_Syntax_node2(x_1225, x_1233, x_1279, x_1451); +if (lean_is_scalar(x_1230)) { + x_1453 = lean_alloc_ctor(0, 2, 0); } else { - x_1429 = x_1210; + x_1453 = x_1230; } -lean_ctor_set(x_1429, 0, x_1428); -lean_ctor_set(x_1429, 1, x_1209); -return x_1429; +lean_ctor_set(x_1453, 0, x_1452); +lean_ctor_set(x_1453, 1, x_1229); +return x_1453; } else { -lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; lean_object* x_1433; -lean_dec(x_1198); -lean_dec(x_1197); -lean_dec(x_1196); -lean_dec(x_1195); -lean_dec(x_1194); -lean_dec(x_1193); -lean_dec(x_1192); +lean_object* x_1454; lean_object* x_1455; lean_object* x_1456; lean_object* x_1457; +lean_dec(x_1218); +lean_dec(x_1217); +lean_dec(x_1216); +lean_dec(x_1215); +lean_dec(x_1214); +lean_dec(x_1213); +lean_dec(x_1212); lean_dec(x_10); lean_dec(x_4); -x_1430 = lean_ctor_get(x_1200, 0); -lean_inc(x_1430); -x_1431 = lean_ctor_get(x_1200, 1); -lean_inc(x_1431); -if (lean_is_exclusive(x_1200)) { - lean_ctor_release(x_1200, 0); - lean_ctor_release(x_1200, 1); - x_1432 = x_1200; +x_1454 = lean_ctor_get(x_1220, 0); +lean_inc(x_1454); +x_1455 = lean_ctor_get(x_1220, 1); +lean_inc(x_1455); +if (lean_is_exclusive(x_1220)) { + lean_ctor_release(x_1220, 0); + lean_ctor_release(x_1220, 1); + x_1456 = x_1220; } else { - lean_dec_ref(x_1200); - x_1432 = lean_box(0); + lean_dec_ref(x_1220); + x_1456 = lean_box(0); } -if (lean_is_scalar(x_1432)) { - x_1433 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1456)) { + x_1457 = lean_alloc_ctor(1, 2, 0); } else { - x_1433 = x_1432; + x_1457 = x_1456; } -lean_ctor_set(x_1433, 0, x_1430); -lean_ctor_set(x_1433, 1, x_1431); -return x_1433; +lean_ctor_set(x_1457, 0, x_1454); +lean_ctor_set(x_1457, 1, x_1455); +return x_1457; } } } @@ -9352,17 +9416,17 @@ lean_inc(x_22); lean_ctor_set_tag(x_66, 2); lean_ctor_set(x_66, 1, x_85); lean_ctor_set(x_66, 0, x_22); -x_86 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_86 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_22); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_22); lean_ctor_set(x_87, 1, x_86); -x_88 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_88 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_23); lean_inc(x_75); x_89 = l_Lean_addMacroScope(x_75, x_88, x_23); -x_90 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_91 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +x_90 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_91 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_22); x_92 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_92, 0, x_22); @@ -9409,7 +9473,7 @@ lean_inc(x_22); x_106 = l_Lean_Syntax_node2(x_22, x_81, x_92, x_105); lean_inc(x_22); x_107 = l_Lean_Syntax_node1(x_22, x_31, x_106); -x_108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +x_108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; lean_inc(x_87); lean_inc(x_22); x_109 = l_Lean_Syntax_node2(x_22, x_108, x_87, x_107); @@ -9587,17 +9651,17 @@ lean_inc(x_22); lean_ctor_set_tag(x_66, 2); lean_ctor_set(x_66, 1, x_169); lean_ctor_set(x_66, 0, x_22); -x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_22); x_171 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_171, 0, x_22); lean_ctor_set(x_171, 1, x_170); -x_172 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_172 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_23); lean_inc(x_158); x_173 = l_Lean_addMacroScope(x_158, x_172, x_23); -x_174 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +x_174 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_22); x_176 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_176, 0, x_22); @@ -9644,7 +9708,7 @@ lean_inc(x_22); x_190 = l_Lean_Syntax_node2(x_22, x_165, x_176, x_189); lean_inc(x_22); x_191 = l_Lean_Syntax_node1(x_22, x_31, x_190); -x_192 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +x_192 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; lean_inc(x_171); lean_inc(x_22); x_193 = l_Lean_Syntax_node2(x_22, x_192, x_171, x_191); @@ -9775,17 +9839,17 @@ lean_inc(x_22); x_233 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_233, 0, x_22); lean_ctor_set(x_233, 1, x_232); -x_234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_22); x_235 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_235, 0, x_22); lean_ctor_set(x_235, 1, x_234); -x_236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_23); lean_inc(x_221); x_237 = l_Lean_addMacroScope(x_221, x_236, x_23); -x_238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_239 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +x_238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_239 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_22); x_240 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_240, 0, x_22); @@ -9832,7 +9896,7 @@ lean_inc(x_22); x_254 = l_Lean_Syntax_node2(x_22, x_228, x_240, x_253); lean_inc(x_22); x_255 = l_Lean_Syntax_node1(x_22, x_31, x_254); -x_256 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +x_256 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; lean_inc(x_235); lean_inc(x_22); x_257 = l_Lean_Syntax_node2(x_22, x_256, x_235, x_255); @@ -10095,17 +10159,17 @@ if (lean_is_scalar(x_326)) { } lean_ctor_set(x_344, 0, x_22); lean_ctor_set(x_344, 1, x_343); -x_345 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; +x_345 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; lean_inc(x_22); x_346 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_346, 0, x_22); lean_ctor_set(x_346, 1, x_345); -x_347 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__94; +x_347 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__96; lean_inc(x_23); lean_inc(x_332); x_348 = l_Lean_addMacroScope(x_332, x_347, x_23); -x_349 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; -x_350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__97; +x_349 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; +x_350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_22); x_351 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_351, 0, x_22); @@ -10152,7 +10216,7 @@ lean_inc(x_22); x_365 = l_Lean_Syntax_node2(x_22, x_339, x_351, x_364); lean_inc(x_22); x_366 = l_Lean_Syntax_node1(x_22, x_287, x_365); -x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__92; lean_inc(x_346); lean_inc(x_22); x_368 = l_Lean_Syntax_node2(x_22, x_367, x_346, x_366); @@ -11111,7 +11175,7 @@ lean_dec(x_10); x_56 = !lean_is_exclusive(x_55); if (x_56 == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; size_t x_193; lean_object* x_194; size_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; size_t x_254; lean_object* x_255; size_t x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; size_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; size_t x_195; lean_object* x_196; size_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; size_t x_256; lean_object* x_257; size_t x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; x_57 = lean_ctor_get(x_55, 0); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); @@ -11314,11 +11378,11 @@ x_140 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__ lean_inc(x_60); lean_inc(x_33); x_141 = l_Lean_Syntax_node2(x_33, x_140, x_139, x_60); -x_142 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; +x_142 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; lean_inc(x_54); lean_inc(x_59); x_143 = l_Lean_addMacroScope(x_59, x_142, x_54); -x_144 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +x_144 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; lean_inc(x_33); x_145 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_145, 0, x_33); @@ -11345,11 +11409,11 @@ lean_ctor_set(x_152, 3, x_151); lean_inc(x_60); lean_inc(x_33); x_153 = l_Lean_Syntax_node2(x_33, x_140, x_152, x_60); -x_154 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; +x_154 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; lean_inc(x_54); lean_inc(x_59); x_155 = l_Lean_addMacroScope(x_59, x_154, x_54); -x_156 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +x_156 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; lean_inc(x_33); x_157 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_157, 0, x_33); @@ -11362,1422 +11426,1431 @@ lean_inc(x_33); x_158 = l_Lean_Syntax_node3(x_33, x_146, x_153, x_132, x_157); lean_inc(x_33); x_159 = l_Lean_Syntax_node3(x_33, x_50, x_147, x_86, x_158); -x_160 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +x_160 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_33); +x_161 = l_Lean_Syntax_node1(x_33, x_160, x_159); +x_162 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; lean_inc(x_60); lean_inc(x_33); -x_161 = l_Lean_Syntax_node1(x_33, x_160, x_60); -x_162 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; +x_163 = l_Lean_Syntax_node1(x_33, x_162, x_60); +x_164 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; lean_inc(x_33); -x_163 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_163, 0, x_33); -lean_ctor_set(x_163, 1, x_162); -x_164 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +x_165 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_165, 0, x_33); +lean_ctor_set(x_165, 1, x_164); +x_166 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; lean_inc_n(x_60, 2); lean_inc(x_33); -x_165 = l_Lean_Syntax_node6(x_33, x_164, x_134, x_60, x_159, x_161, x_60, x_163); -x_166 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; +x_167 = l_Lean_Syntax_node6(x_33, x_166, x_134, x_60, x_161, x_163, x_60, x_165); +x_168 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; lean_inc_n(x_60, 2); lean_inc(x_33); -x_167 = l_Lean_Syntax_node2(x_33, x_166, x_60, x_60); -x_168 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; +x_169 = l_Lean_Syntax_node2(x_33, x_168, x_60, x_60); +x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; lean_inc(x_54); lean_inc(x_59); -x_169 = l_Lean_addMacroScope(x_59, x_168, x_54); -x_170 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; -lean_inc(x_33); -x_171 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_171, 0, x_33); -lean_ctor_set(x_171, 1, x_170); -lean_ctor_set(x_171, 2, x_169); -lean_ctor_set(x_171, 3, x_38); -lean_inc(x_171); -lean_inc(x_33); -x_172 = l_Lean_Syntax_node1(x_33, x_50, x_171); -x_173 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; -lean_inc(x_33); -x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_33); -lean_ctor_set(x_174, 1, x_173); -x_175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +x_171 = l_Lean_addMacroScope(x_59, x_170, x_54); +x_172 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_33); +x_173 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_173, 0, x_33); +lean_ctor_set(x_173, 1, x_172); +lean_ctor_set(x_173, 2, x_171); +lean_ctor_set(x_173, 3, x_38); +lean_inc(x_173); +lean_inc(x_33); +x_174 = l_Lean_Syntax_node1(x_33, x_50, x_173); +x_175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; +lean_inc(x_33); +x_176 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_176, 0, x_33); +lean_ctor_set(x_176, 1, x_175); +x_177 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; lean_inc(x_54); lean_inc(x_59); -x_176 = l_Lean_addMacroScope(x_59, x_175, x_54); -x_177 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +x_178 = l_Lean_addMacroScope(x_59, x_177, x_54); +x_179 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; lean_inc(x_33); -x_178 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_178, 0, x_33); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_178, 2, x_176); -lean_ctor_set(x_178, 3, x_38); -x_179 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; +x_180 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_180, 0, x_33); +lean_ctor_set(x_180, 1, x_179); +lean_ctor_set(x_180, 2, x_178); +lean_ctor_set(x_180, 3, x_38); +x_181 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; lean_inc(x_33); -x_180 = l_Lean_Syntax_node1(x_33, x_179, x_178); +x_182 = l_Lean_Syntax_node1(x_33, x_181, x_180); lean_inc(x_33); -x_181 = l_Lean_Syntax_node1(x_33, x_50, x_128); -x_182 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; -lean_inc(x_165); +x_183 = l_Lean_Syntax_node1(x_33, x_50, x_128); +x_184 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +lean_inc(x_167); lean_inc(x_132); lean_inc(x_60); lean_inc(x_33); -x_183 = l_Lean_Syntax_node5(x_33, x_182, x_180, x_60, x_181, x_132, x_165); -x_184 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +x_185 = l_Lean_Syntax_node5(x_33, x_184, x_182, x_60, x_183, x_132, x_167); +x_186 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_inc(x_33); -x_185 = l_Lean_Syntax_node1(x_33, x_184, x_183); -x_186 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; +x_187 = l_Lean_Syntax_node1(x_33, x_186, x_185); +x_188 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; lean_inc(x_33); -x_187 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_187, 0, x_33); -lean_ctor_set(x_187, 1, x_186); -x_188 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; +x_189 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_189, 0, x_33); +lean_ctor_set(x_189, 1, x_188); +x_190 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; lean_inc(x_60); lean_inc(x_33); -x_189 = l_Lean_Syntax_node2(x_33, x_188, x_60, x_171); -lean_inc(x_33); -x_190 = l_Lean_Syntax_node1(x_33, x_50, x_189); -x_191 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_33); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_33); -lean_ctor_set(x_192, 1, x_191); -x_193 = lean_array_size(x_25); -x_194 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_38, x_193, x_15, x_25); -x_195 = lean_array_size(x_194); -x_196 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_195, x_15, x_194); -x_197 = l_Array_append___rarg(x_48, x_196); -lean_dec(x_196); -lean_inc(x_33); -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_33); -lean_ctor_set(x_198, 1, x_50); -lean_ctor_set(x_198, 2, x_197); -x_199 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; -lean_inc(x_33); -x_200 = l_Lean_Syntax_node1(x_33, x_199, x_198); -x_201 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_192); +x_191 = l_Lean_Syntax_node2(x_33, x_190, x_60, x_173); +lean_inc(x_33); +x_192 = l_Lean_Syntax_node1(x_33, x_50, x_191); +x_193 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +lean_inc(x_33); +x_194 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_194, 0, x_33); +lean_ctor_set(x_194, 1, x_193); +x_195 = lean_array_size(x_25); +x_196 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_38, x_195, x_15, x_25); +x_197 = lean_array_size(x_196); +x_198 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_197, x_15, x_196); +x_199 = l_Array_append___rarg(x_48, x_198); +lean_dec(x_198); +lean_inc(x_33); +x_200 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_200, 0, x_33); +lean_ctor_set(x_200, 1, x_50); +lean_ctor_set(x_200, 2, x_199); +x_201 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +lean_inc(x_33); +x_202 = l_Lean_Syntax_node1(x_33, x_201, x_200); +x_203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_194); lean_inc_n(x_60, 2); -lean_inc(x_187); +lean_inc(x_189); lean_inc(x_33); -x_202 = l_Lean_Syntax_node6(x_33, x_201, x_187, x_60, x_60, x_190, x_192, x_200); -x_203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +x_204 = l_Lean_Syntax_node6(x_33, x_203, x_189, x_60, x_60, x_192, x_194, x_202); +x_205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; lean_inc(x_60); -lean_inc(x_185); -lean_inc(x_174); +lean_inc(x_187); +lean_inc(x_176); lean_inc(x_33); -x_204 = l_Lean_Syntax_node4(x_33, x_203, x_174, x_185, x_60, x_202); -x_205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +x_206 = l_Lean_Syntax_node4(x_33, x_205, x_176, x_187, x_60, x_204); +x_207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; lean_inc(x_132); lean_inc(x_60); lean_inc(x_33); -x_206 = l_Lean_Syntax_node5(x_33, x_205, x_145, x_172, x_60, x_132, x_204); -x_207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +x_208 = l_Lean_Syntax_node5(x_33, x_207, x_145, x_174, x_60, x_132, x_206); +x_209 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; lean_inc(x_33); -x_208 = l_Lean_Syntax_node1(x_33, x_207, x_206); -x_209 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_167); +x_210 = l_Lean_Syntax_node1(x_33, x_209, x_208); +x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_169); lean_inc_n(x_60, 2); lean_inc(x_33); -x_210 = l_Lean_Syntax_node4(x_33, x_209, x_60, x_60, x_208, x_167); -x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; +x_212 = l_Lean_Syntax_node4(x_33, x_211, x_60, x_60, x_210, x_169); +x_213 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_54); lean_inc(x_59); -x_212 = l_Lean_addMacroScope(x_59, x_211, x_54); -x_213 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_214 = l_Lean_addMacroScope(x_59, x_213, x_54); +x_215 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; lean_inc(x_33); -x_214 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_214, 0, x_33); -lean_ctor_set(x_214, 1, x_213); -lean_ctor_set(x_214, 2, x_212); -lean_ctor_set(x_214, 3, x_38); +x_216 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_216, 0, x_33); +lean_ctor_set(x_216, 1, x_215); +lean_ctor_set(x_216, 2, x_214); +lean_ctor_set(x_216, 3, x_38); lean_inc(x_33); -x_215 = l_Lean_Syntax_node1(x_33, x_50, x_214); -x_216 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +x_217 = l_Lean_Syntax_node1(x_33, x_50, x_216); +x_218 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; lean_inc(x_33); -x_217 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_217, 0, x_33); -lean_ctor_set(x_217, 1, x_216); -x_218 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_33); +lean_ctor_set(x_219, 1, x_218); +x_220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; lean_inc(x_33); -x_219 = l_Lean_Syntax_node2(x_33, x_218, x_174, x_185); -x_220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +x_221 = l_Lean_Syntax_node2(x_33, x_220, x_176, x_187); +x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; lean_inc(x_60); lean_inc(x_33); -x_221 = l_Lean_Syntax_node2(x_33, x_220, x_219, x_60); -x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +x_223 = l_Lean_Syntax_node2(x_33, x_222, x_221, x_60); +x_224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; lean_inc(x_33); -x_223 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_223, 0, x_33); -lean_ctor_set(x_223, 1, x_222); -x_224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +x_225 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_225, 0, x_33); +lean_ctor_set(x_225, 1, x_224); +x_226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; lean_inc(x_54); lean_inc(x_59); -x_225 = l_Lean_addMacroScope(x_59, x_224, x_54); -x_226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +x_227 = l_Lean_addMacroScope(x_59, x_226, x_54); +x_228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; lean_inc(x_33); -x_227 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_227, 0, x_33); -lean_ctor_set(x_227, 1, x_226); -lean_ctor_set(x_227, 2, x_225); -lean_ctor_set(x_227, 3, x_38); +x_229 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_229, 0, x_33); +lean_ctor_set(x_229, 1, x_228); +lean_ctor_set(x_229, 2, x_227); +lean_ctor_set(x_229, 3, x_38); lean_inc(x_33); -x_228 = l_Lean_Syntax_node2(x_33, x_127, x_119, x_67); +x_230 = l_Lean_Syntax_node2(x_33, x_127, x_119, x_67); lean_inc(x_33); -x_229 = l_Lean_Syntax_node1(x_33, x_50, x_228); -x_230 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_231 = l_Lean_Syntax_node1(x_33, x_50, x_230); +x_232 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_33); -x_231 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_231, 0, x_33); -lean_ctor_set(x_231, 1, x_230); -x_232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_33); +lean_ctor_set(x_233, 1, x_232); +x_234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; lean_inc(x_54); lean_inc(x_59); -x_233 = l_Lean_addMacroScope(x_59, x_232, x_54); -x_234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_235 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_33); -x_236 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_236, 0, x_33); -lean_ctor_set(x_236, 1, x_234); -lean_ctor_set(x_236, 2, x_233); -lean_ctor_set(x_236, 3, x_235); -lean_inc(x_215); +x_235 = l_Lean_addMacroScope(x_59, x_234, x_54); +x_236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_237 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_33); +x_238 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_238, 0, x_33); +lean_ctor_set(x_238, 1, x_236); +lean_ctor_set(x_238, 2, x_235); +lean_ctor_set(x_238, 3, x_237); +lean_inc(x_217); lean_inc(x_33); -x_237 = l_Lean_Syntax_node2(x_33, x_52, x_236, x_215); -x_238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_239 = l_Lean_Syntax_node2(x_33, x_52, x_238, x_217); +x_240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; lean_inc(x_33); -x_239 = l_Lean_Syntax_node1(x_33, x_238, x_237); -x_240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_227); +x_241 = l_Lean_Syntax_node1(x_33, x_240, x_239); +x_242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_229); lean_inc(x_33); -x_241 = l_Lean_Syntax_node4(x_33, x_240, x_227, x_229, x_231, x_239); -x_242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +x_243 = l_Lean_Syntax_node4(x_33, x_242, x_229, x_231, x_233, x_241); +x_244 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; lean_inc(x_60); lean_inc(x_33); -x_243 = l_Lean_Syntax_node3(x_33, x_242, x_223, x_60, x_241); +x_245 = l_Lean_Syntax_node3(x_33, x_244, x_225, x_60, x_243); lean_inc(x_60); lean_inc(x_33); -x_244 = l_Lean_Syntax_node2(x_33, x_220, x_243, x_60); -x_245 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; -x_246 = l_Lean_addMacroScope(x_59, x_245, x_54); -x_247 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; -x_248 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; +x_246 = l_Lean_Syntax_node2(x_33, x_222, x_245, x_60); +x_247 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; +x_248 = l_Lean_addMacroScope(x_59, x_247, x_54); +x_249 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; +x_250 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; lean_inc(x_33); -x_249 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_249, 0, x_33); -lean_ctor_set(x_249, 1, x_247); -lean_ctor_set(x_249, 2, x_246); -lean_ctor_set(x_249, 3, x_248); -x_250 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; -lean_inc(x_33); -x_251 = lean_alloc_ctor(2, 2, 0); +x_251 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_251, 0, x_33); -lean_ctor_set(x_251, 1, x_250); +lean_ctor_set(x_251, 1, x_249); +lean_ctor_set(x_251, 2, x_248); +lean_ctor_set(x_251, 3, x_250); +x_252 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; +lean_inc(x_33); +x_253 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_253, 0, x_33); +lean_ctor_set(x_253, 1, x_252); lean_inc(x_60); lean_inc(x_33); -x_252 = l_Lean_Syntax_node2(x_33, x_188, x_60, x_227); +x_254 = l_Lean_Syntax_node2(x_33, x_190, x_60, x_229); lean_inc(x_33); -x_253 = l_Lean_Syntax_node1(x_33, x_50, x_252); -x_254 = lean_array_size(x_26); -x_255 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_38, x_254, x_15, x_26); -x_256 = lean_array_size(x_255); -x_257 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_256, x_15, x_255); -x_258 = l_Array_append___rarg(x_48, x_257); -lean_dec(x_257); +x_255 = l_Lean_Syntax_node1(x_33, x_50, x_254); +x_256 = lean_array_size(x_26); +x_257 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_38, x_256, x_15, x_26); +x_258 = lean_array_size(x_257); +x_259 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_258, x_15, x_257); +x_260 = l_Array_append___rarg(x_48, x_259); +lean_dec(x_259); lean_inc(x_33); -x_259 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_259, 0, x_33); -lean_ctor_set(x_259, 1, x_50); -lean_ctor_set(x_259, 2, x_258); +x_261 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_261, 0, x_33); +lean_ctor_set(x_261, 1, x_50); +lean_ctor_set(x_261, 2, x_260); lean_inc(x_33); -x_260 = l_Lean_Syntax_node1(x_33, x_199, x_259); +x_262 = l_Lean_Syntax_node1(x_33, x_201, x_261); lean_inc_n(x_60, 2); lean_inc(x_33); -x_261 = l_Lean_Syntax_node6(x_33, x_201, x_187, x_60, x_60, x_253, x_192, x_260); -x_262 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +x_263 = l_Lean_Syntax_node6(x_33, x_203, x_189, x_60, x_60, x_255, x_194, x_262); +x_264 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; lean_inc(x_33); -x_263 = l_Lean_Syntax_node3(x_33, x_262, x_249, x_251, x_261); +x_265 = l_Lean_Syntax_node3(x_33, x_264, x_251, x_253, x_263); lean_inc(x_33); -x_264 = l_Lean_Syntax_node1(x_33, x_238, x_263); +x_266 = l_Lean_Syntax_node1(x_33, x_240, x_265); lean_inc(x_60); lean_inc(x_33); -x_265 = l_Lean_Syntax_node2(x_33, x_220, x_264, x_60); +x_267 = l_Lean_Syntax_node2(x_33, x_222, x_266, x_60); lean_inc(x_33); -x_266 = l_Lean_Syntax_node3(x_33, x_50, x_221, x_244, x_265); -x_267 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_268 = l_Lean_Syntax_node3(x_33, x_50, x_223, x_246, x_267); +x_269 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_33); -x_268 = l_Lean_Syntax_node1(x_33, x_267, x_266); -x_269 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_270 = l_Lean_Syntax_node1(x_33, x_269, x_268); +x_271 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_33); -x_270 = l_Lean_Syntax_node2(x_33, x_269, x_217, x_268); +x_272 = l_Lean_Syntax_node2(x_33, x_271, x_219, x_270); lean_inc(x_132); lean_inc(x_60); lean_inc(x_33); -x_271 = l_Lean_Syntax_node5(x_33, x_205, x_157, x_215, x_60, x_132, x_270); +x_273 = l_Lean_Syntax_node5(x_33, x_207, x_157, x_217, x_60, x_132, x_272); lean_inc(x_33); -x_272 = l_Lean_Syntax_node1(x_33, x_207, x_271); -lean_inc(x_167); +x_274 = l_Lean_Syntax_node1(x_33, x_209, x_273); +lean_inc(x_169); lean_inc_n(x_60, 2); lean_inc(x_33); -x_273 = l_Lean_Syntax_node4(x_33, x_209, x_60, x_60, x_272, x_167); +x_275 = l_Lean_Syntax_node4(x_33, x_211, x_60, x_60, x_274, x_169); lean_inc(x_60); lean_inc(x_33); -x_274 = l_Lean_Syntax_node3(x_33, x_50, x_210, x_60, x_273); -x_275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_276 = l_Lean_Syntax_node3(x_33, x_50, x_212, x_60, x_275); +x_277 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_33); -x_276 = l_Lean_Syntax_node2(x_33, x_275, x_19, x_274); +x_278 = l_Lean_Syntax_node2(x_33, x_277, x_19, x_276); lean_inc(x_33); -x_277 = l_Lean_Syntax_node1(x_33, x_50, x_276); -x_278 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +x_279 = l_Lean_Syntax_node1(x_33, x_50, x_278); +x_280 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; lean_inc(x_33); -x_279 = l_Lean_Syntax_node4(x_33, x_278, x_132, x_165, x_167, x_277); -x_280 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +x_281 = l_Lean_Syntax_node4(x_33, x_280, x_132, x_167, x_169, x_279); +x_282 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; lean_inc(x_60); lean_inc(x_33); -x_281 = l_Lean_Syntax_node6(x_33, x_280, x_115, x_117, x_60, x_60, x_130, x_279); +x_283 = l_Lean_Syntax_node6(x_33, x_282, x_115, x_117, x_60, x_60, x_130, x_281); lean_inc(x_33); -x_282 = l_Lean_Syntax_node2(x_33, x_98, x_113, x_281); -x_283 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +x_284 = l_Lean_Syntax_node2(x_33, x_98, x_113, x_283); +x_285 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; lean_inc(x_33); -x_284 = l_Lean_Syntax_node3(x_33, x_283, x_105, x_107, x_282); -x_285 = l_Lean_Syntax_node2(x_33, x_50, x_99, x_284); -lean_ctor_set(x_55, 0, x_285); +x_286 = l_Lean_Syntax_node3(x_33, x_285, x_105, x_107, x_284); +x_287 = l_Lean_Syntax_node2(x_33, x_50, x_99, x_286); +lean_ctor_set(x_55, 0, x_287); return x_55; } else { -lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; size_t x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; size_t x_423; lean_object* x_424; size_t x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; size_t x_484; lean_object* x_485; size_t x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; -x_286 = lean_ctor_get(x_55, 0); -x_287 = lean_ctor_get(x_55, 1); -lean_inc(x_287); -lean_inc(x_286); -lean_dec(x_55); -x_288 = lean_ctor_get(x_286, 0); +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; size_t x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; size_t x_427; lean_object* x_428; size_t x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; size_t x_488; lean_object* x_489; size_t x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; +x_288 = lean_ctor_get(x_55, 0); +x_289 = lean_ctor_get(x_55, 1); +lean_inc(x_289); lean_inc(x_288); -lean_dec(x_286); -x_289 = lean_environment_main_module(x_288); +lean_dec(x_55); +x_290 = lean_ctor_get(x_288, 0); +lean_inc(x_290); +lean_dec(x_288); +x_291 = lean_environment_main_module(x_290); lean_inc(x_33); -x_290 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_290, 0, x_33); -lean_ctor_set(x_290, 1, x_50); -lean_ctor_set(x_290, 2, x_48); -x_291 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_290, 6); +x_292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_292, 0, x_33); +lean_ctor_set(x_292, 1, x_50); +lean_ctor_set(x_292, 2, x_48); +x_293 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_292, 6); lean_inc(x_33); -x_292 = l_Lean_Syntax_node6(x_33, x_291, x_290, x_290, x_290, x_290, x_290, x_290); -x_293 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; +x_294 = l_Lean_Syntax_node6(x_33, x_293, x_292, x_292, x_292, x_292, x_292, x_292); +x_295 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; lean_inc(x_33); lean_ctor_set_tag(x_23, 2); -lean_ctor_set(x_23, 1, x_293); +lean_ctor_set(x_23, 1, x_295); lean_ctor_set(x_23, 0, x_33); -x_294 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +x_296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; lean_inc(x_54); -lean_inc(x_289); -x_295 = l_Lean_addMacroScope(x_289, x_294, x_54); -x_296 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_33); -x_297 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_297, 0, x_33); -lean_ctor_set(x_297, 1, x_296); -lean_ctor_set(x_297, 2, x_295); -lean_ctor_set(x_297, 3, x_38); -x_298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_290); -lean_inc(x_297); +lean_inc(x_291); +x_297 = l_Lean_addMacroScope(x_291, x_296, x_54); +x_298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_33); +x_299 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_299, 0, x_33); +lean_ctor_set(x_299, 1, x_298); +lean_ctor_set(x_299, 2, x_297); +lean_ctor_set(x_299, 3, x_38); +x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_292); +lean_inc(x_299); lean_inc(x_33); -x_299 = l_Lean_Syntax_node2(x_33, x_298, x_297, x_290); -x_300 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; -lean_inc_n(x_290, 2); +x_301 = l_Lean_Syntax_node2(x_33, x_300, x_299, x_292); +x_302 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; +lean_inc_n(x_292, 2); lean_inc(x_33); -x_301 = l_Lean_Syntax_node2(x_33, x_300, x_290, x_290); -x_302 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +x_303 = l_Lean_Syntax_node2(x_33, x_302, x_292, x_292); +x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; lean_inc(x_33); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_302); +lean_ctor_set(x_19, 1, x_304); lean_ctor_set(x_19, 0, x_33); lean_inc(x_19); lean_inc(x_33); -x_303 = l_Lean_Syntax_node1(x_33, x_50, x_19); -x_304 = lean_array_size(x_21); -x_305 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_304, x_15, x_21); -x_306 = l_Array_append___rarg(x_48, x_305); -lean_dec(x_305); -lean_inc(x_33); -x_307 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_307, 0, x_33); -lean_ctor_set(x_307, 1, x_50); -lean_ctor_set(x_307, 2, x_306); -x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +x_305 = l_Lean_Syntax_node1(x_33, x_50, x_19); +x_306 = lean_array_size(x_21); +x_307 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_306, x_15, x_21); +x_308 = l_Array_append___rarg(x_48, x_307); +lean_dec(x_307); lean_inc(x_33); -x_309 = lean_alloc_ctor(2, 2, 0); +x_309 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_309, 0, x_33); -lean_ctor_set(x_309, 1, x_308); -x_310 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_54); -lean_inc(x_289); -x_311 = l_Lean_addMacroScope(x_289, x_310, x_54); -x_312 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_313 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_ctor_set(x_309, 1, x_50); +lean_ctor_set(x_309, 2, x_308); +x_310 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; lean_inc(x_33); -x_314 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_314, 0, x_33); -lean_ctor_set(x_314, 1, x_312); -lean_ctor_set(x_314, 2, x_311); -lean_ctor_set(x_314, 3, x_313); -x_315 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +x_311 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_311, 0, x_33); +lean_ctor_set(x_311, 1, x_310); +x_312 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_54); +lean_inc(x_291); +x_313 = l_Lean_addMacroScope(x_291, x_312, x_54); +x_314 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_315 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; lean_inc(x_33); -x_316 = lean_alloc_ctor(2, 2, 0); +x_316 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_316, 0, x_33); -lean_ctor_set(x_316, 1, x_315); -x_317 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_ctor_set(x_316, 1, x_314); +lean_ctor_set(x_316, 2, x_313); +lean_ctor_set(x_316, 3, x_315); +x_317 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_33); +x_318 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_318, 0, x_33); +lean_ctor_set(x_318, 1, x_317); +x_319 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; lean_inc(x_54); -lean_inc(x_289); -x_318 = l_Lean_addMacroScope(x_289, x_317, x_54); -x_319 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_33); -x_321 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_321, 0, x_33); -lean_ctor_set(x_321, 1, x_319); -lean_ctor_set(x_321, 2, x_318); -lean_ctor_set(x_321, 3, x_320); -lean_inc(x_316); -lean_inc(x_33); -x_322 = l_Lean_Syntax_node3(x_33, x_50, x_314, x_316, x_321); +lean_inc(x_291); +x_320 = l_Lean_addMacroScope(x_291, x_319, x_54); +x_321 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_33); +x_323 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_323, 0, x_33); +lean_ctor_set(x_323, 1, x_321); +lean_ctor_set(x_323, 2, x_320); +lean_ctor_set(x_323, 3, x_322); +lean_inc(x_318); lean_inc(x_33); -x_323 = l_Lean_Syntax_node2(x_33, x_50, x_309, x_322); -x_324 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +x_324 = l_Lean_Syntax_node3(x_33, x_50, x_316, x_318, x_323); lean_inc(x_33); -x_325 = l_Lean_Syntax_node1(x_33, x_324, x_323); -x_326 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; -lean_inc(x_290); +x_325 = l_Lean_Syntax_node2(x_33, x_50, x_311, x_324); +x_326 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; lean_inc(x_33); -x_327 = l_Lean_Syntax_node7(x_33, x_326, x_23, x_299, x_301, x_303, x_307, x_290, x_325); -x_328 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +x_327 = l_Lean_Syntax_node1(x_33, x_326, x_325); +x_328 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; +lean_inc(x_292); lean_inc(x_33); -x_329 = l_Lean_Syntax_node2(x_33, x_328, x_292, x_327); -x_330 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +x_329 = l_Lean_Syntax_node7(x_33, x_328, x_23, x_301, x_303, x_305, x_309, x_292, x_327); +x_330 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; lean_inc(x_33); -x_331 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_331, 0, x_33); -lean_ctor_set(x_331, 1, x_330); -x_332 = l_Array_append___rarg(x_48, x_3); +x_331 = l_Lean_Syntax_node2(x_33, x_330, x_294, x_329); +x_332 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; lean_inc(x_33); -x_333 = lean_alloc_ctor(1, 3, 0); +x_333 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_333, 0, x_33); -lean_ctor_set(x_333, 1, x_50); -lean_ctor_set(x_333, 2, x_332); -x_334 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_ctor_set(x_333, 1, x_332); +x_334 = l_Array_append___rarg(x_48, x_3); lean_inc(x_33); -x_335 = l_Lean_Syntax_node2(x_33, x_334, x_331, x_333); -x_336 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +x_335 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_335, 0, x_33); +lean_ctor_set(x_335, 1, x_50); +lean_ctor_set(x_335, 2, x_334); +x_336 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; lean_inc(x_33); -x_337 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_337, 0, x_33); -lean_ctor_set(x_337, 1, x_336); -x_338 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; +x_337 = l_Lean_Syntax_node2(x_33, x_336, x_333, x_335); +x_338 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; lean_inc(x_33); x_339 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_339, 0, x_33); lean_ctor_set(x_339, 1, x_338); -x_340 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; +x_340 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; lean_inc(x_33); -x_341 = l_Lean_Syntax_node1(x_33, x_340, x_339); +x_341 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_341, 0, x_33); +lean_ctor_set(x_341, 1, x_340); +x_342 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; lean_inc(x_33); -x_342 = l_Lean_Syntax_node1(x_33, x_50, x_341); -lean_inc_n(x_290, 5); +x_343 = l_Lean_Syntax_node1(x_33, x_342, x_341); lean_inc(x_33); -x_343 = l_Lean_Syntax_node6(x_33, x_291, x_290, x_290, x_290, x_290, x_290, x_342); -x_344 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_290); +x_344 = l_Lean_Syntax_node1(x_33, x_50, x_343); +lean_inc_n(x_292, 5); lean_inc(x_33); -x_345 = l_Lean_Syntax_node1(x_33, x_344, x_290); -x_346 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +x_345 = l_Lean_Syntax_node6(x_33, x_293, x_292, x_292, x_292, x_292, x_292, x_344); +x_346 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_292); lean_inc(x_33); -x_347 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_347, 0, x_33); -lean_ctor_set(x_347, 1, x_346); -x_348 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +x_347 = l_Lean_Syntax_node1(x_33, x_346, x_292); +x_348 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; lean_inc(x_33); x_349 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_349, 0, x_33); lean_ctor_set(x_349, 1, x_348); -x_350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_54); -lean_inc(x_289); -x_351 = l_Lean_addMacroScope(x_289, x_350, x_54); -x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_353 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +x_350 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; lean_inc(x_33); -x_354 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_354, 0, x_33); -lean_ctor_set(x_354, 1, x_352); -lean_ctor_set(x_354, 2, x_351); -lean_ctor_set(x_354, 3, x_353); +x_351 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_351, 0, x_33); +lean_ctor_set(x_351, 1, x_350); +x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_54); +lean_inc(x_291); +x_353 = l_Lean_addMacroScope(x_291, x_352, x_54); +x_354 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_355 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; lean_inc(x_33); -x_355 = l_Lean_Syntax_node1(x_33, x_50, x_53); +x_356 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_356, 0, x_33); +lean_ctor_set(x_356, 1, x_354); +lean_ctor_set(x_356, 2, x_353); +lean_ctor_set(x_356, 3, x_355); lean_inc(x_33); -x_356 = l_Lean_Syntax_node2(x_33, x_52, x_354, x_355); -x_357 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_349); +x_357 = l_Lean_Syntax_node1(x_33, x_50, x_53); lean_inc(x_33); -x_358 = l_Lean_Syntax_node2(x_33, x_357, x_349, x_356); -x_359 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_358); -lean_inc(x_290); +x_358 = l_Lean_Syntax_node2(x_33, x_52, x_356, x_357); +x_359 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_351); lean_inc(x_33); -x_360 = l_Lean_Syntax_node2(x_33, x_359, x_290, x_358); -x_361 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +x_360 = l_Lean_Syntax_node2(x_33, x_359, x_351, x_358); +x_361 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_360); +lean_inc(x_292); lean_inc(x_33); -x_362 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_362, 0, x_33); -lean_ctor_set(x_362, 1, x_361); -x_363 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +x_362 = l_Lean_Syntax_node2(x_33, x_361, x_292, x_360); +x_363 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; lean_inc(x_33); x_364 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_364, 0, x_33); lean_ctor_set(x_364, 1, x_363); -x_365 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_54); -lean_inc(x_289); -x_366 = l_Lean_addMacroScope(x_289, x_365, x_54); -x_367 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_368 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_33); -x_369 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_369, 0, x_33); -lean_ctor_set(x_369, 1, x_367); -lean_ctor_set(x_369, 2, x_366); -lean_ctor_set(x_369, 3, x_368); -x_370 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_290); +x_365 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; lean_inc(x_33); -x_371 = l_Lean_Syntax_node2(x_33, x_370, x_369, x_290); -x_372 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; +x_366 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_366, 0, x_33); +lean_ctor_set(x_366, 1, x_365); +x_367 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; lean_inc(x_54); -lean_inc(x_289); -x_373 = l_Lean_addMacroScope(x_289, x_372, x_54); -x_374 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_291); +x_368 = l_Lean_addMacroScope(x_291, x_367, x_54); +x_369 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_370 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; lean_inc(x_33); -x_375 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_375, 0, x_33); -lean_ctor_set(x_375, 1, x_374); -lean_ctor_set(x_375, 2, x_373); -lean_ctor_set(x_375, 3, x_38); -x_376 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_375); -lean_inc(x_362); +x_371 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_371, 0, x_33); +lean_ctor_set(x_371, 1, x_369); +lean_ctor_set(x_371, 2, x_368); +lean_ctor_set(x_371, 3, x_370); +x_372 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_292); lean_inc(x_33); -x_377 = l_Lean_Syntax_node3(x_33, x_376, x_371, x_362, x_375); -x_378 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +x_373 = l_Lean_Syntax_node2(x_33, x_372, x_371, x_292); +x_374 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; lean_inc(x_54); -lean_inc(x_289); -x_379 = l_Lean_addMacroScope(x_289, x_378, x_54); -x_380 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_381 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_33); -x_382 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_382, 0, x_33); -lean_ctor_set(x_382, 1, x_380); -lean_ctor_set(x_382, 2, x_379); -lean_ctor_set(x_382, 3, x_381); -lean_inc(x_290); +lean_inc(x_291); +x_375 = l_Lean_addMacroScope(x_291, x_374, x_54); +x_376 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_33); +x_377 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_377, 0, x_33); +lean_ctor_set(x_377, 1, x_376); +lean_ctor_set(x_377, 2, x_375); +lean_ctor_set(x_377, 3, x_38); +x_378 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_377); +lean_inc(x_364); lean_inc(x_33); -x_383 = l_Lean_Syntax_node2(x_33, x_370, x_382, x_290); -x_384 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; +x_379 = l_Lean_Syntax_node3(x_33, x_378, x_373, x_364, x_377); +x_380 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; lean_inc(x_54); -lean_inc(x_289); -x_385 = l_Lean_addMacroScope(x_289, x_384, x_54); -x_386 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_33); -x_387 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_387, 0, x_33); -lean_ctor_set(x_387, 1, x_386); -lean_ctor_set(x_387, 2, x_385); -lean_ctor_set(x_387, 3, x_38); -lean_inc(x_387); -lean_inc(x_362); +lean_inc(x_291); +x_381 = l_Lean_addMacroScope(x_291, x_380, x_54); +x_382 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_383 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_33); +x_384 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_384, 0, x_33); +lean_ctor_set(x_384, 1, x_382); +lean_ctor_set(x_384, 2, x_381); +lean_ctor_set(x_384, 3, x_383); +lean_inc(x_292); lean_inc(x_33); -x_388 = l_Lean_Syntax_node3(x_33, x_376, x_383, x_362, x_387); +x_385 = l_Lean_Syntax_node2(x_33, x_372, x_384, x_292); +x_386 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_54); +lean_inc(x_291); +x_387 = l_Lean_addMacroScope(x_291, x_386, x_54); +x_388 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; lean_inc(x_33); -x_389 = l_Lean_Syntax_node3(x_33, x_50, x_377, x_316, x_388); -x_390 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_290); +x_389 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_389, 0, x_33); +lean_ctor_set(x_389, 1, x_388); +lean_ctor_set(x_389, 2, x_387); +lean_ctor_set(x_389, 3, x_38); +lean_inc(x_389); +lean_inc(x_364); lean_inc(x_33); -x_391 = l_Lean_Syntax_node1(x_33, x_390, x_290); -x_392 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; +x_390 = l_Lean_Syntax_node3(x_33, x_378, x_385, x_364, x_389); lean_inc(x_33); -x_393 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_393, 0, x_33); -lean_ctor_set(x_393, 1, x_392); -x_394 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc_n(x_290, 2); +x_391 = l_Lean_Syntax_node3(x_33, x_50, x_379, x_318, x_390); +x_392 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; lean_inc(x_33); -x_395 = l_Lean_Syntax_node6(x_33, x_394, x_364, x_290, x_389, x_391, x_290, x_393); -x_396 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_290, 2); +x_393 = l_Lean_Syntax_node1(x_33, x_392, x_391); +x_394 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_292); lean_inc(x_33); -x_397 = l_Lean_Syntax_node2(x_33, x_396, x_290, x_290); -x_398 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; -lean_inc(x_54); -lean_inc(x_289); -x_399 = l_Lean_addMacroScope(x_289, x_398, x_54); -x_400 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +x_395 = l_Lean_Syntax_node1(x_33, x_394, x_292); +x_396 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; lean_inc(x_33); -x_401 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_401, 0, x_33); -lean_ctor_set(x_401, 1, x_400); -lean_ctor_set(x_401, 2, x_399); -lean_ctor_set(x_401, 3, x_38); -lean_inc(x_401); +x_397 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_397, 0, x_33); +lean_ctor_set(x_397, 1, x_396); +x_398 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc_n(x_292, 2); lean_inc(x_33); -x_402 = l_Lean_Syntax_node1(x_33, x_50, x_401); -x_403 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; +x_399 = l_Lean_Syntax_node6(x_33, x_398, x_366, x_292, x_393, x_395, x_292, x_397); +x_400 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_292, 2); lean_inc(x_33); -x_404 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_404, 0, x_33); -lean_ctor_set(x_404, 1, x_403); -x_405 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +x_401 = l_Lean_Syntax_node2(x_33, x_400, x_292, x_292); +x_402 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; lean_inc(x_54); -lean_inc(x_289); -x_406 = l_Lean_addMacroScope(x_289, x_405, x_54); -x_407 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +lean_inc(x_291); +x_403 = l_Lean_addMacroScope(x_291, x_402, x_54); +x_404 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_33); +x_405 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_405, 0, x_33); +lean_ctor_set(x_405, 1, x_404); +lean_ctor_set(x_405, 2, x_403); +lean_ctor_set(x_405, 3, x_38); +lean_inc(x_405); +lean_inc(x_33); +x_406 = l_Lean_Syntax_node1(x_33, x_50, x_405); +x_407 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; lean_inc(x_33); -x_408 = lean_alloc_ctor(3, 4, 0); +x_408 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_408, 0, x_33); lean_ctor_set(x_408, 1, x_407); -lean_ctor_set(x_408, 2, x_406); -lean_ctor_set(x_408, 3, x_38); -x_409 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; +x_409 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +lean_inc(x_54); +lean_inc(x_291); +x_410 = l_Lean_addMacroScope(x_291, x_409, x_54); +x_411 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; lean_inc(x_33); -x_410 = l_Lean_Syntax_node1(x_33, x_409, x_408); +x_412 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_412, 0, x_33); +lean_ctor_set(x_412, 1, x_411); +lean_ctor_set(x_412, 2, x_410); +lean_ctor_set(x_412, 3, x_38); +x_413 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; lean_inc(x_33); -x_411 = l_Lean_Syntax_node1(x_33, x_50, x_358); -x_412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; -lean_inc(x_395); -lean_inc(x_362); -lean_inc(x_290); +x_414 = l_Lean_Syntax_node1(x_33, x_413, x_412); lean_inc(x_33); -x_413 = l_Lean_Syntax_node5(x_33, x_412, x_410, x_290, x_411, x_362, x_395); -x_414 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +x_415 = l_Lean_Syntax_node1(x_33, x_50, x_360); +x_416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +lean_inc(x_399); +lean_inc(x_364); +lean_inc(x_292); lean_inc(x_33); -x_415 = l_Lean_Syntax_node1(x_33, x_414, x_413); -x_416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; +x_417 = l_Lean_Syntax_node5(x_33, x_416, x_414, x_292, x_415, x_364, x_399); +x_418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_inc(x_33); -x_417 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_417, 0, x_33); -lean_ctor_set(x_417, 1, x_416); -x_418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; -lean_inc(x_290); +x_419 = l_Lean_Syntax_node1(x_33, x_418, x_417); +x_420 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; lean_inc(x_33); -x_419 = l_Lean_Syntax_node2(x_33, x_418, x_290, x_401); -lean_inc(x_33); -x_420 = l_Lean_Syntax_node1(x_33, x_50, x_419); -x_421 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_33); -x_422 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_422, 0, x_33); -lean_ctor_set(x_422, 1, x_421); -x_423 = lean_array_size(x_25); -x_424 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_38, x_423, x_15, x_25); -x_425 = lean_array_size(x_424); -x_426 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_425, x_15, x_424); -x_427 = l_Array_append___rarg(x_48, x_426); -lean_dec(x_426); -lean_inc(x_33); -x_428 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_428, 0, x_33); -lean_ctor_set(x_428, 1, x_50); -lean_ctor_set(x_428, 2, x_427); -x_429 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; -lean_inc(x_33); -x_430 = l_Lean_Syntax_node1(x_33, x_429, x_428); -x_431 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_422); -lean_inc_n(x_290, 2); -lean_inc(x_417); +x_421 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_421, 0, x_33); +lean_ctor_set(x_421, 1, x_420); +x_422 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; +lean_inc(x_292); lean_inc(x_33); -x_432 = l_Lean_Syntax_node6(x_33, x_431, x_417, x_290, x_290, x_420, x_422, x_430); -x_433 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; -lean_inc(x_290); -lean_inc(x_415); -lean_inc(x_404); +x_423 = l_Lean_Syntax_node2(x_33, x_422, x_292, x_405); lean_inc(x_33); -x_434 = l_Lean_Syntax_node4(x_33, x_433, x_404, x_415, x_290, x_432); -x_435 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_362); -lean_inc(x_290); +x_424 = l_Lean_Syntax_node1(x_33, x_50, x_423); +x_425 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; lean_inc(x_33); -x_436 = l_Lean_Syntax_node5(x_33, x_435, x_375, x_402, x_290, x_362, x_434); -x_437 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +x_426 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_426, 0, x_33); +lean_ctor_set(x_426, 1, x_425); +x_427 = lean_array_size(x_25); +x_428 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_38, x_427, x_15, x_25); +x_429 = lean_array_size(x_428); +x_430 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_429, x_15, x_428); +x_431 = l_Array_append___rarg(x_48, x_430); +lean_dec(x_430); +lean_inc(x_33); +x_432 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_432, 0, x_33); +lean_ctor_set(x_432, 1, x_50); +lean_ctor_set(x_432, 2, x_431); +x_433 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +lean_inc(x_33); +x_434 = l_Lean_Syntax_node1(x_33, x_433, x_432); +x_435 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_426); +lean_inc_n(x_292, 2); +lean_inc(x_421); +lean_inc(x_33); +x_436 = l_Lean_Syntax_node6(x_33, x_435, x_421, x_292, x_292, x_424, x_426, x_434); +x_437 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +lean_inc(x_292); +lean_inc(x_419); +lean_inc(x_408); lean_inc(x_33); -x_438 = l_Lean_Syntax_node1(x_33, x_437, x_436); -x_439 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_397); -lean_inc_n(x_290, 2); +x_438 = l_Lean_Syntax_node4(x_33, x_437, x_408, x_419, x_292, x_436); +x_439 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_364); +lean_inc(x_292); lean_inc(x_33); -x_440 = l_Lean_Syntax_node4(x_33, x_439, x_290, x_290, x_438, x_397); -x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_54); -lean_inc(x_289); -x_442 = l_Lean_addMacroScope(x_289, x_441, x_54); -x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_440 = l_Lean_Syntax_node5(x_33, x_439, x_377, x_406, x_292, x_364, x_438); +x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; lean_inc(x_33); -x_444 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_444, 0, x_33); -lean_ctor_set(x_444, 1, x_443); -lean_ctor_set(x_444, 2, x_442); -lean_ctor_set(x_444, 3, x_38); +x_442 = l_Lean_Syntax_node1(x_33, x_441, x_440); +x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_401); +lean_inc_n(x_292, 2); lean_inc(x_33); -x_445 = l_Lean_Syntax_node1(x_33, x_50, x_444); -x_446 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +x_444 = l_Lean_Syntax_node4(x_33, x_443, x_292, x_292, x_442, x_401); +x_445 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_54); +lean_inc(x_291); +x_446 = l_Lean_addMacroScope(x_291, x_445, x_54); +x_447 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; lean_inc(x_33); -x_447 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_447, 0, x_33); -lean_ctor_set(x_447, 1, x_446); -x_448 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +x_448 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_448, 0, x_33); +lean_ctor_set(x_448, 1, x_447); +lean_ctor_set(x_448, 2, x_446); +lean_ctor_set(x_448, 3, x_38); lean_inc(x_33); -x_449 = l_Lean_Syntax_node2(x_33, x_448, x_404, x_415); -x_450 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_290); +x_449 = l_Lean_Syntax_node1(x_33, x_50, x_448); +x_450 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; lean_inc(x_33); -x_451 = l_Lean_Syntax_node2(x_33, x_450, x_449, x_290); -x_452 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +x_451 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_451, 0, x_33); +lean_ctor_set(x_451, 1, x_450); +x_452 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; lean_inc(x_33); -x_453 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_453, 0, x_33); -lean_ctor_set(x_453, 1, x_452); -x_454 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; -lean_inc(x_54); -lean_inc(x_289); -x_455 = l_Lean_addMacroScope(x_289, x_454, x_54); -x_456 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +x_453 = l_Lean_Syntax_node2(x_33, x_452, x_408, x_419); +x_454 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_292); +lean_inc(x_33); +x_455 = l_Lean_Syntax_node2(x_33, x_454, x_453, x_292); +x_456 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; lean_inc(x_33); -x_457 = lean_alloc_ctor(3, 4, 0); +x_457 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_457, 0, x_33); lean_ctor_set(x_457, 1, x_456); -lean_ctor_set(x_457, 2, x_455); -lean_ctor_set(x_457, 3, x_38); -lean_inc(x_33); -x_458 = l_Lean_Syntax_node2(x_33, x_357, x_349, x_297); -lean_inc(x_33); -x_459 = l_Lean_Syntax_node1(x_33, x_50, x_458); -x_460 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_458 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +lean_inc(x_54); +lean_inc(x_291); +x_459 = l_Lean_addMacroScope(x_291, x_458, x_54); +x_460 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; lean_inc(x_33); -x_461 = lean_alloc_ctor(2, 2, 0); +x_461 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_461, 0, x_33); lean_ctor_set(x_461, 1, x_460); -x_462 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -lean_inc(x_54); -lean_inc(x_289); -x_463 = l_Lean_addMacroScope(x_289, x_462, x_54); -x_464 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_465 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_33); -x_466 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_466, 0, x_33); -lean_ctor_set(x_466, 1, x_464); -lean_ctor_set(x_466, 2, x_463); -lean_ctor_set(x_466, 3, x_465); -lean_inc(x_445); +lean_ctor_set(x_461, 2, x_459); +lean_ctor_set(x_461, 3, x_38); lean_inc(x_33); -x_467 = l_Lean_Syntax_node2(x_33, x_52, x_466, x_445); -x_468 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_462 = l_Lean_Syntax_node2(x_33, x_359, x_351, x_299); lean_inc(x_33); -x_469 = l_Lean_Syntax_node1(x_33, x_468, x_467); -x_470 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_457); +x_463 = l_Lean_Syntax_node1(x_33, x_50, x_462); +x_464 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_33); -x_471 = l_Lean_Syntax_node4(x_33, x_470, x_457, x_459, x_461, x_469); -x_472 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_290); +x_465 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_465, 0, x_33); +lean_ctor_set(x_465, 1, x_464); +x_466 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +lean_inc(x_54); +lean_inc(x_291); +x_467 = l_Lean_addMacroScope(x_291, x_466, x_54); +x_468 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_469 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_33); +x_470 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_470, 0, x_33); +lean_ctor_set(x_470, 1, x_468); +lean_ctor_set(x_470, 2, x_467); +lean_ctor_set(x_470, 3, x_469); +lean_inc(x_449); +lean_inc(x_33); +x_471 = l_Lean_Syntax_node2(x_33, x_52, x_470, x_449); +x_472 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_33); +x_473 = l_Lean_Syntax_node1(x_33, x_472, x_471); +x_474 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_461); +lean_inc(x_33); +x_475 = l_Lean_Syntax_node4(x_33, x_474, x_461, x_463, x_465, x_473); +x_476 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_292); lean_inc(x_33); -x_473 = l_Lean_Syntax_node3(x_33, x_472, x_453, x_290, x_471); -lean_inc(x_290); +x_477 = l_Lean_Syntax_node3(x_33, x_476, x_457, x_292, x_475); +lean_inc(x_292); lean_inc(x_33); -x_474 = l_Lean_Syntax_node2(x_33, x_450, x_473, x_290); -x_475 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; -x_476 = l_Lean_addMacroScope(x_289, x_475, x_54); -x_477 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; -x_478 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; -lean_inc(x_33); -x_479 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_479, 0, x_33); -lean_ctor_set(x_479, 1, x_477); -lean_ctor_set(x_479, 2, x_476); -lean_ctor_set(x_479, 3, x_478); -x_480 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; -lean_inc(x_33); -x_481 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_481, 0, x_33); -lean_ctor_set(x_481, 1, x_480); -lean_inc(x_290); +x_478 = l_Lean_Syntax_node2(x_33, x_454, x_477, x_292); +x_479 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; +x_480 = l_Lean_addMacroScope(x_291, x_479, x_54); +x_481 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; +x_482 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; +lean_inc(x_33); +x_483 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_483, 0, x_33); +lean_ctor_set(x_483, 1, x_481); +lean_ctor_set(x_483, 2, x_480); +lean_ctor_set(x_483, 3, x_482); +x_484 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; +lean_inc(x_33); +x_485 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_485, 0, x_33); +lean_ctor_set(x_485, 1, x_484); +lean_inc(x_292); lean_inc(x_33); -x_482 = l_Lean_Syntax_node2(x_33, x_418, x_290, x_457); +x_486 = l_Lean_Syntax_node2(x_33, x_422, x_292, x_461); lean_inc(x_33); -x_483 = l_Lean_Syntax_node1(x_33, x_50, x_482); -x_484 = lean_array_size(x_26); -x_485 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_38, x_484, x_15, x_26); -x_486 = lean_array_size(x_485); -x_487 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_486, x_15, x_485); -x_488 = l_Array_append___rarg(x_48, x_487); -lean_dec(x_487); +x_487 = l_Lean_Syntax_node1(x_33, x_50, x_486); +x_488 = lean_array_size(x_26); +x_489 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_38, x_488, x_15, x_26); +x_490 = lean_array_size(x_489); +x_491 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_490, x_15, x_489); +x_492 = l_Array_append___rarg(x_48, x_491); +lean_dec(x_491); lean_inc(x_33); -x_489 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_489, 0, x_33); -lean_ctor_set(x_489, 1, x_50); -lean_ctor_set(x_489, 2, x_488); +x_493 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_493, 0, x_33); +lean_ctor_set(x_493, 1, x_50); +lean_ctor_set(x_493, 2, x_492); lean_inc(x_33); -x_490 = l_Lean_Syntax_node1(x_33, x_429, x_489); -lean_inc_n(x_290, 2); +x_494 = l_Lean_Syntax_node1(x_33, x_433, x_493); +lean_inc_n(x_292, 2); lean_inc(x_33); -x_491 = l_Lean_Syntax_node6(x_33, x_431, x_417, x_290, x_290, x_483, x_422, x_490); -x_492 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +x_495 = l_Lean_Syntax_node6(x_33, x_435, x_421, x_292, x_292, x_487, x_426, x_494); +x_496 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; lean_inc(x_33); -x_493 = l_Lean_Syntax_node3(x_33, x_492, x_479, x_481, x_491); +x_497 = l_Lean_Syntax_node3(x_33, x_496, x_483, x_485, x_495); lean_inc(x_33); -x_494 = l_Lean_Syntax_node1(x_33, x_468, x_493); -lean_inc(x_290); +x_498 = l_Lean_Syntax_node1(x_33, x_472, x_497); +lean_inc(x_292); lean_inc(x_33); -x_495 = l_Lean_Syntax_node2(x_33, x_450, x_494, x_290); +x_499 = l_Lean_Syntax_node2(x_33, x_454, x_498, x_292); lean_inc(x_33); -x_496 = l_Lean_Syntax_node3(x_33, x_50, x_451, x_474, x_495); -x_497 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_500 = l_Lean_Syntax_node3(x_33, x_50, x_455, x_478, x_499); +x_501 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_33); -x_498 = l_Lean_Syntax_node1(x_33, x_497, x_496); -x_499 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_502 = l_Lean_Syntax_node1(x_33, x_501, x_500); +x_503 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_33); -x_500 = l_Lean_Syntax_node2(x_33, x_499, x_447, x_498); -lean_inc(x_362); -lean_inc(x_290); +x_504 = l_Lean_Syntax_node2(x_33, x_503, x_451, x_502); +lean_inc(x_364); +lean_inc(x_292); lean_inc(x_33); -x_501 = l_Lean_Syntax_node5(x_33, x_435, x_387, x_445, x_290, x_362, x_500); +x_505 = l_Lean_Syntax_node5(x_33, x_439, x_389, x_449, x_292, x_364, x_504); lean_inc(x_33); -x_502 = l_Lean_Syntax_node1(x_33, x_437, x_501); -lean_inc(x_397); -lean_inc_n(x_290, 2); +x_506 = l_Lean_Syntax_node1(x_33, x_441, x_505); +lean_inc(x_401); +lean_inc_n(x_292, 2); lean_inc(x_33); -x_503 = l_Lean_Syntax_node4(x_33, x_439, x_290, x_290, x_502, x_397); -lean_inc(x_290); +x_507 = l_Lean_Syntax_node4(x_33, x_443, x_292, x_292, x_506, x_401); +lean_inc(x_292); lean_inc(x_33); -x_504 = l_Lean_Syntax_node3(x_33, x_50, x_440, x_290, x_503); -x_505 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_508 = l_Lean_Syntax_node3(x_33, x_50, x_444, x_292, x_507); +x_509 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_33); -x_506 = l_Lean_Syntax_node2(x_33, x_505, x_19, x_504); +x_510 = l_Lean_Syntax_node2(x_33, x_509, x_19, x_508); lean_inc(x_33); -x_507 = l_Lean_Syntax_node1(x_33, x_50, x_506); -x_508 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +x_511 = l_Lean_Syntax_node1(x_33, x_50, x_510); +x_512 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; lean_inc(x_33); -x_509 = l_Lean_Syntax_node4(x_33, x_508, x_362, x_395, x_397, x_507); -x_510 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_290); +x_513 = l_Lean_Syntax_node4(x_33, x_512, x_364, x_399, x_401, x_511); +x_514 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_292); lean_inc(x_33); -x_511 = l_Lean_Syntax_node6(x_33, x_510, x_345, x_347, x_290, x_290, x_360, x_509); +x_515 = l_Lean_Syntax_node6(x_33, x_514, x_347, x_349, x_292, x_292, x_362, x_513); lean_inc(x_33); -x_512 = l_Lean_Syntax_node2(x_33, x_328, x_343, x_511); -x_513 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +x_516 = l_Lean_Syntax_node2(x_33, x_330, x_345, x_515); +x_517 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; lean_inc(x_33); -x_514 = l_Lean_Syntax_node3(x_33, x_513, x_335, x_337, x_512); -x_515 = l_Lean_Syntax_node2(x_33, x_50, x_329, x_514); -x_516 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_516, 0, x_515); -lean_ctor_set(x_516, 1, x_287); -return x_516; +x_518 = l_Lean_Syntax_node3(x_33, x_517, x_337, x_339, x_516); +x_519 = l_Lean_Syntax_node2(x_33, x_50, x_331, x_518); +x_520 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_520, 0, x_519); +lean_ctor_set(x_520, 1, x_289); +return x_520; } } else { -lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; size_t x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; size_t x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; size_t x_675; lean_object* x_676; size_t x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; size_t x_736; lean_object* x_737; size_t x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; -x_517 = lean_ctor_get(x_34, 1); -lean_inc(x_517); -lean_dec(x_34); -x_518 = lean_box(0); -x_519 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_33); -x_520 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_520, 0, x_33); -lean_ctor_set(x_520, 1, x_519); -x_521 = lean_ctor_get(x_1, 0); +lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; size_t x_530; lean_object* x_531; lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; size_t x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; size_t x_681; lean_object* x_682; size_t x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; size_t x_742; lean_object* x_743; size_t x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; +x_521 = lean_ctor_get(x_34, 1); lean_inc(x_521); +lean_dec(x_34); +x_522 = lean_box(0); +x_523 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_33); +x_524 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_524, 0, x_33); +lean_ctor_set(x_524, 1, x_523); +x_525 = lean_ctor_get(x_1, 0); +lean_inc(x_525); lean_dec(x_1); -x_522 = lean_ctor_get(x_521, 0); -lean_inc(x_522); -lean_dec(x_521); -x_523 = lean_mk_syntax_ident(x_522); -x_524 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_33); -x_525 = l_Lean_Syntax_node2(x_33, x_524, x_520, x_523); -x_526 = lean_array_size(x_29); -x_527 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; -x_528 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_527, x_526, x_15, x_29); -x_529 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -x_530 = l_Array_append___rarg(x_529, x_528); -lean_dec(x_528); -x_531 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -lean_inc(x_33); -x_532 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_532, 0, x_33); -lean_ctor_set(x_532, 1, x_531); -lean_ctor_set(x_532, 2, x_530); -x_533 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_33); -x_534 = l_Lean_Syntax_node2(x_33, x_533, x_525, x_532); -x_535 = lean_ctor_get(x_9, 10); -lean_inc(x_535); +x_526 = lean_ctor_get(x_525, 0); +lean_inc(x_526); +lean_dec(x_525); +x_527 = lean_mk_syntax_ident(x_526); +x_528 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_33); +x_529 = l_Lean_Syntax_node2(x_33, x_528, x_524, x_527); +x_530 = lean_array_size(x_29); +x_531 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; +x_532 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_531, x_530, x_15, x_29); +x_533 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +x_534 = l_Array_append___rarg(x_533, x_532); +lean_dec(x_532); +x_535 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +lean_inc(x_33); +x_536 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_536, 0, x_33); +lean_ctor_set(x_536, 1, x_535); +lean_ctor_set(x_536, 2, x_534); +x_537 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_33); +x_538 = l_Lean_Syntax_node2(x_33, x_537, x_529, x_536); +x_539 = lean_ctor_get(x_9, 10); +lean_inc(x_539); lean_dec(x_9); -x_536 = lean_st_ref_get(x_10, x_517); +x_540 = lean_st_ref_get(x_10, x_521); lean_dec(x_10); -x_537 = lean_ctor_get(x_536, 0); -lean_inc(x_537); -x_538 = lean_ctor_get(x_536, 1); -lean_inc(x_538); -if (lean_is_exclusive(x_536)) { - lean_ctor_release(x_536, 0); - lean_ctor_release(x_536, 1); - x_539 = x_536; +x_541 = lean_ctor_get(x_540, 0); +lean_inc(x_541); +x_542 = lean_ctor_get(x_540, 1); +lean_inc(x_542); +if (lean_is_exclusive(x_540)) { + lean_ctor_release(x_540, 0); + lean_ctor_release(x_540, 1); + x_543 = x_540; } else { - lean_dec_ref(x_536); - x_539 = lean_box(0); + lean_dec_ref(x_540); + x_543 = lean_box(0); } -x_540 = lean_ctor_get(x_537, 0); -lean_inc(x_540); -lean_dec(x_537); -x_541 = lean_environment_main_module(x_540); +x_544 = lean_ctor_get(x_541, 0); +lean_inc(x_544); +lean_dec(x_541); +x_545 = lean_environment_main_module(x_544); lean_inc(x_33); -x_542 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_542, 0, x_33); -lean_ctor_set(x_542, 1, x_531); -lean_ctor_set(x_542, 2, x_529); -x_543 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_542, 6); +x_546 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_546, 0, x_33); +lean_ctor_set(x_546, 1, x_535); +lean_ctor_set(x_546, 2, x_533); +x_547 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_546, 6); lean_inc(x_33); -x_544 = l_Lean_Syntax_node6(x_33, x_543, x_542, x_542, x_542, x_542, x_542, x_542); -x_545 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; +x_548 = l_Lean_Syntax_node6(x_33, x_547, x_546, x_546, x_546, x_546, x_546, x_546); +x_549 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; lean_inc(x_33); lean_ctor_set_tag(x_23, 2); -lean_ctor_set(x_23, 1, x_545); +lean_ctor_set(x_23, 1, x_549); lean_ctor_set(x_23, 0, x_33); -x_546 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_535); -lean_inc(x_541); -x_547 = l_Lean_addMacroScope(x_541, x_546, x_535); -x_548 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_33); -x_549 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_549, 0, x_33); -lean_ctor_set(x_549, 1, x_548); -lean_ctor_set(x_549, 2, x_547); -lean_ctor_set(x_549, 3, x_518); -x_550 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_542); -lean_inc(x_549); +x_550 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_539); +lean_inc(x_545); +x_551 = l_Lean_addMacroScope(x_545, x_550, x_539); +x_552 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_33); +x_553 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_553, 0, x_33); +lean_ctor_set(x_553, 1, x_552); +lean_ctor_set(x_553, 2, x_551); +lean_ctor_set(x_553, 3, x_522); +x_554 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_546); +lean_inc(x_553); lean_inc(x_33); -x_551 = l_Lean_Syntax_node2(x_33, x_550, x_549, x_542); -x_552 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; -lean_inc_n(x_542, 2); +x_555 = l_Lean_Syntax_node2(x_33, x_554, x_553, x_546); +x_556 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; +lean_inc_n(x_546, 2); lean_inc(x_33); -x_553 = l_Lean_Syntax_node2(x_33, x_552, x_542, x_542); -x_554 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +x_557 = l_Lean_Syntax_node2(x_33, x_556, x_546, x_546); +x_558 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; lean_inc(x_33); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_554); +lean_ctor_set(x_19, 1, x_558); lean_ctor_set(x_19, 0, x_33); lean_inc(x_19); lean_inc(x_33); -x_555 = l_Lean_Syntax_node1(x_33, x_531, x_19); -x_556 = lean_array_size(x_21); -x_557 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_556, x_15, x_21); -x_558 = l_Array_append___rarg(x_529, x_557); -lean_dec(x_557); -lean_inc(x_33); -x_559 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_559, 0, x_33); -lean_ctor_set(x_559, 1, x_531); -lean_ctor_set(x_559, 2, x_558); -x_560 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_33); -x_561 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_561, 0, x_33); -lean_ctor_set(x_561, 1, x_560); -x_562 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_535); -lean_inc(x_541); -x_563 = l_Lean_addMacroScope(x_541, x_562, x_535); -x_564 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_565 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_33); -x_566 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_566, 0, x_33); -lean_ctor_set(x_566, 1, x_564); -lean_ctor_set(x_566, 2, x_563); -lean_ctor_set(x_566, 3, x_565); -x_567 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_33); -x_568 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_568, 0, x_33); -lean_ctor_set(x_568, 1, x_567); -x_569 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_535); -lean_inc(x_541); -x_570 = l_Lean_addMacroScope(x_541, x_569, x_535); -x_571 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_572 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_33); -x_573 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_573, 0, x_33); -lean_ctor_set(x_573, 1, x_571); -lean_ctor_set(x_573, 2, x_570); -lean_ctor_set(x_573, 3, x_572); -lean_inc(x_568); +x_559 = l_Lean_Syntax_node1(x_33, x_535, x_19); +x_560 = lean_array_size(x_21); +x_561 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_560, x_15, x_21); +x_562 = l_Array_append___rarg(x_533, x_561); +lean_dec(x_561); +lean_inc(x_33); +x_563 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_563, 0, x_33); +lean_ctor_set(x_563, 1, x_535); +lean_ctor_set(x_563, 2, x_562); +x_564 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_33); +x_565 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_565, 0, x_33); +lean_ctor_set(x_565, 1, x_564); +x_566 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_539); +lean_inc(x_545); +x_567 = l_Lean_addMacroScope(x_545, x_566, x_539); +x_568 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_569 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_33); +x_570 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_570, 0, x_33); +lean_ctor_set(x_570, 1, x_568); +lean_ctor_set(x_570, 2, x_567); +lean_ctor_set(x_570, 3, x_569); +x_571 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_33); +x_572 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_572, 0, x_33); +lean_ctor_set(x_572, 1, x_571); +x_573 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_539); +lean_inc(x_545); +x_574 = l_Lean_addMacroScope(x_545, x_573, x_539); +x_575 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_576 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_33); +x_577 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_577, 0, x_33); +lean_ctor_set(x_577, 1, x_575); +lean_ctor_set(x_577, 2, x_574); +lean_ctor_set(x_577, 3, x_576); +lean_inc(x_572); +lean_inc(x_33); +x_578 = l_Lean_Syntax_node3(x_33, x_535, x_570, x_572, x_577); +lean_inc(x_33); +x_579 = l_Lean_Syntax_node2(x_33, x_535, x_565, x_578); +x_580 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_33); +x_581 = l_Lean_Syntax_node1(x_33, x_580, x_579); +x_582 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; +lean_inc(x_546); +lean_inc(x_33); +x_583 = l_Lean_Syntax_node7(x_33, x_582, x_23, x_555, x_557, x_559, x_563, x_546, x_581); +x_584 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_33); +x_585 = l_Lean_Syntax_node2(x_33, x_584, x_548, x_583); +x_586 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; lean_inc(x_33); -x_574 = l_Lean_Syntax_node3(x_33, x_531, x_566, x_568, x_573); -lean_inc(x_33); -x_575 = l_Lean_Syntax_node2(x_33, x_531, x_561, x_574); -x_576 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_33); -x_577 = l_Lean_Syntax_node1(x_33, x_576, x_575); -x_578 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; -lean_inc(x_542); -lean_inc(x_33); -x_579 = l_Lean_Syntax_node7(x_33, x_578, x_23, x_551, x_553, x_555, x_559, x_542, x_577); -x_580 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_33); -x_581 = l_Lean_Syntax_node2(x_33, x_580, x_544, x_579); -x_582 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_33); -x_583 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_583, 0, x_33); -lean_ctor_set(x_583, 1, x_582); -x_584 = l_Array_append___rarg(x_529, x_3); -lean_inc(x_33); -x_585 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_585, 0, x_33); -lean_ctor_set(x_585, 1, x_531); -lean_ctor_set(x_585, 2, x_584); -x_586 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_33); -x_587 = l_Lean_Syntax_node2(x_33, x_586, x_583, x_585); -x_588 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +x_587 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_587, 0, x_33); +lean_ctor_set(x_587, 1, x_586); +x_588 = l_Array_append___rarg(x_533, x_3); lean_inc(x_33); -x_589 = lean_alloc_ctor(2, 2, 0); +x_589 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_589, 0, x_33); -lean_ctor_set(x_589, 1, x_588); -x_590 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; -lean_inc(x_33); -x_591 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_591, 0, x_33); -lean_ctor_set(x_591, 1, x_590); -x_592 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; -lean_inc(x_33); -x_593 = l_Lean_Syntax_node1(x_33, x_592, x_591); -lean_inc(x_33); -x_594 = l_Lean_Syntax_node1(x_33, x_531, x_593); -lean_inc_n(x_542, 5); +lean_ctor_set(x_589, 1, x_535); +lean_ctor_set(x_589, 2, x_588); +x_590 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_33); +x_591 = l_Lean_Syntax_node2(x_33, x_590, x_587, x_589); +x_592 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_33); +x_593 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_593, 0, x_33); +lean_ctor_set(x_593, 1, x_592); +x_594 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; +lean_inc(x_33); +x_595 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_595, 0, x_33); +lean_ctor_set(x_595, 1, x_594); +x_596 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; +lean_inc(x_33); +x_597 = l_Lean_Syntax_node1(x_33, x_596, x_595); +lean_inc(x_33); +x_598 = l_Lean_Syntax_node1(x_33, x_535, x_597); +lean_inc_n(x_546, 5); +lean_inc(x_33); +x_599 = l_Lean_Syntax_node6(x_33, x_547, x_546, x_546, x_546, x_546, x_546, x_598); +x_600 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_546); +lean_inc(x_33); +x_601 = l_Lean_Syntax_node1(x_33, x_600, x_546); +x_602 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_33); +x_603 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_603, 0, x_33); +lean_ctor_set(x_603, 1, x_602); +x_604 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_33); +x_605 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_605, 0, x_33); +lean_ctor_set(x_605, 1, x_604); +x_606 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_539); +lean_inc(x_545); +x_607 = l_Lean_addMacroScope(x_545, x_606, x_539); +x_608 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_609 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; lean_inc(x_33); -x_595 = l_Lean_Syntax_node6(x_33, x_543, x_542, x_542, x_542, x_542, x_542, x_594); -x_596 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_542); -lean_inc(x_33); -x_597 = l_Lean_Syntax_node1(x_33, x_596, x_542); -x_598 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_33); -x_599 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_599, 0, x_33); -lean_ctor_set(x_599, 1, x_598); -x_600 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_33); -x_601 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_601, 0, x_33); -lean_ctor_set(x_601, 1, x_600); -x_602 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_535); -lean_inc(x_541); -x_603 = l_Lean_addMacroScope(x_541, x_602, x_535); -x_604 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_605 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_33); -x_606 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_606, 0, x_33); -lean_ctor_set(x_606, 1, x_604); -lean_ctor_set(x_606, 2, x_603); -lean_ctor_set(x_606, 3, x_605); -lean_inc(x_33); -x_607 = l_Lean_Syntax_node1(x_33, x_531, x_534); -lean_inc(x_33); -x_608 = l_Lean_Syntax_node2(x_33, x_533, x_606, x_607); -x_609 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_601); -lean_inc(x_33); -x_610 = l_Lean_Syntax_node2(x_33, x_609, x_601, x_608); -x_611 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_610); -lean_inc(x_542); -lean_inc(x_33); -x_612 = l_Lean_Syntax_node2(x_33, x_611, x_542, x_610); -x_613 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_33); -x_614 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_614, 0, x_33); -lean_ctor_set(x_614, 1, x_613); -x_615 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_33); -x_616 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_616, 0, x_33); -lean_ctor_set(x_616, 1, x_615); -x_617 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_535); -lean_inc(x_541); -x_618 = l_Lean_addMacroScope(x_541, x_617, x_535); -x_619 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_620 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_33); -x_621 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_621, 0, x_33); -lean_ctor_set(x_621, 1, x_619); -lean_ctor_set(x_621, 2, x_618); -lean_ctor_set(x_621, 3, x_620); -x_622 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_542); -lean_inc(x_33); -x_623 = l_Lean_Syntax_node2(x_33, x_622, x_621, x_542); -x_624 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_535); -lean_inc(x_541); -x_625 = l_Lean_addMacroScope(x_541, x_624, x_535); -x_626 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_33); -x_627 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_627, 0, x_33); -lean_ctor_set(x_627, 1, x_626); -lean_ctor_set(x_627, 2, x_625); -lean_ctor_set(x_627, 3, x_518); -x_628 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_627); -lean_inc(x_614); +x_610 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_610, 0, x_33); +lean_ctor_set(x_610, 1, x_608); +lean_ctor_set(x_610, 2, x_607); +lean_ctor_set(x_610, 3, x_609); lean_inc(x_33); -x_629 = l_Lean_Syntax_node3(x_33, x_628, x_623, x_614, x_627); -x_630 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_535); -lean_inc(x_541); -x_631 = l_Lean_addMacroScope(x_541, x_630, x_535); -x_632 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_633 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +x_611 = l_Lean_Syntax_node1(x_33, x_535, x_538); lean_inc(x_33); -x_634 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_634, 0, x_33); -lean_ctor_set(x_634, 1, x_632); -lean_ctor_set(x_634, 2, x_631); -lean_ctor_set(x_634, 3, x_633); -lean_inc(x_542); +x_612 = l_Lean_Syntax_node2(x_33, x_537, x_610, x_611); +x_613 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_605); lean_inc(x_33); -x_635 = l_Lean_Syntax_node2(x_33, x_622, x_634, x_542); -x_636 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_535); -lean_inc(x_541); -x_637 = l_Lean_addMacroScope(x_541, x_636, x_535); -x_638 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_33); -x_639 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_639, 0, x_33); -lean_ctor_set(x_639, 1, x_638); -lean_ctor_set(x_639, 2, x_637); -lean_ctor_set(x_639, 3, x_518); -lean_inc(x_639); +x_614 = l_Lean_Syntax_node2(x_33, x_613, x_605, x_612); +x_615 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; lean_inc(x_614); +lean_inc(x_546); +lean_inc(x_33); +x_616 = l_Lean_Syntax_node2(x_33, x_615, x_546, x_614); +x_617 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_33); +x_618 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_618, 0, x_33); +lean_ctor_set(x_618, 1, x_617); +x_619 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_33); +x_620 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_620, 0, x_33); +lean_ctor_set(x_620, 1, x_619); +x_621 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_539); +lean_inc(x_545); +x_622 = l_Lean_addMacroScope(x_545, x_621, x_539); +x_623 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_624 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_33); +x_625 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_625, 0, x_33); +lean_ctor_set(x_625, 1, x_623); +lean_ctor_set(x_625, 2, x_622); +lean_ctor_set(x_625, 3, x_624); +x_626 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_546); +lean_inc(x_33); +x_627 = l_Lean_Syntax_node2(x_33, x_626, x_625, x_546); +x_628 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_539); +lean_inc(x_545); +x_629 = l_Lean_addMacroScope(x_545, x_628, x_539); +x_630 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_33); +x_631 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_631, 0, x_33); +lean_ctor_set(x_631, 1, x_630); +lean_ctor_set(x_631, 2, x_629); +lean_ctor_set(x_631, 3, x_522); +x_632 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_631); +lean_inc(x_618); +lean_inc(x_33); +x_633 = l_Lean_Syntax_node3(x_33, x_632, x_627, x_618, x_631); +x_634 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_539); +lean_inc(x_545); +x_635 = l_Lean_addMacroScope(x_545, x_634, x_539); +x_636 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_637 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_33); +x_638 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_638, 0, x_33); +lean_ctor_set(x_638, 1, x_636); +lean_ctor_set(x_638, 2, x_635); +lean_ctor_set(x_638, 3, x_637); +lean_inc(x_546); +lean_inc(x_33); +x_639 = l_Lean_Syntax_node2(x_33, x_626, x_638, x_546); +x_640 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_539); +lean_inc(x_545); +x_641 = l_Lean_addMacroScope(x_545, x_640, x_539); +x_642 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_33); +x_643 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_643, 0, x_33); +lean_ctor_set(x_643, 1, x_642); +lean_ctor_set(x_643, 2, x_641); +lean_ctor_set(x_643, 3, x_522); +lean_inc(x_643); +lean_inc(x_618); +lean_inc(x_33); +x_644 = l_Lean_Syntax_node3(x_33, x_632, x_639, x_618, x_643); +lean_inc(x_33); +x_645 = l_Lean_Syntax_node3(x_33, x_535, x_633, x_572, x_644); +x_646 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_33); +x_647 = l_Lean_Syntax_node1(x_33, x_646, x_645); +x_648 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_546); +lean_inc(x_33); +x_649 = l_Lean_Syntax_node1(x_33, x_648, x_546); +x_650 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_33); +x_651 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_651, 0, x_33); +lean_ctor_set(x_651, 1, x_650); +x_652 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc_n(x_546, 2); +lean_inc(x_33); +x_653 = l_Lean_Syntax_node6(x_33, x_652, x_620, x_546, x_647, x_649, x_546, x_651); +x_654 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_546, 2); +lean_inc(x_33); +x_655 = l_Lean_Syntax_node2(x_33, x_654, x_546, x_546); +x_656 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; +lean_inc(x_539); +lean_inc(x_545); +x_657 = l_Lean_addMacroScope(x_545, x_656, x_539); +x_658 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_33); +x_659 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_659, 0, x_33); +lean_ctor_set(x_659, 1, x_658); +lean_ctor_set(x_659, 2, x_657); +lean_ctor_set(x_659, 3, x_522); +lean_inc(x_659); +lean_inc(x_33); +x_660 = l_Lean_Syntax_node1(x_33, x_535, x_659); +x_661 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; lean_inc(x_33); -x_640 = l_Lean_Syntax_node3(x_33, x_628, x_635, x_614, x_639); -lean_inc(x_33); -x_641 = l_Lean_Syntax_node3(x_33, x_531, x_629, x_568, x_640); -x_642 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_542); -lean_inc(x_33); -x_643 = l_Lean_Syntax_node1(x_33, x_642, x_542); -x_644 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_33); -x_645 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_645, 0, x_33); -lean_ctor_set(x_645, 1, x_644); -x_646 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc_n(x_542, 2); +x_662 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_662, 0, x_33); +lean_ctor_set(x_662, 1, x_661); +x_663 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +lean_inc(x_539); +lean_inc(x_545); +x_664 = l_Lean_addMacroScope(x_545, x_663, x_539); +x_665 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +lean_inc(x_33); +x_666 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_666, 0, x_33); +lean_ctor_set(x_666, 1, x_665); +lean_ctor_set(x_666, 2, x_664); +lean_ctor_set(x_666, 3, x_522); +x_667 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; lean_inc(x_33); -x_647 = l_Lean_Syntax_node6(x_33, x_646, x_616, x_542, x_641, x_643, x_542, x_645); -x_648 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_542, 2); +x_668 = l_Lean_Syntax_node1(x_33, x_667, x_666); lean_inc(x_33); -x_649 = l_Lean_Syntax_node2(x_33, x_648, x_542, x_542); -x_650 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; -lean_inc(x_535); -lean_inc(x_541); -x_651 = l_Lean_addMacroScope(x_541, x_650, x_535); -x_652 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; -lean_inc(x_33); -x_653 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_653, 0, x_33); -lean_ctor_set(x_653, 1, x_652); -lean_ctor_set(x_653, 2, x_651); -lean_ctor_set(x_653, 3, x_518); +x_669 = l_Lean_Syntax_node1(x_33, x_535, x_614); +x_670 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; lean_inc(x_653); +lean_inc(x_618); +lean_inc(x_546); lean_inc(x_33); -x_654 = l_Lean_Syntax_node1(x_33, x_531, x_653); -x_655 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; -lean_inc(x_33); -x_656 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_656, 0, x_33); -lean_ctor_set(x_656, 1, x_655); -x_657 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; -lean_inc(x_535); -lean_inc(x_541); -x_658 = l_Lean_addMacroScope(x_541, x_657, x_535); -x_659 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; -lean_inc(x_33); -x_660 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_660, 0, x_33); -lean_ctor_set(x_660, 1, x_659); -lean_ctor_set(x_660, 2, x_658); -lean_ctor_set(x_660, 3, x_518); -x_661 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; -lean_inc(x_33); -x_662 = l_Lean_Syntax_node1(x_33, x_661, x_660); -lean_inc(x_33); -x_663 = l_Lean_Syntax_node1(x_33, x_531, x_610); -x_664 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; -lean_inc(x_647); -lean_inc(x_614); -lean_inc(x_542); +x_671 = l_Lean_Syntax_node5(x_33, x_670, x_668, x_546, x_669, x_618, x_653); +x_672 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_inc(x_33); -x_665 = l_Lean_Syntax_node5(x_33, x_664, x_662, x_542, x_663, x_614, x_647); -x_666 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +x_673 = l_Lean_Syntax_node1(x_33, x_672, x_671); +x_674 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; lean_inc(x_33); -x_667 = l_Lean_Syntax_node1(x_33, x_666, x_665); -x_668 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; -lean_inc(x_33); -x_669 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_669, 0, x_33); -lean_ctor_set(x_669, 1, x_668); -x_670 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; -lean_inc(x_542); +x_675 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_675, 0, x_33); +lean_ctor_set(x_675, 1, x_674); +x_676 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; +lean_inc(x_546); lean_inc(x_33); -x_671 = l_Lean_Syntax_node2(x_33, x_670, x_542, x_653); +x_677 = l_Lean_Syntax_node2(x_33, x_676, x_546, x_659); lean_inc(x_33); -x_672 = l_Lean_Syntax_node1(x_33, x_531, x_671); -x_673 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +x_678 = l_Lean_Syntax_node1(x_33, x_535, x_677); +x_679 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; lean_inc(x_33); -x_674 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_674, 0, x_33); -lean_ctor_set(x_674, 1, x_673); -x_675 = lean_array_size(x_25); -x_676 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_518, x_675, x_15, x_25); -x_677 = lean_array_size(x_676); -x_678 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_677, x_15, x_676); -x_679 = l_Array_append___rarg(x_529, x_678); -lean_dec(x_678); -lean_inc(x_33); -x_680 = lean_alloc_ctor(1, 3, 0); +x_680 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_680, 0, x_33); -lean_ctor_set(x_680, 1, x_531); -lean_ctor_set(x_680, 2, x_679); -x_681 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; -lean_inc(x_33); -x_682 = l_Lean_Syntax_node1(x_33, x_681, x_680); -x_683 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_674); -lean_inc_n(x_542, 2); -lean_inc(x_669); -lean_inc(x_33); -x_684 = l_Lean_Syntax_node6(x_33, x_683, x_669, x_542, x_542, x_672, x_674, x_682); -x_685 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; -lean_inc(x_542); -lean_inc(x_667); -lean_inc(x_656); -lean_inc(x_33); -x_686 = l_Lean_Syntax_node4(x_33, x_685, x_656, x_667, x_542, x_684); -x_687 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_614); -lean_inc(x_542); -lean_inc(x_33); -x_688 = l_Lean_Syntax_node5(x_33, x_687, x_627, x_654, x_542, x_614, x_686); -x_689 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_33); -x_690 = l_Lean_Syntax_node1(x_33, x_689, x_688); -x_691 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_649); -lean_inc_n(x_542, 2); -lean_inc(x_33); -x_692 = l_Lean_Syntax_node4(x_33, x_691, x_542, x_542, x_690, x_649); -x_693 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_535); -lean_inc(x_541); -x_694 = l_Lean_addMacroScope(x_541, x_693, x_535); -x_695 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_33); -x_696 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_696, 0, x_33); -lean_ctor_set(x_696, 1, x_695); -lean_ctor_set(x_696, 2, x_694); -lean_ctor_set(x_696, 3, x_518); -lean_inc(x_33); -x_697 = l_Lean_Syntax_node1(x_33, x_531, x_696); -x_698 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_33); -x_699 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_699, 0, x_33); -lean_ctor_set(x_699, 1, x_698); -x_700 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; -lean_inc(x_33); -x_701 = l_Lean_Syntax_node2(x_33, x_700, x_656, x_667); -x_702 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_542); -lean_inc(x_33); -x_703 = l_Lean_Syntax_node2(x_33, x_702, x_701, x_542); -x_704 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +lean_ctor_set(x_680, 1, x_679); +x_681 = lean_array_size(x_25); +x_682 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_522, x_681, x_15, x_25); +x_683 = lean_array_size(x_682); +x_684 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_683, x_15, x_682); +x_685 = l_Array_append___rarg(x_533, x_684); +lean_dec(x_684); +lean_inc(x_33); +x_686 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_686, 0, x_33); +lean_ctor_set(x_686, 1, x_535); +lean_ctor_set(x_686, 2, x_685); +x_687 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +lean_inc(x_33); +x_688 = l_Lean_Syntax_node1(x_33, x_687, x_686); +x_689 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_680); +lean_inc_n(x_546, 2); +lean_inc(x_675); +lean_inc(x_33); +x_690 = l_Lean_Syntax_node6(x_33, x_689, x_675, x_546, x_546, x_678, x_680, x_688); +x_691 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +lean_inc(x_546); +lean_inc(x_673); +lean_inc(x_662); +lean_inc(x_33); +x_692 = l_Lean_Syntax_node4(x_33, x_691, x_662, x_673, x_546, x_690); +x_693 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_618); +lean_inc(x_546); +lean_inc(x_33); +x_694 = l_Lean_Syntax_node5(x_33, x_693, x_631, x_660, x_546, x_618, x_692); +x_695 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_33); +x_696 = l_Lean_Syntax_node1(x_33, x_695, x_694); +x_697 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_655); +lean_inc_n(x_546, 2); +lean_inc(x_33); +x_698 = l_Lean_Syntax_node4(x_33, x_697, x_546, x_546, x_696, x_655); +x_699 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_539); +lean_inc(x_545); +x_700 = l_Lean_addMacroScope(x_545, x_699, x_539); +x_701 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_33); +x_702 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_702, 0, x_33); +lean_ctor_set(x_702, 1, x_701); +lean_ctor_set(x_702, 2, x_700); +lean_ctor_set(x_702, 3, x_522); +lean_inc(x_33); +x_703 = l_Lean_Syntax_node1(x_33, x_535, x_702); +x_704 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; lean_inc(x_33); x_705 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_705, 0, x_33); lean_ctor_set(x_705, 1, x_704); -x_706 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; -lean_inc(x_535); -lean_inc(x_541); -x_707 = l_Lean_addMacroScope(x_541, x_706, x_535); -x_708 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +x_706 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +lean_inc(x_33); +x_707 = l_Lean_Syntax_node2(x_33, x_706, x_662, x_673); +x_708 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_546); +lean_inc(x_33); +x_709 = l_Lean_Syntax_node2(x_33, x_708, x_707, x_546); +x_710 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_33); +x_711 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_711, 0, x_33); +lean_ctor_set(x_711, 1, x_710); +x_712 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +lean_inc(x_539); +lean_inc(x_545); +x_713 = l_Lean_addMacroScope(x_545, x_712, x_539); +x_714 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +lean_inc(x_33); +x_715 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_715, 0, x_33); +lean_ctor_set(x_715, 1, x_714); +lean_ctor_set(x_715, 2, x_713); +lean_ctor_set(x_715, 3, x_522); +lean_inc(x_33); +x_716 = l_Lean_Syntax_node2(x_33, x_613, x_605, x_553); +lean_inc(x_33); +x_717 = l_Lean_Syntax_node1(x_33, x_535, x_716); +x_718 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_33); +x_719 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_719, 0, x_33); +lean_ctor_set(x_719, 1, x_718); +x_720 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +lean_inc(x_539); +lean_inc(x_545); +x_721 = l_Lean_addMacroScope(x_545, x_720, x_539); +x_722 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_723 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_33); +x_724 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_724, 0, x_33); +lean_ctor_set(x_724, 1, x_722); +lean_ctor_set(x_724, 2, x_721); +lean_ctor_set(x_724, 3, x_723); +lean_inc(x_703); +lean_inc(x_33); +x_725 = l_Lean_Syntax_node2(x_33, x_537, x_724, x_703); +x_726 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_33); +x_727 = l_Lean_Syntax_node1(x_33, x_726, x_725); +x_728 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_715); lean_inc(x_33); -x_709 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_709, 0, x_33); -lean_ctor_set(x_709, 1, x_708); -lean_ctor_set(x_709, 2, x_707); -lean_ctor_set(x_709, 3, x_518); +x_729 = l_Lean_Syntax_node4(x_33, x_728, x_715, x_717, x_719, x_727); +x_730 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_546); lean_inc(x_33); -x_710 = l_Lean_Syntax_node2(x_33, x_609, x_601, x_549); +x_731 = l_Lean_Syntax_node3(x_33, x_730, x_711, x_546, x_729); +lean_inc(x_546); lean_inc(x_33); -x_711 = l_Lean_Syntax_node1(x_33, x_531, x_710); -x_712 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_732 = l_Lean_Syntax_node2(x_33, x_708, x_731, x_546); +x_733 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; +x_734 = l_Lean_addMacroScope(x_545, x_733, x_539); +x_735 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; +x_736 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; lean_inc(x_33); -x_713 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_713, 0, x_33); -lean_ctor_set(x_713, 1, x_712); -x_714 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -lean_inc(x_535); -lean_inc(x_541); -x_715 = l_Lean_addMacroScope(x_541, x_714, x_535); -x_716 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_717 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_33); -x_718 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_718, 0, x_33); -lean_ctor_set(x_718, 1, x_716); -lean_ctor_set(x_718, 2, x_715); -lean_ctor_set(x_718, 3, x_717); -lean_inc(x_697); -lean_inc(x_33); -x_719 = l_Lean_Syntax_node2(x_33, x_533, x_718, x_697); -x_720 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_33); -x_721 = l_Lean_Syntax_node1(x_33, x_720, x_719); -x_722 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_709); -lean_inc(x_33); -x_723 = l_Lean_Syntax_node4(x_33, x_722, x_709, x_711, x_713, x_721); -x_724 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_542); +x_737 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_737, 0, x_33); +lean_ctor_set(x_737, 1, x_735); +lean_ctor_set(x_737, 2, x_734); +lean_ctor_set(x_737, 3, x_736); +x_738 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; lean_inc(x_33); -x_725 = l_Lean_Syntax_node3(x_33, x_724, x_705, x_542, x_723); -lean_inc(x_542); +x_739 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_739, 0, x_33); +lean_ctor_set(x_739, 1, x_738); +lean_inc(x_546); lean_inc(x_33); -x_726 = l_Lean_Syntax_node2(x_33, x_702, x_725, x_542); -x_727 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; -x_728 = l_Lean_addMacroScope(x_541, x_727, x_535); -x_729 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; -x_730 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; -lean_inc(x_33); -x_731 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_731, 0, x_33); -lean_ctor_set(x_731, 1, x_729); -lean_ctor_set(x_731, 2, x_728); -lean_ctor_set(x_731, 3, x_730); -x_732 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; -lean_inc(x_33); -x_733 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_733, 0, x_33); -lean_ctor_set(x_733, 1, x_732); -lean_inc(x_542); +x_740 = l_Lean_Syntax_node2(x_33, x_676, x_546, x_715); lean_inc(x_33); -x_734 = l_Lean_Syntax_node2(x_33, x_670, x_542, x_709); +x_741 = l_Lean_Syntax_node1(x_33, x_535, x_740); +x_742 = lean_array_size(x_26); +x_743 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_522, x_742, x_15, x_26); +x_744 = lean_array_size(x_743); +x_745 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_744, x_15, x_743); +x_746 = l_Array_append___rarg(x_533, x_745); +lean_dec(x_745); lean_inc(x_33); -x_735 = l_Lean_Syntax_node1(x_33, x_531, x_734); -x_736 = lean_array_size(x_26); -x_737 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_518, x_736, x_15, x_26); -x_738 = lean_array_size(x_737); -x_739 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_738, x_15, x_737); -x_740 = l_Array_append___rarg(x_529, x_739); -lean_dec(x_739); +x_747 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_747, 0, x_33); +lean_ctor_set(x_747, 1, x_535); +lean_ctor_set(x_747, 2, x_746); lean_inc(x_33); -x_741 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_741, 0, x_33); -lean_ctor_set(x_741, 1, x_531); -lean_ctor_set(x_741, 2, x_740); +x_748 = l_Lean_Syntax_node1(x_33, x_687, x_747); +lean_inc_n(x_546, 2); lean_inc(x_33); -x_742 = l_Lean_Syntax_node1(x_33, x_681, x_741); -lean_inc_n(x_542, 2); +x_749 = l_Lean_Syntax_node6(x_33, x_689, x_675, x_546, x_546, x_741, x_680, x_748); +x_750 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; lean_inc(x_33); -x_743 = l_Lean_Syntax_node6(x_33, x_683, x_669, x_542, x_542, x_735, x_674, x_742); -x_744 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +x_751 = l_Lean_Syntax_node3(x_33, x_750, x_737, x_739, x_749); lean_inc(x_33); -x_745 = l_Lean_Syntax_node3(x_33, x_744, x_731, x_733, x_743); +x_752 = l_Lean_Syntax_node1(x_33, x_726, x_751); +lean_inc(x_546); lean_inc(x_33); -x_746 = l_Lean_Syntax_node1(x_33, x_720, x_745); -lean_inc(x_542); +x_753 = l_Lean_Syntax_node2(x_33, x_708, x_752, x_546); lean_inc(x_33); -x_747 = l_Lean_Syntax_node2(x_33, x_702, x_746, x_542); +x_754 = l_Lean_Syntax_node3(x_33, x_535, x_709, x_732, x_753); +x_755 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_33); -x_748 = l_Lean_Syntax_node3(x_33, x_531, x_703, x_726, x_747); -x_749 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_756 = l_Lean_Syntax_node1(x_33, x_755, x_754); +x_757 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_33); -x_750 = l_Lean_Syntax_node1(x_33, x_749, x_748); -x_751 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_758 = l_Lean_Syntax_node2(x_33, x_757, x_705, x_756); +lean_inc(x_618); +lean_inc(x_546); lean_inc(x_33); -x_752 = l_Lean_Syntax_node2(x_33, x_751, x_699, x_750); -lean_inc(x_614); -lean_inc(x_542); +x_759 = l_Lean_Syntax_node5(x_33, x_693, x_643, x_703, x_546, x_618, x_758); lean_inc(x_33); -x_753 = l_Lean_Syntax_node5(x_33, x_687, x_639, x_697, x_542, x_614, x_752); +x_760 = l_Lean_Syntax_node1(x_33, x_695, x_759); +lean_inc(x_655); +lean_inc_n(x_546, 2); lean_inc(x_33); -x_754 = l_Lean_Syntax_node1(x_33, x_689, x_753); -lean_inc(x_649); -lean_inc_n(x_542, 2); +x_761 = l_Lean_Syntax_node4(x_33, x_697, x_546, x_546, x_760, x_655); +lean_inc(x_546); lean_inc(x_33); -x_755 = l_Lean_Syntax_node4(x_33, x_691, x_542, x_542, x_754, x_649); -lean_inc(x_542); +x_762 = l_Lean_Syntax_node3(x_33, x_535, x_698, x_546, x_761); +x_763 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_33); -x_756 = l_Lean_Syntax_node3(x_33, x_531, x_692, x_542, x_755); -x_757 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_764 = l_Lean_Syntax_node2(x_33, x_763, x_19, x_762); lean_inc(x_33); -x_758 = l_Lean_Syntax_node2(x_33, x_757, x_19, x_756); +x_765 = l_Lean_Syntax_node1(x_33, x_535, x_764); +x_766 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; lean_inc(x_33); -x_759 = l_Lean_Syntax_node1(x_33, x_531, x_758); -x_760 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +x_767 = l_Lean_Syntax_node4(x_33, x_766, x_618, x_653, x_655, x_765); +x_768 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_546); lean_inc(x_33); -x_761 = l_Lean_Syntax_node4(x_33, x_760, x_614, x_647, x_649, x_759); -x_762 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_542); +x_769 = l_Lean_Syntax_node6(x_33, x_768, x_601, x_603, x_546, x_546, x_616, x_767); lean_inc(x_33); -x_763 = l_Lean_Syntax_node6(x_33, x_762, x_597, x_599, x_542, x_542, x_612, x_761); +x_770 = l_Lean_Syntax_node2(x_33, x_584, x_599, x_769); +x_771 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; lean_inc(x_33); -x_764 = l_Lean_Syntax_node2(x_33, x_580, x_595, x_763); -x_765 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_33); -x_766 = l_Lean_Syntax_node3(x_33, x_765, x_587, x_589, x_764); -x_767 = l_Lean_Syntax_node2(x_33, x_531, x_581, x_766); -if (lean_is_scalar(x_539)) { - x_768 = lean_alloc_ctor(0, 2, 0); +x_772 = l_Lean_Syntax_node3(x_33, x_771, x_591, x_593, x_770); +x_773 = l_Lean_Syntax_node2(x_33, x_535, x_585, x_772); +if (lean_is_scalar(x_543)) { + x_774 = lean_alloc_ctor(0, 2, 0); } else { - x_768 = x_539; + x_774 = x_543; } -lean_ctor_set(x_768, 0, x_767); -lean_ctor_set(x_768, 1, x_538); -return x_768; +lean_ctor_set(x_774, 0, x_773); +lean_ctor_set(x_774, 1, x_542); +return x_774; } } else { -uint8_t x_769; +uint8_t x_775; lean_free_object(x_23); lean_dec(x_26); lean_dec(x_25); @@ -12786,1366 +12859,1372 @@ lean_dec(x_21); lean_dec(x_10); lean_dec(x_9); lean_dec(x_1); -x_769 = !lean_is_exclusive(x_28); -if (x_769 == 0) +x_775 = !lean_is_exclusive(x_28); +if (x_775 == 0) { return x_28; } else { -lean_object* x_770; lean_object* x_771; lean_object* x_772; -x_770 = lean_ctor_get(x_28, 0); -x_771 = lean_ctor_get(x_28, 1); -lean_inc(x_771); -lean_inc(x_770); +lean_object* x_776; lean_object* x_777; lean_object* x_778; +x_776 = lean_ctor_get(x_28, 0); +x_777 = lean_ctor_get(x_28, 1); +lean_inc(x_777); +lean_inc(x_776); lean_dec(x_28); -x_772 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_772, 0, x_770); -lean_ctor_set(x_772, 1, x_771); -return x_772; +x_778 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_778, 0, x_776); +lean_ctor_set(x_778, 1, x_777); +return x_778; } } } else { -lean_object* x_773; lean_object* x_774; size_t x_775; lean_object* x_776; -x_773 = lean_ctor_get(x_23, 0); -x_774 = lean_ctor_get(x_23, 1); -lean_inc(x_774); -lean_inc(x_773); +lean_object* x_779; lean_object* x_780; size_t x_781; lean_object* x_782; +x_779 = lean_ctor_get(x_23, 0); +x_780 = lean_ctor_get(x_23, 1); +lean_inc(x_780); +lean_inc(x_779); lean_dec(x_23); -x_775 = lean_array_size(x_2); -x_776 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_775, x_15, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +x_781 = lean_array_size(x_2); +x_782 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_781, x_15, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_18); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_776) == 0) +if (lean_obj_tag(x_782) == 0) { -lean_object* x_777; lean_object* x_778; lean_object* x_779; uint8_t x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; size_t x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; size_t x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; size_t x_943; lean_object* x_944; size_t x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; size_t x_1004; lean_object* x_1005; size_t x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; -x_777 = lean_ctor_get(x_776, 0); -lean_inc(x_777); -x_778 = lean_ctor_get(x_776, 1); -lean_inc(x_778); -lean_dec(x_776); -x_779 = lean_ctor_get(x_9, 5); -lean_inc(x_779); -x_780 = 0; -x_781 = l_Lean_SourceInfo_fromRef(x_779, x_780); -lean_dec(x_779); -x_782 = lean_st_ref_get(x_10, x_778); -x_783 = lean_ctor_get(x_782, 1); +lean_object* x_783; lean_object* x_784; lean_object* x_785; uint8_t x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; size_t x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; size_t x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; size_t x_951; lean_object* x_952; size_t x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; size_t x_1012; lean_object* x_1013; size_t x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; +x_783 = lean_ctor_get(x_782, 0); lean_inc(x_783); -if (lean_is_exclusive(x_782)) { - lean_ctor_release(x_782, 0); - lean_ctor_release(x_782, 1); - x_784 = x_782; +x_784 = lean_ctor_get(x_782, 1); +lean_inc(x_784); +lean_dec(x_782); +x_785 = lean_ctor_get(x_9, 5); +lean_inc(x_785); +x_786 = 0; +x_787 = l_Lean_SourceInfo_fromRef(x_785, x_786); +lean_dec(x_785); +x_788 = lean_st_ref_get(x_10, x_784); +x_789 = lean_ctor_get(x_788, 1); +lean_inc(x_789); +if (lean_is_exclusive(x_788)) { + lean_ctor_release(x_788, 0); + lean_ctor_release(x_788, 1); + x_790 = x_788; } else { - lean_dec_ref(x_782); - x_784 = lean_box(0); -} -x_785 = lean_box(0); -x_786 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_781); -if (lean_is_scalar(x_784)) { - x_787 = lean_alloc_ctor(2, 2, 0); + lean_dec_ref(x_788); + x_790 = lean_box(0); +} +x_791 = lean_box(0); +x_792 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_787); +if (lean_is_scalar(x_790)) { + x_793 = lean_alloc_ctor(2, 2, 0); } else { - x_787 = x_784; - lean_ctor_set_tag(x_787, 2); + x_793 = x_790; + lean_ctor_set_tag(x_793, 2); } -lean_ctor_set(x_787, 0, x_781); -lean_ctor_set(x_787, 1, x_786); -x_788 = lean_ctor_get(x_1, 0); -lean_inc(x_788); +lean_ctor_set(x_793, 0, x_787); +lean_ctor_set(x_793, 1, x_792); +x_794 = lean_ctor_get(x_1, 0); +lean_inc(x_794); lean_dec(x_1); -x_789 = lean_ctor_get(x_788, 0); -lean_inc(x_789); -lean_dec(x_788); -x_790 = lean_mk_syntax_ident(x_789); -x_791 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_781); -x_792 = l_Lean_Syntax_node2(x_781, x_791, x_787, x_790); -x_793 = lean_array_size(x_777); -x_794 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; -x_795 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_794, x_793, x_15, x_777); -x_796 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -x_797 = l_Array_append___rarg(x_796, x_795); -lean_dec(x_795); -x_798 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -lean_inc(x_781); -x_799 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_799, 0, x_781); -lean_ctor_set(x_799, 1, x_798); -lean_ctor_set(x_799, 2, x_797); -x_800 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_781); -x_801 = l_Lean_Syntax_node2(x_781, x_800, x_792, x_799); -x_802 = lean_ctor_get(x_9, 10); -lean_inc(x_802); +x_795 = lean_ctor_get(x_794, 0); +lean_inc(x_795); +lean_dec(x_794); +x_796 = lean_mk_syntax_ident(x_795); +x_797 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; +lean_inc(x_787); +x_798 = l_Lean_Syntax_node2(x_787, x_797, x_793, x_796); +x_799 = lean_array_size(x_783); +x_800 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; +x_801 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_800, x_799, x_15, x_783); +x_802 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +x_803 = l_Array_append___rarg(x_802, x_801); +lean_dec(x_801); +x_804 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +lean_inc(x_787); +x_805 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_805, 0, x_787); +lean_ctor_set(x_805, 1, x_804); +lean_ctor_set(x_805, 2, x_803); +x_806 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_787); +x_807 = l_Lean_Syntax_node2(x_787, x_806, x_798, x_805); +x_808 = lean_ctor_get(x_9, 10); +lean_inc(x_808); lean_dec(x_9); -x_803 = lean_st_ref_get(x_10, x_783); +x_809 = lean_st_ref_get(x_10, x_789); lean_dec(x_10); -x_804 = lean_ctor_get(x_803, 0); -lean_inc(x_804); -x_805 = lean_ctor_get(x_803, 1); -lean_inc(x_805); -if (lean_is_exclusive(x_803)) { - lean_ctor_release(x_803, 0); - lean_ctor_release(x_803, 1); - x_806 = x_803; +x_810 = lean_ctor_get(x_809, 0); +lean_inc(x_810); +x_811 = lean_ctor_get(x_809, 1); +lean_inc(x_811); +if (lean_is_exclusive(x_809)) { + lean_ctor_release(x_809, 0); + lean_ctor_release(x_809, 1); + x_812 = x_809; } else { - lean_dec_ref(x_803); - x_806 = lean_box(0); -} -x_807 = lean_ctor_get(x_804, 0); -lean_inc(x_807); -lean_dec(x_804); -x_808 = lean_environment_main_module(x_807); -lean_inc(x_781); -x_809 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_809, 0, x_781); -lean_ctor_set(x_809, 1, x_798); -lean_ctor_set(x_809, 2, x_796); -x_810 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_809, 6); -lean_inc(x_781); -x_811 = l_Lean_Syntax_node6(x_781, x_810, x_809, x_809, x_809, x_809, x_809, x_809); -x_812 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; -lean_inc(x_781); -x_813 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_813, 0, x_781); -lean_ctor_set(x_813, 1, x_812); -x_814 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_802); + lean_dec_ref(x_809); + x_812 = lean_box(0); +} +x_813 = lean_ctor_get(x_810, 0); +lean_inc(x_813); +lean_dec(x_810); +x_814 = lean_environment_main_module(x_813); +lean_inc(x_787); +x_815 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_815, 0, x_787); +lean_ctor_set(x_815, 1, x_804); +lean_ctor_set(x_815, 2, x_802); +x_816 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_815, 6); +lean_inc(x_787); +x_817 = l_Lean_Syntax_node6(x_787, x_816, x_815, x_815, x_815, x_815, x_815, x_815); +x_818 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; +lean_inc(x_787); +x_819 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_819, 0, x_787); +lean_ctor_set(x_819, 1, x_818); +x_820 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; lean_inc(x_808); -x_815 = l_Lean_addMacroScope(x_808, x_814, x_802); -x_816 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_781); -x_817 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_817, 0, x_781); -lean_ctor_set(x_817, 1, x_816); -lean_ctor_set(x_817, 2, x_815); -lean_ctor_set(x_817, 3, x_785); -x_818 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_809); -lean_inc(x_817); -lean_inc(x_781); -x_819 = l_Lean_Syntax_node2(x_781, x_818, x_817, x_809); -x_820 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_821 = l_Lean_Syntax_node2(x_781, x_820, x_809, x_809); -x_822 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_781); +lean_inc(x_814); +x_821 = l_Lean_addMacroScope(x_814, x_820, x_808); +x_822 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_787); +x_823 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_823, 0, x_787); +lean_ctor_set(x_823, 1, x_822); +lean_ctor_set(x_823, 2, x_821); +lean_ctor_set(x_823, 3, x_791); +x_824 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; +lean_inc(x_815); +lean_inc(x_823); +lean_inc(x_787); +x_825 = l_Lean_Syntax_node2(x_787, x_824, x_823, x_815); +x_826 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_827 = l_Lean_Syntax_node2(x_787, x_826, x_815, x_815); +x_828 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_787); lean_ctor_set_tag(x_19, 2); -lean_ctor_set(x_19, 1, x_822); -lean_ctor_set(x_19, 0, x_781); +lean_ctor_set(x_19, 1, x_828); +lean_ctor_set(x_19, 0, x_787); lean_inc(x_19); -lean_inc(x_781); -x_823 = l_Lean_Syntax_node1(x_781, x_798, x_19); -x_824 = lean_array_size(x_21); -x_825 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_824, x_15, x_21); -x_826 = l_Array_append___rarg(x_796, x_825); -lean_dec(x_825); -lean_inc(x_781); -x_827 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_827, 0, x_781); -lean_ctor_set(x_827, 1, x_798); -lean_ctor_set(x_827, 2, x_826); -x_828 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_781); -x_829 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_829, 0, x_781); -lean_ctor_set(x_829, 1, x_828); -x_830 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_802); +lean_inc(x_787); +x_829 = l_Lean_Syntax_node1(x_787, x_804, x_19); +x_830 = lean_array_size(x_21); +x_831 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_830, x_15, x_21); +x_832 = l_Array_append___rarg(x_802, x_831); +lean_dec(x_831); +lean_inc(x_787); +x_833 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_833, 0, x_787); +lean_ctor_set(x_833, 1, x_804); +lean_ctor_set(x_833, 2, x_832); +x_834 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_787); +x_835 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_835, 0, x_787); +lean_ctor_set(x_835, 1, x_834); +x_836 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; lean_inc(x_808); -x_831 = l_Lean_addMacroScope(x_808, x_830, x_802); -x_832 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_833 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_781); -x_834 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_834, 0, x_781); -lean_ctor_set(x_834, 1, x_832); -lean_ctor_set(x_834, 2, x_831); -lean_ctor_set(x_834, 3, x_833); -x_835 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_781); -x_836 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_836, 0, x_781); -lean_ctor_set(x_836, 1, x_835); -x_837 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_802); +lean_inc(x_814); +x_837 = l_Lean_addMacroScope(x_814, x_836, x_808); +x_838 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_839 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_787); +x_840 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_840, 0, x_787); +lean_ctor_set(x_840, 1, x_838); +lean_ctor_set(x_840, 2, x_837); +lean_ctor_set(x_840, 3, x_839); +x_841 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_787); +x_842 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_842, 0, x_787); +lean_ctor_set(x_842, 1, x_841); +x_843 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; lean_inc(x_808); -x_838 = l_Lean_addMacroScope(x_808, x_837, x_802); -x_839 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_840 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_781); -x_841 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_841, 0, x_781); -lean_ctor_set(x_841, 1, x_839); -lean_ctor_set(x_841, 2, x_838); -lean_ctor_set(x_841, 3, x_840); -lean_inc(x_836); -lean_inc(x_781); -x_842 = l_Lean_Syntax_node3(x_781, x_798, x_834, x_836, x_841); -lean_inc(x_781); -x_843 = l_Lean_Syntax_node2(x_781, x_798, x_829, x_842); -x_844 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_781); -x_845 = l_Lean_Syntax_node1(x_781, x_844, x_843); -x_846 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; -lean_inc(x_809); -lean_inc(x_781); -x_847 = l_Lean_Syntax_node7(x_781, x_846, x_813, x_819, x_821, x_823, x_827, x_809, x_845); -x_848 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_781); -x_849 = l_Lean_Syntax_node2(x_781, x_848, x_811, x_847); -x_850 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_781); -x_851 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_851, 0, x_781); -lean_ctor_set(x_851, 1, x_850); -x_852 = l_Array_append___rarg(x_796, x_3); -lean_inc(x_781); -x_853 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_853, 0, x_781); -lean_ctor_set(x_853, 1, x_798); -lean_ctor_set(x_853, 2, x_852); -x_854 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_781); -x_855 = l_Lean_Syntax_node2(x_781, x_854, x_851, x_853); -x_856 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_781); +lean_inc(x_814); +x_844 = l_Lean_addMacroScope(x_814, x_843, x_808); +x_845 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_846 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_787); +x_847 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_847, 0, x_787); +lean_ctor_set(x_847, 1, x_845); +lean_ctor_set(x_847, 2, x_844); +lean_ctor_set(x_847, 3, x_846); +lean_inc(x_842); +lean_inc(x_787); +x_848 = l_Lean_Syntax_node3(x_787, x_804, x_840, x_842, x_847); +lean_inc(x_787); +x_849 = l_Lean_Syntax_node2(x_787, x_804, x_835, x_848); +x_850 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_787); +x_851 = l_Lean_Syntax_node1(x_787, x_850, x_849); +x_852 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; +lean_inc(x_815); +lean_inc(x_787); +x_853 = l_Lean_Syntax_node7(x_787, x_852, x_819, x_825, x_827, x_829, x_833, x_815, x_851); +x_854 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_787); +x_855 = l_Lean_Syntax_node2(x_787, x_854, x_817, x_853); +x_856 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_787); x_857 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_857, 0, x_781); +lean_ctor_set(x_857, 0, x_787); lean_ctor_set(x_857, 1, x_856); -x_858 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; -lean_inc(x_781); -x_859 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_859, 0, x_781); -lean_ctor_set(x_859, 1, x_858); -x_860 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; -lean_inc(x_781); -x_861 = l_Lean_Syntax_node1(x_781, x_860, x_859); -lean_inc(x_781); -x_862 = l_Lean_Syntax_node1(x_781, x_798, x_861); -lean_inc_n(x_809, 5); -lean_inc(x_781); -x_863 = l_Lean_Syntax_node6(x_781, x_810, x_809, x_809, x_809, x_809, x_809, x_862); -x_864 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_809); -lean_inc(x_781); -x_865 = l_Lean_Syntax_node1(x_781, x_864, x_809); -x_866 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_781); -x_867 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_867, 0, x_781); -lean_ctor_set(x_867, 1, x_866); -x_868 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_781); -x_869 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_869, 0, x_781); -lean_ctor_set(x_869, 1, x_868); -x_870 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_802); +x_858 = l_Array_append___rarg(x_802, x_3); +lean_inc(x_787); +x_859 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_859, 0, x_787); +lean_ctor_set(x_859, 1, x_804); +lean_ctor_set(x_859, 2, x_858); +x_860 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_787); +x_861 = l_Lean_Syntax_node2(x_787, x_860, x_857, x_859); +x_862 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_787); +x_863 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_863, 0, x_787); +lean_ctor_set(x_863, 1, x_862); +x_864 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; +lean_inc(x_787); +x_865 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_865, 0, x_787); +lean_ctor_set(x_865, 1, x_864); +x_866 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; +lean_inc(x_787); +x_867 = l_Lean_Syntax_node1(x_787, x_866, x_865); +lean_inc(x_787); +x_868 = l_Lean_Syntax_node1(x_787, x_804, x_867); +lean_inc_n(x_815, 5); +lean_inc(x_787); +x_869 = l_Lean_Syntax_node6(x_787, x_816, x_815, x_815, x_815, x_815, x_815, x_868); +x_870 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_815); +lean_inc(x_787); +x_871 = l_Lean_Syntax_node1(x_787, x_870, x_815); +x_872 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_787); +x_873 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_873, 0, x_787); +lean_ctor_set(x_873, 1, x_872); +x_874 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_787); +x_875 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_875, 0, x_787); +lean_ctor_set(x_875, 1, x_874); +x_876 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; lean_inc(x_808); -x_871 = l_Lean_addMacroScope(x_808, x_870, x_802); -x_872 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_873 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_781); -x_874 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_874, 0, x_781); -lean_ctor_set(x_874, 1, x_872); -lean_ctor_set(x_874, 2, x_871); -lean_ctor_set(x_874, 3, x_873); -lean_inc(x_781); -x_875 = l_Lean_Syntax_node1(x_781, x_798, x_801); -lean_inc(x_781); -x_876 = l_Lean_Syntax_node2(x_781, x_800, x_874, x_875); -x_877 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_869); -lean_inc(x_781); -x_878 = l_Lean_Syntax_node2(x_781, x_877, x_869, x_876); -x_879 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_878); -lean_inc(x_809); -lean_inc(x_781); -x_880 = l_Lean_Syntax_node2(x_781, x_879, x_809, x_878); -x_881 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_781); -x_882 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_882, 0, x_781); -lean_ctor_set(x_882, 1, x_881); -x_883 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_781); -x_884 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_884, 0, x_781); -lean_ctor_set(x_884, 1, x_883); -x_885 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_802); -lean_inc(x_808); -x_886 = l_Lean_addMacroScope(x_808, x_885, x_802); -x_887 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_888 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_781); -x_889 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_889, 0, x_781); -lean_ctor_set(x_889, 1, x_887); -lean_ctor_set(x_889, 2, x_886); -lean_ctor_set(x_889, 3, x_888); -x_890 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_809); -lean_inc(x_781); -x_891 = l_Lean_Syntax_node2(x_781, x_890, x_889, x_809); -x_892 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_802); +lean_inc(x_814); +x_877 = l_Lean_addMacroScope(x_814, x_876, x_808); +x_878 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_879 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_787); +x_880 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_880, 0, x_787); +lean_ctor_set(x_880, 1, x_878); +lean_ctor_set(x_880, 2, x_877); +lean_ctor_set(x_880, 3, x_879); +lean_inc(x_787); +x_881 = l_Lean_Syntax_node1(x_787, x_804, x_807); +lean_inc(x_787); +x_882 = l_Lean_Syntax_node2(x_787, x_806, x_880, x_881); +x_883 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_875); +lean_inc(x_787); +x_884 = l_Lean_Syntax_node2(x_787, x_883, x_875, x_882); +x_885 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_884); +lean_inc(x_815); +lean_inc(x_787); +x_886 = l_Lean_Syntax_node2(x_787, x_885, x_815, x_884); +x_887 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_787); +x_888 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_888, 0, x_787); +lean_ctor_set(x_888, 1, x_887); +x_889 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_787); +x_890 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_890, 0, x_787); +lean_ctor_set(x_890, 1, x_889); +x_891 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; lean_inc(x_808); -x_893 = l_Lean_addMacroScope(x_808, x_892, x_802); -x_894 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_781); +lean_inc(x_814); +x_892 = l_Lean_addMacroScope(x_814, x_891, x_808); +x_893 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_894 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_787); x_895 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_895, 0, x_781); -lean_ctor_set(x_895, 1, x_894); -lean_ctor_set(x_895, 2, x_893); -lean_ctor_set(x_895, 3, x_785); -x_896 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_895); -lean_inc(x_882); -lean_inc(x_781); -x_897 = l_Lean_Syntax_node3(x_781, x_896, x_891, x_882, x_895); -x_898 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_802); +lean_ctor_set(x_895, 0, x_787); +lean_ctor_set(x_895, 1, x_893); +lean_ctor_set(x_895, 2, x_892); +lean_ctor_set(x_895, 3, x_894); +x_896 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_815); +lean_inc(x_787); +x_897 = l_Lean_Syntax_node2(x_787, x_896, x_895, x_815); +x_898 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; lean_inc(x_808); -x_899 = l_Lean_addMacroScope(x_808, x_898, x_802); -x_900 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_901 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_781); -x_902 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_902, 0, x_781); -lean_ctor_set(x_902, 1, x_900); -lean_ctor_set(x_902, 2, x_899); -lean_ctor_set(x_902, 3, x_901); -lean_inc(x_809); -lean_inc(x_781); -x_903 = l_Lean_Syntax_node2(x_781, x_890, x_902, x_809); -x_904 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_802); +lean_inc(x_814); +x_899 = l_Lean_addMacroScope(x_814, x_898, x_808); +x_900 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_787); +x_901 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_901, 0, x_787); +lean_ctor_set(x_901, 1, x_900); +lean_ctor_set(x_901, 2, x_899); +lean_ctor_set(x_901, 3, x_791); +x_902 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_901); +lean_inc(x_888); +lean_inc(x_787); +x_903 = l_Lean_Syntax_node3(x_787, x_902, x_897, x_888, x_901); +x_904 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; lean_inc(x_808); -x_905 = l_Lean_addMacroScope(x_808, x_904, x_802); -x_906 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_781); -x_907 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_907, 0, x_781); -lean_ctor_set(x_907, 1, x_906); -lean_ctor_set(x_907, 2, x_905); -lean_ctor_set(x_907, 3, x_785); -lean_inc(x_907); -lean_inc(x_882); -lean_inc(x_781); -x_908 = l_Lean_Syntax_node3(x_781, x_896, x_903, x_882, x_907); -lean_inc(x_781); -x_909 = l_Lean_Syntax_node3(x_781, x_798, x_897, x_836, x_908); +lean_inc(x_814); +x_905 = l_Lean_addMacroScope(x_814, x_904, x_808); +x_906 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_907 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_787); +x_908 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_908, 0, x_787); +lean_ctor_set(x_908, 1, x_906); +lean_ctor_set(x_908, 2, x_905); +lean_ctor_set(x_908, 3, x_907); +lean_inc(x_815); +lean_inc(x_787); +x_909 = l_Lean_Syntax_node2(x_787, x_896, x_908, x_815); x_910 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_809); -lean_inc(x_781); -x_911 = l_Lean_Syntax_node1(x_781, x_910, x_809); -x_912 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_781); -x_913 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_913, 0, x_781); -lean_ctor_set(x_913, 1, x_912); -x_914 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_915 = l_Lean_Syntax_node6(x_781, x_914, x_884, x_809, x_909, x_911, x_809, x_913); -x_916 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_917 = l_Lean_Syntax_node2(x_781, x_916, x_809, x_809); -x_918 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; -lean_inc(x_802); lean_inc(x_808); -x_919 = l_Lean_addMacroScope(x_808, x_918, x_802); -x_920 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; -lean_inc(x_781); -x_921 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_921, 0, x_781); +lean_inc(x_814); +x_911 = l_Lean_addMacroScope(x_814, x_910, x_808); +x_912 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_787); +x_913 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_913, 0, x_787); +lean_ctor_set(x_913, 1, x_912); +lean_ctor_set(x_913, 2, x_911); +lean_ctor_set(x_913, 3, x_791); +lean_inc(x_913); +lean_inc(x_888); +lean_inc(x_787); +x_914 = l_Lean_Syntax_node3(x_787, x_902, x_909, x_888, x_913); +lean_inc(x_787); +x_915 = l_Lean_Syntax_node3(x_787, x_804, x_903, x_842, x_914); +x_916 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_787); +x_917 = l_Lean_Syntax_node1(x_787, x_916, x_915); +x_918 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_815); +lean_inc(x_787); +x_919 = l_Lean_Syntax_node1(x_787, x_918, x_815); +x_920 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_787); +x_921 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_921, 0, x_787); lean_ctor_set(x_921, 1, x_920); -lean_ctor_set(x_921, 2, x_919); -lean_ctor_set(x_921, 3, x_785); -lean_inc(x_921); -lean_inc(x_781); -x_922 = l_Lean_Syntax_node1(x_781, x_798, x_921); -x_923 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; -lean_inc(x_781); -x_924 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_924, 0, x_781); -lean_ctor_set(x_924, 1, x_923); -x_925 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; -lean_inc(x_802); +x_922 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_923 = l_Lean_Syntax_node6(x_787, x_922, x_890, x_815, x_917, x_919, x_815, x_921); +x_924 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_925 = l_Lean_Syntax_node2(x_787, x_924, x_815, x_815); +x_926 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; lean_inc(x_808); -x_926 = l_Lean_addMacroScope(x_808, x_925, x_802); -x_927 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; -lean_inc(x_781); -x_928 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_928, 0, x_781); -lean_ctor_set(x_928, 1, x_927); -lean_ctor_set(x_928, 2, x_926); -lean_ctor_set(x_928, 3, x_785); -x_929 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; -lean_inc(x_781); -x_930 = l_Lean_Syntax_node1(x_781, x_929, x_928); -lean_inc(x_781); -x_931 = l_Lean_Syntax_node1(x_781, x_798, x_878); -x_932 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; -lean_inc(x_915); -lean_inc(x_882); -lean_inc(x_809); -lean_inc(x_781); -x_933 = l_Lean_Syntax_node5(x_781, x_932, x_930, x_809, x_931, x_882, x_915); -x_934 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; -lean_inc(x_781); -x_935 = l_Lean_Syntax_node1(x_781, x_934, x_933); -x_936 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; -lean_inc(x_781); -x_937 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_937, 0, x_781); -lean_ctor_set(x_937, 1, x_936); -x_938 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; -lean_inc(x_809); -lean_inc(x_781); -x_939 = l_Lean_Syntax_node2(x_781, x_938, x_809, x_921); -lean_inc(x_781); -x_940 = l_Lean_Syntax_node1(x_781, x_798, x_939); -x_941 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_781); -x_942 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_942, 0, x_781); -lean_ctor_set(x_942, 1, x_941); -x_943 = lean_array_size(x_773); -x_944 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_785, x_943, x_15, x_773); -x_945 = lean_array_size(x_944); -x_946 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_945, x_15, x_944); -x_947 = l_Array_append___rarg(x_796, x_946); -lean_dec(x_946); -lean_inc(x_781); -x_948 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_948, 0, x_781); -lean_ctor_set(x_948, 1, x_798); -lean_ctor_set(x_948, 2, x_947); -x_949 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; -lean_inc(x_781); -x_950 = l_Lean_Syntax_node1(x_781, x_949, x_948); -x_951 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_942); -lean_inc_n(x_809, 2); -lean_inc(x_937); -lean_inc(x_781); -x_952 = l_Lean_Syntax_node6(x_781, x_951, x_937, x_809, x_809, x_940, x_942, x_950); -x_953 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; -lean_inc(x_809); -lean_inc(x_935); -lean_inc(x_924); -lean_inc(x_781); -x_954 = l_Lean_Syntax_node4(x_781, x_953, x_924, x_935, x_809, x_952); -x_955 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_882); -lean_inc(x_809); -lean_inc(x_781); -x_956 = l_Lean_Syntax_node5(x_781, x_955, x_895, x_922, x_809, x_882, x_954); -x_957 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_781); -x_958 = l_Lean_Syntax_node1(x_781, x_957, x_956); -x_959 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_917); -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_960 = l_Lean_Syntax_node4(x_781, x_959, x_809, x_809, x_958, x_917); -x_961 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_802); +lean_inc(x_814); +x_927 = l_Lean_addMacroScope(x_814, x_926, x_808); +x_928 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_787); +x_929 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_929, 0, x_787); +lean_ctor_set(x_929, 1, x_928); +lean_ctor_set(x_929, 2, x_927); +lean_ctor_set(x_929, 3, x_791); +lean_inc(x_929); +lean_inc(x_787); +x_930 = l_Lean_Syntax_node1(x_787, x_804, x_929); +x_931 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; +lean_inc(x_787); +x_932 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_932, 0, x_787); +lean_ctor_set(x_932, 1, x_931); +x_933 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; lean_inc(x_808); -x_962 = l_Lean_addMacroScope(x_808, x_961, x_802); -x_963 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_781); -x_964 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_964, 0, x_781); -lean_ctor_set(x_964, 1, x_963); -lean_ctor_set(x_964, 2, x_962); -lean_ctor_set(x_964, 3, x_785); -lean_inc(x_781); -x_965 = l_Lean_Syntax_node1(x_781, x_798, x_964); -x_966 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_781); -x_967 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_967, 0, x_781); -lean_ctor_set(x_967, 1, x_966); -x_968 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; -lean_inc(x_781); -x_969 = l_Lean_Syntax_node2(x_781, x_968, x_924, x_935); -x_970 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_809); -lean_inc(x_781); -x_971 = l_Lean_Syntax_node2(x_781, x_970, x_969, x_809); -x_972 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_781); -x_973 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_973, 0, x_781); -lean_ctor_set(x_973, 1, x_972); -x_974 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; -lean_inc(x_802); +lean_inc(x_814); +x_934 = l_Lean_addMacroScope(x_814, x_933, x_808); +x_935 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +lean_inc(x_787); +x_936 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_936, 0, x_787); +lean_ctor_set(x_936, 1, x_935); +lean_ctor_set(x_936, 2, x_934); +lean_ctor_set(x_936, 3, x_791); +x_937 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; +lean_inc(x_787); +x_938 = l_Lean_Syntax_node1(x_787, x_937, x_936); +lean_inc(x_787); +x_939 = l_Lean_Syntax_node1(x_787, x_804, x_884); +x_940 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +lean_inc(x_923); +lean_inc(x_888); +lean_inc(x_815); +lean_inc(x_787); +x_941 = l_Lean_Syntax_node5(x_787, x_940, x_938, x_815, x_939, x_888, x_923); +x_942 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +lean_inc(x_787); +x_943 = l_Lean_Syntax_node1(x_787, x_942, x_941); +x_944 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; +lean_inc(x_787); +x_945 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_945, 0, x_787); +lean_ctor_set(x_945, 1, x_944); +x_946 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; +lean_inc(x_815); +lean_inc(x_787); +x_947 = l_Lean_Syntax_node2(x_787, x_946, x_815, x_929); +lean_inc(x_787); +x_948 = l_Lean_Syntax_node1(x_787, x_804, x_947); +x_949 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +lean_inc(x_787); +x_950 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_950, 0, x_787); +lean_ctor_set(x_950, 1, x_949); +x_951 = lean_array_size(x_779); +x_952 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_791, x_951, x_15, x_779); +x_953 = lean_array_size(x_952); +x_954 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_953, x_15, x_952); +x_955 = l_Array_append___rarg(x_802, x_954); +lean_dec(x_954); +lean_inc(x_787); +x_956 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_956, 0, x_787); +lean_ctor_set(x_956, 1, x_804); +lean_ctor_set(x_956, 2, x_955); +x_957 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +lean_inc(x_787); +x_958 = l_Lean_Syntax_node1(x_787, x_957, x_956); +x_959 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_950); +lean_inc_n(x_815, 2); +lean_inc(x_945); +lean_inc(x_787); +x_960 = l_Lean_Syntax_node6(x_787, x_959, x_945, x_815, x_815, x_948, x_950, x_958); +x_961 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +lean_inc(x_815); +lean_inc(x_943); +lean_inc(x_932); +lean_inc(x_787); +x_962 = l_Lean_Syntax_node4(x_787, x_961, x_932, x_943, x_815, x_960); +x_963 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_888); +lean_inc(x_815); +lean_inc(x_787); +x_964 = l_Lean_Syntax_node5(x_787, x_963, x_901, x_930, x_815, x_888, x_962); +x_965 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_787); +x_966 = l_Lean_Syntax_node1(x_787, x_965, x_964); +x_967 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_925); +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_968 = l_Lean_Syntax_node4(x_787, x_967, x_815, x_815, x_966, x_925); +x_969 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_808); -x_975 = l_Lean_addMacroScope(x_808, x_974, x_802); -x_976 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; -lean_inc(x_781); -x_977 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_977, 0, x_781); -lean_ctor_set(x_977, 1, x_976); -lean_ctor_set(x_977, 2, x_975); -lean_ctor_set(x_977, 3, x_785); -lean_inc(x_781); -x_978 = l_Lean_Syntax_node2(x_781, x_877, x_869, x_817); -lean_inc(x_781); -x_979 = l_Lean_Syntax_node1(x_781, x_798, x_978); -x_980 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_781); +lean_inc(x_814); +x_970 = l_Lean_addMacroScope(x_814, x_969, x_808); +x_971 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_787); +x_972 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_972, 0, x_787); +lean_ctor_set(x_972, 1, x_971); +lean_ctor_set(x_972, 2, x_970); +lean_ctor_set(x_972, 3, x_791); +lean_inc(x_787); +x_973 = l_Lean_Syntax_node1(x_787, x_804, x_972); +x_974 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_787); +x_975 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_975, 0, x_787); +lean_ctor_set(x_975, 1, x_974); +x_976 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +lean_inc(x_787); +x_977 = l_Lean_Syntax_node2(x_787, x_976, x_932, x_943); +x_978 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_815); +lean_inc(x_787); +x_979 = l_Lean_Syntax_node2(x_787, x_978, x_977, x_815); +x_980 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_787); x_981 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_981, 0, x_781); +lean_ctor_set(x_981, 0, x_787); lean_ctor_set(x_981, 1, x_980); -x_982 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -lean_inc(x_802); +x_982 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; lean_inc(x_808); -x_983 = l_Lean_addMacroScope(x_808, x_982, x_802); -x_984 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_985 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_781); -x_986 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_986, 0, x_781); -lean_ctor_set(x_986, 1, x_984); -lean_ctor_set(x_986, 2, x_983); -lean_ctor_set(x_986, 3, x_985); -lean_inc(x_965); -lean_inc(x_781); -x_987 = l_Lean_Syntax_node2(x_781, x_800, x_986, x_965); -x_988 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_781); -x_989 = l_Lean_Syntax_node1(x_781, x_988, x_987); -x_990 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_977); -lean_inc(x_781); -x_991 = l_Lean_Syntax_node4(x_781, x_990, x_977, x_979, x_981, x_989); -x_992 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_809); -lean_inc(x_781); -x_993 = l_Lean_Syntax_node3(x_781, x_992, x_973, x_809, x_991); -lean_inc(x_809); -lean_inc(x_781); -x_994 = l_Lean_Syntax_node2(x_781, x_970, x_993, x_809); -x_995 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; -x_996 = l_Lean_addMacroScope(x_808, x_995, x_802); -x_997 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; -x_998 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; -lean_inc(x_781); -x_999 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_999, 0, x_781); -lean_ctor_set(x_999, 1, x_997); -lean_ctor_set(x_999, 2, x_996); -lean_ctor_set(x_999, 3, x_998); -x_1000 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; -lean_inc(x_781); -x_1001 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1001, 0, x_781); -lean_ctor_set(x_1001, 1, x_1000); -lean_inc(x_809); -lean_inc(x_781); -x_1002 = l_Lean_Syntax_node2(x_781, x_938, x_809, x_977); -lean_inc(x_781); -x_1003 = l_Lean_Syntax_node1(x_781, x_798, x_1002); -x_1004 = lean_array_size(x_774); -x_1005 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_785, x_1004, x_15, x_774); -x_1006 = lean_array_size(x_1005); -x_1007 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1006, x_15, x_1005); -x_1008 = l_Array_append___rarg(x_796, x_1007); -lean_dec(x_1007); -lean_inc(x_781); -x_1009 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1009, 0, x_781); -lean_ctor_set(x_1009, 1, x_798); -lean_ctor_set(x_1009, 2, x_1008); -lean_inc(x_781); -x_1010 = l_Lean_Syntax_node1(x_781, x_949, x_1009); -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_1011 = l_Lean_Syntax_node6(x_781, x_951, x_937, x_809, x_809, x_1003, x_942, x_1010); -x_1012 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; -lean_inc(x_781); -x_1013 = l_Lean_Syntax_node3(x_781, x_1012, x_999, x_1001, x_1011); -lean_inc(x_781); -x_1014 = l_Lean_Syntax_node1(x_781, x_988, x_1013); -lean_inc(x_809); -lean_inc(x_781); -x_1015 = l_Lean_Syntax_node2(x_781, x_970, x_1014, x_809); -lean_inc(x_781); -x_1016 = l_Lean_Syntax_node3(x_781, x_798, x_971, x_994, x_1015); -x_1017 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_781); -x_1018 = l_Lean_Syntax_node1(x_781, x_1017, x_1016); -x_1019 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_781); -x_1020 = l_Lean_Syntax_node2(x_781, x_1019, x_967, x_1018); -lean_inc(x_882); -lean_inc(x_809); -lean_inc(x_781); -x_1021 = l_Lean_Syntax_node5(x_781, x_955, x_907, x_965, x_809, x_882, x_1020); -lean_inc(x_781); -x_1022 = l_Lean_Syntax_node1(x_781, x_957, x_1021); -lean_inc(x_917); -lean_inc_n(x_809, 2); -lean_inc(x_781); -x_1023 = l_Lean_Syntax_node4(x_781, x_959, x_809, x_809, x_1022, x_917); -lean_inc(x_809); -lean_inc(x_781); -x_1024 = l_Lean_Syntax_node3(x_781, x_798, x_960, x_809, x_1023); -x_1025 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_781); -x_1026 = l_Lean_Syntax_node2(x_781, x_1025, x_19, x_1024); -lean_inc(x_781); -x_1027 = l_Lean_Syntax_node1(x_781, x_798, x_1026); -x_1028 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_781); -x_1029 = l_Lean_Syntax_node4(x_781, x_1028, x_882, x_915, x_917, x_1027); -x_1030 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_809); -lean_inc(x_781); -x_1031 = l_Lean_Syntax_node6(x_781, x_1030, x_865, x_867, x_809, x_809, x_880, x_1029); -lean_inc(x_781); -x_1032 = l_Lean_Syntax_node2(x_781, x_848, x_863, x_1031); -x_1033 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_781); -x_1034 = l_Lean_Syntax_node3(x_781, x_1033, x_855, x_857, x_1032); -x_1035 = l_Lean_Syntax_node2(x_781, x_798, x_849, x_1034); -if (lean_is_scalar(x_806)) { - x_1036 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_814); +x_983 = l_Lean_addMacroScope(x_814, x_982, x_808); +x_984 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +lean_inc(x_787); +x_985 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_985, 0, x_787); +lean_ctor_set(x_985, 1, x_984); +lean_ctor_set(x_985, 2, x_983); +lean_ctor_set(x_985, 3, x_791); +lean_inc(x_787); +x_986 = l_Lean_Syntax_node2(x_787, x_883, x_875, x_823); +lean_inc(x_787); +x_987 = l_Lean_Syntax_node1(x_787, x_804, x_986); +x_988 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_787); +x_989 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_989, 0, x_787); +lean_ctor_set(x_989, 1, x_988); +x_990 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +lean_inc(x_808); +lean_inc(x_814); +x_991 = l_Lean_addMacroScope(x_814, x_990, x_808); +x_992 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_993 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_787); +x_994 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_994, 0, x_787); +lean_ctor_set(x_994, 1, x_992); +lean_ctor_set(x_994, 2, x_991); +lean_ctor_set(x_994, 3, x_993); +lean_inc(x_973); +lean_inc(x_787); +x_995 = l_Lean_Syntax_node2(x_787, x_806, x_994, x_973); +x_996 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_787); +x_997 = l_Lean_Syntax_node1(x_787, x_996, x_995); +x_998 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_985); +lean_inc(x_787); +x_999 = l_Lean_Syntax_node4(x_787, x_998, x_985, x_987, x_989, x_997); +x_1000 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_815); +lean_inc(x_787); +x_1001 = l_Lean_Syntax_node3(x_787, x_1000, x_981, x_815, x_999); +lean_inc(x_815); +lean_inc(x_787); +x_1002 = l_Lean_Syntax_node2(x_787, x_978, x_1001, x_815); +x_1003 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; +x_1004 = l_Lean_addMacroScope(x_814, x_1003, x_808); +x_1005 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; +x_1006 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; +lean_inc(x_787); +x_1007 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1007, 0, x_787); +lean_ctor_set(x_1007, 1, x_1005); +lean_ctor_set(x_1007, 2, x_1004); +lean_ctor_set(x_1007, 3, x_1006); +x_1008 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; +lean_inc(x_787); +x_1009 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1009, 0, x_787); +lean_ctor_set(x_1009, 1, x_1008); +lean_inc(x_815); +lean_inc(x_787); +x_1010 = l_Lean_Syntax_node2(x_787, x_946, x_815, x_985); +lean_inc(x_787); +x_1011 = l_Lean_Syntax_node1(x_787, x_804, x_1010); +x_1012 = lean_array_size(x_780); +x_1013 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_791, x_1012, x_15, x_780); +x_1014 = lean_array_size(x_1013); +x_1015 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1014, x_15, x_1013); +x_1016 = l_Array_append___rarg(x_802, x_1015); +lean_dec(x_1015); +lean_inc(x_787); +x_1017 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1017, 0, x_787); +lean_ctor_set(x_1017, 1, x_804); +lean_ctor_set(x_1017, 2, x_1016); +lean_inc(x_787); +x_1018 = l_Lean_Syntax_node1(x_787, x_957, x_1017); +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_1019 = l_Lean_Syntax_node6(x_787, x_959, x_945, x_815, x_815, x_1011, x_950, x_1018); +x_1020 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +lean_inc(x_787); +x_1021 = l_Lean_Syntax_node3(x_787, x_1020, x_1007, x_1009, x_1019); +lean_inc(x_787); +x_1022 = l_Lean_Syntax_node1(x_787, x_996, x_1021); +lean_inc(x_815); +lean_inc(x_787); +x_1023 = l_Lean_Syntax_node2(x_787, x_978, x_1022, x_815); +lean_inc(x_787); +x_1024 = l_Lean_Syntax_node3(x_787, x_804, x_979, x_1002, x_1023); +x_1025 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_787); +x_1026 = l_Lean_Syntax_node1(x_787, x_1025, x_1024); +x_1027 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_787); +x_1028 = l_Lean_Syntax_node2(x_787, x_1027, x_975, x_1026); +lean_inc(x_888); +lean_inc(x_815); +lean_inc(x_787); +x_1029 = l_Lean_Syntax_node5(x_787, x_963, x_913, x_973, x_815, x_888, x_1028); +lean_inc(x_787); +x_1030 = l_Lean_Syntax_node1(x_787, x_965, x_1029); +lean_inc(x_925); +lean_inc_n(x_815, 2); +lean_inc(x_787); +x_1031 = l_Lean_Syntax_node4(x_787, x_967, x_815, x_815, x_1030, x_925); +lean_inc(x_815); +lean_inc(x_787); +x_1032 = l_Lean_Syntax_node3(x_787, x_804, x_968, x_815, x_1031); +x_1033 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_787); +x_1034 = l_Lean_Syntax_node2(x_787, x_1033, x_19, x_1032); +lean_inc(x_787); +x_1035 = l_Lean_Syntax_node1(x_787, x_804, x_1034); +x_1036 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_787); +x_1037 = l_Lean_Syntax_node4(x_787, x_1036, x_888, x_923, x_925, x_1035); +x_1038 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_815); +lean_inc(x_787); +x_1039 = l_Lean_Syntax_node6(x_787, x_1038, x_871, x_873, x_815, x_815, x_886, x_1037); +lean_inc(x_787); +x_1040 = l_Lean_Syntax_node2(x_787, x_854, x_869, x_1039); +x_1041 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_787); +x_1042 = l_Lean_Syntax_node3(x_787, x_1041, x_861, x_863, x_1040); +x_1043 = l_Lean_Syntax_node2(x_787, x_804, x_855, x_1042); +if (lean_is_scalar(x_812)) { + x_1044 = lean_alloc_ctor(0, 2, 0); } else { - x_1036 = x_806; + x_1044 = x_812; } -lean_ctor_set(x_1036, 0, x_1035); -lean_ctor_set(x_1036, 1, x_805); -return x_1036; +lean_ctor_set(x_1044, 0, x_1043); +lean_ctor_set(x_1044, 1, x_811); +return x_1044; } else { -lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; -lean_dec(x_774); -lean_dec(x_773); +lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; +lean_dec(x_780); +lean_dec(x_779); lean_free_object(x_19); lean_dec(x_21); lean_dec(x_10); lean_dec(x_9); lean_dec(x_1); -x_1037 = lean_ctor_get(x_776, 0); -lean_inc(x_1037); -x_1038 = lean_ctor_get(x_776, 1); -lean_inc(x_1038); -if (lean_is_exclusive(x_776)) { - lean_ctor_release(x_776, 0); - lean_ctor_release(x_776, 1); - x_1039 = x_776; +x_1045 = lean_ctor_get(x_782, 0); +lean_inc(x_1045); +x_1046 = lean_ctor_get(x_782, 1); +lean_inc(x_1046); +if (lean_is_exclusive(x_782)) { + lean_ctor_release(x_782, 0); + lean_ctor_release(x_782, 1); + x_1047 = x_782; } else { - lean_dec_ref(x_776); - x_1039 = lean_box(0); + lean_dec_ref(x_782); + x_1047 = lean_box(0); } -if (lean_is_scalar(x_1039)) { - x_1040 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1047)) { + x_1048 = lean_alloc_ctor(1, 2, 0); } else { - x_1040 = x_1039; + x_1048 = x_1047; } -lean_ctor_set(x_1040, 0, x_1037); -lean_ctor_set(x_1040, 1, x_1038); -return x_1040; +lean_ctor_set(x_1048, 0, x_1045); +lean_ctor_set(x_1048, 1, x_1046); +return x_1048; } } } else { -lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; size_t x_1047; lean_object* x_1048; -x_1041 = lean_ctor_get(x_19, 0); -x_1042 = lean_ctor_get(x_19, 1); -lean_inc(x_1042); -lean_inc(x_1041); +lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; size_t x_1055; lean_object* x_1056; +x_1049 = lean_ctor_get(x_19, 0); +x_1050 = lean_ctor_get(x_19, 1); +lean_inc(x_1050); +lean_inc(x_1049); lean_dec(x_19); -x_1043 = l_Array_unzip___rarg(x_1042); -lean_dec(x_1042); -x_1044 = lean_ctor_get(x_1043, 0); -lean_inc(x_1044); -x_1045 = lean_ctor_get(x_1043, 1); -lean_inc(x_1045); -if (lean_is_exclusive(x_1043)) { - lean_ctor_release(x_1043, 0); - lean_ctor_release(x_1043, 1); - x_1046 = x_1043; +x_1051 = l_Array_unzip___rarg(x_1050); +lean_dec(x_1050); +x_1052 = lean_ctor_get(x_1051, 0); +lean_inc(x_1052); +x_1053 = lean_ctor_get(x_1051, 1); +lean_inc(x_1053); +if (lean_is_exclusive(x_1051)) { + lean_ctor_release(x_1051, 0); + lean_ctor_release(x_1051, 1); + x_1054 = x_1051; } else { - lean_dec_ref(x_1043); - x_1046 = lean_box(0); + lean_dec_ref(x_1051); + x_1054 = lean_box(0); } -x_1047 = lean_array_size(x_2); -x_1048 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_1047, x_15, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +x_1055 = lean_array_size(x_2); +x_1056 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(x_1055, x_15, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_18); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_1048) == 0) -{ -lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; uint8_t x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; size_t x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; size_t x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; size_t x_1216; lean_object* x_1217; size_t x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; size_t x_1277; lean_object* x_1278; size_t x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; -x_1049 = lean_ctor_get(x_1048, 0); -lean_inc(x_1049); -x_1050 = lean_ctor_get(x_1048, 1); -lean_inc(x_1050); -lean_dec(x_1048); -x_1051 = lean_ctor_get(x_9, 5); -lean_inc(x_1051); -x_1052 = 0; -x_1053 = l_Lean_SourceInfo_fromRef(x_1051, x_1052); -lean_dec(x_1051); -x_1054 = lean_st_ref_get(x_10, x_1050); -x_1055 = lean_ctor_get(x_1054, 1); -lean_inc(x_1055); -if (lean_is_exclusive(x_1054)) { - lean_ctor_release(x_1054, 0); - lean_ctor_release(x_1054, 1); - x_1056 = x_1054; +if (lean_obj_tag(x_1056) == 0) +{ +lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; uint8_t x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; size_t x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; size_t x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; size_t x_1226; lean_object* x_1227; size_t x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; size_t x_1287; lean_object* x_1288; size_t x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; +x_1057 = lean_ctor_get(x_1056, 0); +lean_inc(x_1057); +x_1058 = lean_ctor_get(x_1056, 1); +lean_inc(x_1058); +lean_dec(x_1056); +x_1059 = lean_ctor_get(x_9, 5); +lean_inc(x_1059); +x_1060 = 0; +x_1061 = l_Lean_SourceInfo_fromRef(x_1059, x_1060); +lean_dec(x_1059); +x_1062 = lean_st_ref_get(x_10, x_1058); +x_1063 = lean_ctor_get(x_1062, 1); +lean_inc(x_1063); +if (lean_is_exclusive(x_1062)) { + lean_ctor_release(x_1062, 0); + lean_ctor_release(x_1062, 1); + x_1064 = x_1062; } else { - lean_dec_ref(x_1054); - x_1056 = lean_box(0); + lean_dec_ref(x_1062); + x_1064 = lean_box(0); } -x_1057 = lean_box(0); -x_1058 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; -lean_inc(x_1053); -if (lean_is_scalar(x_1056)) { - x_1059 = lean_alloc_ctor(2, 2, 0); +x_1065 = lean_box(0); +x_1066 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__60; +lean_inc(x_1061); +if (lean_is_scalar(x_1064)) { + x_1067 = lean_alloc_ctor(2, 2, 0); } else { - x_1059 = x_1056; - lean_ctor_set_tag(x_1059, 2); + x_1067 = x_1064; + lean_ctor_set_tag(x_1067, 2); } -lean_ctor_set(x_1059, 0, x_1053); -lean_ctor_set(x_1059, 1, x_1058); -x_1060 = lean_ctor_get(x_1, 0); -lean_inc(x_1060); +lean_ctor_set(x_1067, 0, x_1061); +lean_ctor_set(x_1067, 1, x_1066); +x_1068 = lean_ctor_get(x_1, 0); +lean_inc(x_1068); lean_dec(x_1); -x_1061 = lean_ctor_get(x_1060, 0); +x_1069 = lean_ctor_get(x_1068, 0); +lean_inc(x_1069); +lean_dec(x_1068); +x_1070 = lean_mk_syntax_ident(x_1069); +x_1071 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; lean_inc(x_1061); -lean_dec(x_1060); -x_1062 = lean_mk_syntax_ident(x_1061); -x_1063 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__59; -lean_inc(x_1053); -x_1064 = l_Lean_Syntax_node2(x_1053, x_1063, x_1059, x_1062); -x_1065 = lean_array_size(x_1049); -x_1066 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; -x_1067 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_1066, x_1065, x_15, x_1049); -x_1068 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; -x_1069 = l_Array_append___rarg(x_1068, x_1067); -lean_dec(x_1067); -x_1070 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; -lean_inc(x_1053); -x_1071 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1071, 0, x_1053); -lean_ctor_set(x_1071, 1, x_1070); -lean_ctor_set(x_1071, 2, x_1069); -x_1072 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; -lean_inc(x_1053); -x_1073 = l_Lean_Syntax_node2(x_1053, x_1072, x_1064, x_1071); -x_1074 = lean_ctor_get(x_9, 10); -lean_inc(x_1074); +x_1072 = l_Lean_Syntax_node2(x_1061, x_1071, x_1067, x_1070); +x_1073 = lean_array_size(x_1057); +x_1074 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; +x_1075 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_1074, x_1073, x_15, x_1057); +x_1076 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; +x_1077 = l_Array_append___rarg(x_1076, x_1075); +lean_dec(x_1075); +x_1078 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; +lean_inc(x_1061); +x_1079 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1079, 0, x_1061); +lean_ctor_set(x_1079, 1, x_1078); +lean_ctor_set(x_1079, 2, x_1077); +x_1080 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; +lean_inc(x_1061); +x_1081 = l_Lean_Syntax_node2(x_1061, x_1080, x_1072, x_1079); +x_1082 = lean_ctor_get(x_9, 10); +lean_inc(x_1082); lean_dec(x_9); -x_1075 = lean_st_ref_get(x_10, x_1055); +x_1083 = lean_st_ref_get(x_10, x_1063); lean_dec(x_10); -x_1076 = lean_ctor_get(x_1075, 0); -lean_inc(x_1076); -x_1077 = lean_ctor_get(x_1075, 1); -lean_inc(x_1077); -if (lean_is_exclusive(x_1075)) { - lean_ctor_release(x_1075, 0); - lean_ctor_release(x_1075, 1); - x_1078 = x_1075; +x_1084 = lean_ctor_get(x_1083, 0); +lean_inc(x_1084); +x_1085 = lean_ctor_get(x_1083, 1); +lean_inc(x_1085); +if (lean_is_exclusive(x_1083)) { + lean_ctor_release(x_1083, 0); + lean_ctor_release(x_1083, 1); + x_1086 = x_1083; } else { - lean_dec_ref(x_1075); - x_1078 = lean_box(0); + lean_dec_ref(x_1083); + x_1086 = lean_box(0); } -x_1079 = lean_ctor_get(x_1076, 0); -lean_inc(x_1079); -lean_dec(x_1076); -x_1080 = lean_environment_main_module(x_1079); -lean_inc(x_1053); -x_1081 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1081, 0, x_1053); -lean_ctor_set(x_1081, 1, x_1070); -lean_ctor_set(x_1081, 2, x_1068); -x_1082 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; -lean_inc_n(x_1081, 6); -lean_inc(x_1053); -x_1083 = l_Lean_Syntax_node6(x_1053, x_1082, x_1081, x_1081, x_1081, x_1081, x_1081, x_1081); -x_1084 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; -lean_inc(x_1053); -if (lean_is_scalar(x_1046)) { - x_1085 = lean_alloc_ctor(2, 2, 0); +x_1087 = lean_ctor_get(x_1084, 0); +lean_inc(x_1087); +lean_dec(x_1084); +x_1088 = lean_environment_main_module(x_1087); +lean_inc(x_1061); +x_1089 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1089, 0, x_1061); +lean_ctor_set(x_1089, 1, x_1078); +lean_ctor_set(x_1089, 2, x_1076); +x_1090 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__8; +lean_inc_n(x_1089, 6); +lean_inc(x_1061); +x_1091 = l_Lean_Syntax_node6(x_1061, x_1090, x_1089, x_1089, x_1089, x_1089, x_1089, x_1089); +x_1092 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; +lean_inc(x_1061); +if (lean_is_scalar(x_1054)) { + x_1093 = lean_alloc_ctor(2, 2, 0); } else { - x_1085 = x_1046; - lean_ctor_set_tag(x_1085, 2); + x_1093 = x_1054; + lean_ctor_set_tag(x_1093, 2); } -lean_ctor_set(x_1085, 0, x_1053); -lean_ctor_set(x_1085, 1, x_1084); -x_1086 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_1074); -lean_inc(x_1080); -x_1087 = l_Lean_addMacroScope(x_1080, x_1086, x_1074); -x_1088 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; -lean_inc(x_1053); -x_1089 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1089, 0, x_1053); -lean_ctor_set(x_1089, 1, x_1088); -lean_ctor_set(x_1089, 2, x_1087); -lean_ctor_set(x_1089, 3, x_1057); -x_1090 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; -lean_inc(x_1081); +lean_ctor_set(x_1093, 0, x_1061); +lean_ctor_set(x_1093, 1, x_1092); +x_1094 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_1082); +lean_inc(x_1088); +x_1095 = l_Lean_addMacroScope(x_1088, x_1094, x_1082); +x_1096 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +lean_inc(x_1061); +x_1097 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1097, 0, x_1061); +lean_ctor_set(x_1097, 1, x_1096); +lean_ctor_set(x_1097, 2, x_1095); +lean_ctor_set(x_1097, 3, x_1065); +x_1098 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__14; lean_inc(x_1089); -lean_inc(x_1053); -x_1091 = l_Lean_Syntax_node2(x_1053, x_1090, x_1089, x_1081); -x_1092 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1093 = l_Lean_Syntax_node2(x_1053, x_1092, x_1081, x_1081); -x_1094 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; -lean_inc(x_1053); -x_1095 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1095, 0, x_1053); -lean_ctor_set(x_1095, 1, x_1094); -lean_inc(x_1095); -lean_inc(x_1053); -x_1096 = l_Lean_Syntax_node1(x_1053, x_1070, x_1095); -x_1097 = lean_array_size(x_1041); -x_1098 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_1097, x_15, x_1041); -x_1099 = l_Array_append___rarg(x_1068, x_1098); -lean_dec(x_1098); -lean_inc(x_1053); -x_1100 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1100, 0, x_1053); -lean_ctor_set(x_1100, 1, x_1070); -lean_ctor_set(x_1100, 2, x_1099); -x_1101 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; -lean_inc(x_1053); -x_1102 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1102, 0, x_1053); -lean_ctor_set(x_1102, 1, x_1101); -x_1103 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -lean_inc(x_1074); -lean_inc(x_1080); -x_1104 = l_Lean_addMacroScope(x_1080, x_1103, x_1074); -x_1105 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; -x_1106 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; -lean_inc(x_1053); -x_1107 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1107, 0, x_1053); -lean_ctor_set(x_1107, 1, x_1105); -lean_ctor_set(x_1107, 2, x_1104); -lean_ctor_set(x_1107, 3, x_1106); -x_1108 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_1053); -x_1109 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1109, 0, x_1053); -lean_ctor_set(x_1109, 1, x_1108); -x_1110 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -lean_inc(x_1074); -lean_inc(x_1080); -x_1111 = l_Lean_addMacroScope(x_1080, x_1110, x_1074); -x_1112 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; -x_1113 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -lean_inc(x_1053); -x_1114 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1114, 0, x_1053); -lean_ctor_set(x_1114, 1, x_1112); -lean_ctor_set(x_1114, 2, x_1111); -lean_ctor_set(x_1114, 3, x_1113); -lean_inc(x_1109); -lean_inc(x_1053); -x_1115 = l_Lean_Syntax_node3(x_1053, x_1070, x_1107, x_1109, x_1114); -lean_inc(x_1053); -x_1116 = l_Lean_Syntax_node2(x_1053, x_1070, x_1102, x_1115); -x_1117 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_1053); -x_1118 = l_Lean_Syntax_node1(x_1053, x_1117, x_1116); -x_1119 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; -lean_inc(x_1081); -lean_inc(x_1053); -x_1120 = l_Lean_Syntax_node7(x_1053, x_1119, x_1085, x_1091, x_1093, x_1096, x_1100, x_1081, x_1118); -x_1121 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; -lean_inc(x_1053); -x_1122 = l_Lean_Syntax_node2(x_1053, x_1121, x_1083, x_1120); -x_1123 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; -lean_inc(x_1053); -x_1124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1124, 0, x_1053); -lean_ctor_set(x_1124, 1, x_1123); -x_1125 = l_Array_append___rarg(x_1068, x_3); -lean_inc(x_1053); -x_1126 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1126, 0, x_1053); -lean_ctor_set(x_1126, 1, x_1070); -lean_ctor_set(x_1126, 2, x_1125); -x_1127 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; -lean_inc(x_1053); -x_1128 = l_Lean_Syntax_node2(x_1053, x_1127, x_1124, x_1126); -x_1129 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; -lean_inc(x_1053); -x_1130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1130, 0, x_1053); -lean_ctor_set(x_1130, 1, x_1129); -x_1131 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; -lean_inc(x_1053); +lean_inc(x_1097); +lean_inc(x_1061); +x_1099 = l_Lean_Syntax_node2(x_1061, x_1098, x_1097, x_1089); +x_1100 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1101 = l_Lean_Syntax_node2(x_1061, x_1100, x_1089, x_1089); +x_1102 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; +lean_inc(x_1061); +x_1103 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1103, 0, x_1061); +lean_ctor_set(x_1103, 1, x_1102); +lean_inc(x_1103); +lean_inc(x_1061); +x_1104 = l_Lean_Syntax_node1(x_1061, x_1078, x_1103); +x_1105 = lean_array_size(x_1049); +x_1106 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_1105, x_15, x_1049); +x_1107 = l_Array_append___rarg(x_1076, x_1106); +lean_dec(x_1106); +lean_inc(x_1061); +x_1108 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1108, 0, x_1061); +lean_ctor_set(x_1108, 1, x_1078); +lean_ctor_set(x_1108, 2, x_1107); +x_1109 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__23; +lean_inc(x_1061); +x_1110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1110, 0, x_1061); +lean_ctor_set(x_1110, 1, x_1109); +x_1111 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +lean_inc(x_1082); +lean_inc(x_1088); +x_1112 = l_Lean_addMacroScope(x_1088, x_1111, x_1082); +x_1113 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__25; +x_1114 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__31; +lean_inc(x_1061); +x_1115 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1115, 0, x_1061); +lean_ctor_set(x_1115, 1, x_1113); +lean_ctor_set(x_1115, 2, x_1112); +lean_ctor_set(x_1115, 3, x_1114); +x_1116 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_1061); +x_1117 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1117, 0, x_1061); +lean_ctor_set(x_1117, 1, x_1116); +x_1118 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +lean_inc(x_1082); +lean_inc(x_1088); +x_1119 = l_Lean_addMacroScope(x_1088, x_1118, x_1082); +x_1120 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +x_1121 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; +lean_inc(x_1061); +x_1122 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1122, 0, x_1061); +lean_ctor_set(x_1122, 1, x_1120); +lean_ctor_set(x_1122, 2, x_1119); +lean_ctor_set(x_1122, 3, x_1121); +lean_inc(x_1117); +lean_inc(x_1061); +x_1123 = l_Lean_Syntax_node3(x_1061, x_1078, x_1115, x_1117, x_1122); +lean_inc(x_1061); +x_1124 = l_Lean_Syntax_node2(x_1061, x_1078, x_1110, x_1123); +x_1125 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_1061); +x_1126 = l_Lean_Syntax_node1(x_1061, x_1125, x_1124); +x_1127 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; +lean_inc(x_1089); +lean_inc(x_1061); +x_1128 = l_Lean_Syntax_node7(x_1061, x_1127, x_1093, x_1099, x_1101, x_1104, x_1108, x_1089, x_1126); +x_1129 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; +lean_inc(x_1061); +x_1130 = l_Lean_Syntax_node2(x_1061, x_1129, x_1091, x_1128); +x_1131 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_1061); x_1132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1132, 0, x_1053); +lean_ctor_set(x_1132, 0, x_1061); lean_ctor_set(x_1132, 1, x_1131); -x_1133 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; -lean_inc(x_1053); -x_1134 = l_Lean_Syntax_node1(x_1053, x_1133, x_1132); -lean_inc(x_1053); -x_1135 = l_Lean_Syntax_node1(x_1053, x_1070, x_1134); -lean_inc_n(x_1081, 5); -lean_inc(x_1053); -x_1136 = l_Lean_Syntax_node6(x_1053, x_1082, x_1081, x_1081, x_1081, x_1081, x_1081, x_1135); -x_1137 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; -lean_inc(x_1081); -lean_inc(x_1053); -x_1138 = l_Lean_Syntax_node1(x_1053, x_1137, x_1081); -x_1139 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; -lean_inc(x_1053); +x_1133 = l_Array_append___rarg(x_1076, x_3); +lean_inc(x_1061); +x_1134 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1134, 0, x_1061); +lean_ctor_set(x_1134, 1, x_1078); +lean_ctor_set(x_1134, 2, x_1133); +x_1135 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +lean_inc(x_1061); +x_1136 = l_Lean_Syntax_node2(x_1061, x_1135, x_1132, x_1134); +x_1137 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_1061); +x_1138 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1138, 0, x_1061); +lean_ctor_set(x_1138, 1, x_1137); +x_1139 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; +lean_inc(x_1061); x_1140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1140, 0, x_1053); +lean_ctor_set(x_1140, 0, x_1061); lean_ctor_set(x_1140, 1, x_1139); -x_1141 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; -lean_inc(x_1053); -x_1142 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1142, 0, x_1053); -lean_ctor_set(x_1142, 1, x_1141); -x_1143 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -lean_inc(x_1074); -lean_inc(x_1080); -x_1144 = l_Lean_addMacroScope(x_1080, x_1143, x_1074); -x_1145 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -x_1146 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; -lean_inc(x_1053); -x_1147 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1147, 0, x_1053); -lean_ctor_set(x_1147, 1, x_1145); -lean_ctor_set(x_1147, 2, x_1144); -lean_ctor_set(x_1147, 3, x_1146); -lean_inc(x_1053); -x_1148 = l_Lean_Syntax_node1(x_1053, x_1070, x_1073); -lean_inc(x_1053); -x_1149 = l_Lean_Syntax_node2(x_1053, x_1072, x_1147, x_1148); -x_1150 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_1142); -lean_inc(x_1053); -x_1151 = l_Lean_Syntax_node2(x_1053, x_1150, x_1142, x_1149); -x_1152 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; -lean_inc(x_1151); -lean_inc(x_1081); -lean_inc(x_1053); -x_1153 = l_Lean_Syntax_node2(x_1053, x_1152, x_1081, x_1151); -x_1154 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_1053); -x_1155 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1155, 0, x_1053); -lean_ctor_set(x_1155, 1, x_1154); -x_1156 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc(x_1053); -x_1157 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1157, 0, x_1053); -lean_ctor_set(x_1157, 1, x_1156); -x_1158 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; -lean_inc(x_1074); -lean_inc(x_1080); -x_1159 = l_Lean_addMacroScope(x_1080, x_1158, x_1074); -x_1160 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_1161 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_1053); -x_1162 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1162, 0, x_1053); -lean_ctor_set(x_1162, 1, x_1160); -lean_ctor_set(x_1162, 2, x_1159); -lean_ctor_set(x_1162, 3, x_1161); -x_1163 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_1081); -lean_inc(x_1053); -x_1164 = l_Lean_Syntax_node2(x_1053, x_1163, x_1162, x_1081); -x_1165 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; -lean_inc(x_1074); -lean_inc(x_1080); -x_1166 = l_Lean_addMacroScope(x_1080, x_1165, x_1074); -x_1167 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; -lean_inc(x_1053); -x_1168 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1168, 0, x_1053); -lean_ctor_set(x_1168, 1, x_1167); -lean_ctor_set(x_1168, 2, x_1166); -lean_ctor_set(x_1168, 3, x_1057); -x_1169 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_1168); -lean_inc(x_1155); -lean_inc(x_1053); -x_1170 = l_Lean_Syntax_node3(x_1053, x_1169, x_1164, x_1155, x_1168); -x_1171 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; -lean_inc(x_1074); -lean_inc(x_1080); -x_1172 = l_Lean_addMacroScope(x_1080, x_1171, x_1074); -x_1173 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_1174 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_1053); -x_1175 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1175, 0, x_1053); -lean_ctor_set(x_1175, 1, x_1173); -lean_ctor_set(x_1175, 2, x_1172); -lean_ctor_set(x_1175, 3, x_1174); -lean_inc(x_1081); -lean_inc(x_1053); -x_1176 = l_Lean_Syntax_node2(x_1053, x_1163, x_1175, x_1081); -x_1177 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_1074); -lean_inc(x_1080); -x_1178 = l_Lean_addMacroScope(x_1080, x_1177, x_1074); -x_1179 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; -lean_inc(x_1053); -x_1180 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1180, 0, x_1053); -lean_ctor_set(x_1180, 1, x_1179); -lean_ctor_set(x_1180, 2, x_1178); -lean_ctor_set(x_1180, 3, x_1057); -lean_inc(x_1180); -lean_inc(x_1155); -lean_inc(x_1053); -x_1181 = l_Lean_Syntax_node3(x_1053, x_1169, x_1176, x_1155, x_1180); -lean_inc(x_1053); -x_1182 = l_Lean_Syntax_node3(x_1053, x_1070, x_1170, x_1109, x_1181); -x_1183 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; -lean_inc(x_1081); -lean_inc(x_1053); -x_1184 = l_Lean_Syntax_node1(x_1053, x_1183, x_1081); -x_1185 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_1053); -x_1186 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1186, 0, x_1053); -lean_ctor_set(x_1186, 1, x_1185); -x_1187 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1188 = l_Lean_Syntax_node6(x_1053, x_1187, x_1157, x_1081, x_1182, x_1184, x_1081, x_1186); -x_1189 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1190 = l_Lean_Syntax_node2(x_1053, x_1189, x_1081, x_1081); -x_1191 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; -lean_inc(x_1074); -lean_inc(x_1080); -x_1192 = l_Lean_addMacroScope(x_1080, x_1191, x_1074); -x_1193 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; -lean_inc(x_1053); -x_1194 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1194, 0, x_1053); -lean_ctor_set(x_1194, 1, x_1193); -lean_ctor_set(x_1194, 2, x_1192); -lean_ctor_set(x_1194, 3, x_1057); -lean_inc(x_1194); -lean_inc(x_1053); -x_1195 = l_Lean_Syntax_node1(x_1053, x_1070, x_1194); -x_1196 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; -lean_inc(x_1053); -x_1197 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1197, 0, x_1053); -lean_ctor_set(x_1197, 1, x_1196); -x_1198 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; -lean_inc(x_1074); -lean_inc(x_1080); -x_1199 = l_Lean_addMacroScope(x_1080, x_1198, x_1074); -x_1200 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; -lean_inc(x_1053); -x_1201 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1201, 0, x_1053); -lean_ctor_set(x_1201, 1, x_1200); -lean_ctor_set(x_1201, 2, x_1199); -lean_ctor_set(x_1201, 3, x_1057); -x_1202 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; -lean_inc(x_1053); -x_1203 = l_Lean_Syntax_node1(x_1053, x_1202, x_1201); -lean_inc(x_1053); -x_1204 = l_Lean_Syntax_node1(x_1053, x_1070, x_1151); -x_1205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +x_1141 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; +lean_inc(x_1061); +x_1142 = l_Lean_Syntax_node1(x_1061, x_1141, x_1140); +lean_inc(x_1061); +x_1143 = l_Lean_Syntax_node1(x_1061, x_1078, x_1142); +lean_inc_n(x_1089, 5); +lean_inc(x_1061); +x_1144 = l_Lean_Syntax_node6(x_1061, x_1090, x_1089, x_1089, x_1089, x_1089, x_1089, x_1143); +x_1145 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__48; +lean_inc(x_1089); +lean_inc(x_1061); +x_1146 = l_Lean_Syntax_node1(x_1061, x_1145, x_1089); +x_1147 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +lean_inc(x_1061); +x_1148 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1148, 0, x_1061); +lean_ctor_set(x_1148, 1, x_1147); +x_1149 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +lean_inc(x_1061); +x_1150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1150, 0, x_1061); +lean_ctor_set(x_1150, 1, x_1149); +x_1151 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +lean_inc(x_1082); +lean_inc(x_1088); +x_1152 = l_Lean_addMacroScope(x_1088, x_1151, x_1082); +x_1153 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +x_1154 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; +lean_inc(x_1061); +x_1155 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1155, 0, x_1061); +lean_ctor_set(x_1155, 1, x_1153); +lean_ctor_set(x_1155, 2, x_1152); +lean_ctor_set(x_1155, 3, x_1154); +lean_inc(x_1061); +x_1156 = l_Lean_Syntax_node1(x_1061, x_1078, x_1081); +lean_inc(x_1061); +x_1157 = l_Lean_Syntax_node2(x_1061, x_1080, x_1155, x_1156); +x_1158 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_1150); +lean_inc(x_1061); +x_1159 = l_Lean_Syntax_node2(x_1061, x_1158, x_1150, x_1157); +x_1160 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; +lean_inc(x_1159); +lean_inc(x_1089); +lean_inc(x_1061); +x_1161 = l_Lean_Syntax_node2(x_1061, x_1160, x_1089, x_1159); +x_1162 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_1061); +x_1163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1163, 0, x_1061); +lean_ctor_set(x_1163, 1, x_1162); +x_1164 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc(x_1061); +x_1165 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1165, 0, x_1061); +lean_ctor_set(x_1165, 1, x_1164); +x_1166 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +lean_inc(x_1082); +lean_inc(x_1088); +x_1167 = l_Lean_addMacroScope(x_1088, x_1166, x_1082); +x_1168 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_1169 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_1061); +x_1170 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1170, 0, x_1061); +lean_ctor_set(x_1170, 1, x_1168); +lean_ctor_set(x_1170, 2, x_1167); +lean_ctor_set(x_1170, 3, x_1169); +x_1171 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; +lean_inc(x_1089); +lean_inc(x_1061); +x_1172 = l_Lean_Syntax_node2(x_1061, x_1171, x_1170, x_1089); +x_1173 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__73; +lean_inc(x_1082); +lean_inc(x_1088); +x_1174 = l_Lean_addMacroScope(x_1088, x_1173, x_1082); +x_1175 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_1061); +x_1176 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1176, 0, x_1061); +lean_ctor_set(x_1176, 1, x_1175); +lean_ctor_set(x_1176, 2, x_1174); +lean_ctor_set(x_1176, 3, x_1065); +x_1177 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_1176); +lean_inc(x_1163); +lean_inc(x_1061); +x_1178 = l_Lean_Syntax_node3(x_1061, x_1177, x_1172, x_1163, x_1176); +x_1179 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +lean_inc(x_1082); +lean_inc(x_1088); +x_1180 = l_Lean_addMacroScope(x_1088, x_1179, x_1082); +x_1181 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_1182 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_1061); +x_1183 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1183, 0, x_1061); +lean_ctor_set(x_1183, 1, x_1181); +lean_ctor_set(x_1183, 2, x_1180); +lean_ctor_set(x_1183, 3, x_1182); +lean_inc(x_1089); +lean_inc(x_1061); +x_1184 = l_Lean_Syntax_node2(x_1061, x_1171, x_1183, x_1089); +x_1185 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__76; +lean_inc(x_1082); +lean_inc(x_1088); +x_1186 = l_Lean_addMacroScope(x_1088, x_1185, x_1082); +x_1187 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_1061); +x_1188 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1188, 0, x_1061); +lean_ctor_set(x_1188, 1, x_1187); +lean_ctor_set(x_1188, 2, x_1186); +lean_ctor_set(x_1188, 3, x_1065); lean_inc(x_1188); -lean_inc(x_1155); -lean_inc(x_1081); -lean_inc(x_1053); -x_1206 = l_Lean_Syntax_node5(x_1053, x_1205, x_1203, x_1081, x_1204, x_1155, x_1188); -x_1207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; -lean_inc(x_1053); -x_1208 = l_Lean_Syntax_node1(x_1053, x_1207, x_1206); -x_1209 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; -lean_inc(x_1053); -x_1210 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1210, 0, x_1053); -lean_ctor_set(x_1210, 1, x_1209); -x_1211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; -lean_inc(x_1081); -lean_inc(x_1053); -x_1212 = l_Lean_Syntax_node2(x_1053, x_1211, x_1081, x_1194); -lean_inc(x_1053); -x_1213 = l_Lean_Syntax_node1(x_1053, x_1070, x_1212); -x_1214 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_1053); -x_1215 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1215, 0, x_1053); -lean_ctor_set(x_1215, 1, x_1214); -x_1216 = lean_array_size(x_1044); -x_1217 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_1057, x_1216, x_15, x_1044); -x_1218 = lean_array_size(x_1217); -x_1219 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1218, x_15, x_1217); -x_1220 = l_Array_append___rarg(x_1068, x_1219); -lean_dec(x_1219); -lean_inc(x_1053); -x_1221 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1221, 0, x_1053); -lean_ctor_set(x_1221, 1, x_1070); -lean_ctor_set(x_1221, 2, x_1220); -x_1222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; -lean_inc(x_1053); -x_1223 = l_Lean_Syntax_node1(x_1053, x_1222, x_1221); -x_1224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_1215); -lean_inc_n(x_1081, 2); -lean_inc(x_1210); -lean_inc(x_1053); -x_1225 = l_Lean_Syntax_node6(x_1053, x_1224, x_1210, x_1081, x_1081, x_1213, x_1215, x_1223); -x_1226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; -lean_inc(x_1081); -lean_inc(x_1208); -lean_inc(x_1197); -lean_inc(x_1053); -x_1227 = l_Lean_Syntax_node4(x_1053, x_1226, x_1197, x_1208, x_1081, x_1225); -x_1228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; -lean_inc(x_1155); -lean_inc(x_1081); -lean_inc(x_1053); -x_1229 = l_Lean_Syntax_node5(x_1053, x_1228, x_1168, x_1195, x_1081, x_1155, x_1227); -x_1230 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_1053); -x_1231 = l_Lean_Syntax_node1(x_1053, x_1230, x_1229); -x_1232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; -lean_inc(x_1190); -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1233 = l_Lean_Syntax_node4(x_1053, x_1232, x_1081, x_1081, x_1231, x_1190); -x_1234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -lean_inc(x_1074); -lean_inc(x_1080); -x_1235 = l_Lean_addMacroScope(x_1080, x_1234, x_1074); -x_1236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; -lean_inc(x_1053); -x_1237 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1237, 0, x_1053); -lean_ctor_set(x_1237, 1, x_1236); -lean_ctor_set(x_1237, 2, x_1235); -lean_ctor_set(x_1237, 3, x_1057); -lean_inc(x_1053); -x_1238 = l_Lean_Syntax_node1(x_1053, x_1070, x_1237); -x_1239 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; -lean_inc(x_1053); -x_1240 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1240, 0, x_1053); -lean_ctor_set(x_1240, 1, x_1239); -x_1241 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; -lean_inc(x_1053); -x_1242 = l_Lean_Syntax_node2(x_1053, x_1241, x_1197, x_1208); -x_1243 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_1081); -lean_inc(x_1053); -x_1244 = l_Lean_Syntax_node2(x_1053, x_1243, x_1242, x_1081); -x_1245 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_1053); -x_1246 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1246, 0, x_1053); -lean_ctor_set(x_1246, 1, x_1245); -x_1247 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; -lean_inc(x_1074); -lean_inc(x_1080); -x_1248 = l_Lean_addMacroScope(x_1080, x_1247, x_1074); -x_1249 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; -lean_inc(x_1053); -x_1250 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1250, 0, x_1053); +lean_inc(x_1163); +lean_inc(x_1061); +x_1189 = l_Lean_Syntax_node3(x_1061, x_1177, x_1184, x_1163, x_1188); +lean_inc(x_1061); +x_1190 = l_Lean_Syntax_node3(x_1061, x_1078, x_1178, x_1117, x_1189); +x_1191 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__70; +lean_inc(x_1061); +x_1192 = l_Lean_Syntax_node1(x_1061, x_1191, x_1190); +x_1193 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_1089); +lean_inc(x_1061); +x_1194 = l_Lean_Syntax_node1(x_1061, x_1193, x_1089); +x_1195 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__79; +lean_inc(x_1061); +x_1196 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1196, 0, x_1061); +lean_ctor_set(x_1196, 1, x_1195); +x_1197 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__67; +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1198 = l_Lean_Syntax_node6(x_1061, x_1197, x_1165, x_1089, x_1192, x_1194, x_1089, x_1196); +x_1199 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1200 = l_Lean_Syntax_node2(x_1061, x_1199, x_1089, x_1089); +x_1201 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; +lean_inc(x_1082); +lean_inc(x_1088); +x_1202 = l_Lean_addMacroScope(x_1088, x_1201, x_1082); +x_1203 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_1061); +x_1204 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1204, 0, x_1061); +lean_ctor_set(x_1204, 1, x_1203); +lean_ctor_set(x_1204, 2, x_1202); +lean_ctor_set(x_1204, 3, x_1065); +lean_inc(x_1204); +lean_inc(x_1061); +x_1205 = l_Lean_Syntax_node1(x_1061, x_1078, x_1204); +x_1206 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; +lean_inc(x_1061); +x_1207 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1207, 0, x_1061); +lean_ctor_set(x_1207, 1, x_1206); +x_1208 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +lean_inc(x_1082); +lean_inc(x_1088); +x_1209 = l_Lean_addMacroScope(x_1088, x_1208, x_1082); +x_1210 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +lean_inc(x_1061); +x_1211 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1211, 0, x_1061); +lean_ctor_set(x_1211, 1, x_1210); +lean_ctor_set(x_1211, 2, x_1209); +lean_ctor_set(x_1211, 3, x_1065); +x_1212 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; +lean_inc(x_1061); +x_1213 = l_Lean_Syntax_node1(x_1061, x_1212, x_1211); +lean_inc(x_1061); +x_1214 = l_Lean_Syntax_node1(x_1061, x_1078, x_1159); +x_1215 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +lean_inc(x_1198); +lean_inc(x_1163); +lean_inc(x_1089); +lean_inc(x_1061); +x_1216 = l_Lean_Syntax_node5(x_1061, x_1215, x_1213, x_1089, x_1214, x_1163, x_1198); +x_1217 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +lean_inc(x_1061); +x_1218 = l_Lean_Syntax_node1(x_1061, x_1217, x_1216); +x_1219 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__20; +lean_inc(x_1061); +x_1220 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1220, 0, x_1061); +lean_ctor_set(x_1220, 1, x_1219); +x_1221 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__23; +lean_inc(x_1089); +lean_inc(x_1061); +x_1222 = l_Lean_Syntax_node2(x_1061, x_1221, x_1089, x_1204); +lean_inc(x_1061); +x_1223 = l_Lean_Syntax_node1(x_1061, x_1078, x_1222); +x_1224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +lean_inc(x_1061); +x_1225 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1225, 0, x_1061); +lean_ctor_set(x_1225, 1, x_1224); +x_1226 = lean_array_size(x_1052); +x_1227 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_1065, x_1226, x_15, x_1052); +x_1228 = lean_array_size(x_1227); +x_1229 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1228, x_15, x_1227); +x_1230 = l_Array_append___rarg(x_1076, x_1229); +lean_dec(x_1229); +lean_inc(x_1061); +x_1231 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1231, 0, x_1061); +lean_ctor_set(x_1231, 1, x_1078); +lean_ctor_set(x_1231, 2, x_1230); +x_1232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +lean_inc(x_1061); +x_1233 = l_Lean_Syntax_node1(x_1061, x_1232, x_1231); +x_1234 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_1225); +lean_inc_n(x_1089, 2); +lean_inc(x_1220); +lean_inc(x_1061); +x_1235 = l_Lean_Syntax_node6(x_1061, x_1234, x_1220, x_1089, x_1089, x_1223, x_1225, x_1233); +x_1236 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +lean_inc(x_1089); +lean_inc(x_1218); +lean_inc(x_1207); +lean_inc(x_1061); +x_1237 = l_Lean_Syntax_node4(x_1061, x_1236, x_1207, x_1218, x_1089, x_1235); +x_1238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__90; +lean_inc(x_1163); +lean_inc(x_1089); +lean_inc(x_1061); +x_1239 = l_Lean_Syntax_node5(x_1061, x_1238, x_1176, x_1205, x_1089, x_1163, x_1237); +x_1240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; +lean_inc(x_1061); +x_1241 = l_Lean_Syntax_node1(x_1061, x_1240, x_1239); +x_1242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_1200); +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1243 = l_Lean_Syntax_node4(x_1061, x_1242, x_1089, x_1089, x_1241, x_1200); +x_1244 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +lean_inc(x_1082); +lean_inc(x_1088); +x_1245 = l_Lean_addMacroScope(x_1088, x_1244, x_1082); +x_1246 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__102; +lean_inc(x_1061); +x_1247 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1247, 0, x_1061); +lean_ctor_set(x_1247, 1, x_1246); +lean_ctor_set(x_1247, 2, x_1245); +lean_ctor_set(x_1247, 3, x_1065); +lean_inc(x_1061); +x_1248 = l_Lean_Syntax_node1(x_1061, x_1078, x_1247); +x_1249 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; +lean_inc(x_1061); +x_1250 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1250, 0, x_1061); lean_ctor_set(x_1250, 1, x_1249); -lean_ctor_set(x_1250, 2, x_1248); -lean_ctor_set(x_1250, 3, x_1057); -lean_inc(x_1053); -x_1251 = l_Lean_Syntax_node2(x_1053, x_1150, x_1142, x_1089); -lean_inc(x_1053); -x_1252 = l_Lean_Syntax_node1(x_1053, x_1070, x_1251); -x_1253 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; -lean_inc(x_1053); -x_1254 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1254, 0, x_1053); -lean_ctor_set(x_1254, 1, x_1253); -x_1255 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__117; -lean_inc(x_1074); -lean_inc(x_1080); -x_1256 = l_Lean_addMacroScope(x_1080, x_1255, x_1074); -x_1257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; -x_1258 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; -lean_inc(x_1053); -x_1259 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1259, 0, x_1053); -lean_ctor_set(x_1259, 1, x_1257); -lean_ctor_set(x_1259, 2, x_1256); -lean_ctor_set(x_1259, 3, x_1258); -lean_inc(x_1238); -lean_inc(x_1053); -x_1260 = l_Lean_Syntax_node2(x_1053, x_1072, x_1259, x_1238); -x_1261 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -lean_inc(x_1053); -x_1262 = l_Lean_Syntax_node1(x_1053, x_1261, x_1260); -x_1263 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_1250); -lean_inc(x_1053); -x_1264 = l_Lean_Syntax_node4(x_1053, x_1263, x_1250, x_1252, x_1254, x_1262); -x_1265 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; -lean_inc(x_1081); -lean_inc(x_1053); -x_1266 = l_Lean_Syntax_node3(x_1053, x_1265, x_1246, x_1081, x_1264); -lean_inc(x_1081); -lean_inc(x_1053); -x_1267 = l_Lean_Syntax_node2(x_1053, x_1243, x_1266, x_1081); -x_1268 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; -x_1269 = l_Lean_addMacroScope(x_1080, x_1268, x_1074); -x_1270 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; -x_1271 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; -lean_inc(x_1053); -x_1272 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_1272, 0, x_1053); -lean_ctor_set(x_1272, 1, x_1270); -lean_ctor_set(x_1272, 2, x_1269); -lean_ctor_set(x_1272, 3, x_1271); -x_1273 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; -lean_inc(x_1053); -x_1274 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_1274, 0, x_1053); -lean_ctor_set(x_1274, 1, x_1273); -lean_inc(x_1081); -lean_inc(x_1053); -x_1275 = l_Lean_Syntax_node2(x_1053, x_1211, x_1081, x_1250); -lean_inc(x_1053); -x_1276 = l_Lean_Syntax_node1(x_1053, x_1070, x_1275); -x_1277 = lean_array_size(x_1045); -x_1278 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_1057, x_1277, x_15, x_1045); -x_1279 = lean_array_size(x_1278); -x_1280 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1279, x_15, x_1278); -x_1281 = l_Array_append___rarg(x_1068, x_1280); -lean_dec(x_1280); -lean_inc(x_1053); -x_1282 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_1282, 0, x_1053); -lean_ctor_set(x_1282, 1, x_1070); -lean_ctor_set(x_1282, 2, x_1281); -lean_inc(x_1053); -x_1283 = l_Lean_Syntax_node1(x_1053, x_1222, x_1282); -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1284 = l_Lean_Syntax_node6(x_1053, x_1224, x_1210, x_1081, x_1081, x_1276, x_1215, x_1283); -x_1285 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; -lean_inc(x_1053); -x_1286 = l_Lean_Syntax_node3(x_1053, x_1285, x_1272, x_1274, x_1284); -lean_inc(x_1053); -x_1287 = l_Lean_Syntax_node1(x_1053, x_1261, x_1286); -lean_inc(x_1081); -lean_inc(x_1053); -x_1288 = l_Lean_Syntax_node2(x_1053, x_1243, x_1287, x_1081); -lean_inc(x_1053); -x_1289 = l_Lean_Syntax_node3(x_1053, x_1070, x_1244, x_1267, x_1288); -x_1290 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_1053); -x_1291 = l_Lean_Syntax_node1(x_1053, x_1290, x_1289); -x_1292 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; -lean_inc(x_1053); -x_1293 = l_Lean_Syntax_node2(x_1053, x_1292, x_1240, x_1291); -lean_inc(x_1155); -lean_inc(x_1081); -lean_inc(x_1053); -x_1294 = l_Lean_Syntax_node5(x_1053, x_1228, x_1180, x_1238, x_1081, x_1155, x_1293); -lean_inc(x_1053); -x_1295 = l_Lean_Syntax_node1(x_1053, x_1230, x_1294); -lean_inc(x_1190); -lean_inc_n(x_1081, 2); -lean_inc(x_1053); -x_1296 = l_Lean_Syntax_node4(x_1053, x_1232, x_1081, x_1081, x_1295, x_1190); -lean_inc(x_1081); -lean_inc(x_1053); -x_1297 = l_Lean_Syntax_node3(x_1053, x_1070, x_1233, x_1081, x_1296); -x_1298 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc(x_1053); -x_1299 = l_Lean_Syntax_node2(x_1053, x_1298, x_1095, x_1297); -lean_inc(x_1053); -x_1300 = l_Lean_Syntax_node1(x_1053, x_1070, x_1299); -x_1301 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; -lean_inc(x_1053); -x_1302 = l_Lean_Syntax_node4(x_1053, x_1301, x_1155, x_1188, x_1190, x_1300); -x_1303 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; -lean_inc(x_1081); -lean_inc(x_1053); -x_1304 = l_Lean_Syntax_node6(x_1053, x_1303, x_1138, x_1140, x_1081, x_1081, x_1153, x_1302); -lean_inc(x_1053); -x_1305 = l_Lean_Syntax_node2(x_1053, x_1121, x_1136, x_1304); -x_1306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; -lean_inc(x_1053); -x_1307 = l_Lean_Syntax_node3(x_1053, x_1306, x_1128, x_1130, x_1305); -x_1308 = l_Lean_Syntax_node2(x_1053, x_1070, x_1122, x_1307); -if (lean_is_scalar(x_1078)) { - x_1309 = lean_alloc_ctor(0, 2, 0); +x_1251 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +lean_inc(x_1061); +x_1252 = l_Lean_Syntax_node2(x_1061, x_1251, x_1207, x_1218); +x_1253 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__109; +lean_inc(x_1089); +lean_inc(x_1061); +x_1254 = l_Lean_Syntax_node2(x_1061, x_1253, x_1252, x_1089); +x_1255 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_1061); +x_1256 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1256, 0, x_1061); +lean_ctor_set(x_1256, 1, x_1255); +x_1257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +lean_inc(x_1082); +lean_inc(x_1088); +x_1258 = l_Lean_addMacroScope(x_1088, x_1257, x_1082); +x_1259 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__30; +lean_inc(x_1061); +x_1260 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1260, 0, x_1061); +lean_ctor_set(x_1260, 1, x_1259); +lean_ctor_set(x_1260, 2, x_1258); +lean_ctor_set(x_1260, 3, x_1065); +lean_inc(x_1061); +x_1261 = l_Lean_Syntax_node2(x_1061, x_1158, x_1150, x_1097); +lean_inc(x_1061); +x_1262 = l_Lean_Syntax_node1(x_1061, x_1078, x_1261); +x_1263 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +lean_inc(x_1061); +x_1264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1264, 0, x_1061); +lean_ctor_set(x_1264, 1, x_1263); +x_1265 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__119; +lean_inc(x_1082); +lean_inc(x_1088); +x_1266 = l_Lean_addMacroScope(x_1088, x_1265, x_1082); +x_1267 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +x_1268 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122; +lean_inc(x_1061); +x_1269 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1269, 0, x_1061); +lean_ctor_set(x_1269, 1, x_1267); +lean_ctor_set(x_1269, 2, x_1266); +lean_ctor_set(x_1269, 3, x_1268); +lean_inc(x_1248); +lean_inc(x_1061); +x_1270 = l_Lean_Syntax_node2(x_1061, x_1080, x_1269, x_1248); +x_1271 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__116; +lean_inc(x_1061); +x_1272 = l_Lean_Syntax_node1(x_1061, x_1271, x_1270); +x_1273 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +lean_inc(x_1260); +lean_inc(x_1061); +x_1274 = l_Lean_Syntax_node4(x_1061, x_1273, x_1260, x_1262, x_1264, x_1272); +x_1275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; +lean_inc(x_1089); +lean_inc(x_1061); +x_1276 = l_Lean_Syntax_node3(x_1061, x_1275, x_1256, x_1089, x_1274); +lean_inc(x_1089); +lean_inc(x_1061); +x_1277 = l_Lean_Syntax_node2(x_1061, x_1253, x_1276, x_1089); +x_1278 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; +x_1279 = l_Lean_addMacroScope(x_1088, x_1278, x_1082); +x_1280 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; +x_1281 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__40; +lean_inc(x_1061); +x_1282 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1282, 0, x_1061); +lean_ctor_set(x_1282, 1, x_1280); +lean_ctor_set(x_1282, 2, x_1279); +lean_ctor_set(x_1282, 3, x_1281); +x_1283 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__41; +lean_inc(x_1061); +x_1284 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_1284, 0, x_1061); +lean_ctor_set(x_1284, 1, x_1283); +lean_inc(x_1089); +lean_inc(x_1061); +x_1285 = l_Lean_Syntax_node2(x_1061, x_1221, x_1089, x_1260); +lean_inc(x_1061); +x_1286 = l_Lean_Syntax_node1(x_1061, x_1078, x_1285); +x_1287 = lean_array_size(x_1053); +x_1288 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__19(x_1065, x_1287, x_15, x_1053); +x_1289 = lean_array_size(x_1288); +x_1290 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_1289, x_15, x_1288); +x_1291 = l_Array_append___rarg(x_1076, x_1290); +lean_dec(x_1290); +lean_inc(x_1061); +x_1292 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_1292, 0, x_1061); +lean_ctor_set(x_1292, 1, x_1078); +lean_ctor_set(x_1292, 2, x_1291); +lean_inc(x_1061); +x_1293 = l_Lean_Syntax_node1(x_1061, x_1232, x_1292); +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1294 = l_Lean_Syntax_node6(x_1061, x_1234, x_1220, x_1089, x_1089, x_1286, x_1225, x_1293); +x_1295 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +lean_inc(x_1061); +x_1296 = l_Lean_Syntax_node3(x_1061, x_1295, x_1282, x_1284, x_1294); +lean_inc(x_1061); +x_1297 = l_Lean_Syntax_node1(x_1061, x_1271, x_1296); +lean_inc(x_1089); +lean_inc(x_1061); +x_1298 = l_Lean_Syntax_node2(x_1061, x_1253, x_1297, x_1089); +lean_inc(x_1061); +x_1299 = l_Lean_Syntax_node3(x_1061, x_1078, x_1254, x_1277, x_1298); +x_1300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_1061); +x_1301 = l_Lean_Syntax_node1(x_1061, x_1300, x_1299); +x_1302 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_1061); +x_1303 = l_Lean_Syntax_node2(x_1061, x_1302, x_1250, x_1301); +lean_inc(x_1163); +lean_inc(x_1089); +lean_inc(x_1061); +x_1304 = l_Lean_Syntax_node5(x_1061, x_1238, x_1188, x_1248, x_1089, x_1163, x_1303); +lean_inc(x_1061); +x_1305 = l_Lean_Syntax_node1(x_1061, x_1240, x_1304); +lean_inc(x_1200); +lean_inc_n(x_1089, 2); +lean_inc(x_1061); +x_1306 = l_Lean_Syntax_node4(x_1061, x_1242, x_1089, x_1089, x_1305, x_1200); +lean_inc(x_1089); +lean_inc(x_1061); +x_1307 = l_Lean_Syntax_node3(x_1061, x_1078, x_1243, x_1089, x_1306); +x_1308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +lean_inc(x_1061); +x_1309 = l_Lean_Syntax_node2(x_1061, x_1308, x_1103, x_1307); +lean_inc(x_1061); +x_1310 = l_Lean_Syntax_node1(x_1061, x_1078, x_1309); +x_1311 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__65; +lean_inc(x_1061); +x_1312 = l_Lean_Syntax_node4(x_1061, x_1311, x_1163, x_1198, x_1200, x_1310); +x_1313 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +lean_inc(x_1089); +lean_inc(x_1061); +x_1314 = l_Lean_Syntax_node6(x_1061, x_1313, x_1146, x_1148, x_1089, x_1089, x_1161, x_1312); +lean_inc(x_1061); +x_1315 = l_Lean_Syntax_node2(x_1061, x_1129, x_1144, x_1314); +x_1316 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +lean_inc(x_1061); +x_1317 = l_Lean_Syntax_node3(x_1061, x_1316, x_1136, x_1138, x_1315); +x_1318 = l_Lean_Syntax_node2(x_1061, x_1078, x_1130, x_1317); +if (lean_is_scalar(x_1086)) { + x_1319 = lean_alloc_ctor(0, 2, 0); } else { - x_1309 = x_1078; + x_1319 = x_1086; } -lean_ctor_set(x_1309, 0, x_1308); -lean_ctor_set(x_1309, 1, x_1077); -return x_1309; +lean_ctor_set(x_1319, 0, x_1318); +lean_ctor_set(x_1319, 1, x_1085); +return x_1319; } else { -lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; -lean_dec(x_1046); -lean_dec(x_1045); -lean_dec(x_1044); -lean_dec(x_1041); +lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; +lean_dec(x_1054); +lean_dec(x_1053); +lean_dec(x_1052); +lean_dec(x_1049); lean_dec(x_10); lean_dec(x_9); lean_dec(x_1); -x_1310 = lean_ctor_get(x_1048, 0); -lean_inc(x_1310); -x_1311 = lean_ctor_get(x_1048, 1); -lean_inc(x_1311); -if (lean_is_exclusive(x_1048)) { - lean_ctor_release(x_1048, 0); - lean_ctor_release(x_1048, 1); - x_1312 = x_1048; +x_1320 = lean_ctor_get(x_1056, 0); +lean_inc(x_1320); +x_1321 = lean_ctor_get(x_1056, 1); +lean_inc(x_1321); +if (lean_is_exclusive(x_1056)) { + lean_ctor_release(x_1056, 0); + lean_ctor_release(x_1056, 1); + x_1322 = x_1056; } else { - lean_dec_ref(x_1048); - x_1312 = lean_box(0); + lean_dec_ref(x_1056); + x_1322 = lean_box(0); } -if (lean_is_scalar(x_1312)) { - x_1313 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_1322)) { + x_1323 = lean_alloc_ctor(1, 2, 0); } else { - x_1313 = x_1312; + x_1323 = x_1322; } -lean_ctor_set(x_1313, 0, x_1310); -lean_ctor_set(x_1313, 1, x_1311); -return x_1313; +lean_ctor_set(x_1323, 0, x_1320); +lean_ctor_set(x_1323, 1, x_1321); +return x_1323; } } } else { -uint8_t x_1314; +uint8_t x_1324; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -14154,23 +14233,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_1314 = !lean_is_exclusive(x_16); -if (x_1314 == 0) +x_1324 = !lean_is_exclusive(x_16); +if (x_1324 == 0) { return x_16; } else { -lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; -x_1315 = lean_ctor_get(x_16, 0); -x_1316 = lean_ctor_get(x_16, 1); -lean_inc(x_1316); -lean_inc(x_1315); +lean_object* x_1325; lean_object* x_1326; lean_object* x_1327; +x_1325 = lean_ctor_get(x_16, 0); +x_1326 = lean_ctor_get(x_16, 1); +lean_inc(x_1326); +lean_inc(x_1325); lean_dec(x_16); -x_1317 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_1317, 0, x_1315); -lean_ctor_set(x_1317, 1, x_1316); -return x_1317; +x_1327 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1327, 0, x_1325); +lean_ctor_set(x_1327, 1, x_1326); +return x_1327; } } } @@ -15915,7 +15994,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1() { _start: { lean_object* x_1; @@ -15923,7 +16002,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Se return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15933,27 +16012,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5() { _start: { lean_object* x_1; @@ -15961,17 +16040,17 @@ x_1 = lean_mk_string_unchecked("initFn", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7() { _start: { lean_object* x_1; @@ -15979,37 +16058,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9; x_2 = l_Array_forIn_x27Unsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11() { _start: { lean_object* x_1; @@ -16017,27 +16096,27 @@ x_1 = lean_mk_string_unchecked("Rpc", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14() { _start: { lean_object* x_1; @@ -16045,32 +16124,32 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15; -x_2 = lean_unsigned_to_nat(4819u); +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15; +x_2 = lean_unsigned_to_nat(4837u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__53; -x_3 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1; +x_3 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -16080,7 +16159,7 @@ lean_inc(x_5); lean_dec(x_4); x_6 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__3; x_7 = 0; -x_8 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16; +x_8 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16; x_9 = l_Lean_registerTraceClass(x_6, x_7, x_8, x_5); return x_9; } @@ -16521,6 +16600,10 @@ l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructure lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__121); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__122); +l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123(); +lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__123); +l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124(); +lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__124); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__1 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__1(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__1); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2(); @@ -16697,39 +16780,39 @@ l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance_ lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__1); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__1); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__2); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__3); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__4); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__5); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__6); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__7); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__8); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__9); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__10); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__11); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__12); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__13); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__14); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__15); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819____closed__16); -res = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4819_(lean_io_mk_world()); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__1); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__2); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__3); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__4); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__5); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__6); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__7); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__8); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__9); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__10); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__11); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__12); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__13); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__14); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__15); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837____closed__16); +res = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4837_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index 4f6553e3699b..30c0e26d56fb 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -59,6 +59,7 @@ static lean_object* l_IO_FS_Stream_readLspNotificationAs___at_Lean_Server_Watchd lean_object* l_Lean_JsonNumber_toString(lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleRequest___spec__8___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518_(lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_queuedMsgs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleWorkspaceSymbol___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -101,12 +102,12 @@ size_t lean_uint64_to_usize(uint64_t); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages___closed__1; lean_object* lean_io_as_task(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311_(lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleRequest___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_findWorkerPath___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__1(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__18; static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__20; uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); @@ -126,7 +127,6 @@ lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__9; LEAN_EXPORT lean_object* l_IO_Mutex_atomically___at_Lean_Server_Watchdog_FileWorker_waitForProc___spec__3(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2503_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardNotification___at_Lean_Server_Watchdog_handleNotification___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_FileWorker_errorPendingRequests___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonInitializeParams___spec__6(lean_object*, lean_object*); @@ -140,6 +140,7 @@ uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__3; lean_object* l_System_FilePath_extension(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__18(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Server_Watchdog_startLoadingReferences___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___spec__1___closed__2; @@ -166,18 +167,19 @@ lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanDidOpenTextDo static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__13; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_Watchdog_handleCancelRequest___spec__1(lean_object*, lean_object*); lean_object* lean_io_getenv(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7825_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__24; uint8_t lean_float_decLt(double, double); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1754_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_toJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4276_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1080_(lean_object*); +uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__15; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__4; lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_References_allRefsFor___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_runClientTask___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_RBNode_balLeft___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__18; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713_(lean_object*); lean_object* lean_io_basemutex_lock(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_startLoadingReferences___closed__1; static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__1; @@ -206,7 +208,6 @@ lean_object* l_List_flatMapTR_go___at___private_Lean_Server_Rpc_Basic_0__Lean_Ls LEAN_EXPORT lean_object* l_Array_groupByKey___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_mainLoop___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Server_Watchdog_handleRename___spec__5(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_Server_Watchdog_FileWorker_waitForProc___spec__1___boxed(lean_object*, lean_object*); @@ -318,6 +319,7 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspRequest___at_Lean_Server_Watchdo static lean_object* l_Lean_Server_Watchdog_mainLoop___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_startLoadingReferences___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_forwardRequestToWorker___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286_(lean_object*); static lean_object* l_Lean_Server_Watchdog_parseParams___rarg___closed__2; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4170____closed__3; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4170____closed__5; @@ -384,7 +386,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleImportClosure___boxed(lean LEAN_EXPORT lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_callHierarchyItemOf_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__5___lambda__1(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__5; -uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_toJsonDidChangeTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_529_(lean_object*); static lean_object* l_Lean_Server_Watchdog_terminateFileWorker___closed__2; lean_object* lean_stream_of_handle(lean_object*); @@ -445,7 +446,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_mainLoop___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_handleDidChange___lambda__1___closed__1; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(lean_object*); lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonDidSaveTextDocumentParams____x40_Lean_Data_Lsp_TextSync___hyg_750_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardRequestToWorker(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__6(lean_object*, lean_object*, lean_object*); @@ -462,6 +462,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_startLoadingReferences___lambda_ lean_object* l_Lean_Server_Ilean_load(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__1(size_t, size_t, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963_(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Server_Watchdog_ImportData_update___spec__2(lean_object*); static lean_object* l_Lean_Server_Watchdog_instFromJsonCallHierarchyItemData___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -471,7 +472,6 @@ static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__1; static lean_object* l_Lean_Server_Watchdog_watchdogMain___closed__1; static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__3; static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__22; -uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_startFileWorker___spec__2(lean_object*); lean_object* l_System_SearchPath_searchModuleNameOfUri(lean_object*, lean_object*, lean_object*); @@ -526,7 +526,6 @@ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_Watchdog_handleCall LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___spec__6(lean_object*, lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Server_Watchdog_findFileWorker_x3f___spec__1___boxed(lean_object*, lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_stdout(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__4; LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_get_x3f___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__2(lean_object*, lean_object*); @@ -541,7 +540,6 @@ lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fro LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleCancelRequest___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_handleRequest___lambda__1___closed__3; static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__1; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390_(lean_object*); static lean_object* l_Lean_Server_Watchdog_handleNotification___closed__5; LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_FS_Stream_readMessage(lean_object*, lean_object*, lean_object*); @@ -586,6 +584,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls LEAN_EXPORT lean_object* l_StateRefT_x27_get___at_Lean_Server_Watchdog_FileWorker_waitForProc___spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Std_DHashMap_Internal_AssocList_contains___at_Lean_Server_Watchdog_handleRename___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4170____closed__14; LEAN_EXPORT lean_object* l_Array_groupByKey___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_tryDischargeQueuedMessages(lean_object*, lean_object*, lean_object*, lean_object*); @@ -654,6 +653,7 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspNotification___at_Lean_Server_Wa uint64_t lean_uint64_xor(uint64_t, uint64_t); static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__7; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_findFileWorker_x3f___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593_(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_handleRename___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_terminateFileWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -666,7 +666,6 @@ static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__4; lean_object* l_Lean_Server_foldDocumentChanges(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_Watchdog_notifyAboutStaleDependency___spec__1(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__6; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardRequestToWorker___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5319_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__59; @@ -719,7 +718,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__2; size_t lean_usize_sub(size_t, size_t); lean_object* lean_array_mk(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7252_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleNotification___spec__13___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_Raw_u2080_expand___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_ImportData_update___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); @@ -749,12 +747,12 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_CallHierarchyItemData_fromItem_x lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_FileWorker_errorPendingRequests(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); size_t lean_array_size(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020_(lean_object*); static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_fromJsonCallHierarchyItemData____x40_Lean_Server_Watchdog___hyg_4170____closed__8; static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdogAux___closed__4; LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleIleanInfoFinal(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__16___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Watchdog_handleRequest___spec__17___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(lean_object*); static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__9; lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lean_Server_Watchdog_startFileWorker___closed__2; @@ -792,6 +790,7 @@ lean_object* l_Lean_Server_References_addIlean(lean_object*, lean_object*, lean_ static lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_Watchdog_initAndRunWatchdogAux___spec__2___closed__55; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_124_(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__9(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Watchdog_handleRequest___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_callHierarchyItemOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_References_finalizeWorkerRefs(lean_object*, lean_object*, lean_object*, lean_object*); @@ -829,6 +828,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleDidChangeWatchedFiles(lean static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_Watchdog_ImportData_update___spec__7(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311_(lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_Watchdog_ImportData_update___spec__11___closed__2; static lean_object* l_Lean_Server_Watchdog_mkLeanServerCapabilities___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); @@ -848,12 +848,12 @@ LEAN_EXPORT lean_object* l_Std_DHashMap_Internal_AssocList_foldlM___at_Lean_Serv lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_handleRename___lambda__1___closed__3; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_handleIleanInfoUpdate(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls_collapseSameOutgoingCalls___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Server_Watchdog_handleDidChangeWatchedFiles___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Server_Watchdog_ImportData_update___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handlePrepareCallHierarchy___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_Watchdog_handleCallHierarchyIncomingCalls_collapseSameIncomingCalls___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Server_Watchdog_handleCallHierarchyOutgoingCalls___spec__5(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SearchPath_findAllWithExt(lean_object*, lean_object*, lean_object*); @@ -19202,7 +19202,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_ctor_get(x_2, 2); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_4, x_1); +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_4, x_1); if (x_7 == 0) { x_2 = x_6; @@ -19233,7 +19233,7 @@ else lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_4, x_1); +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_4, x_1); if (x_6 == 0) { x_2 = x_5; @@ -19265,7 +19265,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8 x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); x_6 = lean_array_get_size(x_1); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_4); +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_4); x_8 = 32; x_9 = lean_uint64_shift_right(x_7, x_8); x_10 = lean_uint64_xor(x_7, x_9); @@ -19296,7 +19296,7 @@ lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); x_25 = lean_array_get_size(x_1); -x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_22); +x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_22); x_27 = 32; x_28 = lean_uint64_shift_right(x_26, x_27); x_29 = lean_uint64_xor(x_26, x_28); @@ -19388,7 +19388,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_6, x_1); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_6, x_1); if (x_9 == 0) { lean_object* x_10; @@ -19415,7 +19415,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_3); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_11, x_1); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -19460,7 +19460,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); x_7 = lean_ctor_get(x_2, 2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_5, x_1); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_5, x_1); if (x_8 == 0) { lean_object* x_9; @@ -19486,7 +19486,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_dec(x_2); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_10, x_1); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_10, x_1); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -19532,7 +19532,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_13 = lean_ctor_get(x_8, 0); x_14 = lean_ctor_get(x_8, 1); x_15 = lean_array_get_size(x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_11); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_11); x_17 = 32; x_18 = lean_uint64_shift_right(x_16, x_17); x_19 = lean_uint64_xor(x_16, x_18); @@ -19794,7 +19794,7 @@ lean_inc(x_124); lean_inc(x_123); lean_dec(x_8); x_125 = lean_array_get_size(x_124); -x_126 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_11); +x_126 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_11); x_127 = 32; x_128 = lean_uint64_shift_right(x_126, x_127); x_129 = lean_uint64_xor(x_126, x_128); @@ -21925,7 +21925,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); x_6 = lean_ctor_get(x_2, 2); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_4, x_1); +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_4, x_1); if (x_7 == 0) { x_2 = x_6; @@ -21956,7 +21956,7 @@ else lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_4, x_1); +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_4, x_1); if (x_6 == 0) { x_2 = x_5; @@ -21988,7 +21988,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; uint64_t x_8 x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 2); x_6 = lean_array_get_size(x_1); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_4); +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_4); x_8 = 32; x_9 = lean_uint64_shift_right(x_7, x_8); x_10 = lean_uint64_xor(x_7, x_9); @@ -22019,7 +22019,7 @@ lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); x_25 = lean_array_get_size(x_1); -x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_22); +x_26 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_22); x_27 = 32; x_28 = lean_uint64_shift_right(x_26, x_27); x_29 = lean_uint64_xor(x_26, x_28); @@ -22111,7 +22111,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_6, x_1); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_6, x_1); if (x_9 == 0) { lean_object* x_10; @@ -22138,7 +22138,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_3); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_11, x_1); +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; @@ -22183,7 +22183,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; x_5 = lean_ctor_get(x_2, 0); x_6 = lean_ctor_get(x_2, 1); x_7 = lean_ctor_get(x_2, 2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_5, x_1); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_5, x_1); if (x_8 == 0) { lean_object* x_9; @@ -22209,7 +22209,7 @@ lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_dec(x_2); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(x_10, x_1); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6992_(x_10, x_1); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -22255,7 +22255,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; uint64_t x_13 = lean_ctor_get(x_8, 0); x_14 = lean_ctor_get(x_8, 1); x_15 = lean_array_get_size(x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_11); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_11); x_17 = 32; x_18 = lean_uint64_shift_right(x_16, x_17); x_19 = lean_uint64_xor(x_16, x_18); @@ -22517,7 +22517,7 @@ lean_inc(x_124); lean_inc(x_123); lean_dec(x_8); x_125 = lean_array_get_size(x_124); -x_126 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6589_(x_11); +x_126 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7162_(x_11); x_127 = 32; x_128 = lean_uint64_shift_right(x_126, x_127); x_129 = lean_uint64_xor(x_126, x_128); @@ -37065,7 +37065,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10390_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10963_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37136,7 +37136,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10613_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_11186_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37247,7 +37247,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7020_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7593_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37296,7 +37296,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7252_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7825_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -37351,7 +37351,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6713_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7286_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37400,7 +37400,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6945_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7518_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -37455,7 +37455,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5738_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6311_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37504,7 +37504,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6313_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6886_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -37559,7 +37559,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3738_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4311_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -37608,7 +37608,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5642_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6215_(x_5); x_9 = 1; x_10 = lean_usize_add(x_2, x_9); x_11 = lean_array_uset(x_7, x_2, x_8); @@ -37663,7 +37663,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat { lean_object* x_4; lean_inc(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3515_(x_1); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4088_(x_1); if (lean_obj_tag(x_4) == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index dd3666f44806..605ddcc8bdfe 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -13,47 +13,50 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19; lean_object* l_Lean_Server_RequestM_mapTask___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonGetWidgetSourceParams; static lean_object* l_Lean_Widget_instToJsonGetWidgetSourceParams___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_builtinModulesRef; LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget___rarg(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_eraseWidgetSpec; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__6; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038_(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6; lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__6; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__2___closed__2; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__2; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__6; extern lean_object* l_Lean_Server_builtinRpcProcedures; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_addBuiltinModule___spec__2(lean_object*, uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_addWidgetSpec___closed__4; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4903_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__12(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__5; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__13; @@ -61,6 +64,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJso static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__1; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__6; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7; lean_object* l_Lean_Json_mkObj(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__19; LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -68,10 +72,12 @@ LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hy LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425_(lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__9; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal___rarg(lean_object*, lean_object*, lean_object*, uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Lean_Widget_evalPanelWidgets___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Server_registerBuiltinRpcProcedure___spec__5(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__1; @@ -84,16 +90,15 @@ lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_WidgetInstance_ofHash___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905_; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__7___rarg(lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__3___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__4(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__18; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__1___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__8___closed__2; lean_object* l_Lean_Name_toString(lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__1___closed__2; @@ -102,24 +107,27 @@ lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableWidgetInsta static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__4; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_WidgetInstance_ofHash___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__2___rarg(lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__4; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__2(lean_object*, uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Widget_getWidgets___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____closed__3; lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__1(lean_object*, uint64_t); +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823_(lean_object*); static lean_object* l_Lean_Widget_instFromJsonUserWidgetDefinition___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedWidgetSource; static lean_object* l_Lean_Widget_getWidgetSource___lambda__5___closed__1; uint8_t lean_uint64_dec_lt(uint64_t, uint64_t); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__1___boxed(lean_object*, lean_object*); @@ -132,15 +140,18 @@ static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__9; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__8; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__3; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedUserWidgetDefinition; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__8; static lean_object* l_Lean_Widget_addWidgetSpec___closed__2; lean_object* l_Lean_Server_RequestM_withWaitFindSnap___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2; static lean_object* l_Lean_Widget_showWidgetSpec___closed__3; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__5; @@ -149,7 +160,6 @@ lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__L LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_Widget_showWidgetSpec___closed__2; -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406_(lean_object*); lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__3(lean_object*, uint64_t); lean_object* l_Lean_RBNode_balLeft___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -160,43 +170,42 @@ static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__15; uint64_t lean_string_hash(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(lean_object*); lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__15; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__9(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____closed__5; lean_object* l_Lean_RBNode_balRight___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1871____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__6; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__35; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__5; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_elabWidgetCmd(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__8; LEAN_EXPORT lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__6; LEAN_EXPORT lean_object* l_Lean_Widget_elabWidgetCmd___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__11; -static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34; static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1000____closed__1; lean_object* l_Lean_Server_RequestM_findInfoTreeAtPosWithTrailingWhitespace(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__4; static lean_object* l_Lean_Widget_showWidgetSpec___closed__6; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__21; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_820____closed__2; static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__7; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44; @@ -212,9 +221,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Widget_erasePanelWidget__ uint8_t l_Lean_Expr_hasMVar(lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__11; lean_object* l_Lean_initializing(lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4; -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797_(lean_object*, lean_object*); lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_getWidgets___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -225,20 +232,21 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGe LEAN_EXPORT lean_object* l_Lean_Widget_instToModuleModule; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__7; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__3; static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__11; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_Widget_getWidgetSource___lambda__2(lean_object*, uint64_t, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__3; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4; +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__5(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_task_pure(lean_object*); lean_object* l_Lean_FileMap_lspPosToUtf8Pos(lean_object*, lean_object*); @@ -249,14 +257,13 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240 LEAN_EXPORT lean_object* l_Lean_Widget_addWidgetSpec; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___at_Lean_Widget_getWidgets___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28; lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__17; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_widgetCmd; @@ -267,14 +274,16 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_toArray___rarg(lean_object*); lean_object* l_Lean_registerSimpleScopedEnvExtension___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__18; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd___boxed__const__1; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__7; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_panelWidgetsExt; @@ -283,41 +292,42 @@ LEAN_EXPORT lean_object* l_Lean_ofExcept___at___private_Lean_Widget_UserWidget_0 lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonUserWidgetDefinition; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153_(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____lambda__1(lean_object*, lean_object*); lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; static lean_object* l_Lean_Widget_getWidgetSource___lambda__6___closed__2; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6; lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23; +static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgets___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4; static lean_object* l_Lean_Widget_instToJsonUserWidgetDefinition___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471____boxed(lean_object*); lean_object* l_Lean_MessageData_ofFormat(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__2; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__5; +static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1; lean_object* l_Lean_Server_Snapshots_Snapshot_env(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1; lean_object* lean_uint64_to_nat(uint64_t); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__5; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__12; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__2; lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182_(lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__1; +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411_(lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveTermGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_2227____spec__1(lean_object*, lean_object*); @@ -326,29 +336,31 @@ lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__21; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Widget_addBuiltinModule___spec__1(lean_object*, uint64_t, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689_(lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; LEAN_EXPORT lean_object* l_Lean_Widget_WidgetInstance_ofHash(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17; lean_object* lean_st_mk_ref(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__17; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158_(lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_utf8PosToLspPos(lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Widget_addWidgetSpec___closed__7; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58; lean_object* l_List_flatMapTR_go___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCancelParams____x40_Lean_Data_Lsp_Basic___hyg_86____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____spec__1(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___boxed(lean_object*, lean_object*); @@ -356,17 +368,17 @@ LEAN_EXPORT lean_object* l_Lean_Widget_elabWidgetInstanceSpec(lean_object*, lean LEAN_EXPORT uint8_t l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgets(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__14; +static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__7; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalWidgetInstanceUnsafe___closed__1; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__1; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__5___closed__2; -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonWidgetSource; static lean_object* l_Lean_Widget_addBuiltinModule___closed__1; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686_(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgets___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -378,28 +390,23 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__5___closed__1; static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9; extern lean_object* l_Lean_warningAsError; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__13; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__5; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818_(lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4908_(lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__3; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1; lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____lambda__2(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalWidgetInstanceUnsafe___closed__2; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2; lean_object* l_Lean_bignumFromJson_x3f(lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__24; static lean_object* l_Lean_Widget_addWidgetSpec___closed__1; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__9; static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__3; @@ -410,11 +417,7 @@ static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__14; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__31; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__7; -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466____boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_addWidgetSpec___closed__6; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -425,33 +428,32 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240 LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__3___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__49; lean_object* l_String_Range_toLspRange(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofConstName(lean_object*, uint8_t); lean_object* l_Array_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__30; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__4; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped(lean_object*); -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3; lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableWidgetInstance_dec____x40_Lean_Widget_Types___hyg_3____spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__6; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__14; lean_object* l_Lean_ScopedEnvExtension_add___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12; LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Widget_evalPanelWidgets___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19; static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__3; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__2; lean_object* l_Lean_Server_RequestM_readDoc___at_Lean_Server_RequestM_withWaitFindSnapAtPos___spec__1(lean_object*, lean_object*); @@ -461,25 +463,27 @@ static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Widget_elabShowPanel LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__2; lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__16; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetGlobal___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__8(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943_; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__12; LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__11(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__13; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__10; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____closed__2; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__2; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__7; +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_savePanelWidgetInfo(uint64_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); @@ -497,6 +501,7 @@ LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Widget_initFn____x40_Lean_W lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_406_(lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__1___closed__1; static lean_object* l_Lean_Widget_showWidgetSpec___closed__5; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__1; static lean_object* l_Lean_Widget_widgetCmd___closed__1; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__10; @@ -506,20 +511,18 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__3; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__11; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4074_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16; LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27; LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2; static lean_object* l_Lean_Widget_addWidgetSpec___closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__12; @@ -527,10 +530,12 @@ static lean_object* l_Lean_Widget_instInhabitedUserWidgetDefinition___closed__1; lean_object* l_id___rarg___boxed(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__1; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; +LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542_; static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__7; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__11; +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5; lean_object* l_Lean_indentExpr(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__14; LEAN_EXPORT lean_object* l_Lean_Widget_WidgetInstance_ofHash___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -542,35 +547,35 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWi LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped___rarg(lean_object*, lean_object*, lean_object*, uint64_t, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900_; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5; lean_object* l_Lean_Server_RequestM_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__8; static lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__2; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_showPanelWidgetsCmd; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__30; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Widget_getWidgetSource___spec__1(lean_object*, lean_object*); lean_object* l_String_toName(lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__18; LEAN_EXPORT lean_object* l_Lean_Widget_widgetInstanceSpec; static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__9; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonWidgetSource; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableWidgetInstance_dec____x40_Lean_Widget_Types___hyg_3____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__9(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634_(lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpec___closed__8; uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -586,12 +591,10 @@ lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object* static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__2; lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__35; lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074_(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463_; static lean_object* l_Lean_Widget_instToModuleModule___closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -602,6 +605,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_showWidgetSpec; LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____spec__3(lean_object*, uint64_t, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51; lean_object* l_Lean_bignumToJson(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget___rarg___boxed(lean_object*, lean_object*); @@ -613,35 +617,35 @@ lean_object* l___private_Lean_ToExpr_0__Lean_Name_toExprAux(lean_object*); uint8_t lean_uint64_dec_eq(uint64_t, uint64_t); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__9; lean_object* l_String_intercalate(lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__24; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468_; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__5; static lean_object* l_Lean_Widget_evalPanelWidgets___closed__1; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__13; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Widget_getWidgetSource___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__1___closed__1; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__5; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__7; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_addBuiltinModule(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Server_registerBuiltinRpcProcedure___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_getWidgets___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__18; size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__1; -static lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe(lean_object*); -static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_erasePanelWidget___rarg___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_352____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Widget_erasePanelWidget___spec__1(uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__6___closed__3; LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938_; +static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Widget_widgetModuleAttrImpl; @@ -656,11 +660,10 @@ static lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelW LEAN_EXPORT lean_object* l_Lean_Widget_moduleRegistry; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__9; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1425____closed__4; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__2; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_logWarning___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__12; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Widget_addBuiltinModule___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -673,6 +676,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_ela static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__10; +static lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__3; lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1683____spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__14; @@ -688,22 +692,20 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240 static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27; lean_object* l_Lean_Meta_evalExpr_x27___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__6; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__2(lean_object*, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__1; static lean_object* l_Lean_Widget_instFromJsonWidgetSource___closed__1; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Widget_instFromJsonGetWidgetSourceParams___closed__1; lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_551____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16; LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_89_(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__3; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792_(lean_object*, lean_object*); static lean_object* l_Lean_Widget_showWidgetSpec___closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8; static lean_object* l_Lean_Widget_instToJsonWidgetSource___closed__1; extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; -static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12; +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetLocal(lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; @@ -711,15 +713,16 @@ static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__8___closed__1; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___at_Lean_Widget_getWidgets___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_RequestError_ofIoError(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23; static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____closed__3; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_getWidgets___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_widgetInstanceSpec___closed__11; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1038____closed__8; @@ -727,15 +730,15 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_142 lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_ensureNoArgs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__8; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_elabShowPanelWidgetsCmd___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -746,18 +749,17 @@ static lean_object* l_Lean_Widget_eraseWidgetSpec___closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonUserWidgetDefinition; lean_object* l_Lean_Server_Snapshots_Snapshot_endPos(lean_object*); static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537_; -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_794_(lean_object*); +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4; +static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__4; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4079_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__1___rarg___closed__1; static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__4___closed__8; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_StateT_pure___at_Lean_Server_instRpcEncodableStateMRpcObjectStore___spec__1___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5; lean_object* l_Lean_throwError___at___private_Lean_Meta_InferType_0__Lean_Meta_inferProjType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_182____spec__1(lean_object*, uint64_t, lean_object*); static lean_object* l_Lean_Widget_widgetCmd___closed__2; @@ -13210,7 +13212,7 @@ lean_dec(x_7); x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_3); -x_11 = lean_alloc_ctor(7, 1, 0); +x_11 = lean_alloc_ctor(8, 1, 0); lean_ctor_set(x_11, 0, x_10); x_12 = l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_realizeGlobalConstNoOverloadWithInfo___spec__3(x_11, x_4, x_5, x_9); return x_12; @@ -13520,7 +13522,7 @@ static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__9() _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("structInstField", 15, 15); +x_1 = lean_mk_string_unchecked("structInstFields", 16, 16); return x_1; } } @@ -13540,7 +13542,7 @@ static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__11() _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("structInstLVal", 14, 14); +x_1 = lean_mk_string_unchecked("structInstField", 15, 15); return x_1; } } @@ -13560,76 +13562,96 @@ static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13() _start: { lean_object* x_1; -x_1 = lean_mk_string_unchecked("id", 2, 2); +x_1 = lean_mk_string_unchecked("structInstLVal", 14, 14); return x_1; } } static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__14() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; +x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("id", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16; -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__18; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__20; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__20() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22() { _start: { lean_object* x_1; @@ -13637,7 +13659,7 @@ x_1 = lean_mk_string_unchecked(":=", 2, 2); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23() { _start: { lean_object* x_1; @@ -13645,26 +13667,26 @@ x_1 = lean_mk_string_unchecked("javascriptHash", 14, 14); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__26() { _start: { lean_object* x_1; @@ -13672,19 +13694,19 @@ x_1 = lean_mk_string_unchecked("proj", 4, 4); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__26; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__26() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28() { _start: { lean_object* x_1; @@ -13692,19 +13714,19 @@ x_1 = lean_mk_string_unchecked("paren", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__26; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30() { _start: { lean_object* x_1; @@ -13712,7 +13734,7 @@ x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__31() { _start: { lean_object* x_1; @@ -13720,19 +13742,19 @@ x_1 = lean_mk_string_unchecked("app", 3, 3); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__31; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__31() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33() { _start: { lean_object* x_1; @@ -13740,16 +13762,16 @@ x_1 = lean_mk_string_unchecked("ToModule.toModule", 17, 17); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__31; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13759,7 +13781,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13771,19 +13793,19 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38() { _start: { lean_object* x_1; @@ -13791,7 +13813,7 @@ x_1 = lean_mk_string_unchecked(")", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39() { _start: { lean_object* x_1; @@ -13799,26 +13821,26 @@ x_1 = lean_mk_string_unchecked("props", 5, 5); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42() { _start: { lean_object* x_1; @@ -13826,16 +13848,16 @@ x_1 = lean_mk_string_unchecked("Server.RpcEncodable.rpcEncode", 29, 29); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44() { _start: { lean_object* x_1; @@ -13843,7 +13865,7 @@ x_1 = lean_mk_string_unchecked("Server", 6, 6); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45() { _start: { lean_object* x_1; @@ -13851,7 +13873,7 @@ x_1 = lean_mk_string_unchecked("RpcEncodable", 12, 12); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__46() { _start: { lean_object* x_1; @@ -13859,54 +13881,54 @@ x_1 = lean_mk_string_unchecked("rpcEncode", 9, 9); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42; -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43; -x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45; +x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__46; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__46() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__42; -x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__44; +x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__46; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__46; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__49; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__49() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51() { _start: { lean_object* x_1; @@ -13914,19 +13936,19 @@ x_1 = lean_mk_string_unchecked("optEllipsis", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__49; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53() { _start: { lean_object* x_1; @@ -13934,7 +13956,7 @@ x_1 = lean_mk_string_unchecked("}", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13944,17 +13966,17 @@ x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56() { _start: { lean_object* x_1; @@ -13962,19 +13984,19 @@ x_1 = lean_mk_string_unchecked("quotedName", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__1; x_3 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__2; -x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54; +x_4 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56() { +static lean_object* _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58() { _start: { lean_object* x_1; @@ -14016,24 +14038,24 @@ x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_12); lean_ctor_set(x_23, 1, x_21); lean_ctor_set(x_23, 2, x_22); -x_24 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; +x_24 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; lean_inc(x_13); lean_inc(x_19); x_25 = l_Lean_addMacroScope(x_19, x_24, x_13); x_26 = lean_box(0); -x_27 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__14; -x_28 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19; +x_27 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16; +x_28 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; lean_inc(x_12); x_29 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_27); lean_ctor_set(x_29, 2, x_25); lean_ctor_set(x_29, 3, x_28); -x_30 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__12; +x_30 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__14; lean_inc(x_23); lean_inc(x_12); x_31 = l_Lean_Syntax_node2(x_12, x_30, x_29, x_23); -x_32 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__20; +x_32 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22; lean_inc(x_12); x_33 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_33, 0, x_12); @@ -14041,11 +14063,11 @@ lean_ctor_set(x_33, 1, x_32); x_34 = l_Lean_Syntax_getId(x_1); lean_inc(x_34); x_35 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_26, x_34); -x_36 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; +x_36 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; lean_inc(x_13); lean_inc(x_19); x_37 = l_Lean_addMacroScope(x_19, x_36, x_13); -x_38 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22; +x_38 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24; lean_inc(x_12); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_12); @@ -14056,17 +14078,17 @@ lean_inc(x_23); lean_inc(x_39); lean_inc(x_12); x_40 = l_Lean_Syntax_node2(x_12, x_30, x_39, x_23); -x_41 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; +x_41 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; lean_inc(x_12); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_12); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33; +x_43 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; lean_inc(x_13); lean_inc(x_19); x_44 = l_Lean_addMacroScope(x_19, x_43, x_13); -x_45 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; -x_46 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; +x_45 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34; +x_46 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; lean_inc(x_12); x_47 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_47, 0, x_12); @@ -14075,15 +14097,15 @@ lean_ctor_set(x_47, 2, x_44); lean_ctor_set(x_47, 3, x_46); lean_inc(x_12); x_48 = l_Lean_Syntax_node1(x_12, x_21, x_1); -x_49 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; +x_49 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; lean_inc(x_12); x_50 = l_Lean_Syntax_node2(x_12, x_49, x_47, x_48); -x_51 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; +x_51 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38; lean_inc(x_12); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_12); lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27; +x_53 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29; lean_inc(x_12); x_54 = l_Lean_Syntax_node3(x_12, x_53, x_42, x_50, x_52); x_55 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; @@ -14091,18 +14113,18 @@ lean_inc(x_12); x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_12); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; +x_57 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27; lean_inc(x_12); x_58 = l_Lean_Syntax_node3(x_12, x_57, x_54, x_56, x_39); -x_59 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; +x_59 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__12; lean_inc(x_33); lean_inc(x_12); x_60 = l_Lean_Syntax_node3(x_12, x_59, x_40, x_33, x_58); -x_61 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; +x_61 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; lean_inc(x_13); lean_inc(x_19); x_62 = l_Lean_addMacroScope(x_19, x_61, x_13); -x_63 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38; +x_63 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40; lean_inc(x_12); x_64 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_64, 0, x_12); @@ -14112,10 +14134,10 @@ lean_ctor_set(x_64, 3, x_26); lean_inc(x_23); lean_inc(x_12); x_65 = l_Lean_Syntax_node2(x_12, x_30, x_64, x_23); -x_66 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45; +x_66 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47; x_67 = l_Lean_addMacroScope(x_19, x_66, x_13); -x_68 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; -x_69 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48; +x_68 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43; +x_69 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50; lean_inc(x_12); x_70 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_70, 0, x_12); @@ -14129,265 +14151,277 @@ x_72 = l_Lean_Syntax_node2(x_12, x_49, x_70, x_71); lean_inc(x_33); lean_inc(x_12); x_73 = l_Lean_Syntax_node3(x_12, x_59, x_65, x_33, x_72); -x_74 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50; +x_74 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52; lean_inc(x_23); lean_inc(x_12); x_75 = l_Lean_Syntax_node1(x_12, x_74, x_23); -x_76 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51; +x_76 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; lean_inc(x_12); x_77 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_77, 0, x_12); lean_ctor_set(x_77, 1, x_76); if (lean_obj_tag(x_35) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; x_78 = l_Lean_quoteNameMk(x_34); lean_inc(x_12); x_79 = l_Lean_Syntax_node3(x_12, x_59, x_31, x_33, x_78); lean_inc_n(x_23, 2); lean_inc(x_12); x_80 = l_Lean_Syntax_node5(x_12, x_21, x_79, x_23, x_60, x_23, x_73); -x_81 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; +x_81 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; +lean_inc(x_12); +x_82 = l_Lean_Syntax_node1(x_12, x_81, x_80); +x_83 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; lean_inc(x_23); -x_82 = l_Lean_Syntax_node6(x_12, x_81, x_14, x_23, x_80, x_75, x_23, x_77); -x_83 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; -x_84 = 1; -x_85 = l_Lean_Elab_Term_elabTerm(x_82, x_83, x_84, x_84, x_3, x_4, x_5, x_6, x_7, x_8, x_17); -return x_85; +x_84 = l_Lean_Syntax_node6(x_12, x_83, x_14, x_23, x_82, x_75, x_23, x_77); +x_85 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; +x_86 = 1; +x_87 = l_Lean_Elab_Term_elabTerm(x_84, x_85, x_86, x_86, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +return x_87; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_dec(x_34); -x_86 = lean_ctor_get(x_35, 0); -lean_inc(x_86); +x_88 = lean_ctor_get(x_35, 0); +lean_inc(x_88); lean_dec(x_35); -x_87 = l_String_intercalate(x_55, x_86); -x_88 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56; -x_89 = lean_string_append(x_88, x_87); -lean_dec(x_87); -x_90 = lean_box(2); -x_91 = l_Lean_Syntax_mkNameLit(x_89, x_90); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_26); -x_93 = lean_array_mk(x_92); -x_94 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; -x_95 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_95, 0, x_90); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_95, 2, x_93); +x_89 = l_String_intercalate(x_55, x_88); +x_90 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58; +x_91 = lean_string_append(x_90, x_89); +lean_dec(x_89); +x_92 = lean_box(2); +x_93 = l_Lean_Syntax_mkNameLit(x_91, x_92); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_26); +x_95 = lean_array_mk(x_94); +x_96 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57; +x_97 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_97, 0, x_92); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_95); lean_inc(x_12); -x_96 = l_Lean_Syntax_node3(x_12, x_59, x_31, x_33, x_95); +x_98 = l_Lean_Syntax_node3(x_12, x_59, x_31, x_33, x_97); lean_inc_n(x_23, 2); lean_inc(x_12); -x_97 = l_Lean_Syntax_node5(x_12, x_21, x_96, x_23, x_60, x_23, x_73); -x_98 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; +x_99 = l_Lean_Syntax_node5(x_12, x_21, x_98, x_23, x_60, x_23, x_73); +x_100 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; +lean_inc(x_12); +x_101 = l_Lean_Syntax_node1(x_12, x_100, x_99); +x_102 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; lean_inc(x_23); -x_99 = l_Lean_Syntax_node6(x_12, x_98, x_14, x_23, x_97, x_75, x_23, x_77); -x_100 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; -x_101 = 1; -x_102 = l_Lean_Elab_Term_elabTerm(x_99, x_100, x_101, x_101, x_3, x_4, x_5, x_6, x_7, x_8, x_17); -return x_102; +x_103 = l_Lean_Syntax_node6(x_12, x_102, x_14, x_23, x_101, x_75, x_23, x_77); +x_104 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; +x_105 = 1; +x_106 = l_Lean_Elab_Term_elabTerm(x_103, x_104, x_105, x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +return x_106; } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_103 = lean_ctor_get(x_14, 0); -x_104 = lean_ctor_get(x_14, 1); -lean_inc(x_104); -lean_inc(x_103); +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_107 = lean_ctor_get(x_14, 0); +x_108 = lean_ctor_get(x_14, 1); +lean_inc(x_108); +lean_inc(x_107); lean_dec(x_14); -x_105 = lean_ctor_get(x_103, 0); -lean_inc(x_105); -lean_dec(x_103); -x_106 = lean_environment_main_module(x_105); -x_107 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__5; +x_109 = lean_ctor_get(x_107, 0); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_environment_main_module(x_109); +x_111 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__5; lean_inc(x_12); -x_108 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_108, 0, x_12); -lean_ctor_set(x_108, 1, x_107); -x_109 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__7; -x_110 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__8; +x_112 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_112, 0, x_12); +lean_ctor_set(x_112, 1, x_111); +x_113 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__7; +x_114 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__8; lean_inc(x_12); -x_111 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_111, 0, x_12); -lean_ctor_set(x_111, 1, x_109); -lean_ctor_set(x_111, 2, x_110); -x_112 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; +x_115 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_115, 0, x_12); +lean_ctor_set(x_115, 1, x_113); +lean_ctor_set(x_115, 2, x_114); +x_116 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; lean_inc(x_13); -lean_inc(x_106); -x_113 = l_Lean_addMacroScope(x_106, x_112, x_13); -x_114 = lean_box(0); -x_115 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__14; -x_116 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__19; +lean_inc(x_110); +x_117 = l_Lean_addMacroScope(x_110, x_116, x_13); +x_118 = lean_box(0); +x_119 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__16; +x_120 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; lean_inc(x_12); -x_117 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_117, 0, x_12); -lean_ctor_set(x_117, 1, x_115); -lean_ctor_set(x_117, 2, x_113); -lean_ctor_set(x_117, 3, x_116); -x_118 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__12; -lean_inc(x_111); +x_121 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_121, 0, x_12); +lean_ctor_set(x_121, 1, x_119); +lean_ctor_set(x_121, 2, x_117); +lean_ctor_set(x_121, 3, x_120); +x_122 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__14; +lean_inc(x_115); lean_inc(x_12); -x_119 = l_Lean_Syntax_node2(x_12, x_118, x_117, x_111); -x_120 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__20; +x_123 = l_Lean_Syntax_node2(x_12, x_122, x_121, x_115); +x_124 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22; lean_inc(x_12); -x_121 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_121, 0, x_12); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_Syntax_getId(x_1); -lean_inc(x_122); -x_123 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_114, x_122); -x_124 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; +x_125 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_125, 0, x_12); +lean_ctor_set(x_125, 1, x_124); +x_126 = l_Lean_Syntax_getId(x_1); +lean_inc(x_126); +x_127 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_118, x_126); +x_128 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; lean_inc(x_13); -lean_inc(x_106); -x_125 = l_Lean_addMacroScope(x_106, x_124, x_13); -x_126 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__22; +lean_inc(x_110); +x_129 = l_Lean_addMacroScope(x_110, x_128, x_13); +x_130 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__24; lean_inc(x_12); -x_127 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_127, 0, x_12); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_127, 2, x_125); -lean_ctor_set(x_127, 3, x_114); -lean_inc(x_111); -lean_inc(x_127); +x_131 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_131, 0, x_12); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_131, 2, x_129); +lean_ctor_set(x_131, 3, x_118); +lean_inc(x_115); +lean_inc(x_131); lean_inc(x_12); -x_128 = l_Lean_Syntax_node2(x_12, x_118, x_127, x_111); -x_129 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__28; +x_132 = l_Lean_Syntax_node2(x_12, x_122, x_131, x_115); +x_133 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; lean_inc(x_12); -x_130 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_130, 0, x_12); -lean_ctor_set(x_130, 1, x_129); -x_131 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33; +x_134 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_134, 0, x_12); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; lean_inc(x_13); -lean_inc(x_106); -x_132 = l_Lean_addMacroScope(x_106, x_131, x_13); -x_133 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; -x_134 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; -lean_inc(x_12); -x_135 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_135, 0, x_12); -lean_ctor_set(x_135, 1, x_133); -lean_ctor_set(x_135, 2, x_132); -lean_ctor_set(x_135, 3, x_134); +lean_inc(x_110); +x_136 = l_Lean_addMacroScope(x_110, x_135, x_13); +x_137 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34; +x_138 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; lean_inc(x_12); -x_136 = l_Lean_Syntax_node1(x_12, x_109, x_1); -x_137 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; +x_139 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_139, 0, x_12); +lean_ctor_set(x_139, 1, x_137); +lean_ctor_set(x_139, 2, x_136); +lean_ctor_set(x_139, 3, x_138); lean_inc(x_12); -x_138 = l_Lean_Syntax_node2(x_12, x_137, x_135, x_136); -x_139 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__36; +x_140 = l_Lean_Syntax_node1(x_12, x_113, x_1); +x_141 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; lean_inc(x_12); -x_140 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_140, 0, x_12); -lean_ctor_set(x_140, 1, x_139); -x_141 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27; -lean_inc(x_12); -x_142 = l_Lean_Syntax_node3(x_12, x_141, x_130, x_138, x_140); -x_143 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; +x_142 = l_Lean_Syntax_node2(x_12, x_141, x_139, x_140); +x_143 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38; lean_inc(x_12); x_144 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_144, 0, x_12); lean_ctor_set(x_144, 1, x_143); -x_145 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; +x_145 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__29; lean_inc(x_12); -x_146 = l_Lean_Syntax_node3(x_12, x_145, x_142, x_144, x_127); -x_147 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; -lean_inc(x_121); +x_146 = l_Lean_Syntax_node3(x_12, x_145, x_134, x_142, x_144); +x_147 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; lean_inc(x_12); -x_148 = l_Lean_Syntax_node3(x_12, x_147, x_128, x_121, x_146); -x_149 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; +x_148 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_148, 0, x_12); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__27; +lean_inc(x_12); +x_150 = l_Lean_Syntax_node3(x_12, x_149, x_146, x_148, x_131); +x_151 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__12; +lean_inc(x_125); +lean_inc(x_12); +x_152 = l_Lean_Syntax_node3(x_12, x_151, x_132, x_125, x_150); +x_153 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; lean_inc(x_13); -lean_inc(x_106); -x_150 = l_Lean_addMacroScope(x_106, x_149, x_13); -x_151 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__38; +lean_inc(x_110); +x_154 = l_Lean_addMacroScope(x_110, x_153, x_13); +x_155 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__40; lean_inc(x_12); -x_152 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_152, 0, x_12); -lean_ctor_set(x_152, 1, x_151); -lean_ctor_set(x_152, 2, x_150); -lean_ctor_set(x_152, 3, x_114); -lean_inc(x_111); +x_156 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_156, 0, x_12); +lean_ctor_set(x_156, 1, x_155); +lean_ctor_set(x_156, 2, x_154); +lean_ctor_set(x_156, 3, x_118); +lean_inc(x_115); lean_inc(x_12); -x_153 = l_Lean_Syntax_node2(x_12, x_118, x_152, x_111); -x_154 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__45; -x_155 = l_Lean_addMacroScope(x_106, x_154, x_13); -x_156 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; -x_157 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__48; +x_157 = l_Lean_Syntax_node2(x_12, x_122, x_156, x_115); +x_158 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__47; +x_159 = l_Lean_addMacroScope(x_110, x_158, x_13); +x_160 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__43; +x_161 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50; lean_inc(x_12); -x_158 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_158, 0, x_12); -lean_ctor_set(x_158, 1, x_156); -lean_ctor_set(x_158, 2, x_155); -lean_ctor_set(x_158, 3, x_157); +x_162 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_162, 0, x_12); +lean_ctor_set(x_162, 1, x_160); +lean_ctor_set(x_162, 2, x_159); +lean_ctor_set(x_162, 3, x_161); lean_inc(x_12); -x_159 = l_Lean_Syntax_node1(x_12, x_109, x_2); +x_163 = l_Lean_Syntax_node1(x_12, x_113, x_2); lean_inc(x_12); -x_160 = l_Lean_Syntax_node2(x_12, x_137, x_158, x_159); -lean_inc(x_121); +x_164 = l_Lean_Syntax_node2(x_12, x_141, x_162, x_163); +lean_inc(x_125); lean_inc(x_12); -x_161 = l_Lean_Syntax_node3(x_12, x_147, x_153, x_121, x_160); -x_162 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__50; -lean_inc(x_111); +x_165 = l_Lean_Syntax_node3(x_12, x_151, x_157, x_125, x_164); +x_166 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52; +lean_inc(x_115); lean_inc(x_12); -x_163 = l_Lean_Syntax_node1(x_12, x_162, x_111); -x_164 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__51; +x_167 = l_Lean_Syntax_node1(x_12, x_166, x_115); +x_168 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; lean_inc(x_12); -x_165 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_165, 0, x_12); -lean_ctor_set(x_165, 1, x_164); -if (lean_obj_tag(x_123) == 0) +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_12); +lean_ctor_set(x_169, 1, x_168); +if (lean_obj_tag(x_127) == 0) { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; -x_166 = l_Lean_quoteNameMk(x_122); +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint8_t x_178; lean_object* x_179; +x_170 = l_Lean_quoteNameMk(x_126); lean_inc(x_12); -x_167 = l_Lean_Syntax_node3(x_12, x_147, x_119, x_121, x_166); -lean_inc_n(x_111, 2); +x_171 = l_Lean_Syntax_node3(x_12, x_151, x_123, x_125, x_170); +lean_inc_n(x_115, 2); lean_inc(x_12); -x_168 = l_Lean_Syntax_node5(x_12, x_109, x_167, x_111, x_148, x_111, x_161); -x_169 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; -lean_inc(x_111); -x_170 = l_Lean_Syntax_node6(x_12, x_169, x_108, x_111, x_168, x_163, x_111, x_165); -x_171 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; -x_172 = 1; -x_173 = l_Lean_Elab_Term_elabTerm(x_170, x_171, x_172, x_172, x_3, x_4, x_5, x_6, x_7, x_8, x_104); -return x_173; +x_172 = l_Lean_Syntax_node5(x_12, x_113, x_171, x_115, x_152, x_115, x_165); +x_173 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; +lean_inc(x_12); +x_174 = l_Lean_Syntax_node1(x_12, x_173, x_172); +x_175 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; +lean_inc(x_115); +x_176 = l_Lean_Syntax_node6(x_12, x_175, x_112, x_115, x_174, x_167, x_115, x_169); +x_177 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; +x_178 = 1; +x_179 = l_Lean_Elab_Term_elabTerm(x_176, x_177, x_178, x_178, x_3, x_4, x_5, x_6, x_7, x_8, x_108); +return x_179; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; -lean_dec(x_122); -x_174 = lean_ctor_get(x_123, 0); -lean_inc(x_174); -lean_dec(x_123); -x_175 = l_String_intercalate(x_143, x_174); -x_176 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56; -x_177 = lean_string_append(x_176, x_175); -lean_dec(x_175); -x_178 = lean_box(2); -x_179 = l_Lean_Syntax_mkNameLit(x_177, x_178); -x_180 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_180, 0, x_179); -lean_ctor_set(x_180, 1, x_114); -x_181 = lean_array_mk(x_180); -x_182 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; -x_183 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_183, 0, x_178); -lean_ctor_set(x_183, 1, x_182); -lean_ctor_set(x_183, 2, x_181); +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; +lean_dec(x_126); +x_180 = lean_ctor_get(x_127, 0); +lean_inc(x_180); +lean_dec(x_127); +x_181 = l_String_intercalate(x_147, x_180); +x_182 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58; +x_183 = lean_string_append(x_182, x_181); +lean_dec(x_181); +x_184 = lean_box(2); +x_185 = l_Lean_Syntax_mkNameLit(x_183, x_184); +x_186 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_118); +x_187 = lean_array_mk(x_186); +x_188 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57; +x_189 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_189, 0, x_184); +lean_ctor_set(x_189, 1, x_188); +lean_ctor_set(x_189, 2, x_187); lean_inc(x_12); -x_184 = l_Lean_Syntax_node3(x_12, x_147, x_119, x_121, x_183); -lean_inc_n(x_111, 2); +x_190 = l_Lean_Syntax_node3(x_12, x_151, x_123, x_125, x_189); +lean_inc_n(x_115, 2); lean_inc(x_12); -x_185 = l_Lean_Syntax_node5(x_12, x_109, x_184, x_111, x_148, x_111, x_161); -x_186 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; -lean_inc(x_111); -x_187 = l_Lean_Syntax_node6(x_12, x_186, x_108, x_111, x_185, x_163, x_111, x_165); -x_188 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__53; -x_189 = 1; -x_190 = l_Lean_Elab_Term_elabTerm(x_187, x_188, x_189, x_189, x_3, x_4, x_5, x_6, x_7, x_8, x_104); -return x_190; +x_191 = l_Lean_Syntax_node5(x_12, x_113, x_190, x_115, x_152, x_115, x_165); +x_192 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__10; +lean_inc(x_12); +x_193 = l_Lean_Syntax_node1(x_12, x_192, x_191); +x_194 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__4; +lean_inc(x_115); +x_195 = l_Lean_Syntax_node6(x_12, x_194, x_112, x_115, x_193, x_167, x_115, x_169); +x_196 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55; +x_197 = 1; +x_198 = l_Lean_Elab_Term_elabTerm(x_195, x_196, x_197, x_197, x_3, x_4, x_5, x_6, x_7, x_8, x_108); +return x_198; } } } @@ -14659,7 +14693,7 @@ lean_inc(x_27); x_47 = l_Lean_Syntax_node3(x_27, x_46, x_29, x_43, x_45); lean_inc(x_27); x_48 = l_Lean_Syntax_node1(x_27, x_41, x_47); -x_49 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; +x_49 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; x_50 = l_Lean_Syntax_node2(x_27, x_49, x_39, x_48); x_51 = l_Lean_Widget_elabWidgetInstanceSpecAux(x_13, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_32); return x_51; @@ -14708,7 +14742,7 @@ lean_inc(x_27); x_69 = l_Lean_Syntax_node3(x_27, x_68, x_62, x_65, x_67); lean_inc(x_27); x_70 = l_Lean_Syntax_node1(x_27, x_63, x_69); -x_71 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; +x_71 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; x_72 = l_Lean_Syntax_node2(x_27, x_71, x_60, x_70); x_73 = l_Lean_Widget_elabWidgetInstanceSpecAux(x_13, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_53); return x_73; @@ -16641,7 +16675,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowP { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_12 = lean_box(0); -x_13 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52; +x_13 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54; lean_inc(x_1); x_14 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_14, 0, x_1); @@ -16728,7 +16762,7 @@ LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowP { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_12 = lean_box(0); -x_13 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__52; +x_13 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__54; lean_inc(x_1); x_14 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_14, 0, x_1); @@ -17981,9 +18015,9 @@ x_254 = lean_ctor_get(x_252, 0); lean_inc(x_254); lean_dec(x_252); x_255 = lean_environment_main_module(x_254); -x_256 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__33; +x_256 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__35; x_257 = l_Lean_addMacroScope(x_255, x_256, x_250); -x_258 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; +x_258 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__34; x_259 = l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__7; lean_inc(x_249); x_260 = lean_alloc_ctor(3, 4, 0); @@ -17994,7 +18028,7 @@ lean_ctor_set(x_260, 3, x_259); x_261 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__7; lean_inc(x_249); x_262 = l_Lean_Syntax_node1(x_249, x_261, x_239); -x_263 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__30; +x_263 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__32; x_264 = l_Lean_Syntax_node2(x_249, x_263, x_260, x_262); x_265 = l_Array_forIn_x27Unsafe_loop___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__13___closed__9; x_266 = 1; @@ -18737,7 +18771,7 @@ x_1 = l_Lean_Widget_instInhabitedUserWidgetDefinition___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1() { _start: { lean_object* x_1; @@ -18745,7 +18779,7 @@ x_1 = lean_mk_string_unchecked("name", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2() { _start: { lean_object* x_1; @@ -18753,7 +18787,7 @@ x_1 = lean_mk_string_unchecked("javascript", 10, 10); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637_(lean_object* x_1) { _start: { uint8_t x_2; @@ -18765,7 +18799,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = lean_ctor_get(x_1, 1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_3); -x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; lean_ctor_set(x_1, 1, x_5); lean_ctor_set(x_1, 0, x_6); x_7 = lean_box(0); @@ -18774,7 +18808,7 @@ lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_4); -x_10 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2; +x_10 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -18802,7 +18836,7 @@ lean_inc(x_18); lean_dec(x_1); x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_18); -x_21 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_21 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -18812,7 +18846,7 @@ lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); x_25 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_25, 0, x_19); -x_26 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2; +x_26 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -18836,7 +18870,7 @@ static lean_object* _init_l_Lean_Widget_instToJsonUserWidgetDefinition___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637_), 1, 0); return x_1; } } @@ -18848,7 +18882,7 @@ x_1 = l_Lean_Widget_instToJsonUserWidgetDefinition___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1() { _start: { lean_object* x_1; @@ -18856,125 +18890,125 @@ x_1 = lean_mk_string_unchecked("UserWidgetDefinition", 20, 20); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; -x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1; +x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3; x_2 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) @@ -18986,7 +19020,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8; +x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -18998,7 +19032,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8; +x_9 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -19012,7 +19046,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2; +x_13 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -19023,7 +19057,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12; +x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -19035,7 +19069,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12; +x_20 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -19078,7 +19112,7 @@ static lean_object* _init_l_Lean_Widget_instFromJsonUserWidgetDefinition___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689_), 1, 0); return x_1; } } @@ -19157,7 +19191,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalU _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_6 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_7 = l_Lean_Environment_evalConstCheck___rarg(x_1, x_5, x_6, x_2); x_8 = l_Lean_ofExcept___at___private_Lean_Widget_UserWidget_0__Lean_Widget_evalUserWidgetDefinitionUnsafe___spec__1___rarg(x_3, x_4, x_7); return x_8; @@ -19216,7 +19250,7 @@ return x_6; LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -if (lean_obj_tag(x_4) == 7) +if (lean_obj_tag(x_4) == 8) { lean_object* x_6; lean_object* x_7; x_6 = lean_ctor_get(x_4, 0); @@ -19376,7 +19410,7 @@ lean_dec(x_2); return x_6; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -19386,17 +19420,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3() { _start: { lean_object* x_1; @@ -19404,17 +19438,17 @@ x_1 = lean_mk_string_unchecked("RpcEncodablePacket", 18, 18); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5() { _start: { lean_object* x_1; @@ -19422,37 +19456,37 @@ x_1 = lean_mk_string_unchecked("_@", 2, 2); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9() { _start: { lean_object* x_1; @@ -19460,17 +19494,17 @@ x_1 = lean_mk_string_unchecked("UserWidget", 10, 10); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11() { _start: { lean_object* x_1; @@ -19478,141 +19512,141 @@ x_1 = lean_mk_string_unchecked("_hyg", 4, 4); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12; -x_2 = lean_unsigned_to_nat(4074u); +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12; +x_2 = lean_unsigned_to_nat(4079u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14; x_2 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__17; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__18() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__25; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__21() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; +x_1 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__41; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__24() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25() { _start: { lean_object* x_1; @@ -19620,7 +19654,7 @@ x_1 = lean_mk_string_unchecked("range", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26() { _start: { lean_object* x_1; @@ -19628,48 +19662,48 @@ x_1 = lean_mk_string_unchecked("range\?", 6, 6); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__30() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31() { _start: { lean_object* x_1; @@ -19677,76 +19711,76 @@ x_1 = lean_mk_string_unchecked("name\?", 5, 5); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__35() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; +x_2 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; lean_inc(x_1); x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_2); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; +x_5 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; lean_inc(x_1); x_6 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_5); x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); lean_dec(x_6); -x_8 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; +x_8 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; lean_inc(x_1); x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); -x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25; +x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25; lean_inc(x_1); x_12 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_352____spec__1(x_1, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); lean_dec(x_12); -x_14 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_14 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_352____spec__1(x_1, x_14); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) @@ -19780,23 +19814,23 @@ return x_21; } } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463_() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1; +x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -19805,7 +19839,7 @@ x_3 = lean_ctor_get(x_1, 1); x_4 = lean_ctor_get(x_1, 2); x_5 = lean_ctor_get(x_1, 3); x_6 = lean_ctor_get(x_1, 4); -x_7 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__13; +x_7 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__15; lean_inc(x_2); x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); @@ -19814,7 +19848,7 @@ x_9 = lean_box(0); x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__21; +x_11 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__23; lean_inc(x_3); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); @@ -19822,7 +19856,7 @@ lean_ctor_set(x_12, 1, x_3); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_9); -x_14 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__37; +x_14 = l_Lean_Widget_elabWidgetInstanceSpecAux___closed__39; lean_inc(x_4); x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); @@ -19830,9 +19864,9 @@ lean_ctor_set(x_15, 1, x_4); x_16 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_9); -x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25; +x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25; x_18 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_551____spec__1(x_17, x_5); -x_19 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1; +x_19 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1; x_20 = l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_551____spec__1(x_19, x_6); x_21 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_21, 0, x_20); @@ -19855,32 +19889,32 @@ x_28 = l_Lean_Json_mkObj(x_27); return x_28; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_1); +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_1); lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471____boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537_() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1; +x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4074_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4079_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint64_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -19923,7 +19957,7 @@ lean_ctor_set(x_19, 1, x_11); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_17); lean_ctor_set(x_19, 4, x_17); -x_20 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_19); +x_20 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_19); lean_dec(x_19); lean_ctor_set(x_13, 0, x_20); return x_13; @@ -19945,7 +19979,7 @@ lean_ctor_set(x_24, 1, x_11); lean_ctor_set(x_24, 2, x_16); lean_ctor_set(x_24, 3, x_17); lean_ctor_set(x_24, 4, x_18); -x_25 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_24); +x_25 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_24); lean_dec(x_24); lean_ctor_set(x_13, 0, x_25); return x_13; @@ -19966,7 +20000,7 @@ lean_ctor_set(x_29, 1, x_11); lean_ctor_set(x_29, 2, x_16); lean_ctor_set(x_29, 3, x_17); lean_ctor_set(x_29, 4, x_28); -x_30 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_29); +x_30 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_29); lean_dec(x_29); lean_ctor_set(x_13, 0, x_30); return x_13; @@ -19994,7 +20028,7 @@ lean_ctor_set(x_35, 1, x_11); lean_ctor_set(x_35, 2, x_31); lean_ctor_set(x_35, 3, x_33); lean_ctor_set(x_35, 4, x_33); -x_36 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_35); +x_36 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_35); lean_dec(x_35); x_37 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_37, 0, x_36); @@ -20027,7 +20061,7 @@ lean_ctor_set(x_42, 1, x_11); lean_ctor_set(x_42, 2, x_31); lean_ctor_set(x_42, 3, x_33); lean_ctor_set(x_42, 4, x_41); -x_43 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_42); +x_43 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_42); lean_dec(x_42); x_44 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_44, 0, x_43); @@ -20064,7 +20098,7 @@ lean_ctor_set(x_52, 1, x_11); lean_ctor_set(x_52, 2, x_47); lean_ctor_set(x_52, 3, x_14); lean_ctor_set(x_52, 4, x_51); -x_53 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_52); +x_53 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_52); lean_dec(x_52); lean_ctor_set(x_13, 0, x_53); return x_13; @@ -20086,7 +20120,7 @@ lean_ctor_set(x_57, 1, x_11); lean_ctor_set(x_57, 2, x_47); lean_ctor_set(x_57, 3, x_14); lean_ctor_set(x_57, 4, x_50); -x_58 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_57); +x_58 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_57); lean_dec(x_57); lean_ctor_set(x_13, 0, x_58); return x_13; @@ -20107,7 +20141,7 @@ lean_ctor_set(x_62, 1, x_11); lean_ctor_set(x_62, 2, x_47); lean_ctor_set(x_62, 3, x_14); lean_ctor_set(x_62, 4, x_61); -x_63 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_62); +x_63 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_62); lean_dec(x_62); lean_ctor_set(x_13, 0, x_63); return x_13; @@ -20137,7 +20171,7 @@ lean_ctor_set(x_70, 1, x_11); lean_ctor_set(x_70, 2, x_64); lean_ctor_set(x_70, 3, x_67); lean_ctor_set(x_70, 4, x_69); -x_71 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_70); +x_71 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_70); lean_dec(x_70); lean_ctor_set(x_13, 0, x_71); return x_13; @@ -20168,7 +20202,7 @@ lean_ctor_set(x_76, 1, x_11); lean_ctor_set(x_76, 2, x_64); lean_ctor_set(x_76, 3, x_67); lean_ctor_set(x_76, 4, x_75); -x_77 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_76); +x_77 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_76); lean_dec(x_76); lean_ctor_set(x_13, 0, x_77); return x_13; @@ -20212,7 +20246,7 @@ lean_ctor_set(x_86, 1, x_11); lean_ctor_set(x_86, 2, x_78); lean_ctor_set(x_86, 3, x_83); lean_ctor_set(x_86, 4, x_85); -x_87 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_86); +x_87 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_86); lean_dec(x_86); x_88 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_88, 0, x_87); @@ -20245,7 +20279,7 @@ lean_ctor_set(x_93, 1, x_11); lean_ctor_set(x_93, 2, x_78); lean_ctor_set(x_93, 3, x_83); lean_ctor_set(x_93, 4, x_92); -x_94 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4466_(x_93); +x_94 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4471_(x_93); lean_dec(x_93); x_95 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_95, 0, x_94); @@ -20256,7 +20290,7 @@ return x_95; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { uint64_t x_3; lean_object* x_4; lean_object* x_5; @@ -20267,7 +20301,7 @@ lean_ctor_set(x_5, 0, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -20276,7 +20310,7 @@ lean_ctor_set(x_3, 0, x_1); return x_3; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1() { _start: { lean_object* x_1; @@ -20284,7 +20318,7 @@ x_1 = lean_cstr_to_nat("18446744073709551616"); return x_1; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2() { _start: { lean_object* x_1; @@ -20292,17 +20326,17 @@ x_1 = lean_mk_string_unchecked("value '{j}' is too large for `UInt64`", 37, 37); return x_1; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2; +x_1 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4() { _start: { lean_object* x_1; @@ -20310,7 +20344,7 @@ x_1 = lean_mk_string_unchecked("[anonymous]", 11, 11); return x_1; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5() { _start: { lean_object* x_1; @@ -20318,7 +20352,7 @@ x_1 = lean_mk_string_unchecked("expected a `Name`, got '", 24, 24); return x_1; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6() { +static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -20328,11 +20362,11 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_133; lean_object* x_134; -x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153_(x_1); +x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158_(x_1); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); lean_dec(x_3); @@ -20428,7 +20462,7 @@ if (x_148 == 0) { lean_object* x_149; lean_object* x_150; uint8_t x_151; x_149 = lean_ctor_get(x_134, 0); -x_150 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4; +x_150 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4; x_151 = lean_string_dec_eq(x_149, x_150); if (x_151 == 0) { @@ -20477,7 +20511,7 @@ lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_dec(x_152); x_159 = lean_unsigned_to_nat(80u); x_160 = l_Lean_Json_pretty(x_133, x_159); -x_161 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5; +x_161 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5; x_162 = lean_string_append(x_161, x_160); lean_dec(x_160); x_163 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__2; @@ -20522,7 +20556,7 @@ lean_object* x_170; lean_object* x_171; lean_free_object(x_134); lean_dec(x_149); lean_dec(x_133); -x_170 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6; +x_170 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6; x_171 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableWidgetInstance_dec____x40_Lean_Widget_Types___hyg_3____spec__3(x_170, x_2); if (lean_obj_tag(x_171) == 0) { @@ -20561,7 +20595,7 @@ lean_object* x_176; lean_object* x_177; uint8_t x_178; x_176 = lean_ctor_get(x_134, 0); lean_inc(x_176); lean_dec(x_134); -x_177 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4; +x_177 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4; x_178 = lean_string_dec_eq(x_176, x_177); if (x_178 == 0) { @@ -20612,7 +20646,7 @@ lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_dec(x_179); x_187 = lean_unsigned_to_nat(80u); x_188 = l_Lean_Json_pretty(x_133, x_187); -x_189 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5; +x_189 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5; x_190 = lean_string_append(x_189, x_188); lean_dec(x_188); x_191 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__2; @@ -20657,7 +20691,7 @@ else lean_object* x_199; lean_object* x_200; lean_dec(x_176); lean_dec(x_133); -x_199 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6; +x_199 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6; x_200 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableWidgetInstance_dec____x40_Lean_Widget_Types___hyg_3____spec__3(x_199, x_2); if (lean_obj_tag(x_200) == 0) { @@ -20725,13 +20759,13 @@ lean_object* x_126; lean_object* x_127; uint8_t x_128; x_126 = lean_ctor_get(x_122, 0); lean_inc(x_126); lean_dec(x_122); -x_127 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1; +x_127 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1; x_128 = lean_nat_dec_le(x_127, x_126); if (x_128 == 0) { lean_object* x_129; lean_object* x_130; x_129 = lean_box(0); -x_130 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1(x_126, x_129); +x_130 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1(x_126, x_129); lean_dec(x_126); x_7 = x_130; goto block_121; @@ -20740,7 +20774,7 @@ else { lean_object* x_131; lean_dec(x_126); -x_131 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3; +x_131 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3; x_7 = x_131; goto block_121; } @@ -21322,30 +21356,30 @@ return x_108; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__1(x_1, x_2); +x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____lambda__2(x_1, x_2); +x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____lambda__2(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074_(x_1, x_2); +x_3 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21354,7 +21388,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4074_), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4079_), 2, 0); return x_1; } } @@ -21362,7 +21396,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____boxed), 2, 0); return x_1; } } @@ -21386,7 +21420,7 @@ x_1 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__3; return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1() { _start: { lean_object* x_1; @@ -21394,83 +21428,83 @@ x_1 = lean_mk_string_unchecked("widgets", 7, 7); return x_1; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12; -x_2 = lean_unsigned_to_nat(4792u); +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12; +x_2 = lean_unsigned_to_nat(4797u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3; x_2 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_240____lambda__2___closed__4; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5; x_2 = 1; x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__3; x_4 = l_Lean_Name_toString(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__8() { +static lean_object* _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7; +x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_872____closed__9; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_2012____spec__2(x_1, x_2); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -21489,27 +21523,27 @@ return x_6; } } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900_() { +static lean_object* _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1; +x_1 = l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4903_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4908_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1; +x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -21526,23 +21560,23 @@ x_9 = l_Lean_Json_mkObj(x_8); return x_9; } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4903_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4908_), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938_() { +static lean_object* _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943_() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1; +x_1 = l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -21561,7 +21595,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_array_uget(x_3, x_2); x_8 = lean_unsigned_to_nat(0u); x_9 = lean_array_uset(x_3, x_2, x_8); -x_10 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4074_(x_7, x_4); +x_10 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_enc____x40_Lean_Widget_UserWidget___hyg_4079_(x_7, x_4); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); @@ -21577,13 +21611,13 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797_(lean_object* x_1, lean_object* x_2) { _start: { size_t x_3; size_t x_4; lean_object* x_5; uint8_t x_6; x_3 = lean_array_size(x_1); x_4 = 0; -x_5 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(x_3, x_4, x_1, x_2); +x_5 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(x_3, x_4, x_1, x_2); x_6 = !lean_is_exclusive(x_5); if (x_6 == 0) { @@ -21593,7 +21627,7 @@ x_8 = lean_array_size(x_7); x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_8, x_4, x_7); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); -x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4903_(x_10); +x_11 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4908_(x_10); lean_ctor_set(x_5, 0, x_11); return x_5; } @@ -21609,7 +21643,7 @@ x_14 = lean_array_size(x_12); x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1627____spec__2(x_14, x_4, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4903_(x_16); +x_17 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4908_(x_16); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_13); @@ -21617,7 +21651,7 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -21625,11 +21659,11 @@ x_5 = lean_unbox_usize(x_1); lean_dec(x_1); x_6 = lean_unbox_usize(x_2); lean_dec(x_2); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(x_5, x_6, x_3, x_4); +x_7 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(x_5, x_6, x_3, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -21647,7 +21681,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_7 = lean_array_uget(x_3, x_2); x_8 = lean_unsigned_to_nat(0u); x_9 = lean_array_uset(x_3, x_2, x_8); -x_10 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074_(x_7, x_4); +x_10 = l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079_(x_7, x_4); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -21684,7 +21718,7 @@ goto _start; } } } -static lean_object* _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1() { +static lean_object* _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1() { _start: { lean_object* x_1; @@ -21692,11 +21726,11 @@ x_1 = lean_mk_string_unchecked("expected JSON array, got '", 26, 26); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; -x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818_(x_1); +x_3 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823_(x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) { @@ -21716,7 +21750,7 @@ x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); x_11 = lean_array_size(x_10); -x_12 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(x_11, x_8, x_10, x_2); +x_12 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(x_11, x_8, x_10, x_2); if (lean_obj_tag(x_12) == 0) { uint8_t x_13; @@ -21761,7 +21795,7 @@ else lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_19 = lean_unsigned_to_nat(80u); x_20 = l_Lean_Json_pretty(x_5, x_19); -x_21 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1; +x_21 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1; x_22 = lean_string_append(x_21, x_20); lean_dec(x_20); x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__2; @@ -21790,7 +21824,7 @@ x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); lean_dec(x_29); x_31 = lean_array_size(x_30); -x_32 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(x_31, x_28, x_30, x_2); +x_32 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(x_31, x_28, x_30, x_2); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -21837,7 +21871,7 @@ else lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_39 = lean_unsigned_to_nat(80u); x_40 = l_Lean_Json_pretty(x_25, x_39); -x_41 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1; +x_41 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1; x_42 = lean_string_append(x_41, x_40); lean_dec(x_40); x_43 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___closed__2; @@ -21849,7 +21883,7 @@ return x_45; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -21857,16 +21891,16 @@ x_5 = lean_unbox_usize(x_1); lean_dec(x_1); x_6 = lean_unbox_usize(x_2); lean_dec(x_2); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____spec__1(x_5, x_6, x_3, x_4); +x_7 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____spec__1(x_5, x_6, x_3, x_4); lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792_(x_1, x_2); +x_3 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21875,7 +21909,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4792_), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_enc____x40_Lean_Widget_UserWidget___hyg_4797_), 2, 0); return x_1; } } @@ -21883,7 +21917,7 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____boxed), 2, 0); return x_1; } } @@ -21982,7 +22016,7 @@ x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); x_11 = lean_ctor_get(x_4, 2); -x_12 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_12 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_13 = l_Lean_Environment_evalConstCheck___rarg(x_10, x_11, x_12, x_1); x_14 = l_Lean_ofExcept___at_Lean_Widget_getWidgets___spec__2(x_13, x_2, x_3, x_4, x_5, x_9); lean_dec(x_13); @@ -22040,7 +22074,7 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; x_24 = lean_ctor_get(x_16, 0); x_25 = l_Lean_ConstantInfo_type(x_24); lean_dec(x_24); -x_26 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_26 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_27 = l_Lean_Expr_isConstOf(x_25, x_26); lean_dec(x_25); if (x_27 == 0) @@ -22125,7 +22159,7 @@ lean_inc(x_48); lean_dec(x_16); x_49 = l_Lean_ConstantInfo_type(x_48); lean_dec(x_48); -x_50 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_50 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_51 = l_Lean_Expr_isConstOf(x_49, x_50); lean_dec(x_49); if (x_51 == 0) @@ -22256,7 +22290,7 @@ lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; x_49 = lean_ctor_get(x_46, 0); x_50 = l_Lean_ConstantInfo_type(x_49); lean_dec(x_49); -x_51 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_51 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_52 = l_Lean_Expr_isConstOf(x_50, x_51); lean_dec(x_50); if (x_52 == 0) @@ -22392,7 +22426,7 @@ lean_inc(x_86); lean_dec(x_46); x_87 = l_Lean_ConstantInfo_type(x_86); lean_dec(x_86); -x_88 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2; +x_88 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2; x_89 = l_Lean_Expr_isConstOf(x_87, x_88); lean_dec(x_87); if (x_89 == 0) @@ -23023,7 +23057,7 @@ lean_dec(x_4); return x_6; } } -static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1() { +static lean_object* _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23033,11 +23067,11 @@ x_3 = l_Lean_Server_instRpcEncodableOfFromJsonOfToJson___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1; +x_3 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1; x_4 = l_Lean_Widget_instRpcEncodableGetWidgetsResponse; x_5 = lean_alloc_closure((void*)(l_Lean_Server_wrapRpcProcedure___elambda__1___boxed), 10, 6); lean_closure_set(x_5, 0, x_1); @@ -23049,12 +23083,12 @@ lean_closure_set(x_5, 5, x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_1); -x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2(x_1, x_2); +x_5 = l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2(x_1, x_2); x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_1367____spec__1___lambda__1___closed__1; x_7 = lean_st_ref_take(x_6, x_4); x_8 = lean_ctor_get(x_7, 0); @@ -23084,7 +23118,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; lean_object* x_7; uint8_t x_8; @@ -23102,7 +23136,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_7); x_12 = lean_box(0); -x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1(x_1, x_2, x_12, x_10); +x_13 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1(x_1, x_2, x_12, x_10); return x_13; } else @@ -23134,7 +23168,7 @@ if (x_21 == 0) { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1(x_1, x_2, x_22, x_20); +x_23 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1(x_1, x_2, x_22, x_20); return x_23; } else @@ -23156,7 +23190,7 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -23222,13 +23256,13 @@ x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); x_29 = lean_box(0); -x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); +x_30 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2(x_1, x_2, x_10, x_29, x_28); lean_dec(x_10); return x_30; } } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1() { _start: { lean_object* x_1; @@ -23236,18 +23270,18 @@ x_1 = lean_mk_string_unchecked("getWidgets", 10, 10); return x_1; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__1; x_2 = l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalModuleUnsafe___closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1; +x_3 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3() { +static lean_object* _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3() { _start: { lean_object* x_1; @@ -23255,30 +23289,30 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Widget_getWidgets), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2; -x_3 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3; -x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1(x_2, x_3, x_1); +x_2 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2; +x_3 = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3; +x_4 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { lean_object* x_6; -x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Server_registerBuiltinRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; @@ -23700,6 +23734,10 @@ l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55 = _init_l_Lean_Widget_elabW lean_mark_persistent(l_Lean_Widget_elabWidgetInstanceSpecAux___closed__55); l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56 = _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56(); lean_mark_persistent(l_Lean_Widget_elabWidgetInstanceSpecAux___closed__56); +l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57 = _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57(); +lean_mark_persistent(l_Lean_Widget_elabWidgetInstanceSpecAux___closed__57); +l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58 = _init_l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58(); +lean_mark_persistent(l_Lean_Widget_elabWidgetInstanceSpecAux___closed__58); l_Lean_Widget_elabWidgetInstanceSpec___closed__1 = _init_l_Lean_Widget_elabWidgetInstanceSpec___closed__1(); lean_mark_persistent(l_Lean_Widget_elabWidgetInstanceSpec___closed__1); l_Lean_Widget_elabWidgetInstanceSpec___closed__2 = _init_l_Lean_Widget_elabWidgetInstanceSpec___closed__2(); @@ -23858,132 +23896,132 @@ l_Lean_Widget_instInhabitedUserWidgetDefinition___closed__1 = _init_l_Lean_Widge lean_mark_persistent(l_Lean_Widget_instInhabitedUserWidgetDefinition___closed__1); l_Lean_Widget_instInhabitedUserWidgetDefinition = _init_l_Lean_Widget_instInhabitedUserWidgetDefinition(); lean_mark_persistent(l_Lean_Widget_instInhabitedUserWidgetDefinition); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__1); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3634____closed__2); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__1); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3637____closed__2); l_Lean_Widget_instToJsonUserWidgetDefinition___closed__1 = _init_l_Lean_Widget_instToJsonUserWidgetDefinition___closed__1(); lean_mark_persistent(l_Lean_Widget_instToJsonUserWidgetDefinition___closed__1); l_Lean_Widget_instToJsonUserWidgetDefinition = _init_l_Lean_Widget_instToJsonUserWidgetDefinition(); lean_mark_persistent(l_Lean_Widget_instToJsonUserWidgetDefinition); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__1); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__2); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__3); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__4); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__5); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__6); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__7); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__8); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__9); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__10); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__11); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3686____closed__12); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__1); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__2); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__3); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__4); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__5); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__6); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__7); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__8); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__9); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__10); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__11); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_3689____closed__12); l_Lean_Widget_instFromJsonUserWidgetDefinition___closed__1 = _init_l_Lean_Widget_instFromJsonUserWidgetDefinition___closed__1(); lean_mark_persistent(l_Lean_Widget_instFromJsonUserWidgetDefinition___closed__1); l_Lean_Widget_instFromJsonUserWidgetDefinition = _init_l_Lean_Widget_instFromJsonUserWidgetDefinition(); lean_mark_persistent(l_Lean_Widget_instFromJsonUserWidgetDefinition); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__1); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__2); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__3); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__4); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__5); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__6); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__7); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__8); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__9); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__10); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__11); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__12); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__13); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__14); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__15); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__16); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__17); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__18 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__18(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__18); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__19); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__20); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__21 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__21(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__21); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__22); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__23); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__24 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__24(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__24); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__25); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__26); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__27); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__28); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__29); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__30 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__30(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__30); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__31); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__32); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__33); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__34); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__35 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__35(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4153____closed__35); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463____closed__1); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463_(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4463_); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537____closed__1); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537_(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4537_); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__1); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__2); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__3); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__4); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__5); -l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4074____closed__6); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__1); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__2); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__3); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__4); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__5); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__6); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__7); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__8); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__9); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__10); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__11); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__12); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__13); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__14); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__15); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__16); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__17); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__18 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__18(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__18); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__19); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__20); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__21 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__21(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__21); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__22); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__23); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__24 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__24(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__24); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__25); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__26); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__27); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__28); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__29); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__30 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__30(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__30); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__31); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__32); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__33); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__34); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__35 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__35(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4158____closed__35); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468____closed__1); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468_(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4468_); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542____closed__1); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542_(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4542_); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__1); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__2); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__3); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__4); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__5); +l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance_dec____x40_Lean_Widget_UserWidget___hyg_4079____closed__6); l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1(); lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__1); l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__2 = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__2(); @@ -23992,32 +24030,32 @@ l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__3 = _init_l_Lean_Wid lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance___closed__3); l_Lean_Widget_instRpcEncodablePanelWidgetInstance = _init_l_Lean_Widget_instRpcEncodablePanelWidgetInstance(); lean_mark_persistent(l_Lean_Widget_instRpcEncodablePanelWidgetInstance); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__1); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__2); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__3); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__4); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__5); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__6); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__7); -l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__8(); -lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4818____closed__8); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900____closed__1); -l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900_(); -lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4900_); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938____closed__1); -l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938_(); -lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4938_); -l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1 = _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4792____closed__1); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__1); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__2); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__3); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__4); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__5); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__6); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__7); +l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__8 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__8(); +lean_mark_persistent(l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4823____closed__8); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1 = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905____closed__1); +l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905_ = _init_l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905_(); +lean_mark_persistent(l_Lean_Widget_instFromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4905_); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1 = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943____closed__1); +l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943_ = _init_l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943_(); +lean_mark_persistent(l_Lean_Widget_instToJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_4943_); +l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1 = _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1(); +lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetWidgetsResponse_dec____x40_Lean_Widget_UserWidget___hyg_4797____closed__1); l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__1 = _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__1(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__1); l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__2 = _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__2(); @@ -24026,15 +24064,15 @@ l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__3 = _init_l_Lean_Widg lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetWidgetsResponse___closed__3); l_Lean_Widget_instRpcEncodableGetWidgetsResponse = _init_l_Lean_Widget_instRpcEncodableGetWidgetsResponse(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableGetWidgetsResponse); -l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1(); -lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____spec__2___closed__1); -l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__1); -l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__2); -l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3(); -lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406____closed__3); -if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5406_(lean_io_mk_world()); +l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1 = _init_l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1(); +lean_mark_persistent(l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____spec__2___closed__1); +l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__1); +l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__2); +l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3 = _init_l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3(); +lean_mark_persistent(l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411____closed__3); +if (builtin) {res = l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_5411_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Std.c b/stage0/stdlib/Std.c index e6192377bdf4..62f8c620a8f2 100644 --- a/stage0/stdlib/Std.c +++ b/stage0/stdlib/Std.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Std -// Imports: Std.Data Std.Sat Std.Tactic Std.Internal +// Imports: Std.Data Std.Sat Std.Time Std.Tactic Std.Internal #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -15,6 +15,7 @@ extern "C" { #endif lean_object* initialize_Std_Data(uint8_t builtin, lean_object*); lean_object* initialize_Std_Sat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time(uint8_t builtin, lean_object*); lean_object* initialize_Std_Tactic(uint8_t builtin, lean_object*); lean_object* initialize_Std_Internal(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -28,6 +29,9 @@ lean_dec_ref(res); res = initialize_Std_Sat(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Std_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Std_Tactic(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Data/Rat.c b/stage0/stdlib/Std/Internal/Rat.c similarity index 63% rename from stage0/stdlib/Lean/Data/Rat.c rename to stage0/stdlib/Std/Internal/Rat.c index d27b829b758b..5bd37d5a59a2 100644 --- a/stage0/stdlib/Lean/Data/Rat.c +++ b/stage0/stdlib/Std/Internal/Rat.c @@ -1,5 +1,5 @@ // Lean compiler output -// Module: Lean.Data.Rat +// Module: Std.Internal.Rat // Imports: Init.NotationExtra Init.Data.ToString.Macro Init.Data.Int.DivMod Init.Data.Nat.Gcd #include #if defined(__clang__) @@ -14,85 +14,85 @@ extern "C" { #endif lean_object* lean_nat_gcd(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_instBEqRat; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_add(lean_object*, lean_object*); lean_object* lean_int_mod(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instLE; -LEAN_EXPORT lean_object* l_Lean_Rat_instOfNat(lean_object*); -static lean_object* l_Lean_instBEqRat___closed__1; -LEAN_EXPORT lean_object* l_Lean_instReprRat(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_neg(lean_object*); -static lean_object* l_Lean_instToStringRat___closed__2; -LEAN_EXPORT uint8_t l_Lean_instDecidableEqRat(lean_object*, lean_object*); -static lean_object* l_Lean_Rat_instSub___closed__1; -LEAN_EXPORT lean_object* l_Lean_Rat_instSub; -LEAN_EXPORT lean_object* l_Lean_mkRat(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instDecidableEqRat___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_div___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_normalize(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instCoeInt(lean_object*); -LEAN_EXPORT uint8_t l_Lean_Rat_instDecidableLe(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_normalize___boxed(lean_object*); -LEAN_EXPORT uint8_t l_Lean_Rat_instDecidableLt(lean_object*, lean_object*); -static lean_object* l_Lean_Rat_instMul___closed__1; -LEAN_EXPORT lean_object* l_Lean_Rat_lt___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_mkRat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDiv___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instLE; +LEAN_EXPORT lean_object* l_Std_Internal_instReprRat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_div___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Internal_Rat_floor___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instSub; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instAdd; +LEAN_EXPORT uint8_t l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_instToStringRat(lean_object*); uint8_t lean_int_dec_le(lean_object*, lean_object*); -static lean_object* l_Lean_instInhabitedRat___closed__2; -static lean_object* l_Lean_instInhabitedRat___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110_(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instMul; -LEAN_EXPORT lean_object* l_Lean_Rat_instLT; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instMul; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instLT; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDecidableLt___boxed(lean_object*, lean_object*); uint8_t l_instDecidableNot___rarg(uint8_t); -LEAN_EXPORT lean_object* l_Lean_Rat_instNeg; -LEAN_EXPORT lean_object* l_Lean_Rat_instDecidableLt___boxed(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_inv(lean_object*); -static lean_object* l_Lean_Rat_instNeg___closed__1; -LEAN_EXPORT lean_object* l_Lean_instInhabitedRat; -static lean_object* l_Lean_Rat_instAdd___closed__1; -static lean_object* l_Lean_instReprRat___closed__2; -static lean_object* l_Lean_instReprRat___closed__1; +static lean_object* l_Std_Internal_instInhabitedRat___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_normalize(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instCoeInt(lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_Rat_instDecidableLt(lean_object*, lean_object*); +static lean_object* l_Std_Internal_instToStringRat___closed__2; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDecidableLe___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_mul(lean_object*, lean_object*); lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_div(lean_object*, lean_object*); lean_object* lean_int_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instToStringRat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instDecidableLe___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_instToStringRat___closed__1; -LEAN_EXPORT lean_object* l_Lean_Rat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_instReprRat___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110____boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_Rat_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instOfNat(lean_object*); +static lean_object* l_Std_Internal_Rat_instSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_ceil(lean_object*); +static lean_object* l_Std_Internal_instInhabitedRat___closed__2; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_lt___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_Rat_instDecidableLe(lean_object*, lean_object*); +static lean_object* l_Std_Internal_instToStringRat___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_instReprRat___boxed(lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_isInt___boxed(lean_object*); +static lean_object* l_Std_Internal_instToStringRat___closed__3; +static lean_object* l_Std_Internal_mkRat___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_isInt___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111____boxed(lean_object*, lean_object*); lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDiv(lean_object*, lean_object*); lean_object* lean_int_mul(lean_object*, lean_object*); -static lean_object* l_Lean_Rat_floor___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instAdd; -LEAN_EXPORT uint8_t l_Lean_Rat_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_mul(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_mul___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_sub(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_Rat_isInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_mul___boxed(lean_object*, lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -static lean_object* l_Lean_instToStringRat___closed__3; uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Internal_instBEqRat___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_normalize___boxed(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_floor(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_ceil(lean_object*); -static lean_object* l_Lean_mkRat___closed__1; +static lean_object* l_Std_Internal_Rat_instAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Rat_inv(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instNeg; +static lean_object* l_Std_Internal_Rat_instNeg___closed__1; +static lean_object* l_Std_Internal_Rat_instMul___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_instInhabitedRat; lean_object* lean_int_add(lean_object*, lean_object*); uint8_t lean_int_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_Rat_isInt(lean_object*); +static lean_object* l_Std_Internal_instReprRat___closed__2; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_sub(lean_object*, lean_object*); lean_object* lean_int_ediv(lean_object*, lean_object*); lean_object* lean_int_neg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instDiv___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_instDecidableEqRat___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_instDecidableEqRat(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Rat_instDiv(lean_object*, lean_object*); +static lean_object* l_Std_Internal_instReprRat___closed__1; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instBEqRat; -static lean_object* _init_l_Lean_instInhabitedRat___closed__1() { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_floor(lean_object*); +static lean_object* _init_l_Std_Internal_instInhabitedRat___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -101,11 +101,11 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_instInhabitedRat___closed__2() { +static lean_object* _init_l_Std_Internal_instInhabitedRat___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedRat___closed__1; +x_1 = l_Std_Internal_instInhabitedRat___closed__1; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -113,15 +113,15 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_instInhabitedRat() { +static lean_object* _init_l_Std_Internal_instInhabitedRat() { _start: { lean_object* x_1; -x_1 = l_Lean_instInhabitedRat___closed__2; +x_1 = l_Std_Internal_instInhabitedRat___closed__2; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -144,34 +144,34 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36_(x_1, x_2); +x_3 = l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_instBEqRat___closed__1() { +static lean_object* _init_l_Std_Internal_instBEqRat___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Rat_0__Lean_beqRat____x40_Lean_Data_Rat___hyg_36____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Std_Internal_Rat_0__Std_Internal_beqRat____x40_Std_Internal_Rat___hyg_37____boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_instBEqRat() { +static lean_object* _init_l_Std_Internal_instBEqRat() { _start: { lean_object* x_1; -x_1 = l_Lean_instBEqRat___closed__1; +x_1 = l_Std_Internal_instBEqRat___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -194,37 +194,37 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110_(x_1, x_2); +x_3 = l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT uint8_t l_Lean_instDecidableEqRat(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_Internal_instDecidableEqRat(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; -x_3 = l___private_Lean_Data_Rat_0__Lean_decEqRat____x40_Lean_Data_Rat___hyg_110_(x_1, x_2); +x_3 = l___private_Std_Internal_Rat_0__Std_Internal_decEqRat____x40_Std_Internal_Rat___hyg_111_(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instDecidableEqRat___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_instDecidableEqRat___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_instDecidableEqRat(x_1, x_2); +x_3 = l_Std_Internal_instDecidableEqRat(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_instToStringRat___closed__1() { +static lean_object* _init_l_Std_Internal_instToStringRat___closed__1() { _start: { lean_object* x_1; @@ -232,7 +232,7 @@ x_1 = lean_mk_string_unchecked("", 0, 0); return x_1; } } -static lean_object* _init_l_Lean_instToStringRat___closed__2() { +static lean_object* _init_l_Std_Internal_instToStringRat___closed__2() { _start: { lean_object* x_1; @@ -240,7 +240,7 @@ x_1 = lean_mk_string_unchecked("/", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_instToStringRat___closed__3() { +static lean_object* _init_l_Std_Internal_instToStringRat___closed__3() { _start: { lean_object* x_1; @@ -248,7 +248,7 @@ x_1 = lean_mk_string_unchecked("-", 1, 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_instToStringRat(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_instToStringRat(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -263,7 +263,7 @@ x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); lean_dec(x_1); x_6 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); -x_7 = l_Lean_instInhabitedRat___closed__1; +x_7 = l_Std_Internal_instInhabitedRat___closed__1; x_8 = lean_int_dec_lt(x_5, x_7); if (x_8 == 0) { @@ -271,10 +271,10 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ x_9 = lean_nat_abs(x_5); lean_dec(x_5); x_10 = l___private_Init_Data_Repr_0__Nat_reprFast(x_9); -x_11 = l_Lean_instToStringRat___closed__1; +x_11 = l_Std_Internal_instToStringRat___closed__1; x_12 = lean_string_append(x_11, x_10); lean_dec(x_10); -x_13 = l_Lean_instToStringRat___closed__2; +x_13 = l_Std_Internal_instToStringRat___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = lean_string_append(x_14, x_6); lean_dec(x_6); @@ -291,13 +291,13 @@ lean_dec(x_17); x_19 = lean_nat_add(x_18, x_3); lean_dec(x_18); x_20 = l___private_Init_Data_Repr_0__Nat_reprFast(x_19); -x_21 = l_Lean_instToStringRat___closed__3; +x_21 = l_Std_Internal_instToStringRat___closed__3; x_22 = lean_string_append(x_21, x_20); lean_dec(x_20); -x_23 = l_Lean_instToStringRat___closed__1; +x_23 = l_Std_Internal_instToStringRat___closed__1; x_24 = lean_string_append(x_23, x_22); lean_dec(x_22); -x_25 = l_Lean_instToStringRat___closed__2; +x_25 = l_Std_Internal_instToStringRat___closed__2; x_26 = lean_string_append(x_24, x_25); x_27 = lean_string_append(x_26, x_6); lean_dec(x_6); @@ -312,7 +312,7 @@ lean_dec(x_2); x_29 = lean_ctor_get(x_1, 0); lean_inc(x_29); lean_dec(x_1); -x_30 = l_Lean_instInhabitedRat___closed__1; +x_30 = l_Std_Internal_instInhabitedRat___closed__1; x_31 = lean_int_dec_lt(x_29, x_30); if (x_31 == 0) { @@ -332,7 +332,7 @@ lean_dec(x_34); x_36 = lean_nat_add(x_35, x_3); lean_dec(x_35); x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_36); -x_38 = l_Lean_instToStringRat___closed__3; +x_38 = l_Std_Internal_instToStringRat___closed__3; x_39 = lean_string_append(x_38, x_37); lean_dec(x_37); return x_39; @@ -340,7 +340,7 @@ return x_39; } } } -static lean_object* _init_l_Lean_instReprRat___closed__1() { +static lean_object* _init_l_Std_Internal_instReprRat___closed__1() { _start: { lean_object* x_1; @@ -348,7 +348,7 @@ x_1 = lean_mk_string_unchecked("(", 1, 1); return x_1; } } -static lean_object* _init_l_Lean_instReprRat___closed__2() { +static lean_object* _init_l_Std_Internal_instReprRat___closed__2() { _start: { lean_object* x_1; @@ -356,7 +356,7 @@ x_1 = lean_mk_string_unchecked(" : Rat)/", 8, 8); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_instReprRat(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_instReprRat(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -371,7 +371,7 @@ x_6 = lean_ctor_get(x_1, 0); lean_inc(x_6); lean_dec(x_1); x_7 = l___private_Init_Data_Repr_0__Nat_reprFast(x_3); -x_8 = l_Lean_instInhabitedRat___closed__1; +x_8 = l_Std_Internal_instInhabitedRat___closed__1; x_9 = lean_int_dec_lt(x_6, x_8); if (x_9 == 0) { @@ -379,14 +379,14 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean x_10 = lean_nat_abs(x_6); lean_dec(x_6); x_11 = l___private_Init_Data_Repr_0__Nat_reprFast(x_10); -x_12 = l_Lean_instReprRat___closed__1; +x_12 = l_Std_Internal_instReprRat___closed__1; x_13 = lean_string_append(x_12, x_11); lean_dec(x_11); -x_14 = l_Lean_instReprRat___closed__2; +x_14 = l_Std_Internal_instReprRat___closed__2; x_15 = lean_string_append(x_13, x_14); x_16 = lean_string_append(x_15, x_7); lean_dec(x_7); -x_17 = l_Lean_instToStringRat___closed__1; +x_17 = l_Std_Internal_instToStringRat___closed__1; x_18 = lean_string_append(x_16, x_17); x_19 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_19, 0, x_18); @@ -402,17 +402,17 @@ lean_dec(x_20); x_22 = lean_nat_add(x_21, x_4); lean_dec(x_21); x_23 = l___private_Init_Data_Repr_0__Nat_reprFast(x_22); -x_24 = l_Lean_instToStringRat___closed__3; +x_24 = l_Std_Internal_instToStringRat___closed__3; x_25 = lean_string_append(x_24, x_23); lean_dec(x_23); -x_26 = l_Lean_instReprRat___closed__1; +x_26 = l_Std_Internal_instReprRat___closed__1; x_27 = lean_string_append(x_26, x_25); lean_dec(x_25); -x_28 = l_Lean_instReprRat___closed__2; +x_28 = l_Std_Internal_instReprRat___closed__2; x_29 = lean_string_append(x_27, x_28); x_30 = lean_string_append(x_29, x_7); lean_dec(x_7); -x_31 = l_Lean_instToStringRat___closed__1; +x_31 = l_Std_Internal_instToStringRat___closed__1; x_32 = lean_string_append(x_30, x_31); x_33 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_33, 0, x_32); @@ -426,7 +426,7 @@ lean_dec(x_3); x_34 = lean_ctor_get(x_1, 0); lean_inc(x_34); lean_dec(x_1); -x_35 = l_Lean_instInhabitedRat___closed__1; +x_35 = l_Std_Internal_instInhabitedRat___closed__1; x_36 = lean_int_dec_lt(x_34, x_35); if (x_36 == 0) { @@ -451,16 +451,16 @@ return x_42; } } } -LEAN_EXPORT lean_object* l_Lean_instReprRat___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_instReprRat___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_instReprRat(x_1, x_2); +x_3 = l_Std_Internal_instReprRat(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Rat_normalize(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_normalize(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -493,20 +493,20 @@ return x_1; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_normalize___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_normalize___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Rat_normalize(x_1); +x_2 = l_Std_Internal_Rat_normalize(x_1); lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_mkRat___closed__1() { +static lean_object* _init_l_Std_Internal_mkRat___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_instInhabitedRat___closed__1; +x_1 = l_Std_Internal_instInhabitedRat___closed__1; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -514,7 +514,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkRat(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_mkRat(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; @@ -563,12 +563,12 @@ else lean_object* x_14; lean_dec(x_2); lean_dec(x_1); -x_14 = l_Lean_mkRat___closed__1; +x_14 = l_Std_Internal_mkRat___closed__1; return x_14; } } } -LEAN_EXPORT uint8_t l_Lean_Rat_isInt(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Std_Internal_Rat_isInt(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -578,23 +578,23 @@ x_4 = lean_nat_dec_eq(x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_isInt___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_isInt___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_Rat_isInt(x_1); +x_2 = l_Std_Internal_Rat_isInt(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Lean_Rat_lt(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_Internal_Rat_lt(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_26; uint8_t x_27; x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); -x_26 = l_Lean_instInhabitedRat___closed__1; +x_26 = l_Std_Internal_instInhabitedRat___closed__1; x_27 = lean_int_dec_lt(x_3, x_26); if (x_27 == 0) { @@ -664,7 +664,7 @@ return x_37; { lean_object* x_5; uint8_t x_6; lean_dec(x_4); -x_5 = l_Lean_instInhabitedRat___closed__1; +x_5 = l_Std_Internal_instInhabitedRat___closed__1; x_6 = lean_int_dec_lt(x_5, x_3); if (x_6 == 0) { @@ -732,16 +732,16 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_lt___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_lt___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Rat_lt(x_1, x_2); +x_3 = l_Std_Internal_Rat_lt(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_mul(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_mul(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -779,23 +779,23 @@ lean_ctor_set(x_19, 1, x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Rat_mul___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_mul___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Rat_mul(x_1, x_2); +x_3 = l_Std_Internal_Rat_mul(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Rat_inv(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_inv(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_instInhabitedRat___closed__1; +x_3 = l_Std_Internal_instInhabitedRat___closed__1; x_4 = lean_int_dec_lt(x_2, x_3); if (x_4 == 0) { @@ -839,26 +839,26 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_div(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_div(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_Rat_inv(x_2); -x_4 = l_Lean_Rat_mul(x_1, x_3); +x_3 = l_Std_Internal_Rat_inv(x_2); +x_4 = l_Std_Internal_Rat_mul(x_1, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_div___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_div___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Rat_div(x_1, x_2); +x_3 = l_Std_Internal_Rat_div(x_1, x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Rat_add(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_add(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -958,7 +958,7 @@ return x_34; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_sub(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_sub(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; @@ -1063,7 +1063,7 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_neg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_neg(lean_object* x_1) { _start: { uint8_t x_2; @@ -1094,7 +1094,7 @@ return x_8; } } } -static lean_object* _init_l_Lean_Rat_floor___closed__1() { +static lean_object* _init_l_Std_Internal_Rat_floor___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -1103,7 +1103,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Rat_floor(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_floor(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1120,7 +1120,7 @@ lean_dec(x_1); x_6 = lean_nat_to_int(x_2); x_7 = lean_int_mod(x_5, x_6); lean_dec(x_6); -x_8 = l_Lean_instInhabitedRat___closed__1; +x_8 = l_Std_Internal_instInhabitedRat___closed__1; x_9 = lean_int_dec_lt(x_5, x_8); lean_dec(x_5); if (x_9 == 0) @@ -1130,7 +1130,7 @@ return x_7; else { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Rat_floor___closed__1; +x_10 = l_Std_Internal_Rat_floor___closed__1; x_11 = lean_int_sub(x_7, x_10); lean_dec(x_7); return x_11; @@ -1147,7 +1147,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_Rat_ceil(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_ceil(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1164,7 +1164,7 @@ lean_dec(x_1); x_6 = lean_nat_to_int(x_2); x_7 = lean_int_mod(x_5, x_6); lean_dec(x_6); -x_8 = l_Lean_instInhabitedRat___closed__1; +x_8 = l_Std_Internal_instInhabitedRat___closed__1; x_9 = lean_int_dec_lt(x_8, x_5); lean_dec(x_5); if (x_9 == 0) @@ -1174,7 +1174,7 @@ return x_7; else { lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_Rat_floor___closed__1; +x_10 = l_Std_Internal_Rat_floor___closed__1; x_11 = lean_int_add(x_7, x_10); lean_dec(x_7); return x_11; @@ -1191,7 +1191,7 @@ return x_12; } } } -static lean_object* _init_l_Lean_Rat_instLT() { +static lean_object* _init_l_Std_Internal_Rat_instLT() { _start: { lean_object* x_1; @@ -1199,24 +1199,24 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT uint8_t l_Lean_Rat_instDecidableLt(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_Internal_Rat_instDecidableLt(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; -x_3 = l_Lean_Rat_lt(x_1, x_2); +x_3 = l_Std_Internal_Rat_lt(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instDecidableLt___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDecidableLt___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Rat_instDecidableLt(x_1, x_2); +x_3 = l_Std_Internal_Rat_instDecidableLt(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_Rat_instLE() { +static lean_object* _init_l_Std_Internal_Rat_instLE() { _start: { lean_object* x_1; @@ -1224,108 +1224,108 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT uint8_t l_Lean_Rat_instDecidableLe(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Std_Internal_Rat_instDecidableLe(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; -x_3 = l_Lean_Rat_lt(x_2, x_1); +x_3 = l_Std_Internal_Rat_lt(x_2, x_1); x_4 = l_instDecidableNot___rarg(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instDecidableLe___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDecidableLe___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Rat_instDecidableLe(x_1, x_2); +x_3 = l_Std_Internal_Rat_instDecidableLe(x_1, x_2); x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_Rat_instAdd___closed__1() { +static lean_object* _init_l_Std_Internal_Rat_instAdd___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Rat_add), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Rat_add), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Rat_instAdd() { +static lean_object* _init_l_Std_Internal_Rat_instAdd() { _start: { lean_object* x_1; -x_1 = l_Lean_Rat_instAdd___closed__1; +x_1 = l_Std_Internal_Rat_instAdd___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Rat_instSub___closed__1() { +static lean_object* _init_l_Std_Internal_Rat_instSub___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Rat_sub), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Rat_sub), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Rat_instSub() { +static lean_object* _init_l_Std_Internal_Rat_instSub() { _start: { lean_object* x_1; -x_1 = l_Lean_Rat_instSub___closed__1; +x_1 = l_Std_Internal_Rat_instSub___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Rat_instNeg___closed__1() { +static lean_object* _init_l_Std_Internal_Rat_instNeg___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Rat_neg), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Rat_neg), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Rat_instNeg() { +static lean_object* _init_l_Std_Internal_Rat_instNeg() { _start: { lean_object* x_1; -x_1 = l_Lean_Rat_instNeg___closed__1; +x_1 = l_Std_Internal_Rat_instNeg___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Rat_instMul___closed__1() { +static lean_object* _init_l_Std_Internal_Rat_instMul___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Rat_mul___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Rat_mul___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Rat_instMul() { +static lean_object* _init_l_Std_Internal_Rat_instMul() { _start: { lean_object* x_1; -x_1 = l_Lean_Rat_instMul___closed__1; +x_1 = l_Std_Internal_Rat_instMul___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instDiv(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDiv(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_Lean_Rat_inv(x_2); -x_4 = l_Lean_Rat_mul(x_1, x_3); +x_3 = l_Std_Internal_Rat_inv(x_2); +x_4 = l_Std_Internal_Rat_mul(x_1, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instDiv___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instDiv___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Rat_instDiv(x_1, x_2); +x_3 = l_Std_Internal_Rat_instDiv(x_1, x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instOfNat(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instOfNat(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -1337,7 +1337,7 @@ lean_ctor_set(x_4, 1, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Rat_instCoeInt(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Std_Internal_Rat_instCoeInt(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -1353,7 +1353,7 @@ lean_object* initialize_Init_Data_ToString_Macro(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Int_DivMod(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Gcd(uint8_t builtin, lean_object*); static bool _G_initialized = false; -LEAN_EXPORT lean_object* initialize_Lean_Data_Rat(uint8_t builtin, lean_object* w) { +LEAN_EXPORT lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object* w) { lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; @@ -1369,50 +1369,50 @@ lean_dec_ref(res); res = initialize_Init_Data_Nat_Gcd(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_instInhabitedRat___closed__1 = _init_l_Lean_instInhabitedRat___closed__1(); -lean_mark_persistent(l_Lean_instInhabitedRat___closed__1); -l_Lean_instInhabitedRat___closed__2 = _init_l_Lean_instInhabitedRat___closed__2(); -lean_mark_persistent(l_Lean_instInhabitedRat___closed__2); -l_Lean_instInhabitedRat = _init_l_Lean_instInhabitedRat(); -lean_mark_persistent(l_Lean_instInhabitedRat); -l_Lean_instBEqRat___closed__1 = _init_l_Lean_instBEqRat___closed__1(); -lean_mark_persistent(l_Lean_instBEqRat___closed__1); -l_Lean_instBEqRat = _init_l_Lean_instBEqRat(); -lean_mark_persistent(l_Lean_instBEqRat); -l_Lean_instToStringRat___closed__1 = _init_l_Lean_instToStringRat___closed__1(); -lean_mark_persistent(l_Lean_instToStringRat___closed__1); -l_Lean_instToStringRat___closed__2 = _init_l_Lean_instToStringRat___closed__2(); -lean_mark_persistent(l_Lean_instToStringRat___closed__2); -l_Lean_instToStringRat___closed__3 = _init_l_Lean_instToStringRat___closed__3(); -lean_mark_persistent(l_Lean_instToStringRat___closed__3); -l_Lean_instReprRat___closed__1 = _init_l_Lean_instReprRat___closed__1(); -lean_mark_persistent(l_Lean_instReprRat___closed__1); -l_Lean_instReprRat___closed__2 = _init_l_Lean_instReprRat___closed__2(); -lean_mark_persistent(l_Lean_instReprRat___closed__2); -l_Lean_mkRat___closed__1 = _init_l_Lean_mkRat___closed__1(); -lean_mark_persistent(l_Lean_mkRat___closed__1); -l_Lean_Rat_floor___closed__1 = _init_l_Lean_Rat_floor___closed__1(); -lean_mark_persistent(l_Lean_Rat_floor___closed__1); -l_Lean_Rat_instLT = _init_l_Lean_Rat_instLT(); -lean_mark_persistent(l_Lean_Rat_instLT); -l_Lean_Rat_instLE = _init_l_Lean_Rat_instLE(); -lean_mark_persistent(l_Lean_Rat_instLE); -l_Lean_Rat_instAdd___closed__1 = _init_l_Lean_Rat_instAdd___closed__1(); -lean_mark_persistent(l_Lean_Rat_instAdd___closed__1); -l_Lean_Rat_instAdd = _init_l_Lean_Rat_instAdd(); -lean_mark_persistent(l_Lean_Rat_instAdd); -l_Lean_Rat_instSub___closed__1 = _init_l_Lean_Rat_instSub___closed__1(); -lean_mark_persistent(l_Lean_Rat_instSub___closed__1); -l_Lean_Rat_instSub = _init_l_Lean_Rat_instSub(); -lean_mark_persistent(l_Lean_Rat_instSub); -l_Lean_Rat_instNeg___closed__1 = _init_l_Lean_Rat_instNeg___closed__1(); -lean_mark_persistent(l_Lean_Rat_instNeg___closed__1); -l_Lean_Rat_instNeg = _init_l_Lean_Rat_instNeg(); -lean_mark_persistent(l_Lean_Rat_instNeg); -l_Lean_Rat_instMul___closed__1 = _init_l_Lean_Rat_instMul___closed__1(); -lean_mark_persistent(l_Lean_Rat_instMul___closed__1); -l_Lean_Rat_instMul = _init_l_Lean_Rat_instMul(); -lean_mark_persistent(l_Lean_Rat_instMul); +l_Std_Internal_instInhabitedRat___closed__1 = _init_l_Std_Internal_instInhabitedRat___closed__1(); +lean_mark_persistent(l_Std_Internal_instInhabitedRat___closed__1); +l_Std_Internal_instInhabitedRat___closed__2 = _init_l_Std_Internal_instInhabitedRat___closed__2(); +lean_mark_persistent(l_Std_Internal_instInhabitedRat___closed__2); +l_Std_Internal_instInhabitedRat = _init_l_Std_Internal_instInhabitedRat(); +lean_mark_persistent(l_Std_Internal_instInhabitedRat); +l_Std_Internal_instBEqRat___closed__1 = _init_l_Std_Internal_instBEqRat___closed__1(); +lean_mark_persistent(l_Std_Internal_instBEqRat___closed__1); +l_Std_Internal_instBEqRat = _init_l_Std_Internal_instBEqRat(); +lean_mark_persistent(l_Std_Internal_instBEqRat); +l_Std_Internal_instToStringRat___closed__1 = _init_l_Std_Internal_instToStringRat___closed__1(); +lean_mark_persistent(l_Std_Internal_instToStringRat___closed__1); +l_Std_Internal_instToStringRat___closed__2 = _init_l_Std_Internal_instToStringRat___closed__2(); +lean_mark_persistent(l_Std_Internal_instToStringRat___closed__2); +l_Std_Internal_instToStringRat___closed__3 = _init_l_Std_Internal_instToStringRat___closed__3(); +lean_mark_persistent(l_Std_Internal_instToStringRat___closed__3); +l_Std_Internal_instReprRat___closed__1 = _init_l_Std_Internal_instReprRat___closed__1(); +lean_mark_persistent(l_Std_Internal_instReprRat___closed__1); +l_Std_Internal_instReprRat___closed__2 = _init_l_Std_Internal_instReprRat___closed__2(); +lean_mark_persistent(l_Std_Internal_instReprRat___closed__2); +l_Std_Internal_mkRat___closed__1 = _init_l_Std_Internal_mkRat___closed__1(); +lean_mark_persistent(l_Std_Internal_mkRat___closed__1); +l_Std_Internal_Rat_floor___closed__1 = _init_l_Std_Internal_Rat_floor___closed__1(); +lean_mark_persistent(l_Std_Internal_Rat_floor___closed__1); +l_Std_Internal_Rat_instLT = _init_l_Std_Internal_Rat_instLT(); +lean_mark_persistent(l_Std_Internal_Rat_instLT); +l_Std_Internal_Rat_instLE = _init_l_Std_Internal_Rat_instLE(); +lean_mark_persistent(l_Std_Internal_Rat_instLE); +l_Std_Internal_Rat_instAdd___closed__1 = _init_l_Std_Internal_Rat_instAdd___closed__1(); +lean_mark_persistent(l_Std_Internal_Rat_instAdd___closed__1); +l_Std_Internal_Rat_instAdd = _init_l_Std_Internal_Rat_instAdd(); +lean_mark_persistent(l_Std_Internal_Rat_instAdd); +l_Std_Internal_Rat_instSub___closed__1 = _init_l_Std_Internal_Rat_instSub___closed__1(); +lean_mark_persistent(l_Std_Internal_Rat_instSub___closed__1); +l_Std_Internal_Rat_instSub = _init_l_Std_Internal_Rat_instSub(); +lean_mark_persistent(l_Std_Internal_Rat_instSub); +l_Std_Internal_Rat_instNeg___closed__1 = _init_l_Std_Internal_Rat_instNeg___closed__1(); +lean_mark_persistent(l_Std_Internal_Rat_instNeg___closed__1); +l_Std_Internal_Rat_instNeg = _init_l_Std_Internal_Rat_instNeg(); +lean_mark_persistent(l_Std_Internal_Rat_instNeg); +l_Std_Internal_Rat_instMul___closed__1 = _init_l_Std_Internal_Rat_instMul___closed__1(); +lean_mark_persistent(l_Std_Internal_Rat_instMul___closed__1); +l_Std_Internal_Rat_instMul = _init_l_Std_Internal_Rat_instMul(); +lean_mark_persistent(l_Std_Internal_Rat_instMul); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Std/Time.c b/stage0/stdlib/Std/Time.c new file mode 100644 index 000000000000..5bf099c30312 --- /dev/null +++ b/stage0/stdlib/Std/Time.c @@ -0,0 +1,57 @@ +// Lean compiler output +// Module: Std.Time +// Imports: Std.Time.Time Std.Time.Date Std.Time.Zoned Std.Time.Format Std.Time.DateTime Std.Time.Notation Std.Time.Duration Std.Time.Zoned.Database +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Format(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Notation(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Duration(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Format(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Notation(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Duration(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date.c b/stage0/stdlib/Std/Time/Date.c new file mode 100644 index 000000000000..d5e99f97e440 --- /dev/null +++ b/stage0/stdlib/Std/Time/Date.c @@ -0,0 +1,33 @@ +// Lean compiler output +// Module: Std.Time.Date +// Imports: Std.Time.Date.Basic Std.Time.Date.PlainDate +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Time_Date_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_PlainDate(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_PlainDate(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Basic.c b/stage0/stdlib/Std/Time/Date/Basic.c new file mode 100644 index 000000000000..8cb434012b4e --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Basic.c @@ -0,0 +1,2894 @@ +// Lean compiler output +// Module: Std.Time.Date.Basic +// Imports: Std.Time.Date.Unit.Basic Std.Time.Date.ValidDate Std.Time.Time.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__36(lean_object*, lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toDays___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__7(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__17___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__41(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__11___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__40___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__40(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__33___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__13(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__31(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__7___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__7(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__23___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__16___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__8(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__23(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__28___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__20___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__30(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__14___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__14___closed__1; +static lean_object* l_Std_Time_instHAddOffsetOffset__9___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__15(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__9___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__34(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__33___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__26___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__24___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__24(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__29___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__13___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__13(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__9(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__28___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__9(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__37___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__16(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__31___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__41___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__24___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__38___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__2(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_Offset_toWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toDays___boxed(lean_object*); +static lean_object* l_Std_Time_Minute_Offset_toDays___closed__1; +static lean_object* l_Std_Time_Second_Offset_toDays___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__19___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__41___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__15___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__9___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__32(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__14(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__17(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__22(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__40(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__37(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__25___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__5___boxed(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__13___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__15(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__40___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__34___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__32(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__8(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__21(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__29(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__10___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__31(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__27(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__35(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__14(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofWeeks(lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__7___closed__1; +lean_object* lean_int_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__12___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__27___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__23___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__32___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__38___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__35___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__33(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__27___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__19(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__14___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__21___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__19(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__12___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__4___boxed(lean_object*, lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__18(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__8___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__30___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__19___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__37___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__25(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__35___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__28(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__30___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__27(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__15___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__18___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__16___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__36(lean_object*, lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__11(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__22(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_Offset_toWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__25___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__23(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__39___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__4___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__11___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__29(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__29___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__8___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__18___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__2(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_Offset_toDays___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__33(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__15___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__17(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__4___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__26___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__34(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__35___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__35(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__30(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__41(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__26(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toDays(lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__16(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__5(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__22___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__17___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__25(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__21___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__18(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__39(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__34___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__32___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__39(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__28(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__38(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__31___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__5___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__7___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__4___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_Offset_toWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__20(lean_object*, lean_object*); +static lean_object* l_Std_Time_Millisecond_Offset_toDays___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofWeeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__10(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__37(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__39___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__20(lean_object*, lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__38(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toDays___boxed(lean_object*); +static lean_object* l_Std_Time_instHAddOffsetOffset__8___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__6(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__12(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__11(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__22___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__12(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__36___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__36___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__4(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__10(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__26(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__24(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__20___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__6___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__10___boxed(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("86400000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("604800000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_Offset_toWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(604800000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toDays___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_Offset_toWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(604800u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Minute_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1440u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Minute_Offset_toWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10080u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_Offset_toWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(168u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__1___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("60000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__2___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("3600000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__3___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__7___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__7___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__7___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__7(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__8___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__8___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__8___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__8(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__9___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__9___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__9___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__9(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__10(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__10___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__10(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__11(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__11___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__11(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__12(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__1___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__12___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__12(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__13(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__7___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__13___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__13(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__14___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__14(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__14___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__14(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__15___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__15(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__15___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__15___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__15(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__16(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__16___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__16(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__17(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__17___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__17(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__18(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__2___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__18___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__18(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__19(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__8___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__19___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__19(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__20(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__20___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__20(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__21(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__21(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__22(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__22___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__22(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__23(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__23___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__23(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__24(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__3___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__24___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__24(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__25(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__9___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__25___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__25(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__26(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__15___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__26___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__26(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__27(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__27___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__27(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__28(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__28___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__28(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__29(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__29___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__29(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__30(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__30___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__30(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__31(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__31___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__31(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__32(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__32___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__32(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__33(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__33___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__33(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__34(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__34___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__34(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instHAddOffsetOffset__35___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__35(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__35___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__35___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__35(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__36(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__36___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__36(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__37(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__37___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__37(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__38(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__38___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__38(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__39(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__39___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__39(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__40(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__40___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__40(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__41(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__35___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffsetOffset__41___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffsetOffset__41(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHAddOffset__6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHAddOffset__6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__1___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__2___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__3___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__7(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__7___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__7___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__7(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__8___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__8___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__8(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__9___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__9___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__9(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__10(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__10___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__10(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__11(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__11___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__11(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__12(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__1___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__12___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__12(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__13(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__7___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__13___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__13(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__14(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__14___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__14(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__15(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__15___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__15___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__15(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__16(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__16___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__16(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__17(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__17___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__17(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__18(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__2___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__18___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__18(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__19(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__8___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__19___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__19(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__20(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__20___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__20(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__3(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__21(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__21(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__22(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__22___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__22(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__23(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__23___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__23(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__24(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__3___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__24___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__24(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__25(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__9___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__25___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__25(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__26(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__15___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__26___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__26(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__27(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__14___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__27___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__27(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__4(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__4___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__4(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__28(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__28___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__28(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__29(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__29___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__29(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__30(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__30___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__30(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__31(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__31___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__31(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__32(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__32___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__32(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__33(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__33___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__33(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__34(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toDays___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__34___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__34(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__5___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__5(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__35(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__35___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_sub(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__35___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__35(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__36(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Nanosecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__36___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__36(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__37(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Millisecond_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__37___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__37(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__38(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__38___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__38(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__39(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Minute_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__39___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__39(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__40(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Hour_Offset_toWeeks___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__40___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__40(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__41(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instHAddOffsetOffset__35___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_sub(x_4, x_2); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffsetOffset__41___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffsetOffset__41(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instHSubOffset__6___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instHSubOffset__6(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Date_Unit_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_ValidDate(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date_Unit_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_ValidDate(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Nanosecond_Offset_toDays___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toDays___closed__1); +l_Std_Time_Nanosecond_Offset_toWeeks___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toWeeks___closed__1); +l_Std_Time_Millisecond_Offset_toDays___closed__1 = _init_l_Std_Time_Millisecond_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_Offset_toDays___closed__1); +l_Std_Time_Millisecond_Offset_toWeeks___closed__1 = _init_l_Std_Time_Millisecond_Offset_toWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_Offset_toWeeks___closed__1); +l_Std_Time_Second_Offset_toDays___closed__1 = _init_l_Std_Time_Second_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Second_Offset_toDays___closed__1); +l_Std_Time_Second_Offset_toWeeks___closed__1 = _init_l_Std_Time_Second_Offset_toWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Second_Offset_toWeeks___closed__1); +l_Std_Time_Minute_Offset_toDays___closed__1 = _init_l_Std_Time_Minute_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_Offset_toDays___closed__1); +l_Std_Time_Minute_Offset_toWeeks___closed__1 = _init_l_Std_Time_Minute_Offset_toWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_Offset_toWeeks___closed__1); +l_Std_Time_Hour_Offset_toDays___closed__1 = _init_l_Std_Time_Hour_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_Offset_toDays___closed__1); +l_Std_Time_Hour_Offset_toWeeks___closed__1 = _init_l_Std_Time_Hour_Offset_toWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_Offset_toWeeks___closed__1); +l_Std_Time_instHAddOffsetOffset___closed__1 = _init_l_Std_Time_instHAddOffsetOffset___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset___closed__1); +l_Std_Time_instHAddOffsetOffset__1___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__1___closed__1); +l_Std_Time_instHAddOffsetOffset__2___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__2___closed__1); +l_Std_Time_instHAddOffsetOffset__3___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__3___closed__1); +l_Std_Time_instHAddOffsetOffset__7___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__7___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__7___closed__1); +l_Std_Time_instHAddOffsetOffset__8___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__8___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__8___closed__1); +l_Std_Time_instHAddOffsetOffset__9___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__9___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__9___closed__1); +l_Std_Time_instHAddOffsetOffset__14___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__14___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__14___closed__1); +l_Std_Time_instHAddOffsetOffset__15___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__15___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__15___closed__1); +l_Std_Time_instHAddOffsetOffset__35___closed__1 = _init_l_Std_Time_instHAddOffsetOffset__35___closed__1(); +lean_mark_persistent(l_Std_Time_instHAddOffsetOffset__35___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/PlainDate.c b/stage0/stdlib/Std/Time/Date/PlainDate.c new file mode 100644 index 000000000000..5e84cb554e9d --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/PlainDate.c @@ -0,0 +1,4831 @@ +// Lean compiler output +// Module: Std.Time.Date.PlainDate +// Imports: Std.Time.Internal Std.Time.Date.Basic Std.Internal.Rat +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* l_Std_Time_ValidDate_ofOrdinal(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subDays___boxed(lean_object*, lean_object*); +lean_object* lean_int_mod(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_instBEqPlainDate___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12; +uint8_t l_Std_Time_Weekday_ofOrdinal(lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__12; +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_dayOfYear(lean_object*); +static lean_object* l_Std_Time_PlainDate_rollOver___closed__3; +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__9; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHSubOffset__1; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_era___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsClip___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__18; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__1; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_inLeapYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +static lean_object* l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__16; +uint8_t l_instDecidableNot___rarg(uint8_t); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withWeekday___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7; +static lean_object* l_Std_Time_instReprPlainDate___closed__1; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13; +static lean_object* l_Std_Time_PlainDate_instHSubOffset___closed__1; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8; +lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekday___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addDays(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9; +static lean_object* l_Std_Time_PlainDate_rollOver___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__4; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_instHAddOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_quarter___boxed(lean_object*); +lean_object* l_Int_repr(lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__10; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__13; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addDays___boxed(lean_object*, lean_object*); +lean_object* lean_int_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subWeeks___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Weekday_toOrdinal(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfMonth___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instInhabited; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_inLeapYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_dayOfYear___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHAddOffset__1; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__7; +static lean_object* l_Std_Time_PlainDate_weekday___closed__3; +static lean_object* l_Std_Time_PlainDate_weekday___closed__2; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5; +static lean_object* l_Std_Time_PlainDate_rollOver___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHSubOffset; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addWeeks(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10; +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsRollOver___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_rollOver___closed__4; +static lean_object* l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearMonthDay_x3f(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_Year_Offset_weeks(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_rollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprPlainDate; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearMonthDayClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +static lean_object* l_Std_Time_PlainDate_weekday___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedPlainDate; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__2; +lean_object* lean_int_mul(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subWeeks(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24; +static lean_object* l_Std_Time_PlainDate_weekday___closed__4; +static lean_object* l_Std_Time_PlainDate_instInhabited___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__6; +static lean_object* l_Std_Time_PlainDate_instHSubOffset__1___closed__1; +uint8_t l_Std_Time_Year_Offset_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subDays(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsRollOver___boxed(lean_object*, lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__17; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__3; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instBEqPlainDate(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withWeekday(lean_object*, uint8_t); +static lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsClip(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23; +static lean_object* l_Std_Time_PlainDate_instHAddOffset__1___closed__1; +static lean_object* l_Std_Time_PlainDate_weekOfMonth___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withYearClip(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsRollOver___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12; +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7; +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHAddOffset; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_rollOver___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withMonthClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearOrdinal(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25; +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withMonthRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2; +static lean_object* l_Std_Time_instInhabitedPlainDate___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withYearRollOver(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20; +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("year", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3; +x_2 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("month", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("day", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("valid", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_dec_lt(x_6, x_4); +x_8 = lean_ctor_get(x_1, 2); +x_9 = lean_int_dec_lt(x_8, x_4); +if (x_5 == 0) +{ +lean_object* x_85; lean_object* x_86; +x_85 = l_Int_repr(x_3); +x_86 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_10 = x_86; +goto block_84; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_87 = l_Int_repr(x_3); +x_88 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_88, 0, x_87); +x_89 = lean_unsigned_to_nat(0u); +x_90 = l_Repr_addAppParen(x_88, x_89); +x_10 = x_90; +goto block_84; +} +block_84: +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_11 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7; +x_12 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = 0; +x_14 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set_uint8(x_14, sizeof(void*)*1, x_13); +x_15 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6; +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10; +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_box(1); +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12; +x_22 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5; +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +if (x_7 == 0) +{ +lean_object* x_78; lean_object* x_79; +x_78 = l_Int_repr(x_6); +x_79 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_79, 0, x_78); +x_25 = x_79; +goto block_77; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_80 = l_Int_repr(x_6); +x_81 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_81, 0, x_80); +x_82 = lean_unsigned_to_nat(0u); +x_83 = l_Repr_addAppParen(x_81, x_82); +x_25 = x_83; +goto block_77; +} +block_77: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_26 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13; +x_27 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_13); +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_19); +x_32 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15; +x_33 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_23); +if (x_9 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_35 = l_Int_repr(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_38 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +x_39 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set_uint8(x_39, sizeof(void*)*1, x_13); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_17); +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_19); +x_43 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18; +x_44 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_23); +x_46 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20; +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24; +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26; +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23; +x_53 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_13); +return x_54; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_55 = l_Int_repr(x_8); +x_56 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_unsigned_to_nat(0u); +x_58 = l_Repr_addAppParen(x_56, x_57); +x_59 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_60 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_58); +x_61 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set_uint8(x_61, sizeof(void*)*1, x_13); +x_62 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_62, 0, x_34); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_17); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_19); +x_65 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18; +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_23); +x_68 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20; +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24; +x_71 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_69); +x_72 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26; +x_73 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +x_74 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23; +x_75 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_73); +x_76 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set_uint8(x_76, sizeof(void*)*1, x_13); +return x_76; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprPlainDate___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprPlainDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprPlainDate___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__3; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__4; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__6; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__7; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__8; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__9; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__11; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__12; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__13; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__6; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__15; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__16; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__17; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__10; +x_3 = l_Std_Time_instInhabitedPlainDate___closed__18; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__19; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_instBEqPlainDate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 2); +x_4 = lean_ctor_get(x_2, 2); +x_5 = lean_int_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_int_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_int_dec_eq(x_11, x_12); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqPlainDate___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_instBEqPlainDate(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearMonthDayClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_5 = lean_int_mod(x_1, x_4); +x_6 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; uint8_t x_10; +x_8 = 0; +x_9 = l_Std_Time_Month_Ordinal_days(x_8, x_2); +x_10 = lean_int_dec_lt(x_9, x_3); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_9); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_2); +lean_ctor_set(x_11, 2, x_3); +return x_11; +} +else +{ +lean_object* x_12; +lean_dec(x_3); +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +lean_ctor_set(x_12, 2, x_9); +return x_12; +} +} +else +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; +x_13 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_14 = lean_int_mod(x_1, x_13); +x_15 = lean_int_dec_eq(x_14, x_6); +lean_dec(x_14); +x_16 = l_instDecidableNot___rarg(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_18 = lean_int_mod(x_1, x_17); +x_19 = lean_int_dec_eq(x_18, x_6); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; lean_object* x_21; uint8_t x_22; +x_20 = 0; +x_21 = l_Std_Time_Month_Ordinal_days(x_20, x_2); +x_22 = lean_int_dec_lt(x_21, x_3); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_2); +lean_ctor_set(x_23, 2, x_3); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_3); +x_24 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_24, 0, x_1); +lean_ctor_set(x_24, 1, x_2); +lean_ctor_set(x_24, 2, x_21); +return x_24; +} +} +else +{ +uint8_t x_25; lean_object* x_26; uint8_t x_27; +x_25 = 1; +x_26 = l_Std_Time_Month_Ordinal_days(x_25, x_2); +x_27 = lean_int_dec_lt(x_26, x_3); +if (x_27 == 0) +{ +lean_object* x_28; +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_2); +lean_ctor_set(x_28, 2, x_3); +return x_28; +} +else +{ +lean_object* x_29; +lean_dec(x_3); +x_29 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_2); +lean_ctor_set(x_29, 2, x_26); +return x_29; +} +} +} +else +{ +uint8_t x_30; lean_object* x_31; uint8_t x_32; +x_30 = 1; +x_31 = l_Std_Time_Month_Ordinal_days(x_30, x_2); +x_32 = lean_int_dec_lt(x_31, x_3); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_31); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_2); +lean_ctor_set(x_33, 2, x_3); +return x_33; +} +else +{ +lean_object* x_34; +lean_dec(x_3); +x_34 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_34, 0, x_1); +lean_ctor_set(x_34, 1, x_2); +lean_ctor_set(x_34, 2, x_31); +return x_34; +} +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDate_instInhabited___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__10; +x_3 = l_Std_Time_instInhabitedPlainDate___closed__18; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instInhabited___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearMonthDay_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_5 = lean_int_mod(x_1, x_4); +x_6 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; uint8_t x_10; +x_8 = 0; +x_9 = l_Std_Time_Month_Ordinal_days(x_8, x_2); +x_10 = lean_int_dec_le(x_3, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box(0); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +lean_ctor_set(x_12, 2, x_3); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_14 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_15 = lean_int_mod(x_1, x_14); +x_16 = lean_int_dec_eq(x_15, x_6); +lean_dec(x_15); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_19 = lean_int_mod(x_1, x_18); +x_20 = lean_int_dec_eq(x_19, x_6); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; uint8_t x_23; +x_21 = 0; +x_22 = l_Std_Time_Month_Ordinal_days(x_21, x_2); +x_23 = lean_int_dec_le(x_3, x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_box(0); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_2); +lean_ctor_set(x_25, 2, x_3); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +} +else +{ +uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_27 = 1; +x_28 = l_Std_Time_Month_Ordinal_days(x_27, x_2); +x_29 = lean_int_dec_le(x_3, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_30 = lean_box(0); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_31, 0, x_1); +lean_ctor_set(x_31, 1, x_2); +lean_ctor_set(x_31, 2, x_3); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; +} +} +} +else +{ +uint8_t x_33; lean_object* x_34; uint8_t x_35; +x_33 = 1; +x_34 = l_Std_Time_Month_Ordinal_days(x_33, x_2); +x_35 = lean_int_dec_le(x_3, x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_37, 0, x_1); +lean_ctor_set(x_37, 1, x_2); +lean_ctor_set(x_37, 2, x_3); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +return x_38; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_4 = lean_int_mod(x_1, x_3); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_6 = lean_int_dec_eq(x_4, x_5); +lean_dec(x_4); +if (x_6 == 0) +{ +uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = 0; +x_8 = l_Std_Time_ValidDate_ofOrdinal(x_7, x_2); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_9); +lean_ctor_set(x_11, 2, x_10); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +x_12 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_13 = lean_int_mod(x_1, x_12); +x_14 = lean_int_dec_eq(x_13, x_5); +lean_dec(x_13); +x_15 = l_instDecidableNot___rarg(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_17 = lean_int_mod(x_1, x_16); +x_18 = lean_int_dec_eq(x_17, x_5); +lean_dec(x_17); +if (x_18 == 0) +{ +uint8_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_19 = 0; +x_20 = l_Std_Time_ValidDate_ofOrdinal(x_19, x_2); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_22); +return x_23; +} +else +{ +uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = 1; +x_25 = l_Std_Time_ValidDate_ofOrdinal(x_24, x_2); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_1); +lean_ctor_set(x_28, 1, x_26); +lean_ctor_set(x_28, 2, x_27); +return x_28; +} +} +else +{ +uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_29 = 1; +x_30 = l_Std_Time_ValidDate_ofOrdinal(x_29, x_2); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_31); +lean_ctor_set(x_33, 2, x_32); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofYearOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_ofYearOrdinal(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(719468u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(146097u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1460u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(36524u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(146096u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(365u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(153u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(31u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; +x_2 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1; +x_3 = lean_int_add(x_1, x_2); +x_4 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_5 = lean_int_dec_le(x_4, x_3); +if (x_5 == 0) +{ +lean_object* x_94; lean_object* x_95; +x_94 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5; +x_95 = lean_int_sub(x_3, x_94); +x_6 = x_95; +goto block_93; +} +else +{ +lean_inc(x_3); +x_6 = x_3; +goto block_93; +} +block_93: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; uint8_t x_47; lean_object* x_48; +x_7 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2; +x_8 = lean_int_div(x_6, x_7); +lean_dec(x_6); +x_9 = lean_int_mul(x_8, x_7); +x_10 = lean_int_sub(x_3, x_9); +lean_dec(x_9); +lean_dec(x_3); +x_11 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3; +x_12 = lean_int_div(x_10, x_11); +x_13 = lean_int_sub(x_10, x_12); +lean_dec(x_12); +x_14 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4; +x_15 = lean_int_div(x_10, x_14); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5; +x_18 = lean_int_div(x_10, x_17); +x_19 = lean_int_sub(x_16, x_18); +lean_dec(x_18); +lean_dec(x_16); +x_20 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6; +x_21 = lean_int_div(x_19, x_20); +lean_dec(x_19); +x_22 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_23 = lean_int_mul(x_8, x_22); +lean_dec(x_8); +x_24 = lean_int_add(x_21, x_23); +lean_dec(x_23); +x_25 = lean_int_mul(x_20, x_21); +x_26 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_27 = lean_int_div(x_21, x_26); +x_28 = lean_int_add(x_25, x_27); +lean_dec(x_27); +lean_dec(x_25); +x_29 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_30 = lean_int_div(x_21, x_29); +lean_dec(x_21); +x_31 = lean_int_sub(x_28, x_30); +lean_dec(x_30); +lean_dec(x_28); +x_32 = lean_int_sub(x_10, x_31); +lean_dec(x_31); +lean_dec(x_10); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7; +x_34 = lean_int_mul(x_33, x_32); +x_35 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8; +x_36 = lean_int_add(x_34, x_35); +lean_dec(x_34); +x_37 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9; +x_38 = lean_int_div(x_36, x_37); +lean_dec(x_36); +x_39 = lean_int_mul(x_37, x_38); +x_40 = lean_int_add(x_39, x_35); +lean_dec(x_39); +x_41 = lean_int_div(x_40, x_33); +lean_dec(x_40); +x_42 = lean_int_sub(x_32, x_41); +lean_dec(x_41); +lean_dec(x_32); +x_43 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_44 = lean_int_add(x_42, x_43); +lean_dec(x_42); +x_45 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10; +x_46 = lean_int_dec_lt(x_38, x_45); +x_47 = lean_int_dec_le(x_43, x_44); +if (x_46 == 0) +{ +lean_object* x_91; +x_91 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13; +x_48 = x_91; +goto block_90; +} +else +{ +lean_object* x_92; +x_92 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14; +x_48 = x_92; +goto block_90; +} +block_90: +{ +lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; +x_49 = lean_int_add(x_38, x_48); +lean_dec(x_48); +lean_dec(x_38); +x_50 = lean_int_dec_le(x_49, x_35); +x_51 = lean_int_dec_le(x_43, x_49); +if (x_50 == 0) +{ +x_52 = x_4; +goto block_89; +} +else +{ +x_52 = x_43; +goto block_89; +} +block_89: +{ +lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; +x_53 = lean_int_add(x_24, x_52); +lean_dec(x_52); +lean_dec(x_24); +x_54 = lean_int_mod(x_53, x_26); +x_55 = lean_int_dec_eq(x_54, x_4); +lean_dec(x_54); +if (x_51 == 0) +{ +lean_dec(x_49); +x_56 = x_43; +goto block_86; +} +else +{ +lean_object* x_87; uint8_t x_88; +x_87 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_88 = lean_int_dec_le(x_49, x_87); +if (x_88 == 0) +{ +lean_dec(x_49); +x_56 = x_87; +goto block_86; +} +else +{ +x_56 = x_49; +goto block_86; +} +} +block_86: +{ +lean_object* x_57; +if (x_47 == 0) +{ +lean_dec(x_44); +x_57 = x_43; +goto block_83; +} +else +{ +lean_object* x_84; uint8_t x_85; +x_84 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12; +x_85 = lean_int_dec_le(x_44, x_84); +if (x_85 == 0) +{ +lean_dec(x_44); +x_57 = x_84; +goto block_83; +} +else +{ +x_57 = x_44; +goto block_83; +} +} +block_83: +{ +if (x_55 == 0) +{ +uint8_t x_58; lean_object* x_59; uint8_t x_60; +x_58 = 0; +x_59 = l_Std_Time_Month_Ordinal_days(x_58, x_56); +x_60 = lean_int_dec_lt(x_59, x_57); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_59); +x_61 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_61, 0, x_53); +lean_ctor_set(x_61, 1, x_56); +lean_ctor_set(x_61, 2, x_57); +return x_61; +} +else +{ +lean_object* x_62; +lean_dec(x_57); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_53); +lean_ctor_set(x_62, 1, x_56); +lean_ctor_set(x_62, 2, x_59); +return x_62; +} +} +else +{ +lean_object* x_63; uint8_t x_64; uint8_t x_65; +x_63 = lean_int_mod(x_53, x_29); +x_64 = lean_int_dec_eq(x_63, x_4); +lean_dec(x_63); +x_65 = l_instDecidableNot___rarg(x_64); +if (x_65 == 0) +{ +lean_object* x_66; uint8_t x_67; +x_66 = lean_int_mod(x_53, x_22); +x_67 = lean_int_dec_eq(x_66, x_4); +lean_dec(x_66); +if (x_67 == 0) +{ +uint8_t x_68; lean_object* x_69; uint8_t x_70; +x_68 = 0; +x_69 = l_Std_Time_Month_Ordinal_days(x_68, x_56); +x_70 = lean_int_dec_lt(x_69, x_57); +if (x_70 == 0) +{ +lean_object* x_71; +lean_dec(x_69); +x_71 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_71, 0, x_53); +lean_ctor_set(x_71, 1, x_56); +lean_ctor_set(x_71, 2, x_57); +return x_71; +} +else +{ +lean_object* x_72; +lean_dec(x_57); +x_72 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_72, 0, x_53); +lean_ctor_set(x_72, 1, x_56); +lean_ctor_set(x_72, 2, x_69); +return x_72; +} +} +else +{ +uint8_t x_73; lean_object* x_74; uint8_t x_75; +x_73 = 1; +x_74 = l_Std_Time_Month_Ordinal_days(x_73, x_56); +x_75 = lean_int_dec_lt(x_74, x_57); +if (x_75 == 0) +{ +lean_object* x_76; +lean_dec(x_74); +x_76 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_76, 0, x_53); +lean_ctor_set(x_76, 1, x_56); +lean_ctor_set(x_76, 2, x_57); +return x_76; +} +else +{ +lean_object* x_77; +lean_dec(x_57); +x_77 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_77, 0, x_53); +lean_ctor_set(x_77, 1, x_56); +lean_ctor_set(x_77, 2, x_74); +return x_77; +} +} +} +else +{ +uint8_t x_78; lean_object* x_79; uint8_t x_80; +x_78 = 1; +x_79 = l_Std_Time_Month_Ordinal_days(x_78, x_56); +x_80 = lean_int_dec_lt(x_79, x_57); +if (x_80 == 0) +{ +lean_object* x_81; +lean_dec(x_79); +x_81 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_81, 0, x_53); +lean_ctor_set(x_81, 1, x_56); +lean_ctor_set(x_81, 2, x_57); +return x_81; +} +else +{ +lean_object* x_82; +lean_dec(x_57); +x_82 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_82, 0, x_53); +lean_ctor_set(x_82, 1, x_56); +lean_ctor_set(x_82, 2, x_79); +return x_82; +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_weekOfMonth___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 2); +x_3 = l_Std_Time_PlainDate_weekOfMonth___closed__1; +x_4 = lean_int_add(x_2, x_3); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_6 = lean_int_ediv(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDate_weekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_quarter(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 1); +x_3 = l_Std_Time_PlainDate_weekOfMonth___closed__1; +x_4 = lean_int_add(x_2, x_3); +x_5 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14; +x_6 = lean_int_ediv(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_quarter___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDate_quarter(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_dayOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_4 = lean_int_mod(x_2, x_3); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_6 = lean_int_dec_eq(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_inc(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +if (x_6 == 0) +{ +uint8_t x_10; lean_object* x_11; +x_10 = 0; +x_11 = l_Std_Time_ValidDate_dayOfYear(x_10, x_9); +lean_dec(x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; +x_12 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_13 = lean_int_mod(x_2, x_12); +x_14 = lean_int_dec_eq(x_13, x_5); +lean_dec(x_13); +x_15 = l_instDecidableNot___rarg(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_17 = lean_int_mod(x_2, x_16); +x_18 = lean_int_dec_eq(x_17, x_5); +lean_dec(x_17); +if (x_18 == 0) +{ +uint8_t x_19; lean_object* x_20; +x_19 = 0; +x_20 = l_Std_Time_ValidDate_dayOfYear(x_19, x_9); +lean_dec(x_9); +return x_20; +} +else +{ +uint8_t x_21; lean_object* x_22; +x_21 = 1; +x_22 = l_Std_Time_ValidDate_dayOfYear(x_21, x_9); +lean_dec(x_9); +return x_22; +} +} +else +{ +uint8_t x_23; lean_object* x_24; +x_23 = 1; +x_24 = l_Std_Time_ValidDate_dayOfYear(x_23, x_9); +lean_dec(x_9); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_dayOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDate_dayOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_era(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Year_Offset_era(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_era___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDate_era(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_inLeapYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_4 = lean_int_mod(x_2, x_3); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_6 = lean_int_dec_eq(x_4, x_5); +lean_dec(x_4); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = 0; +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; +x_8 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_9 = lean_int_mod(x_2, x_8); +x_10 = lean_int_dec_eq(x_9, x_5); +lean_dec(x_9); +x_11 = l_instDecidableNot___rarg(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_13 = lean_int_mod(x_2, x_12); +x_14 = lean_int_dec_eq(x_13, x_5); +lean_dec(x_13); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = 1; +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_inLeapYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDate_inLeapYear(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(399u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8; +x_4 = lean_int_dec_lt(x_3, x_2); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_5); +if (x_4 == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_1, 0); +lean_inc(x_56); +lean_dec(x_1); +x_57 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_58 = lean_int_sub(x_56, x_57); +lean_dec(x_56); +x_6 = x_58; +goto block_55; +} +else +{ +lean_object* x_59; +x_59 = lean_ctor_get(x_1, 0); +lean_inc(x_59); +lean_dec(x_1); +x_6 = x_59; +goto block_55; +} +block_55: +{ +lean_object* x_7; uint8_t x_8; lean_object* x_9; +x_7 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_8 = lean_int_dec_le(x_7, x_6); +if (x_8 == 0) +{ +lean_object* x_53; lean_object* x_54; +x_53 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2; +x_54 = lean_int_sub(x_6, x_53); +x_9 = x_54; +goto block_52; +} +else +{ +lean_inc(x_6); +x_9 = x_6; +goto block_52; +} +block_52: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_10 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_11 = lean_int_div(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mul(x_11, x_10); +x_13 = lean_int_sub(x_6, x_12); +lean_dec(x_12); +lean_dec(x_6); +x_14 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6; +x_15 = lean_int_mul(x_13, x_14); +x_16 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_17 = lean_int_div(x_13, x_16); +x_18 = lean_int_add(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +x_19 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_20 = lean_int_div(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_sub(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2; +x_23 = lean_int_mul(x_11, x_22); +lean_dec(x_11); +if (x_4 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_24 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13; +x_25 = lean_int_add(x_2, x_24); +lean_dec(x_2); +x_26 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9; +x_27 = lean_int_mul(x_26, x_25); +lean_dec(x_25); +x_28 = lean_int_add(x_27, x_3); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7; +x_30 = lean_int_div(x_28, x_29); +lean_dec(x_28); +x_31 = lean_int_add(x_30, x_5); +lean_dec(x_5); +lean_dec(x_30); +x_32 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_33 = lean_int_sub(x_31, x_32); +lean_dec(x_31); +x_34 = lean_int_add(x_21, x_33); +lean_dec(x_33); +lean_dec(x_21); +x_35 = lean_int_add(x_23, x_34); +lean_dec(x_34); +lean_dec(x_23); +x_36 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1; +x_37 = lean_int_sub(x_35, x_36); +lean_dec(x_35); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_38 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1; +x_39 = lean_int_add(x_2, x_38); +lean_dec(x_2); +x_40 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9; +x_41 = lean_int_mul(x_40, x_39); +lean_dec(x_39); +x_42 = lean_int_add(x_41, x_3); +lean_dec(x_41); +x_43 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7; +x_44 = lean_int_div(x_42, x_43); +lean_dec(x_42); +x_45 = lean_int_add(x_44, x_5); +lean_dec(x_5); +lean_dec(x_44); +x_46 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_47 = lean_int_sub(x_45, x_46); +lean_dec(x_45); +x_48 = lean_int_add(x_21, x_47); +lean_dec(x_47); +lean_dec(x_21); +x_49 = lean_int_add(x_23, x_48); +lean_dec(x_48); +lean_dec(x_23); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1; +x_51 = lean_int_sub(x_49, x_50); +lean_dec(x_49); +return x_51; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_4 = lean_int_add(x_3, x_2); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_5 = lean_int_add(x_4, x_3); +lean_dec(x_3); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_4 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_5 = lean_int_mul(x_2, x_4); +x_6 = lean_int_add(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_7 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_8 = lean_int_sub(x_5, x_7); +lean_dec(x_5); +x_9 = lean_int_add(x_8, x_2); +lean_dec(x_8); +x_10 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_11 = lean_int_emod(x_9, x_10); +x_12 = lean_int_add(x_11, x_7); +lean_dec(x_11); +x_13 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_14 = lean_int_add(x_4, x_13); +lean_dec(x_13); +lean_dec(x_4); +x_15 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_16 = lean_int_mod(x_14, x_15); +x_17 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_18 = lean_int_dec_eq(x_16, x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +uint8_t x_19; lean_object* x_20; uint8_t x_21; +x_19 = 0; +x_20 = l_Std_Time_Month_Ordinal_days(x_19, x_12); +x_21 = lean_int_dec_lt(x_20, x_6); +if (x_21 == 0) +{ +lean_dec(x_20); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +else +{ +lean_dec(x_6); +lean_ctor_set(x_1, 2, x_20); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t x_25; +x_22 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_23 = lean_int_mod(x_14, x_22); +x_24 = lean_int_dec_eq(x_23, x_17); +lean_dec(x_23); +x_25 = l_instDecidableNot___rarg(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_27 = lean_int_mod(x_14, x_26); +x_28 = lean_int_dec_eq(x_27, x_17); +lean_dec(x_27); +if (x_28 == 0) +{ +uint8_t x_29; lean_object* x_30; uint8_t x_31; +x_29 = 0; +x_30 = l_Std_Time_Month_Ordinal_days(x_29, x_12); +x_31 = lean_int_dec_lt(x_30, x_6); +if (x_31 == 0) +{ +lean_dec(x_30); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +else +{ +lean_dec(x_6); +lean_ctor_set(x_1, 2, x_30); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +} +else +{ +uint8_t x_32; lean_object* x_33; uint8_t x_34; +x_32 = 1; +x_33 = l_Std_Time_Month_Ordinal_days(x_32, x_12); +x_34 = lean_int_dec_lt(x_33, x_6); +if (x_34 == 0) +{ +lean_dec(x_33); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +else +{ +lean_dec(x_6); +lean_ctor_set(x_1, 2, x_33); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +} +} +else +{ +uint8_t x_35; lean_object* x_36; uint8_t x_37; +x_35 = 1; +x_36 = l_Std_Time_Month_Ordinal_days(x_35, x_12); +x_37 = lean_int_dec_lt(x_36, x_6); +if (x_37 == 0) +{ +lean_dec(x_36); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +else +{ +lean_dec(x_6); +lean_ctor_set(x_1, 2, x_36); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_14); +return x_1; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_38 = lean_ctor_get(x_1, 0); +x_39 = lean_ctor_get(x_1, 1); +x_40 = lean_ctor_get(x_1, 2); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_1); +x_41 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_42 = lean_int_sub(x_39, x_41); +lean_dec(x_39); +x_43 = lean_int_add(x_42, x_2); +lean_dec(x_42); +x_44 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_45 = lean_int_emod(x_43, x_44); +x_46 = lean_int_add(x_45, x_41); +lean_dec(x_45); +x_47 = lean_int_ediv(x_43, x_44); +lean_dec(x_43); +x_48 = lean_int_add(x_38, x_47); +lean_dec(x_47); +lean_dec(x_38); +x_49 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_50 = lean_int_mod(x_48, x_49); +x_51 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_52 = lean_int_dec_eq(x_50, x_51); +lean_dec(x_50); +if (x_52 == 0) +{ +uint8_t x_53; lean_object* x_54; uint8_t x_55; +x_53 = 0; +x_54 = l_Std_Time_Month_Ordinal_days(x_53, x_46); +x_55 = lean_int_dec_lt(x_54, x_40); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_48); +lean_ctor_set(x_56, 1, x_46); +lean_ctor_set(x_56, 2, x_40); +return x_56; +} +else +{ +lean_object* x_57; +lean_dec(x_40); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_48); +lean_ctor_set(x_57, 1, x_46); +lean_ctor_set(x_57, 2, x_54); +return x_57; +} +} +else +{ +lean_object* x_58; lean_object* x_59; uint8_t x_60; uint8_t x_61; +x_58 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_59 = lean_int_mod(x_48, x_58); +x_60 = lean_int_dec_eq(x_59, x_51); +lean_dec(x_59); +x_61 = l_instDecidableNot___rarg(x_60); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_62 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_63 = lean_int_mod(x_48, x_62); +x_64 = lean_int_dec_eq(x_63, x_51); +lean_dec(x_63); +if (x_64 == 0) +{ +uint8_t x_65; lean_object* x_66; uint8_t x_67; +x_65 = 0; +x_66 = l_Std_Time_Month_Ordinal_days(x_65, x_46); +x_67 = lean_int_dec_lt(x_66, x_40); +if (x_67 == 0) +{ +lean_object* x_68; +lean_dec(x_66); +x_68 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_68, 0, x_48); +lean_ctor_set(x_68, 1, x_46); +lean_ctor_set(x_68, 2, x_40); +return x_68; +} +else +{ +lean_object* x_69; +lean_dec(x_40); +x_69 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_69, 0, x_48); +lean_ctor_set(x_69, 1, x_46); +lean_ctor_set(x_69, 2, x_66); +return x_69; +} +} +else +{ +uint8_t x_70; lean_object* x_71; uint8_t x_72; +x_70 = 1; +x_71 = l_Std_Time_Month_Ordinal_days(x_70, x_46); +x_72 = lean_int_dec_lt(x_71, x_40); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_71); +x_73 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_73, 0, x_48); +lean_ctor_set(x_73, 1, x_46); +lean_ctor_set(x_73, 2, x_40); +return x_73; +} +else +{ +lean_object* x_74; +lean_dec(x_40); +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_48); +lean_ctor_set(x_74, 1, x_46); +lean_ctor_set(x_74, 2, x_71); +return x_74; +} +} +} +else +{ +uint8_t x_75; lean_object* x_76; uint8_t x_77; +x_75 = 1; +x_76 = l_Std_Time_Month_Ordinal_days(x_75, x_46); +x_77 = lean_int_dec_lt(x_76, x_40); +if (x_77 == 0) +{ +lean_object* x_78; +lean_dec(x_76); +x_78 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_78, 0, x_48); +lean_ctor_set(x_78, 1, x_46); +lean_ctor_set(x_78, 2, x_40); +return x_78; +} +else +{ +lean_object* x_79; +lean_dec(x_40); +x_79 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_79, 0, x_48); +lean_ctor_set(x_79, 1, x_46); +lean_ctor_set(x_79, 2, x_76); +return x_79; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainDate_addMonthsClip(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_rollOver___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDate___closed__6; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_rollOver___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDate_rollOver___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_rollOver___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDate_rollOver___closed__2; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_rollOver___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDate_rollOver___closed__3; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_rollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_5 = lean_int_mod(x_1, x_4); +x_6 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_9 = lean_int_sub(x_3, x_8); +if (x_7 == 0) +{ +uint8_t x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = 0; +x_16 = l_Std_Time_Month_Ordinal_days(x_15, x_2); +x_17 = l_Std_Time_PlainDate_rollOver___closed__4; +x_18 = lean_int_dec_lt(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_16); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_2); +lean_ctor_set(x_19, 2, x_17); +x_10 = x_19; +goto block_14; +} +else +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_2); +lean_ctor_set(x_20, 2, x_16); +x_10 = x_20; +goto block_14; +} +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; +x_21 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_22 = lean_int_mod(x_1, x_21); +x_23 = lean_int_dec_eq(x_22, x_6); +lean_dec(x_22); +x_24 = l_instDecidableNot___rarg(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_26 = lean_int_mod(x_1, x_25); +x_27 = lean_int_dec_eq(x_26, x_6); +lean_dec(x_26); +if (x_27 == 0) +{ +uint8_t x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = 0; +x_29 = l_Std_Time_Month_Ordinal_days(x_28, x_2); +x_30 = l_Std_Time_PlainDate_rollOver___closed__4; +x_31 = lean_int_dec_lt(x_29, x_30); +if (x_31 == 0) +{ +lean_object* x_32; +lean_dec(x_29); +x_32 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_2); +lean_ctor_set(x_32, 2, x_30); +x_10 = x_32; +goto block_14; +} +else +{ +lean_object* x_33; +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_2); +lean_ctor_set(x_33, 2, x_29); +x_10 = x_33; +goto block_14; +} +} +else +{ +uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = 1; +x_35 = l_Std_Time_Month_Ordinal_days(x_34, x_2); +x_36 = l_Std_Time_PlainDate_rollOver___closed__4; +x_37 = lean_int_dec_lt(x_35, x_36); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_35); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_2); +lean_ctor_set(x_38, 2, x_36); +x_10 = x_38; +goto block_14; +} +else +{ +lean_object* x_39; +x_39 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_39, 0, x_1); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_35); +x_10 = x_39; +goto block_14; +} +} +} +else +{ +uint8_t x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = 1; +x_41 = l_Std_Time_Month_Ordinal_days(x_40, x_2); +x_42 = l_Std_Time_PlainDate_rollOver___closed__4; +x_43 = lean_int_dec_lt(x_41, x_42); +if (x_43 == 0) +{ +lean_object* x_44; +lean_dec(x_41); +x_44 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_44, 0, x_1); +lean_ctor_set(x_44, 1, x_2); +lean_ctor_set(x_44, 2, x_42); +x_10 = x_44; +goto block_14; +} +else +{ +lean_object* x_45; +x_45 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_45, 0, x_1); +lean_ctor_set(x_45, 1, x_2); +lean_ctor_set(x_45, 2, x_41); +x_10 = x_45; +goto block_14; +} +} +} +block_14: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_10); +x_12 = lean_int_add(x_11, x_9); +lean_dec(x_9); +lean_dec(x_11); +x_13 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_12); +lean_dec(x_12); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_rollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_PlainDate_rollOver(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withYearClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_1, 0); +lean_dec(x_6); +x_7 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_8 = lean_int_mod(x_2, x_7); +x_9 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_10 = lean_int_dec_eq(x_8, x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +uint8_t x_11; lean_object* x_12; uint8_t x_13; +x_11 = 0; +x_12 = l_Std_Time_Month_Ordinal_days(x_11, x_4); +x_13 = lean_int_dec_lt(x_12, x_5); +if (x_13 == 0) +{ +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_12); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_14 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_15 = lean_int_mod(x_2, x_14); +x_16 = lean_int_dec_eq(x_15, x_9); +lean_dec(x_15); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_19 = lean_int_mod(x_2, x_18); +x_20 = lean_int_dec_eq(x_19, x_9); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; uint8_t x_23; +x_21 = 0; +x_22 = l_Std_Time_Month_Ordinal_days(x_21, x_4); +x_23 = lean_int_dec_lt(x_22, x_5); +if (x_23 == 0) +{ +lean_dec(x_22); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_22); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +} +else +{ +uint8_t x_24; lean_object* x_25; uint8_t x_26; +x_24 = 1; +x_25 = l_Std_Time_Month_Ordinal_days(x_24, x_4); +x_26 = lean_int_dec_lt(x_25, x_5); +if (x_26 == 0) +{ +lean_dec(x_25); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_25); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_27 = 1; +x_28 = l_Std_Time_Month_Ordinal_days(x_27, x_4); +x_29 = lean_int_dec_lt(x_28, x_5); +if (x_29 == 0) +{ +lean_dec(x_28); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_28); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_30 = lean_ctor_get(x_1, 1); +x_31 = lean_ctor_get(x_1, 2); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_1); +x_32 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_33 = lean_int_mod(x_2, x_32); +x_34 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_35 = lean_int_dec_eq(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; uint8_t x_38; +x_36 = 0; +x_37 = l_Std_Time_Month_Ordinal_days(x_36, x_30); +x_38 = lean_int_dec_lt(x_37, x_31); +if (x_38 == 0) +{ +lean_object* x_39; +lean_dec(x_37); +x_39 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_39, 0, x_2); +lean_ctor_set(x_39, 1, x_30); +lean_ctor_set(x_39, 2, x_31); +return x_39; +} +else +{ +lean_object* x_40; +lean_dec(x_31); +x_40 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_40, 0, x_2); +lean_ctor_set(x_40, 1, x_30); +lean_ctor_set(x_40, 2, x_37); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; +x_41 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_42 = lean_int_mod(x_2, x_41); +x_43 = lean_int_dec_eq(x_42, x_34); +lean_dec(x_42); +x_44 = l_instDecidableNot___rarg(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_46 = lean_int_mod(x_2, x_45); +x_47 = lean_int_dec_eq(x_46, x_34); +lean_dec(x_46); +if (x_47 == 0) +{ +uint8_t x_48; lean_object* x_49; uint8_t x_50; +x_48 = 0; +x_49 = l_Std_Time_Month_Ordinal_days(x_48, x_30); +x_50 = lean_int_dec_lt(x_49, x_31); +if (x_50 == 0) +{ +lean_object* x_51; +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_2); +lean_ctor_set(x_51, 1, x_30); +lean_ctor_set(x_51, 2, x_31); +return x_51; +} +else +{ +lean_object* x_52; +lean_dec(x_31); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_2); +lean_ctor_set(x_52, 1, x_30); +lean_ctor_set(x_52, 2, x_49); +return x_52; +} +} +else +{ +uint8_t x_53; lean_object* x_54; uint8_t x_55; +x_53 = 1; +x_54 = l_Std_Time_Month_Ordinal_days(x_53, x_30); +x_55 = lean_int_dec_lt(x_54, x_31); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_2); +lean_ctor_set(x_56, 1, x_30); +lean_ctor_set(x_56, 2, x_31); +return x_56; +} +else +{ +lean_object* x_57; +lean_dec(x_31); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_2); +lean_ctor_set(x_57, 1, x_30); +lean_ctor_set(x_57, 2, x_54); +return x_57; +} +} +} +else +{ +uint8_t x_58; lean_object* x_59; uint8_t x_60; +x_58 = 1; +x_59 = l_Std_Time_Month_Ordinal_days(x_58, x_30); +x_60 = lean_int_dec_lt(x_59, x_31); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_59); +x_61 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_61, 0, x_2); +lean_ctor_set(x_61, 1, x_30); +lean_ctor_set(x_61, 2, x_31); +return x_61; +} +else +{ +lean_object* x_62; +lean_dec(x_31); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_2); +lean_ctor_set(x_62, 1, x_30); +lean_ctor_set(x_62, 2, x_59); +return x_62; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withYearRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_rollOver(x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_8 = lean_int_mod(x_4, x_7); +x_9 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_10 = lean_int_dec_eq(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_12 = lean_int_sub(x_6, x_11); +lean_dec(x_6); +if (x_10 == 0) +{ +uint8_t x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = 0; +x_20 = l_Std_Time_Month_Ordinal_days(x_19, x_5); +x_21 = l_Std_Time_PlainDate_rollOver___closed__4; +x_22 = lean_int_dec_lt(x_20, x_21); +if (x_22 == 0) +{ +lean_dec(x_20); +lean_ctor_set(x_1, 2, x_21); +x_13 = x_1; +goto block_18; +} +else +{ +lean_ctor_set(x_1, 2, x_20); +x_13 = x_1; +goto block_18; +} +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; +x_23 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_24 = lean_int_mod(x_4, x_23); +x_25 = lean_int_dec_eq(x_24, x_9); +lean_dec(x_24); +x_26 = l_instDecidableNot___rarg(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_27 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_28 = lean_int_mod(x_4, x_27); +x_29 = lean_int_dec_eq(x_28, x_9); +lean_dec(x_28); +if (x_29 == 0) +{ +uint8_t x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = 0; +x_31 = l_Std_Time_Month_Ordinal_days(x_30, x_5); +x_32 = l_Std_Time_PlainDate_rollOver___closed__4; +x_33 = lean_int_dec_lt(x_31, x_32); +if (x_33 == 0) +{ +lean_dec(x_31); +lean_ctor_set(x_1, 2, x_32); +x_13 = x_1; +goto block_18; +} +else +{ +lean_ctor_set(x_1, 2, x_31); +x_13 = x_1; +goto block_18; +} +} +else +{ +uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_34 = 1; +x_35 = l_Std_Time_Month_Ordinal_days(x_34, x_5); +x_36 = l_Std_Time_PlainDate_rollOver___closed__4; +x_37 = lean_int_dec_lt(x_35, x_36); +if (x_37 == 0) +{ +lean_dec(x_35); +lean_ctor_set(x_1, 2, x_36); +x_13 = x_1; +goto block_18; +} +else +{ +lean_ctor_set(x_1, 2, x_35); +x_13 = x_1; +goto block_18; +} +} +} +else +{ +uint8_t x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = 1; +x_39 = l_Std_Time_Month_Ordinal_days(x_38, x_5); +x_40 = l_Std_Time_PlainDate_rollOver___closed__4; +x_41 = lean_int_dec_lt(x_39, x_40); +if (x_41 == 0) +{ +lean_dec(x_39); +lean_ctor_set(x_1, 2, x_40); +x_13 = x_1; +goto block_18; +} +else +{ +lean_ctor_set(x_1, 2, x_39); +x_13 = x_1; +goto block_18; +} +} +} +block_18: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = l_Std_Time_PlainDate_addMonthsClip(x_13, x_2); +x_15 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +x_16 = lean_int_add(x_15, x_12); +lean_dec(x_12); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_16); +lean_dec(x_16); +return x_17; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_ctor_get(x_1, 1); +x_44 = lean_ctor_get(x_1, 2); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_1); +x_45 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_46 = lean_int_mod(x_42, x_45); +x_47 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_48 = lean_int_dec_eq(x_46, x_47); +lean_dec(x_46); +x_49 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_50 = lean_int_sub(x_44, x_49); +lean_dec(x_44); +if (x_48 == 0) +{ +uint8_t x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_57 = 0; +x_58 = l_Std_Time_Month_Ordinal_days(x_57, x_43); +x_59 = l_Std_Time_PlainDate_rollOver___closed__4; +x_60 = lean_int_dec_lt(x_58, x_59); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_58); +x_61 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_61, 0, x_42); +lean_ctor_set(x_61, 1, x_43); +lean_ctor_set(x_61, 2, x_59); +x_51 = x_61; +goto block_56; +} +else +{ +lean_object* x_62; +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_42); +lean_ctor_set(x_62, 1, x_43); +lean_ctor_set(x_62, 2, x_58); +x_51 = x_62; +goto block_56; +} +} +else +{ +lean_object* x_63; lean_object* x_64; uint8_t x_65; uint8_t x_66; +x_63 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_64 = lean_int_mod(x_42, x_63); +x_65 = lean_int_dec_eq(x_64, x_47); +lean_dec(x_64); +x_66 = l_instDecidableNot___rarg(x_65); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_68 = lean_int_mod(x_42, x_67); +x_69 = lean_int_dec_eq(x_68, x_47); +lean_dec(x_68); +if (x_69 == 0) +{ +uint8_t x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_70 = 0; +x_71 = l_Std_Time_Month_Ordinal_days(x_70, x_43); +x_72 = l_Std_Time_PlainDate_rollOver___closed__4; +x_73 = lean_int_dec_lt(x_71, x_72); +if (x_73 == 0) +{ +lean_object* x_74; +lean_dec(x_71); +x_74 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_74, 0, x_42); +lean_ctor_set(x_74, 1, x_43); +lean_ctor_set(x_74, 2, x_72); +x_51 = x_74; +goto block_56; +} +else +{ +lean_object* x_75; +x_75 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_75, 0, x_42); +lean_ctor_set(x_75, 1, x_43); +lean_ctor_set(x_75, 2, x_71); +x_51 = x_75; +goto block_56; +} +} +else +{ +uint8_t x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = 1; +x_77 = l_Std_Time_Month_Ordinal_days(x_76, x_43); +x_78 = l_Std_Time_PlainDate_rollOver___closed__4; +x_79 = lean_int_dec_lt(x_77, x_78); +if (x_79 == 0) +{ +lean_object* x_80; +lean_dec(x_77); +x_80 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_80, 0, x_42); +lean_ctor_set(x_80, 1, x_43); +lean_ctor_set(x_80, 2, x_78); +x_51 = x_80; +goto block_56; +} +else +{ +lean_object* x_81; +x_81 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_81, 0, x_42); +lean_ctor_set(x_81, 1, x_43); +lean_ctor_set(x_81, 2, x_77); +x_51 = x_81; +goto block_56; +} +} +} +else +{ +uint8_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_82 = 1; +x_83 = l_Std_Time_Month_Ordinal_days(x_82, x_43); +x_84 = l_Std_Time_PlainDate_rollOver___closed__4; +x_85 = lean_int_dec_lt(x_83, x_84); +if (x_85 == 0) +{ +lean_object* x_86; +lean_dec(x_83); +x_86 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_86, 0, x_42); +lean_ctor_set(x_86, 1, x_43); +lean_ctor_set(x_86, 2, x_84); +x_51 = x_86; +goto block_56; +} +else +{ +lean_object* x_87; +x_87 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_87, 0, x_42); +lean_ctor_set(x_87, 1, x_43); +lean_ctor_set(x_87, 2, x_83); +x_51 = x_87; +goto block_56; +} +} +} +block_56: +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = l_Std_Time_PlainDate_addMonthsClip(x_51, x_2); +x_53 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_52); +x_54 = lean_int_add(x_53, x_50); +lean_dec(x_50); +lean_dec(x_53); +x_55 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_54); +lean_dec(x_54); +return x_55; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainDate_addMonthsRollOver(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_PlainDate_addMonthsRollOver(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDate_addMonthsRollOver(x_1, x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_PlainDate_addMonthsClip(x_1, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_addYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_addYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDate_addMonthsClip(x_1, x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_subYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_subYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +lean_dec(x_6); +x_7 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_8 = lean_int_mod(x_4, x_7); +x_9 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_10 = lean_int_dec_eq(x_8, x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +uint8_t x_11; lean_object* x_12; uint8_t x_13; +x_11 = 0; +x_12 = l_Std_Time_Month_Ordinal_days(x_11, x_5); +x_13 = lean_int_dec_lt(x_12, x_2); +if (x_13 == 0) +{ +lean_dec(x_12); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_1, 2, x_12); +return x_1; +} +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_14 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_15 = lean_int_mod(x_4, x_14); +x_16 = lean_int_dec_eq(x_15, x_9); +lean_dec(x_15); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_19 = lean_int_mod(x_4, x_18); +x_20 = lean_int_dec_eq(x_19, x_9); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; uint8_t x_23; +x_21 = 0; +x_22 = l_Std_Time_Month_Ordinal_days(x_21, x_5); +x_23 = lean_int_dec_lt(x_22, x_2); +if (x_23 == 0) +{ +lean_dec(x_22); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_1, 2, x_22); +return x_1; +} +} +else +{ +uint8_t x_24; lean_object* x_25; uint8_t x_26; +x_24 = 1; +x_25 = l_Std_Time_Month_Ordinal_days(x_24, x_5); +x_26 = lean_int_dec_lt(x_25, x_2); +if (x_26 == 0) +{ +lean_dec(x_25); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_1, 2, x_25); +return x_1; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_27 = 1; +x_28 = l_Std_Time_Month_Ordinal_days(x_27, x_5); +x_29 = lean_int_dec_lt(x_28, x_2); +if (x_29 == 0) +{ +lean_dec(x_28); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_1, 2, x_28); +return x_1; +} +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_1); +x_32 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_33 = lean_int_mod(x_30, x_32); +x_34 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_35 = lean_int_dec_eq(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; uint8_t x_38; +x_36 = 0; +x_37 = l_Std_Time_Month_Ordinal_days(x_36, x_31); +x_38 = lean_int_dec_lt(x_37, x_2); +if (x_38 == 0) +{ +lean_object* x_39; +lean_dec(x_37); +x_39 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_39, 0, x_30); +lean_ctor_set(x_39, 1, x_31); +lean_ctor_set(x_39, 2, x_2); +return x_39; +} +else +{ +lean_object* x_40; +lean_dec(x_2); +x_40 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_40, 0, x_30); +lean_ctor_set(x_40, 1, x_31); +lean_ctor_set(x_40, 2, x_37); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; +x_41 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_42 = lean_int_mod(x_30, x_41); +x_43 = lean_int_dec_eq(x_42, x_34); +lean_dec(x_42); +x_44 = l_instDecidableNot___rarg(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_46 = lean_int_mod(x_30, x_45); +x_47 = lean_int_dec_eq(x_46, x_34); +lean_dec(x_46); +if (x_47 == 0) +{ +uint8_t x_48; lean_object* x_49; uint8_t x_50; +x_48 = 0; +x_49 = l_Std_Time_Month_Ordinal_days(x_48, x_31); +x_50 = lean_int_dec_lt(x_49, x_2); +if (x_50 == 0) +{ +lean_object* x_51; +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_31); +lean_ctor_set(x_51, 2, x_2); +return x_51; +} +else +{ +lean_object* x_52; +lean_dec(x_2); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_30); +lean_ctor_set(x_52, 1, x_31); +lean_ctor_set(x_52, 2, x_49); +return x_52; +} +} +else +{ +uint8_t x_53; lean_object* x_54; uint8_t x_55; +x_53 = 1; +x_54 = l_Std_Time_Month_Ordinal_days(x_53, x_31); +x_55 = lean_int_dec_lt(x_54, x_2); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_30); +lean_ctor_set(x_56, 1, x_31); +lean_ctor_set(x_56, 2, x_2); +return x_56; +} +else +{ +lean_object* x_57; +lean_dec(x_2); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_30); +lean_ctor_set(x_57, 1, x_31); +lean_ctor_set(x_57, 2, x_54); +return x_57; +} +} +} +else +{ +uint8_t x_58; lean_object* x_59; uint8_t x_60; +x_58 = 1; +x_59 = l_Std_Time_Month_Ordinal_days(x_58, x_31); +x_60 = lean_int_dec_lt(x_59, x_2); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_59); +x_61 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_61, 0, x_30); +lean_ctor_set(x_61, 1, x_31); +lean_ctor_set(x_61, 2, x_2); +return x_61; +} +else +{ +lean_object* x_62; +lean_dec(x_2); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_30); +lean_ctor_set(x_62, 1, x_31); +lean_ctor_set(x_62, 2, x_59); +return x_62; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_rollOver(x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withDaysRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_withDaysRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withMonthClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_1, 1); +lean_dec(x_6); +x_7 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_8 = lean_int_mod(x_4, x_7); +x_9 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_10 = lean_int_dec_eq(x_8, x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +uint8_t x_11; lean_object* x_12; uint8_t x_13; +x_11 = 0; +x_12 = l_Std_Time_Month_Ordinal_days(x_11, x_2); +x_13 = lean_int_dec_lt(x_12, x_5); +if (x_13 == 0) +{ +lean_dec(x_12); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_12); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_14 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_15 = lean_int_mod(x_4, x_14); +x_16 = lean_int_dec_eq(x_15, x_9); +lean_dec(x_15); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_19 = lean_int_mod(x_4, x_18); +x_20 = lean_int_dec_eq(x_19, x_9); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; uint8_t x_23; +x_21 = 0; +x_22 = l_Std_Time_Month_Ordinal_days(x_21, x_2); +x_23 = lean_int_dec_lt(x_22, x_5); +if (x_23 == 0) +{ +lean_dec(x_22); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_22); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +} +else +{ +uint8_t x_24; lean_object* x_25; uint8_t x_26; +x_24 = 1; +x_25 = l_Std_Time_Month_Ordinal_days(x_24, x_2); +x_26 = lean_int_dec_lt(x_25, x_5); +if (x_26 == 0) +{ +lean_dec(x_25); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_25); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +} +} +else +{ +uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_27 = 1; +x_28 = l_Std_Time_Month_Ordinal_days(x_27, x_2); +x_29 = lean_int_dec_lt(x_28, x_5); +if (x_29 == 0) +{ +lean_dec(x_28); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_5); +lean_ctor_set(x_1, 2, x_28); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 2); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_1); +x_32 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_33 = lean_int_mod(x_30, x_32); +x_34 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_35 = lean_int_dec_eq(x_33, x_34); +lean_dec(x_33); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; uint8_t x_38; +x_36 = 0; +x_37 = l_Std_Time_Month_Ordinal_days(x_36, x_2); +x_38 = lean_int_dec_lt(x_37, x_31); +if (x_38 == 0) +{ +lean_object* x_39; +lean_dec(x_37); +x_39 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_39, 0, x_30); +lean_ctor_set(x_39, 1, x_2); +lean_ctor_set(x_39, 2, x_31); +return x_39; +} +else +{ +lean_object* x_40; +lean_dec(x_31); +x_40 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_40, 0, x_30); +lean_ctor_set(x_40, 1, x_2); +lean_ctor_set(x_40, 2, x_37); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; +x_41 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_42 = lean_int_mod(x_30, x_41); +x_43 = lean_int_dec_eq(x_42, x_34); +lean_dec(x_42); +x_44 = l_instDecidableNot___rarg(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_46 = lean_int_mod(x_30, x_45); +x_47 = lean_int_dec_eq(x_46, x_34); +lean_dec(x_46); +if (x_47 == 0) +{ +uint8_t x_48; lean_object* x_49; uint8_t x_50; +x_48 = 0; +x_49 = l_Std_Time_Month_Ordinal_days(x_48, x_2); +x_50 = lean_int_dec_lt(x_49, x_31); +if (x_50 == 0) +{ +lean_object* x_51; +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_30); +lean_ctor_set(x_51, 1, x_2); +lean_ctor_set(x_51, 2, x_31); +return x_51; +} +else +{ +lean_object* x_52; +lean_dec(x_31); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_30); +lean_ctor_set(x_52, 1, x_2); +lean_ctor_set(x_52, 2, x_49); +return x_52; +} +} +else +{ +uint8_t x_53; lean_object* x_54; uint8_t x_55; +x_53 = 1; +x_54 = l_Std_Time_Month_Ordinal_days(x_53, x_2); +x_55 = lean_int_dec_lt(x_54, x_31); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_30); +lean_ctor_set(x_56, 1, x_2); +lean_ctor_set(x_56, 2, x_31); +return x_56; +} +else +{ +lean_object* x_57; +lean_dec(x_31); +x_57 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_57, 0, x_30); +lean_ctor_set(x_57, 1, x_2); +lean_ctor_set(x_57, 2, x_54); +return x_57; +} +} +} +else +{ +uint8_t x_58; lean_object* x_59; uint8_t x_60; +x_58 = 1; +x_59 = l_Std_Time_Month_Ordinal_days(x_58, x_2); +x_60 = lean_int_dec_lt(x_59, x_31); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_59); +x_61 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_61, 0, x_30); +lean_ctor_set(x_61, 1, x_2); +lean_ctor_set(x_61, 2, x_31); +return x_61; +} +else +{ +lean_object* x_62; +lean_dec(x_31); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_30); +lean_ctor_set(x_62, 1, x_2); +lean_ctor_set(x_62, 2, x_59); +return x_62; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withMonthRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_rollOver(x_3, x_2, x_4); +lean_dec(x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_PlainDate_weekday___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_weekday___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_weekday___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDate_weekday___closed__2; +x_2 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDate_weekday___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDate_weekday(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_3 = l_Std_Time_PlainDate_weekday___closed__1; +x_4 = lean_int_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_5 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7; +x_6 = lean_int_add(x_2, x_5); +lean_dec(x_2); +x_7 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_8 = lean_int_emod(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_PlainDate_weekday___closed__4; +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_12 = lean_int_sub(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainDate_weekday___closed__3; +x_14 = lean_int_emod(x_12, x_13); +lean_dec(x_12); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = lean_int_emod(x_15, x_13); +lean_dec(x_15); +x_17 = lean_int_add(x_16, x_11); +lean_dec(x_16); +x_18 = l_Std_Time_Weekday_ofOrdinal(x_17); +lean_dec(x_17); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_19 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_20 = lean_int_add(x_2, x_19); +lean_dec(x_2); +x_21 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_22 = lean_int_emod(x_20, x_21); +lean_dec(x_20); +x_23 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_24 = lean_int_sub(x_22, x_23); +lean_dec(x_22); +x_25 = l_Std_Time_PlainDate_weekday___closed__3; +x_26 = lean_int_emod(x_24, x_25); +lean_dec(x_24); +x_27 = lean_int_add(x_26, x_25); +lean_dec(x_26); +x_28 = lean_int_emod(x_27, x_25); +lean_dec(x_27); +x_29 = lean_int_add(x_28, x_23); +lean_dec(x_28); +x_30 = l_Std_Time_Weekday_ofOrdinal(x_29); +lean_dec(x_29); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekday___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDate_weekday(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = !lean_is_exclusive(x_1); +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_7 = lean_int_mod(x_3, x_6); +x_8 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_9 = lean_int_dec_eq(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainDate_weekOfMonth___closed__1; +x_11 = lean_int_add(x_5, x_10); +lean_dec(x_5); +if (x_9 == 0) +{ +uint8_t x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = 0; +x_23 = l_Std_Time_Month_Ordinal_days(x_22, x_4); +x_24 = l_Std_Time_PlainDate_rollOver___closed__4; +x_25 = lean_int_dec_lt(x_23, x_24); +if (x_25 == 0) +{ +lean_dec(x_23); +lean_ctor_set(x_1, 2, x_24); +x_12 = x_1; +goto block_21; +} +else +{ +lean_ctor_set(x_1, 2, x_23); +x_12 = x_1; +goto block_21; +} +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; +x_26 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_27 = lean_int_mod(x_3, x_26); +x_28 = lean_int_dec_eq(x_27, x_8); +lean_dec(x_27); +x_29 = l_instDecidableNot___rarg(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_31 = lean_int_mod(x_3, x_30); +x_32 = lean_int_dec_eq(x_31, x_8); +lean_dec(x_31); +if (x_32 == 0) +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = 0; +x_34 = l_Std_Time_Month_Ordinal_days(x_33, x_4); +x_35 = l_Std_Time_PlainDate_rollOver___closed__4; +x_36 = lean_int_dec_lt(x_34, x_35); +if (x_36 == 0) +{ +lean_dec(x_34); +lean_ctor_set(x_1, 2, x_35); +x_12 = x_1; +goto block_21; +} +else +{ +lean_ctor_set(x_1, 2, x_34); +x_12 = x_1; +goto block_21; +} +} +else +{ +uint8_t x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_37 = 1; +x_38 = l_Std_Time_Month_Ordinal_days(x_37, x_4); +x_39 = l_Std_Time_PlainDate_rollOver___closed__4; +x_40 = lean_int_dec_lt(x_38, x_39); +if (x_40 == 0) +{ +lean_dec(x_38); +lean_ctor_set(x_1, 2, x_39); +x_12 = x_1; +goto block_21; +} +else +{ +lean_ctor_set(x_1, 2, x_38); +x_12 = x_1; +goto block_21; +} +} +} +else +{ +uint8_t x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = 1; +x_42 = l_Std_Time_Month_Ordinal_days(x_41, x_4); +x_43 = l_Std_Time_PlainDate_rollOver___closed__4; +x_44 = lean_int_dec_lt(x_42, x_43); +if (x_44 == 0) +{ +lean_dec(x_42); +lean_ctor_set(x_1, 2, x_43); +x_12 = x_1; +goto block_21; +} +else +{ +lean_ctor_set(x_1, 2, x_42); +x_12 = x_1; +goto block_21; +} +} +} +block_21: +{ +uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_13 = l_Std_Time_PlainDate_weekday(x_12); +x_14 = l_Std_Time_Weekday_toOrdinal(x_13); +x_15 = lean_int_add(x_14, x_10); +lean_dec(x_14); +x_16 = lean_int_add(x_11, x_15); +lean_dec(x_15); +lean_dec(x_11); +x_17 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_18 = lean_int_ediv(x_16, x_17); +lean_dec(x_16); +x_19 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_18); +return x_20; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_45 = lean_ctor_get(x_1, 0); +x_46 = lean_ctor_get(x_1, 1); +x_47 = lean_ctor_get(x_1, 2); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_1); +x_48 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_49 = lean_int_mod(x_45, x_48); +x_50 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_51 = lean_int_dec_eq(x_49, x_50); +lean_dec(x_49); +x_52 = l_Std_Time_PlainDate_weekOfMonth___closed__1; +x_53 = lean_int_add(x_47, x_52); +lean_dec(x_47); +if (x_51 == 0) +{ +uint8_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = 0; +x_65 = l_Std_Time_Month_Ordinal_days(x_64, x_46); +x_66 = l_Std_Time_PlainDate_rollOver___closed__4; +x_67 = lean_int_dec_lt(x_65, x_66); +if (x_67 == 0) +{ +lean_object* x_68; +lean_dec(x_65); +x_68 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_68, 0, x_45); +lean_ctor_set(x_68, 1, x_46); +lean_ctor_set(x_68, 2, x_66); +x_54 = x_68; +goto block_63; +} +else +{ +lean_object* x_69; +x_69 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_69, 0, x_45); +lean_ctor_set(x_69, 1, x_46); +lean_ctor_set(x_69, 2, x_65); +x_54 = x_69; +goto block_63; +} +} +else +{ +lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; +x_70 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_71 = lean_int_mod(x_45, x_70); +x_72 = lean_int_dec_eq(x_71, x_50); +lean_dec(x_71); +x_73 = l_instDecidableNot___rarg(x_72); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_74 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_75 = lean_int_mod(x_45, x_74); +x_76 = lean_int_dec_eq(x_75, x_50); +lean_dec(x_75); +if (x_76 == 0) +{ +uint8_t x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_77 = 0; +x_78 = l_Std_Time_Month_Ordinal_days(x_77, x_46); +x_79 = l_Std_Time_PlainDate_rollOver___closed__4; +x_80 = lean_int_dec_lt(x_78, x_79); +if (x_80 == 0) +{ +lean_object* x_81; +lean_dec(x_78); +x_81 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_81, 0, x_45); +lean_ctor_set(x_81, 1, x_46); +lean_ctor_set(x_81, 2, x_79); +x_54 = x_81; +goto block_63; +} +else +{ +lean_object* x_82; +x_82 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_82, 0, x_45); +lean_ctor_set(x_82, 1, x_46); +lean_ctor_set(x_82, 2, x_78); +x_54 = x_82; +goto block_63; +} +} +else +{ +uint8_t x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_83 = 1; +x_84 = l_Std_Time_Month_Ordinal_days(x_83, x_46); +x_85 = l_Std_Time_PlainDate_rollOver___closed__4; +x_86 = lean_int_dec_lt(x_84, x_85); +if (x_86 == 0) +{ +lean_object* x_87; +lean_dec(x_84); +x_87 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_87, 0, x_45); +lean_ctor_set(x_87, 1, x_46); +lean_ctor_set(x_87, 2, x_85); +x_54 = x_87; +goto block_63; +} +else +{ +lean_object* x_88; +x_88 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_88, 0, x_45); +lean_ctor_set(x_88, 1, x_46); +lean_ctor_set(x_88, 2, x_84); +x_54 = x_88; +goto block_63; +} +} +} +else +{ +uint8_t x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = 1; +x_90 = l_Std_Time_Month_Ordinal_days(x_89, x_46); +x_91 = l_Std_Time_PlainDate_rollOver___closed__4; +x_92 = lean_int_dec_lt(x_90, x_91); +if (x_92 == 0) +{ +lean_object* x_93; +lean_dec(x_90); +x_93 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_93, 0, x_45); +lean_ctor_set(x_93, 1, x_46); +lean_ctor_set(x_93, 2, x_91); +x_54 = x_93; +goto block_63; +} +else +{ +lean_object* x_94; +x_94 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_94, 0, x_45); +lean_ctor_set(x_94, 1, x_46); +lean_ctor_set(x_94, 2, x_90); +x_54 = x_94; +goto block_63; +} +} +} +block_63: +{ +uint8_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_55 = l_Std_Time_PlainDate_weekday(x_54); +x_56 = l_Std_Time_Weekday_toOrdinal(x_55); +x_57 = lean_int_add(x_56, x_52); +lean_dec(x_56); +x_58 = lean_int_add(x_53, x_57); +lean_dec(x_57); +lean_dec(x_53); +x_59 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_60 = lean_int_ediv(x_58, x_59); +lean_dec(x_58); +x_61 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_62 = lean_int_add(x_60, x_61); +lean_dec(x_60); +return x_62; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withWeekday(lean_object* x_1, uint8_t x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; +lean_inc(x_1); +x_3 = l_Std_Time_PlainDate_weekday(x_1); +x_4 = l_Std_Time_Weekday_toOrdinal(x_3); +x_5 = l_Std_Time_Weekday_toOrdinal(x_2); +x_6 = lean_int_neg(x_4); +lean_dec(x_4); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_6); +lean_dec(x_5); +x_8 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_9 = lean_int_dec_lt(x_7, x_8); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +if (x_9 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_int_add(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_11); +lean_dec(x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_14 = lean_int_add(x_7, x_13); +lean_dec(x_7); +x_15 = lean_int_add(x_10, x_14); +lean_dec(x_14); +lean_dec(x_10); +x_16 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_15); +lean_dec(x_15); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_withWeekday___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Std_Time_PlainDate_withWeekday(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1; +x_4 = lean_int_mod(x_2, x_3); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8; +x_6 = lean_int_dec_eq(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_dayOfYear(x_1); +x_8 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10; +x_9 = lean_int_add(x_8, x_7); +lean_dec(x_7); +x_10 = l_Std_Time_PlainDate_weekday(x_1); +x_11 = l_Std_Time_Weekday_toOrdinal(x_10); +x_12 = lean_int_neg(x_11); +lean_dec(x_11); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16; +x_15 = lean_int_ediv(x_13, x_14); +lean_dec(x_13); +x_16 = l_Std_Time_instInhabitedPlainDate___closed__1; +x_17 = lean_int_dec_lt(x_15, x_16); +if (x_6 == 0) +{ +lean_object* x_25; +x_25 = lean_unsigned_to_nat(365u); +x_18 = x_25; +goto block_24; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; +x_26 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2; +x_27 = lean_int_mod(x_2, x_26); +x_28 = lean_int_dec_eq(x_27, x_5); +lean_dec(x_27); +x_29 = l_instDecidableNot___rarg(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3; +x_31 = lean_int_mod(x_2, x_30); +x_32 = lean_int_dec_eq(x_31, x_5); +lean_dec(x_31); +if (x_32 == 0) +{ +lean_object* x_33; +x_33 = lean_unsigned_to_nat(365u); +x_18 = x_33; +goto block_24; +} +else +{ +lean_object* x_34; +x_34 = lean_unsigned_to_nat(366u); +x_18 = x_34; +goto block_24; +} +} +else +{ +lean_object* x_35; +x_35 = lean_unsigned_to_nat(366u); +x_18 = x_35; +goto block_24; +} +} +block_24: +{ +if (x_17 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = l_Std_Time_Year_Offset_weeks(x_2); +lean_dec(x_2); +x_20 = lean_int_dec_lt(x_19, x_15); +lean_dec(x_19); +if (x_20 == 0) +{ +return x_15; +} +else +{ +lean_object* x_21; +lean_dec(x_15); +x_21 = l_Std_Time_instInhabitedPlainDate___closed__1; +return x_21; +} +} +else +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_15); +x_22 = lean_int_sub(x_2, x_16); +lean_dec(x_2); +x_23 = l_Std_Time_Year_Offset_weeks(x_22); +lean_dec(x_22); +return x_23; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_addDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_subDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_addWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_subWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instHSubOffset__1___closed__1; +return x_1; +} +} +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_PlainDate(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__1); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__2); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__3); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__4); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__5); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__6); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__7); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__8); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__9); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__10); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__11); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__12); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__13); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__14); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__15); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__16); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__17); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__18); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__19); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__20); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__21); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__22); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__23); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__24); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__25); +l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26 = _init_l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26(); +lean_mark_persistent(l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40____closed__26); +l_Std_Time_instReprPlainDate___closed__1 = _init_l_Std_Time_instReprPlainDate___closed__1(); +lean_mark_persistent(l_Std_Time_instReprPlainDate___closed__1); +l_Std_Time_instReprPlainDate = _init_l_Std_Time_instReprPlainDate(); +lean_mark_persistent(l_Std_Time_instReprPlainDate); +l_Std_Time_instInhabitedPlainDate___closed__1 = _init_l_Std_Time_instInhabitedPlainDate___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__1); +l_Std_Time_instInhabitedPlainDate___closed__2 = _init_l_Std_Time_instInhabitedPlainDate___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__2); +l_Std_Time_instInhabitedPlainDate___closed__3 = _init_l_Std_Time_instInhabitedPlainDate___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__3); +l_Std_Time_instInhabitedPlainDate___closed__4 = _init_l_Std_Time_instInhabitedPlainDate___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__4); +l_Std_Time_instInhabitedPlainDate___closed__5 = _init_l_Std_Time_instInhabitedPlainDate___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__5); +l_Std_Time_instInhabitedPlainDate___closed__6 = _init_l_Std_Time_instInhabitedPlainDate___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__6); +l_Std_Time_instInhabitedPlainDate___closed__7 = _init_l_Std_Time_instInhabitedPlainDate___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__7); +l_Std_Time_instInhabitedPlainDate___closed__8 = _init_l_Std_Time_instInhabitedPlainDate___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__8); +l_Std_Time_instInhabitedPlainDate___closed__9 = _init_l_Std_Time_instInhabitedPlainDate___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__9); +l_Std_Time_instInhabitedPlainDate___closed__10 = _init_l_Std_Time_instInhabitedPlainDate___closed__10(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__10); +l_Std_Time_instInhabitedPlainDate___closed__11 = _init_l_Std_Time_instInhabitedPlainDate___closed__11(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__11); +l_Std_Time_instInhabitedPlainDate___closed__12 = _init_l_Std_Time_instInhabitedPlainDate___closed__12(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__12); +l_Std_Time_instInhabitedPlainDate___closed__13 = _init_l_Std_Time_instInhabitedPlainDate___closed__13(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__13); +l_Std_Time_instInhabitedPlainDate___closed__14 = _init_l_Std_Time_instInhabitedPlainDate___closed__14(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__14); +l_Std_Time_instInhabitedPlainDate___closed__15 = _init_l_Std_Time_instInhabitedPlainDate___closed__15(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__15); +l_Std_Time_instInhabitedPlainDate___closed__16 = _init_l_Std_Time_instInhabitedPlainDate___closed__16(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__16); +l_Std_Time_instInhabitedPlainDate___closed__17 = _init_l_Std_Time_instInhabitedPlainDate___closed__17(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__17); +l_Std_Time_instInhabitedPlainDate___closed__18 = _init_l_Std_Time_instInhabitedPlainDate___closed__18(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__18); +l_Std_Time_instInhabitedPlainDate___closed__19 = _init_l_Std_Time_instInhabitedPlainDate___closed__19(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate___closed__19); +l_Std_Time_instInhabitedPlainDate = _init_l_Std_Time_instInhabitedPlainDate(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDate); +l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1 = _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_ofYearMonthDayClip___closed__1); +l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2 = _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_ofYearMonthDayClip___closed__2); +l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3 = _init_l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_ofYearMonthDayClip___closed__3); +l_Std_Time_PlainDate_instInhabited___closed__1 = _init_l_Std_Time_PlainDate_instInhabited___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instInhabited___closed__1); +l_Std_Time_PlainDate_instInhabited = _init_l_Std_Time_PlainDate_instInhabited(); +lean_mark_persistent(l_Std_Time_PlainDate_instInhabited); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__1); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__2); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__3); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__4); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__5); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__6); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__7); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__8); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__9); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__10); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__11); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__12); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__13); +l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14 = _init_l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14(); +lean_mark_persistent(l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch___closed__14); +l_Std_Time_PlainDate_weekOfMonth___closed__1 = _init_l_Std_Time_PlainDate_weekOfMonth___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_weekOfMonth___closed__1); +l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1 = _init_l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__1); +l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2 = _init_l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_toDaysSinceUNIXEpoch___closed__2); +l_Std_Time_PlainDate_rollOver___closed__1 = _init_l_Std_Time_PlainDate_rollOver___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_rollOver___closed__1); +l_Std_Time_PlainDate_rollOver___closed__2 = _init_l_Std_Time_PlainDate_rollOver___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_rollOver___closed__2); +l_Std_Time_PlainDate_rollOver___closed__3 = _init_l_Std_Time_PlainDate_rollOver___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_rollOver___closed__3); +l_Std_Time_PlainDate_rollOver___closed__4 = _init_l_Std_Time_PlainDate_rollOver___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDate_rollOver___closed__4); +l_Std_Time_PlainDate_weekday___closed__1 = _init_l_Std_Time_PlainDate_weekday___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_weekday___closed__1); +l_Std_Time_PlainDate_weekday___closed__2 = _init_l_Std_Time_PlainDate_weekday___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_weekday___closed__2); +l_Std_Time_PlainDate_weekday___closed__3 = _init_l_Std_Time_PlainDate_weekday___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_weekday___closed__3); +l_Std_Time_PlainDate_weekday___closed__4 = _init_l_Std_Time_PlainDate_weekday___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDate_weekday___closed__4); +l_Std_Time_PlainDate_instHAddOffset___closed__1 = _init_l_Std_Time_PlainDate_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHAddOffset___closed__1); +l_Std_Time_PlainDate_instHAddOffset = _init_l_Std_Time_PlainDate_instHAddOffset(); +lean_mark_persistent(l_Std_Time_PlainDate_instHAddOffset); +l_Std_Time_PlainDate_instHSubOffset___closed__1 = _init_l_Std_Time_PlainDate_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHSubOffset___closed__1); +l_Std_Time_PlainDate_instHSubOffset = _init_l_Std_Time_PlainDate_instHSubOffset(); +lean_mark_persistent(l_Std_Time_PlainDate_instHSubOffset); +l_Std_Time_PlainDate_instHAddOffset__1___closed__1 = _init_l_Std_Time_PlainDate_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHAddOffset__1___closed__1); +l_Std_Time_PlainDate_instHAddOffset__1 = _init_l_Std_Time_PlainDate_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHAddOffset__1); +l_Std_Time_PlainDate_instHSubOffset__1___closed__1 = _init_l_Std_Time_PlainDate_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHSubOffset__1___closed__1); +l_Std_Time_PlainDate_instHSubOffset__1 = _init_l_Std_Time_PlainDate_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHSubOffset__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Basic.c b/stage0/stdlib/Std/Time/Date/Unit/Basic.c new file mode 100644 index 000000000000..d3307c82dee6 --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Basic.c @@ -0,0 +1,100 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Basic +// Imports: Std.Time.Date.Unit.Day Std.Time.Date.Unit.Month Std.Time.Date.Unit.Year Std.Time.Date.Unit.Weekday Std.Time.Date.Unit.Week +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Day_Offset_ofWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toWeeks___boxed(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofWeeks___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toWeeks(lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofWeeks(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Day_Offset_ofWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_ofWeeks___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toWeeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_ofWeeks___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toWeeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toWeeks(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Month(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Year(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Weekday(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Week(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Month(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Year(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Weekday(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Week(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Day_Offset_ofWeeks___closed__1 = _init_l_Std_Time_Day_Offset_ofWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_ofWeeks___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Day.c b/stage0/stdlib/Std/Time/Date/Unit/Day.c new file mode 100644 index 000000000000..99b598e9bbb6 --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Day.c @@ -0,0 +1,1554 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Day +// Imports: Std.Time.Time +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLtOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetBEq; +static lean_object* l_Std_Time_Day_instOrdinalBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetLE; +static lean_object* l_Std_Time_Day_Offset_toNanoseconds___closed__1; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOfNatOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14; +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263_; +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOfNatOrdinal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOrdinalBEq; +static lean_object* l_Std_Time_Day_instOffsetAdd___closed__1; +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__6; +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instInhabitedOfYear(uint8_t); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13; +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__8; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24; +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetAdd; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMinutes___boxed(lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat___rarg(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOffsetSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear(uint8_t); +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOrdinalRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetToString; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___rarg(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset(uint8_t); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2; +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_toOffset(lean_object*); +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__7; +static lean_object* l_Std_Time_Day_Offset_toMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instOfNatOfYear___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_Offset_toMinutes___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toHours___boxed(lean_object*); +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_Offset_toSeconds___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4; +lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat___boxed(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOffsetNeg___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toSeconds(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMinutes(lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__4; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOffsetBEq___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMilliseconds(lean_object*); +static lean_object* l_Std_Time_Day_instOfNatOrdinal___closed__1; +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOrdinalLE; +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLeOffset(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOffsetInhabited___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instOfNatOfYear(uint8_t, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toOrdinal(lean_object*, lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_Day_instOffsetToString___closed__1; +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOrdinalRepr; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetRepr; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofInt___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetNeg; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat(uint8_t); +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLtOffset(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofSeconds(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg(lean_object*); +lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMilliseconds(lean_object*); +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__2; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9; +static lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLeOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetSub; +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLtOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetLT; +static lean_object* l_Std_Time_Day_instOffsetRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_instInhabitedOrdinal; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofFin(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLeOffset___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instInhabitedOfYear___boxed(lean_object*); +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__9; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26; +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofNat(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_413_; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21; +static lean_object* l_Std_Time_Day_instInhabitedOrdinal___closed__3; +lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toNanoseconds(lean_object*); +static lean_object* l_Std_Time_Day_Offset_toHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofInt___boxed(lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofInt(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOffsetInhabited; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22; +LEAN_EXPORT lean_object* l_Std_Time_Day_instOrdinalLT; +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_toOffset___boxed(lean_object*); +static lean_object* _init_l_Std_Time_Day_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(30u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Day_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Day_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Day_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__5; +x_2 = l_Std_Time_Day_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Day_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Day_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instInhabitedOrdinal___closed__9; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetInhabited___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetInhabited___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Day_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Day_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Day_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Day_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Day_instDecidableLtOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Day_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_Day_instOffsetInhabited___closed__1; +x_4 = lean_int_dec_lt(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Int_repr(x_1); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Int_repr(x_1); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Day_Ordinal_instReprOfYear___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Day_Ordinal_instReprOfYear___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instReprOfYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Day_Ordinal_instReprOfYear(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_Day_instOffsetInhabited___closed__1; +x_3 = lean_int_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_nat_abs(x_1); +x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_nat_abs(x_1); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_6, x_7); +lean_dec(x_6); +x_9 = lean_nat_add(x_8, x_7); +lean_dec(x_8); +x_10 = l___private_Init_Data_Repr_0__Nat_reprFast(x_9); +x_11 = l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1; +x_12 = lean_string_append(x_11, x_10); +lean_dec(x_10); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Ordinal_instToStringOfYear___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instToStringOfYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Day_Ordinal_instToStringOfYear(x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("decide", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11; +x_3 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5; +x_3 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Day_Ordinal_OfYear_ofNat___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_ofNat___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Day_Ordinal_OfYear_ofNat(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instOfNatOfYear(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_4 = lean_unsigned_to_nat(364u); +x_5 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_3, x_2, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +x_7 = lean_unsigned_to_nat(365u); +x_8 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_6, x_2, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instOfNatOfYear___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Day_Ordinal_instOfNatOfYear(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instInhabitedOfYear(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_instInhabitedOfYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Day_Ordinal_instInhabitedOfYear(x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_413_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_le(x_2, x_1); +if (x_3 == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_Std_Time_Day_instOfNatOrdinal___closed__1; +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_nat_to_int(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Ordinal_OfYear_toOffset___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Ordinal_OfYear_toOffset___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Day_Ordinal_OfYear_toOffset(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Day_Offset_toOrdinal(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_Offset_toNanoseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("86400000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toNanoseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toNanoseconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_Offset_toMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toMilliseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toMilliseconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_Offset_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toSeconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_Offset_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1440u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toMinutes___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Day_Offset_toHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Day_Offset_toHours___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Day_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Day_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Day_instOrdinalRepr___closed__1 = _init_l_Std_Time_Day_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalRepr___closed__1); +l_Std_Time_Day_instOrdinalRepr = _init_l_Std_Time_Day_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalRepr); +l_Std_Time_Day_instOrdinalBEq___closed__1 = _init_l_Std_Time_Day_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalBEq___closed__1); +l_Std_Time_Day_instOrdinalBEq = _init_l_Std_Time_Day_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalBEq); +l_Std_Time_Day_instOrdinalLE = _init_l_Std_Time_Day_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalLE); +l_Std_Time_Day_instOrdinalLT = _init_l_Std_Time_Day_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Day_instOrdinalLT); +l_Std_Time_Day_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Day_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOfNatOrdinal___closed__1); +l_Std_Time_Day_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__1); +l_Std_Time_Day_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__2); +l_Std_Time_Day_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__3); +l_Std_Time_Day_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__4); +l_Std_Time_Day_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__5); +l_Std_Time_Day_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__6); +l_Std_Time_Day_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__7); +l_Std_Time_Day_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__8); +l_Std_Time_Day_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Day_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal___closed__9); +l_Std_Time_Day_instInhabitedOrdinal = _init_l_Std_Time_Day_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Day_instInhabitedOrdinal); +l_Std_Time_Day_instOffsetRepr___closed__1 = _init_l_Std_Time_Day_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetRepr___closed__1); +l_Std_Time_Day_instOffsetRepr = _init_l_Std_Time_Day_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Day_instOffsetRepr); +l_Std_Time_Day_instOffsetBEq___closed__1 = _init_l_Std_Time_Day_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetBEq___closed__1); +l_Std_Time_Day_instOffsetBEq = _init_l_Std_Time_Day_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Day_instOffsetBEq); +l_Std_Time_Day_instOffsetInhabited___closed__1 = _init_l_Std_Time_Day_instOffsetInhabited___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetInhabited___closed__1); +l_Std_Time_Day_instOffsetInhabited = _init_l_Std_Time_Day_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Day_instOffsetInhabited); +l_Std_Time_Day_instOffsetAdd___closed__1 = _init_l_Std_Time_Day_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetAdd___closed__1); +l_Std_Time_Day_instOffsetAdd = _init_l_Std_Time_Day_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Day_instOffsetAdd); +l_Std_Time_Day_instOffsetSub___closed__1 = _init_l_Std_Time_Day_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetSub___closed__1); +l_Std_Time_Day_instOffsetSub = _init_l_Std_Time_Day_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Day_instOffsetSub); +l_Std_Time_Day_instOffsetNeg___closed__1 = _init_l_Std_Time_Day_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetNeg___closed__1); +l_Std_Time_Day_instOffsetNeg = _init_l_Std_Time_Day_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Day_instOffsetNeg); +l_Std_Time_Day_instOffsetLE = _init_l_Std_Time_Day_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Day_instOffsetLE); +l_Std_Time_Day_instOffsetLT = _init_l_Std_Time_Day_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Day_instOffsetLT); +l_Std_Time_Day_instOffsetToString___closed__1 = _init_l_Std_Time_Day_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Day_instOffsetToString___closed__1); +l_Std_Time_Day_instOffsetToString = _init_l_Std_Time_Day_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Day_instOffsetToString); +l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1 = _init_l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Ordinal_instToStringOfYear___rarg___closed__1); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__1); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__2); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__3); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__4); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__5); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__6); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__7); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__8); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__9); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__10); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__11); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__12); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__13); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__14); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__15); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__16); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__17); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__18); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__19); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__20); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__21); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__22); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__23); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__24); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__25); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__26); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263____closed__27); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_263_ = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_263_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_263_); +l___auto____x40_Std_Time_Date_Unit_Day___hyg_413_ = _init_l___auto____x40_Std_Time_Date_Unit_Day___hyg_413_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Day___hyg_413_); +l_Std_Time_Day_Offset_toNanoseconds___closed__1 = _init_l_Std_Time_Day_Offset_toNanoseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_toNanoseconds___closed__1); +l_Std_Time_Day_Offset_toMilliseconds___closed__1 = _init_l_Std_Time_Day_Offset_toMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_toMilliseconds___closed__1); +l_Std_Time_Day_Offset_toSeconds___closed__1 = _init_l_Std_Time_Day_Offset_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_toSeconds___closed__1); +l_Std_Time_Day_Offset_toMinutes___closed__1 = _init_l_Std_Time_Day_Offset_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_toMinutes___closed__1); +l_Std_Time_Day_Offset_toHours___closed__1 = _init_l_Std_Time_Day_Offset_toHours___closed__1(); +lean_mark_persistent(l_Std_Time_Day_Offset_toHours___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Month.c b/stage0/stdlib/Std/Time/Date/Unit/Month.c new file mode 100644 index 000000000000..1e368a236137 --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Month.c @@ -0,0 +1,3303 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Month +// Imports: Std.Internal.Rat Std.Time.Date.Unit.Day +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__1; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__4; +static lean_object* l_Std_Time_Month_instOffsetRepr___closed__1; +static lean_object* l_Std_Time_Month_Ordinal_days___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap; +static lean_object* l_Std_Time_Month_Quarter_ofMonth___closed__1; +static lean_object* l_Std_Time_Month_Ordinal_december___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_october___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalLE; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__3; +lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_march___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Month_Quarter_ofMonth___boxed(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetSub; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetBEq; +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLtOrdinal(lean_object*, lean_object*); +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_february___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toSeconds(uint8_t, lean_object*); +extern lean_object* l_Int_instDiv; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9; +static lean_object* l_Std_Time_Month_Ordinal_november___closed__5; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__5; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__23; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11; +extern lean_object* l_Int_instInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset___boxed(lean_object*); +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalRepr; +static lean_object* l_Std_Time_Month_Ordinal_february___closed__4; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_march___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29; +static lean_object* l_Std_Time_Month_Ordinal_august___closed__5; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Month_instInhabitedOrdinal; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__3; +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7; +static lean_object* l_Std_Time_Month_Ordinal_october___closed__4; +static lean_object* l_Std_Time_Month_Ordinal_may___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_august___closed__1; +lean_object* lean_array_fget(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__26; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__9; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5; +static lean_object* l_Std_Time_Month_instOffsetBEq___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; +static lean_object* l_Std_Time_Month_Ordinal_august___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetAdd; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17; +static lean_object* l_Std_Time_Month_Ordinal_october___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19; +static lean_object* l_Std_Time_Month_Ordinal_may___closed__5; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__6; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_days___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toMinutes(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetRepr; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__10; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofInt___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_december___closed__3; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8; +static lean_object* l_Std_Time_Month_Ordinal_days___closed__3; +static lean_object* l_Std_Time_Month_instOffsetToString___closed__1; +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_instOrdinalBEq___closed__1; +static lean_object* l_Std_Time_Month_instOfNatOrdinal___closed__1; +static lean_object* l_Std_Time_Month_Ordinal_days___closed__5; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__24; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_november___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetLE; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toNat___boxed(lean_object*); +extern lean_object* l_Int_instNegInt; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__4; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__20; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__5; +static lean_object* l_Std_Time_Month_Ordinal_december___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toSeconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toHours(uint8_t, lean_object*); +static lean_object* l_Std_Time_Month_instOffsetBEq___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_august___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_clipDay___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_march___closed__5; +lean_object* l_outOfBounds___rarg(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_july___closed__6; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7; +static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__1; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_july; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10; +static lean_object* l_Std_Time_Month_Ordinal_may___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13; +extern lean_object* l_Int_instMul; +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofNat(lean_object*); +lean_object* l_instReprInt___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_december___closed__4; +extern lean_object* l_Int_instLTInt; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__25; +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_november___closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22; +static lean_object* l_Std_Time_Month_Ordinal_october___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalLT; +lean_object* l_Std_Internal_Rat_div(lean_object*, lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toDays(uint8_t, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLeOrdinal(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__11; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_march; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_clipDay(uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetToString; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toHours___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_october; +static lean_object* l_Std_Time_Month_Ordinal_december___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Month_Quarter_ofMonth(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toNat___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetDecidableEq___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_may; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__8; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__1; +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_december; +static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_april; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25; +static lean_object* l_Std_Time_Month_Ordinal_november___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_june; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12; +static lean_object* l_Std_Time_Month_Ordinal_february___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__6; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27; +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_; +static lean_object* l_Std_Time_Month_Ordinal_july___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__2; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +static lean_object* l_Std_Time_Month_Ordinal_may___closed__4; +lean_object* lean_nat_abs(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__2; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_november; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_july___closed__4; +lean_object* l_Int_instDecidableEq___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__22; +lean_object* l_instToStringInt___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetLT; +static lean_object* l_Std_Time_Month_Ordinal_april___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__18; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofFin(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_march___closed__2; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_october___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_february; +extern lean_object* l_Int_instAdd; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_august; +static lean_object* l_Std_Time_Month_Ordinal_october___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetNeg; +static lean_object* l_Std_Time_Month_Ordinal_july___closed__5; +static lean_object* l_Std_Time_Month_Ordinal_february___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__14; +extern lean_object* l_Int_instSub; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__21; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__15; +lean_object* l_Int_toNat(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_may___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetDiv; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__17; +static lean_object* l_Std_Time_Month_instOrdinalRepr___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23; +static lean_object* l_Std_Time_Month_Ordinal_july___closed__2; +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_september; +static lean_object* l_Std_Time_Month_Ordinal_toSeconds___closed__12; +static lean_object* l_Std_Time_Month_Ordinal_august___closed__6; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_february___closed__6; +static lean_object* l_Std_Time_Month_Ordinal_days___closed__2; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOrdinalBEq; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofNat(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toMinutes___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_december___closed__1; +lean_object* lean_int_neg(lean_object*); +static lean_object* l_Std_Time_Month_Quarter_ofMonth___closed__2; +static lean_object* l_Std_Time_Month_Ordinal_march___closed__3; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_days___closed__1; +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_july___closed__1; +static lean_object* l_Std_Time_Month_Ordinal_may___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__4; +static lean_object* l_Std_Time_Month_Ordinal_toMinutes___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; +extern lean_object* l_Int_instLEInt; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__5; +static lean_object* l_Std_Time_Month_instInhabitedOrdinal___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_january; +extern lean_object* l_Std_Time_Day_instOffsetInhabited; +static lean_object* l_Std_Time_Month_Ordinal_september___closed__4; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15; +static lean_object* l_Std_Time_Month_Ordinal_june___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOfNatOffset(lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_february___closed__1; +static lean_object* l_Std_Time_Month_Ordinal_november___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOfNatOrdinal(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Month_instOffsetDecidableEq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toNat(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28; +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetMul; +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Month_Ordinal_august___closed__3; +static lean_object* l_Std_Time_Month_Ordinal_toDays___closed__2; +static lean_object* _init_l_Std_Time_Month_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(11u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__5; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__9; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Month_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Month_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Month_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instReprInt___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Int_instDecidableEq___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetBEq___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Month_instOffsetBEq___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instOffsetBEq___closed__2; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instInhabited; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instAdd; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instSub; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetMul() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instMul; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetDiv() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instDiv; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instNegInt; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instToStringInt___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instOffsetToString___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instLTInt; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instLEInt; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Month_instOffsetDecidableEq(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instOffsetDecidableEq___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Month_instOffsetDecidableEq(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Quarter_ofMonth___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Quarter_ofMonth___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Quarter_ofMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l_Std_Time_Month_Quarter_ofMonth___closed__2; +x_3 = lean_int_add(x_1, x_2); +x_4 = l_Std_Time_Month_Quarter_ofMonth___closed__1; +x_5 = lean_int_ediv(x_3, x_4); +lean_dec(x_3); +x_6 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Quarter_ofMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Month_Quarter_ofMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Month_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_january() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__9; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_february___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_february___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_february___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_february___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_february___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_february() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_february___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Quarter_ofMonth___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_march___closed__1; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_march___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_march___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_march___closed__4; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_march() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_march___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_april___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_april___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_april___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_april___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_april___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_april() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_april___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_may___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_may___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_may___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_may___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_may___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_may() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_may___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_june___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_june___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_june___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_june___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_june___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_june() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_june___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_july___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_july___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_july___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_july___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_july___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_july() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_july___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_august___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_august___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_august___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_august___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_august___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_august() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_august___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_september___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_september___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_september___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_september___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_september___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_september() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_september___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_october___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_october___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_october___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_october___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_october___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_october() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_october___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instInhabitedOrdinal___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_november___closed__1; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_november___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_november___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_november___closed__4; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_november() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_november___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_december___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_december___closed__2; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_december___closed__3; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_december___closed__4; +x_2 = l_Std_Time_Month_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_december___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_december() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_december___closed__6; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Month_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Month_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("decide", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11; +x_3 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5; +x_3 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toNat___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_abs(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toNat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Month_Ordinal_toNat(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_le(x_2, x_1); +if (x_3 == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_nat_to_int(x_1); +return x_5; +} +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(31u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(90u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(120u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(151u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(181u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(212u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(243u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(273u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(304u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(334u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__11; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__10; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__9; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__8; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__7; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__6; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__5; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__19; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__4; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__3; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__2; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__22; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toNat___closed__1; +x_2 = l_Std_Time_Month_Ordinal_toSeconds___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__24; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toSeconds___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__25; +x_2 = lean_array_get_size(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toSeconds(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_nat_abs(x_2); +x_4 = l_Std_Time_Month_Ordinal_toSeconds___closed__26; +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = l_Std_Time_Day_instOffsetInhabited; +x_7 = l_outOfBounds___rarg(x_6); +x_8 = l_Std_Time_Month_Ordinal_toSeconds___closed__1; +x_9 = lean_int_mul(x_7, x_8); +lean_dec(x_7); +if (x_1 == 0) +{ +lean_dec(x_3); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(2u); +x_11 = lean_nat_dec_le(x_10, x_3); +lean_dec(x_3); +if (x_11 == 0) +{ +return x_9; +} +else +{ +lean_object* x_12; +x_12 = lean_int_add(x_9, x_8); +lean_dec(x_9); +return x_12; +} +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = l_Std_Time_Month_Ordinal_toSeconds___closed__25; +x_14 = lean_array_fget(x_13, x_3); +x_15 = l_Std_Time_Month_Ordinal_toSeconds___closed__1; +x_16 = lean_int_mul(x_14, x_15); +lean_dec(x_14); +if (x_1 == 0) +{ +lean_dec(x_3); +return x_16; +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_unsigned_to_nat(2u); +x_18 = lean_nat_dec_le(x_17, x_3); +lean_dec(x_3); +if (x_18 == 0) +{ +return x_16; +} +else +{ +lean_object* x_19; +x_19 = lean_int_add(x_16, x_15); +lean_dec(x_16); +return x_19; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_toSeconds(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toMinutes(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Month_Ordinal_toSeconds(x_1, x_2); +x_4 = l_Std_Time_Month_Ordinal_toMinutes___closed__1; +x_5 = lean_int_ediv(x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_toMinutes(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toHours(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Std_Time_Month_Ordinal_toSeconds(x_1, x_2); +x_4 = l_Std_Time_Month_Ordinal_toMinutes___closed__1; +x_5 = lean_int_ediv(x_3, x_4); +lean_dec(x_3); +x_6 = lean_int_ediv(x_5, x_4); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_toHours(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toDays___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__1; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toDays___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toDays___closed__1; +x_2 = l_Std_Time_Month_Ordinal_toDays___closed__2; +x_3 = l_Std_Internal_Rat_div(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_toDays___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toDays___closed__3; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = lean_nat_to_int(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toDays(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = l_Std_Time_Month_Ordinal_toSeconds(x_1, x_2); +x_4 = l_Std_Time_Month_Ordinal_toDays___closed__3; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_7 = l_Std_Time_Month_Ordinal_toDays___closed__4; +x_8 = lean_int_ediv(x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_toDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_toDays(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__2; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(28u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32; +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Month_Ordinal_toSeconds___closed__25; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(29u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_days___closed__1; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_days___closed__2; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_days___closed__3; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_days___closed__4; +x_2 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Month_Ordinal_days___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Month_Ordinal_days___closed__5; +x_2 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_Month_Ordinal_february___closed__1; +x_4 = lean_int_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = l_Std_Time_Month_Quarter_ofMonth___closed__2; +x_6 = lean_int_add(x_2, x_5); +x_7 = l_Int_toNat(x_6); +lean_dec(x_6); +x_8 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap; +x_9 = lean_array_fget(x_8, x_7); +lean_dec(x_7); +return x_9; +} +else +{ +if (x_1 == 0) +{ +lean_object* x_10; +x_10 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15; +return x_10; +} +else +{ +lean_object* x_11; +x_11 = l_Std_Time_Month_Ordinal_days___closed__6; +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_days___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_days(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_Month_Quarter_ofMonth___closed__2; +x_4 = lean_int_add(x_2, x_3); +x_5 = l_Int_toNat(x_4); +lean_dec(x_4); +x_6 = l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes; +x_7 = lean_array_fget(x_6, x_5); +lean_dec(x_5); +if (x_1 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = l_Std_Time_Month_Ordinal_toNat___closed__1; +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +x_10 = l_Std_Time_Month_Ordinal_february___closed__1; +x_11 = lean_int_dec_lt(x_10, x_2); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = l_Std_Time_Month_Ordinal_toNat___closed__1; +x_13 = lean_int_add(x_7, x_12); +lean_dec(x_7); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Std_Time_Month_instOfNatOrdinal___closed__1; +x_15 = lean_int_add(x_7, x_14); +lean_dec(x_7); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_cumulativeDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Month_Ordinal_cumulativeDays(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_clipDay(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_Month_Ordinal_days(x_1, x_2); +x_5 = lean_int_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +lean_dec(x_4); +lean_inc(x_3); +return x_3; +} +else +{ +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Month_Ordinal_clipDay___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l_Std_Time_Month_Ordinal_clipDay(x_4, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Month(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Month_instOrdinalRepr___closed__1 = _init_l_Std_Time_Month_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalRepr___closed__1); +l_Std_Time_Month_instOrdinalRepr = _init_l_Std_Time_Month_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalRepr); +l_Std_Time_Month_instOrdinalBEq___closed__1 = _init_l_Std_Time_Month_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalBEq___closed__1); +l_Std_Time_Month_instOrdinalBEq = _init_l_Std_Time_Month_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalBEq); +l_Std_Time_Month_instOrdinalLE = _init_l_Std_Time_Month_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalLE); +l_Std_Time_Month_instOrdinalLT = _init_l_Std_Time_Month_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Month_instOrdinalLT); +l_Std_Time_Month_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Month_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOfNatOrdinal___closed__1); +l_Std_Time_Month_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__1); +l_Std_Time_Month_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__2); +l_Std_Time_Month_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__3); +l_Std_Time_Month_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__4); +l_Std_Time_Month_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__5); +l_Std_Time_Month_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__6); +l_Std_Time_Month_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__7); +l_Std_Time_Month_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__8); +l_Std_Time_Month_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Month_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal___closed__9); +l_Std_Time_Month_instInhabitedOrdinal = _init_l_Std_Time_Month_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Month_instInhabitedOrdinal); +l_Std_Time_Month_instOffsetRepr___closed__1 = _init_l_Std_Time_Month_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOffsetRepr___closed__1); +l_Std_Time_Month_instOffsetRepr = _init_l_Std_Time_Month_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Month_instOffsetRepr); +l_Std_Time_Month_instOffsetBEq___closed__1 = _init_l_Std_Time_Month_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOffsetBEq___closed__1); +l_Std_Time_Month_instOffsetBEq___closed__2 = _init_l_Std_Time_Month_instOffsetBEq___closed__2(); +lean_mark_persistent(l_Std_Time_Month_instOffsetBEq___closed__2); +l_Std_Time_Month_instOffsetBEq = _init_l_Std_Time_Month_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Month_instOffsetBEq); +l_Std_Time_Month_instOffsetInhabited = _init_l_Std_Time_Month_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Month_instOffsetInhabited); +l_Std_Time_Month_instOffsetAdd = _init_l_Std_Time_Month_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Month_instOffsetAdd); +l_Std_Time_Month_instOffsetSub = _init_l_Std_Time_Month_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Month_instOffsetSub); +l_Std_Time_Month_instOffsetMul = _init_l_Std_Time_Month_instOffsetMul(); +lean_mark_persistent(l_Std_Time_Month_instOffsetMul); +l_Std_Time_Month_instOffsetDiv = _init_l_Std_Time_Month_instOffsetDiv(); +lean_mark_persistent(l_Std_Time_Month_instOffsetDiv); +l_Std_Time_Month_instOffsetNeg = _init_l_Std_Time_Month_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Month_instOffsetNeg); +l_Std_Time_Month_instOffsetToString___closed__1 = _init_l_Std_Time_Month_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Month_instOffsetToString___closed__1); +l_Std_Time_Month_instOffsetToString = _init_l_Std_Time_Month_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Month_instOffsetToString); +l_Std_Time_Month_instOffsetLT = _init_l_Std_Time_Month_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Month_instOffsetLT); +l_Std_Time_Month_instOffsetLE = _init_l_Std_Time_Month_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Month_instOffsetLE); +l_Std_Time_Month_Quarter_ofMonth___closed__1 = _init_l_Std_Time_Month_Quarter_ofMonth___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Quarter_ofMonth___closed__1); +l_Std_Time_Month_Quarter_ofMonth___closed__2 = _init_l_Std_Time_Month_Quarter_ofMonth___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Quarter_ofMonth___closed__2); +l_Std_Time_Month_Ordinal_january = _init_l_Std_Time_Month_Ordinal_january(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_january); +l_Std_Time_Month_Ordinal_february___closed__1 = _init_l_Std_Time_Month_Ordinal_february___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__1); +l_Std_Time_Month_Ordinal_february___closed__2 = _init_l_Std_Time_Month_Ordinal_february___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__2); +l_Std_Time_Month_Ordinal_february___closed__3 = _init_l_Std_Time_Month_Ordinal_february___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__3); +l_Std_Time_Month_Ordinal_february___closed__4 = _init_l_Std_Time_Month_Ordinal_february___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__4); +l_Std_Time_Month_Ordinal_february___closed__5 = _init_l_Std_Time_Month_Ordinal_february___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__5); +l_Std_Time_Month_Ordinal_february___closed__6 = _init_l_Std_Time_Month_Ordinal_february___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february___closed__6); +l_Std_Time_Month_Ordinal_february = _init_l_Std_Time_Month_Ordinal_february(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_february); +l_Std_Time_Month_Ordinal_march___closed__1 = _init_l_Std_Time_Month_Ordinal_march___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march___closed__1); +l_Std_Time_Month_Ordinal_march___closed__2 = _init_l_Std_Time_Month_Ordinal_march___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march___closed__2); +l_Std_Time_Month_Ordinal_march___closed__3 = _init_l_Std_Time_Month_Ordinal_march___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march___closed__3); +l_Std_Time_Month_Ordinal_march___closed__4 = _init_l_Std_Time_Month_Ordinal_march___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march___closed__4); +l_Std_Time_Month_Ordinal_march___closed__5 = _init_l_Std_Time_Month_Ordinal_march___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march___closed__5); +l_Std_Time_Month_Ordinal_march = _init_l_Std_Time_Month_Ordinal_march(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_march); +l_Std_Time_Month_Ordinal_april___closed__1 = _init_l_Std_Time_Month_Ordinal_april___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__1); +l_Std_Time_Month_Ordinal_april___closed__2 = _init_l_Std_Time_Month_Ordinal_april___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__2); +l_Std_Time_Month_Ordinal_april___closed__3 = _init_l_Std_Time_Month_Ordinal_april___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__3); +l_Std_Time_Month_Ordinal_april___closed__4 = _init_l_Std_Time_Month_Ordinal_april___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__4); +l_Std_Time_Month_Ordinal_april___closed__5 = _init_l_Std_Time_Month_Ordinal_april___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__5); +l_Std_Time_Month_Ordinal_april___closed__6 = _init_l_Std_Time_Month_Ordinal_april___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april___closed__6); +l_Std_Time_Month_Ordinal_april = _init_l_Std_Time_Month_Ordinal_april(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_april); +l_Std_Time_Month_Ordinal_may___closed__1 = _init_l_Std_Time_Month_Ordinal_may___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__1); +l_Std_Time_Month_Ordinal_may___closed__2 = _init_l_Std_Time_Month_Ordinal_may___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__2); +l_Std_Time_Month_Ordinal_may___closed__3 = _init_l_Std_Time_Month_Ordinal_may___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__3); +l_Std_Time_Month_Ordinal_may___closed__4 = _init_l_Std_Time_Month_Ordinal_may___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__4); +l_Std_Time_Month_Ordinal_may___closed__5 = _init_l_Std_Time_Month_Ordinal_may___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__5); +l_Std_Time_Month_Ordinal_may___closed__6 = _init_l_Std_Time_Month_Ordinal_may___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may___closed__6); +l_Std_Time_Month_Ordinal_may = _init_l_Std_Time_Month_Ordinal_may(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_may); +l_Std_Time_Month_Ordinal_june___closed__1 = _init_l_Std_Time_Month_Ordinal_june___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__1); +l_Std_Time_Month_Ordinal_june___closed__2 = _init_l_Std_Time_Month_Ordinal_june___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__2); +l_Std_Time_Month_Ordinal_june___closed__3 = _init_l_Std_Time_Month_Ordinal_june___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__3); +l_Std_Time_Month_Ordinal_june___closed__4 = _init_l_Std_Time_Month_Ordinal_june___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__4); +l_Std_Time_Month_Ordinal_june___closed__5 = _init_l_Std_Time_Month_Ordinal_june___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__5); +l_Std_Time_Month_Ordinal_june___closed__6 = _init_l_Std_Time_Month_Ordinal_june___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june___closed__6); +l_Std_Time_Month_Ordinal_june = _init_l_Std_Time_Month_Ordinal_june(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_june); +l_Std_Time_Month_Ordinal_july___closed__1 = _init_l_Std_Time_Month_Ordinal_july___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__1); +l_Std_Time_Month_Ordinal_july___closed__2 = _init_l_Std_Time_Month_Ordinal_july___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__2); +l_Std_Time_Month_Ordinal_july___closed__3 = _init_l_Std_Time_Month_Ordinal_july___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__3); +l_Std_Time_Month_Ordinal_july___closed__4 = _init_l_Std_Time_Month_Ordinal_july___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__4); +l_Std_Time_Month_Ordinal_july___closed__5 = _init_l_Std_Time_Month_Ordinal_july___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__5); +l_Std_Time_Month_Ordinal_july___closed__6 = _init_l_Std_Time_Month_Ordinal_july___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july___closed__6); +l_Std_Time_Month_Ordinal_july = _init_l_Std_Time_Month_Ordinal_july(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_july); +l_Std_Time_Month_Ordinal_august___closed__1 = _init_l_Std_Time_Month_Ordinal_august___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__1); +l_Std_Time_Month_Ordinal_august___closed__2 = _init_l_Std_Time_Month_Ordinal_august___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__2); +l_Std_Time_Month_Ordinal_august___closed__3 = _init_l_Std_Time_Month_Ordinal_august___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__3); +l_Std_Time_Month_Ordinal_august___closed__4 = _init_l_Std_Time_Month_Ordinal_august___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__4); +l_Std_Time_Month_Ordinal_august___closed__5 = _init_l_Std_Time_Month_Ordinal_august___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__5); +l_Std_Time_Month_Ordinal_august___closed__6 = _init_l_Std_Time_Month_Ordinal_august___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august___closed__6); +l_Std_Time_Month_Ordinal_august = _init_l_Std_Time_Month_Ordinal_august(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_august); +l_Std_Time_Month_Ordinal_september___closed__1 = _init_l_Std_Time_Month_Ordinal_september___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__1); +l_Std_Time_Month_Ordinal_september___closed__2 = _init_l_Std_Time_Month_Ordinal_september___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__2); +l_Std_Time_Month_Ordinal_september___closed__3 = _init_l_Std_Time_Month_Ordinal_september___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__3); +l_Std_Time_Month_Ordinal_september___closed__4 = _init_l_Std_Time_Month_Ordinal_september___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__4); +l_Std_Time_Month_Ordinal_september___closed__5 = _init_l_Std_Time_Month_Ordinal_september___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__5); +l_Std_Time_Month_Ordinal_september___closed__6 = _init_l_Std_Time_Month_Ordinal_september___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september___closed__6); +l_Std_Time_Month_Ordinal_september = _init_l_Std_Time_Month_Ordinal_september(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_september); +l_Std_Time_Month_Ordinal_october___closed__1 = _init_l_Std_Time_Month_Ordinal_october___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__1); +l_Std_Time_Month_Ordinal_october___closed__2 = _init_l_Std_Time_Month_Ordinal_october___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__2); +l_Std_Time_Month_Ordinal_october___closed__3 = _init_l_Std_Time_Month_Ordinal_october___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__3); +l_Std_Time_Month_Ordinal_october___closed__4 = _init_l_Std_Time_Month_Ordinal_october___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__4); +l_Std_Time_Month_Ordinal_october___closed__5 = _init_l_Std_Time_Month_Ordinal_october___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__5); +l_Std_Time_Month_Ordinal_october___closed__6 = _init_l_Std_Time_Month_Ordinal_october___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october___closed__6); +l_Std_Time_Month_Ordinal_october = _init_l_Std_Time_Month_Ordinal_october(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_october); +l_Std_Time_Month_Ordinal_november___closed__1 = _init_l_Std_Time_Month_Ordinal_november___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november___closed__1); +l_Std_Time_Month_Ordinal_november___closed__2 = _init_l_Std_Time_Month_Ordinal_november___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november___closed__2); +l_Std_Time_Month_Ordinal_november___closed__3 = _init_l_Std_Time_Month_Ordinal_november___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november___closed__3); +l_Std_Time_Month_Ordinal_november___closed__4 = _init_l_Std_Time_Month_Ordinal_november___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november___closed__4); +l_Std_Time_Month_Ordinal_november___closed__5 = _init_l_Std_Time_Month_Ordinal_november___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november___closed__5); +l_Std_Time_Month_Ordinal_november = _init_l_Std_Time_Month_Ordinal_november(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_november); +l_Std_Time_Month_Ordinal_december___closed__1 = _init_l_Std_Time_Month_Ordinal_december___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__1); +l_Std_Time_Month_Ordinal_december___closed__2 = _init_l_Std_Time_Month_Ordinal_december___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__2); +l_Std_Time_Month_Ordinal_december___closed__3 = _init_l_Std_Time_Month_Ordinal_december___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__3); +l_Std_Time_Month_Ordinal_december___closed__4 = _init_l_Std_Time_Month_Ordinal_december___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__4); +l_Std_Time_Month_Ordinal_december___closed__5 = _init_l_Std_Time_Month_Ordinal_december___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__5); +l_Std_Time_Month_Ordinal_december___closed__6 = _init_l_Std_Time_Month_Ordinal_december___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december___closed__6); +l_Std_Time_Month_Ordinal_december = _init_l_Std_Time_Month_Ordinal_december(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_december); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__1); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__2); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__3); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__4); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__5); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__6); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__7); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__8); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__9); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__10); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__11); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__12); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__13); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__14); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__15); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__16); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__17); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__18); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__19); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__20); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__21); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__22); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__23); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__24); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__25); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__26); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226____closed__27); +l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_ = _init_l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Month___hyg_226_); +l_Std_Time_Month_Ordinal_toNat___closed__1 = _init_l_Std_Time_Month_Ordinal_toNat___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toNat___closed__1); +l_Std_Time_Month_Ordinal_toSeconds___closed__1 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__1); +l_Std_Time_Month_Ordinal_toSeconds___closed__2 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__2); +l_Std_Time_Month_Ordinal_toSeconds___closed__3 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__3); +l_Std_Time_Month_Ordinal_toSeconds___closed__4 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__4); +l_Std_Time_Month_Ordinal_toSeconds___closed__5 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__5); +l_Std_Time_Month_Ordinal_toSeconds___closed__6 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__6); +l_Std_Time_Month_Ordinal_toSeconds___closed__7 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__7(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__7); +l_Std_Time_Month_Ordinal_toSeconds___closed__8 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__8(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__8); +l_Std_Time_Month_Ordinal_toSeconds___closed__9 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__9(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__9); +l_Std_Time_Month_Ordinal_toSeconds___closed__10 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__10(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__10); +l_Std_Time_Month_Ordinal_toSeconds___closed__11 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__11(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__11); +l_Std_Time_Month_Ordinal_toSeconds___closed__12 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__12(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__12); +l_Std_Time_Month_Ordinal_toSeconds___closed__13 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__13(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__13); +l_Std_Time_Month_Ordinal_toSeconds___closed__14 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__14(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__14); +l_Std_Time_Month_Ordinal_toSeconds___closed__15 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__15(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__15); +l_Std_Time_Month_Ordinal_toSeconds___closed__16 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__16(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__16); +l_Std_Time_Month_Ordinal_toSeconds___closed__17 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__17(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__17); +l_Std_Time_Month_Ordinal_toSeconds___closed__18 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__18(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__18); +l_Std_Time_Month_Ordinal_toSeconds___closed__19 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__19(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__19); +l_Std_Time_Month_Ordinal_toSeconds___closed__20 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__20(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__20); +l_Std_Time_Month_Ordinal_toSeconds___closed__21 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__21(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__21); +l_Std_Time_Month_Ordinal_toSeconds___closed__22 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__22(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__22); +l_Std_Time_Month_Ordinal_toSeconds___closed__23 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__23(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__23); +l_Std_Time_Month_Ordinal_toSeconds___closed__24 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__24(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__24); +l_Std_Time_Month_Ordinal_toSeconds___closed__25 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__25(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__25); +l_Std_Time_Month_Ordinal_toSeconds___closed__26 = _init_l_Std_Time_Month_Ordinal_toSeconds___closed__26(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toSeconds___closed__26); +l_Std_Time_Month_Ordinal_toMinutes___closed__1 = _init_l_Std_Time_Month_Ordinal_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toMinutes___closed__1); +l_Std_Time_Month_Ordinal_toDays___closed__1 = _init_l_Std_Time_Month_Ordinal_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toDays___closed__1); +l_Std_Time_Month_Ordinal_toDays___closed__2 = _init_l_Std_Time_Month_Ordinal_toDays___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toDays___closed__2); +l_Std_Time_Month_Ordinal_toDays___closed__3 = _init_l_Std_Time_Month_Ordinal_toDays___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toDays___closed__3); +l_Std_Time_Month_Ordinal_toDays___closed__4 = _init_l_Std_Time_Month_Ordinal_toDays___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_toDays___closed__4); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__1); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__2); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__3); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__4); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__5); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__6); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__7); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__8); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__9); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__10); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__11); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__12); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__13); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__14); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__15); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__16); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__17); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__18); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__19); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__20); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__21); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__22); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__23); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__24); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__25); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__26); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__27); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__28); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__29); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__30); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__31); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__32); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33 = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap___closed__33); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_monthSizesNonLeap); +l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes = _init_l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Month_0__Std_Time_Month_Ordinal_cumulativeSizes); +l_Std_Time_Month_Ordinal_days___closed__1 = _init_l_Std_Time_Month_Ordinal_days___closed__1(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__1); +l_Std_Time_Month_Ordinal_days___closed__2 = _init_l_Std_Time_Month_Ordinal_days___closed__2(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__2); +l_Std_Time_Month_Ordinal_days___closed__3 = _init_l_Std_Time_Month_Ordinal_days___closed__3(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__3); +l_Std_Time_Month_Ordinal_days___closed__4 = _init_l_Std_Time_Month_Ordinal_days___closed__4(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__4); +l_Std_Time_Month_Ordinal_days___closed__5 = _init_l_Std_Time_Month_Ordinal_days___closed__5(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__5); +l_Std_Time_Month_Ordinal_days___closed__6 = _init_l_Std_Time_Month_Ordinal_days___closed__6(); +lean_mark_persistent(l_Std_Time_Month_Ordinal_days___closed__6); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Week.c b/stage0/stdlib/Std/Time/Date/Unit/Week.c new file mode 100644 index 000000000000..d05fb893529b --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Week.c @@ -0,0 +1,1294 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Week +// Imports: Std.Internal.Rat Std.Time.Date.Unit.Day +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Week_Offset_toHours___closed__1; +static lean_object* l_Std_Time_Week_instOffsetNeg___closed__1; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalLE; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLtOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetLT; +static lean_object* l_Std_Time_Week_Offset_toMinutes___closed__1; +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt___boxed(lean_object*); +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26; +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__4; +static lean_object* l_Std_Time_Week_Offset_toDays___closed__1; +lean_object* lean_int_emod(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds(lean_object*); +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours(lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Week_instOffsetInhabited___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_instInhabitedOrdinal; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13; +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetBEq; +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalLT; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetToString; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_Week_instOfNatOrdinal___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4; +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalBEq; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15; +static lean_object* l_Std_Time_Week_instOffsetBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset___boxed(lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Std_Time_Week_Offset_toNanoseconds___closed__1; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24; +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetRepr; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes___boxed(lean_object*); +static lean_object* l_Std_Time_Week_instOffsetRepr___closed__1; +static lean_object* l_Std_Time_Week_instOrdinalBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds___boxed(lean_object*); +static lean_object* l_Std_Time_Week_instOrdinalRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_instOfMonthRepr; +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOrdinalRepr; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetLE; +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetSub; +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__2; +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofFin(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22; +static lean_object* l_Std_Time_Week_Offset_toSeconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds___boxed(lean_object*); +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__5; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetAdd; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLeOrdinal(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOrdinal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Week_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_; +LEAN_EXPORT lean_object* l_Std_Time_Week_instOffsetNeg; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt(lean_object*); +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Week_Offset_toMilliseconds___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18; +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27; +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +static lean_object* l_Std_Time_Week_instInhabitedOrdinal___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds___boxed(lean_object*); +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5; +static lean_object* l_Std_Time_Week_instOffsetAdd___closed__1; +static lean_object* l_Std_Time_Week_instOffsetSub___closed__1; +static lean_object* l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19; +static lean_object* _init_l_Std_Time_Week_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(52u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Week_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Week_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Week_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(52u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Week_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__5; +x_2 = l_Std_Time_Week_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Week_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Week_instInhabitedOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instInhabitedOrdinal___closed__9; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetInhabited___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetInhabited___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Week_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Week_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Week_Ordinal_instOfMonthRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Week_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("decide", 6, 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11; +x_3 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3; +x_4 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6; +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(2); +x_2 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5; +x_3 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26; +x_4 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_le(x_2, x_1); +if (x_3 == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = l_Std_Time_Week_instOfNatOrdinal___closed__1; +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_nat_to_int(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(604800000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toMilliseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toMilliseconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toNanoseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("604800000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toNanoseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toNanoseconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(604800u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toSeconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10080u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toMinutes___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(168u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toHours___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Week_Offset_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toDays___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Week_Offset_toDays___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Week_Offset_ofDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Week_Offset_ofDays(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Week(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Week_instOrdinalRepr___closed__1 = _init_l_Std_Time_Week_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalRepr___closed__1); +l_Std_Time_Week_instOrdinalRepr = _init_l_Std_Time_Week_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalRepr); +l_Std_Time_Week_instOrdinalBEq___closed__1 = _init_l_Std_Time_Week_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalBEq___closed__1); +l_Std_Time_Week_instOrdinalBEq = _init_l_Std_Time_Week_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalBEq); +l_Std_Time_Week_instOrdinalLE = _init_l_Std_Time_Week_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalLE); +l_Std_Time_Week_instOrdinalLT = _init_l_Std_Time_Week_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Week_instOrdinalLT); +l_Std_Time_Week_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Week_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOfNatOrdinal___closed__1); +l_Std_Time_Week_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__1); +l_Std_Time_Week_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__2); +l_Std_Time_Week_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__3); +l_Std_Time_Week_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__4); +l_Std_Time_Week_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__5); +l_Std_Time_Week_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__6); +l_Std_Time_Week_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__7); +l_Std_Time_Week_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__8); +l_Std_Time_Week_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Week_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal___closed__9); +l_Std_Time_Week_instInhabitedOrdinal = _init_l_Std_Time_Week_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Week_instInhabitedOrdinal); +l_Std_Time_Week_instOffsetRepr___closed__1 = _init_l_Std_Time_Week_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetRepr___closed__1); +l_Std_Time_Week_instOffsetRepr = _init_l_Std_Time_Week_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Week_instOffsetRepr); +l_Std_Time_Week_instOffsetBEq___closed__1 = _init_l_Std_Time_Week_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetBEq___closed__1); +l_Std_Time_Week_instOffsetBEq = _init_l_Std_Time_Week_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Week_instOffsetBEq); +l_Std_Time_Week_instOffsetInhabited___closed__1 = _init_l_Std_Time_Week_instOffsetInhabited___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetInhabited___closed__1); +l_Std_Time_Week_instOffsetInhabited = _init_l_Std_Time_Week_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Week_instOffsetInhabited); +l_Std_Time_Week_instOffsetAdd___closed__1 = _init_l_Std_Time_Week_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetAdd___closed__1); +l_Std_Time_Week_instOffsetAdd = _init_l_Std_Time_Week_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Week_instOffsetAdd); +l_Std_Time_Week_instOffsetSub___closed__1 = _init_l_Std_Time_Week_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetSub___closed__1); +l_Std_Time_Week_instOffsetSub = _init_l_Std_Time_Week_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Week_instOffsetSub); +l_Std_Time_Week_instOffsetNeg___closed__1 = _init_l_Std_Time_Week_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetNeg___closed__1); +l_Std_Time_Week_instOffsetNeg = _init_l_Std_Time_Week_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Week_instOffsetNeg); +l_Std_Time_Week_instOffsetLE = _init_l_Std_Time_Week_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Week_instOffsetLE); +l_Std_Time_Week_instOffsetLT = _init_l_Std_Time_Week_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Week_instOffsetLT); +l_Std_Time_Week_instOffsetToString___closed__1 = _init_l_Std_Time_Week_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Week_instOffsetToString___closed__1); +l_Std_Time_Week_instOffsetToString = _init_l_Std_Time_Week_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Week_instOffsetToString); +l_Std_Time_Week_Ordinal_instOfMonthRepr = _init_l_Std_Time_Week_Ordinal_instOfMonthRepr(); +lean_mark_persistent(l_Std_Time_Week_Ordinal_instOfMonthRepr); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__1); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__2); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__3); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__4); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__5); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__6); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__7); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__8); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__9); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__10); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__11); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__12); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__13); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__14); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__15); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__16); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__17); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__18); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__19); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__20); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__21); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__22); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__23); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__24); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__25); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__26); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27 = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149____closed__27); +l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_ = _init_l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_(); +lean_mark_persistent(l___auto____x40_Std_Time_Date_Unit_Week___hyg_149_); +l_Std_Time_Week_Offset_toMilliseconds___closed__1 = _init_l_Std_Time_Week_Offset_toMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toMilliseconds___closed__1); +l_Std_Time_Week_Offset_toNanoseconds___closed__1 = _init_l_Std_Time_Week_Offset_toNanoseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toNanoseconds___closed__1); +l_Std_Time_Week_Offset_toSeconds___closed__1 = _init_l_Std_Time_Week_Offset_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toSeconds___closed__1); +l_Std_Time_Week_Offset_toMinutes___closed__1 = _init_l_Std_Time_Week_Offset_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toMinutes___closed__1); +l_Std_Time_Week_Offset_toHours___closed__1 = _init_l_Std_Time_Week_Offset_toHours___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toHours___closed__1); +l_Std_Time_Week_Offset_toDays___closed__1 = _init_l_Std_Time_Week_Offset_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Week_Offset_toDays___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Weekday.c b/stage0/stdlib/Std/Time/Date/Unit/Weekday.c new file mode 100644 index 000000000000..4a8b01a340e3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Weekday.c @@ -0,0 +1,2221 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Weekday +// Imports: Std.Internal.Rat Std.Time.Date.Unit.Day +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__13; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__20; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__17; +LEAN_EXPORT uint8_t l_Std_Time_Weekday_ofOrdinal(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27; +static lean_object* l_Std_Time_Weekday_ofNat_x21___closed__3; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37; +LEAN_EXPORT uint8_t l_Std_Time_Weekday_next(uint8_t); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__16; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__22; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x3f(lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__27; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__10; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__30; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__19; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +LEAN_EXPORT lean_object* l_Std_Time_instBEqWeekday; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__23; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__3; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__4; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__2; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__43; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_next___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__36; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___lambda__1(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__41; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__25; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14; +static lean_object* l_Std_Time_instBEqWeekday___closed__1; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__3; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__37; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__21; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__2; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__33; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Weekday_isWeekend___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__24; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toOrdinal(uint8_t); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_ofNat_x21___closed__4; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x21(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__31; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toNat___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__34; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__7; +static lean_object* l_Std_Time_Weekday_ofOrdinal___closed__1; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__8; +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Std_Time_Weekday_ofNat_x21___spec__1(lean_object*); +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofOrdinal___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Weekday_isWeekend(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_instReprWeekday; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* lean_nat_abs(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__6; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__4; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__32; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36; +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toNat(uint8_t); +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l_Std_Time_Weekday_ofNat_x21___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +LEAN_EXPORT uint8_t l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toOrdinal___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__29; +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedWeekday; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__11; +static lean_object* l_Std_Time_instReprWeekday___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22; +static lean_object* l_Std_Time_Weekday_ofNat_x3f___closed__6; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__26; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__39; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__14; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toCtorIdx(uint8_t); +static lean_object* l_Std_Time_Weekday_noConfusion___rarg___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__35; +static lean_object* l_Std_Time_Weekday_ofNat_x21___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_instOfNatOrdinal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x3f___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__40; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__5; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__12; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__38; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x21___boxed(lean_object*); +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__28; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__7; +static lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__18; +static lean_object* l_Std_Time_Weekday_toOrdinal___closed__42; +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; +} +case 4: +{ +lean_object* x_6; +x_6 = lean_unsigned_to_nat(4u); +return x_6; +} +case 5: +{ +lean_object* x_7; +x_7 = lean_unsigned_to_nat(5u); +return x_7; +} +default: +{ +lean_object* x_8; +x_8 = lean_unsigned_to_nat(6u); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Weekday_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Weekday_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Weekday_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Weekday_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Weekday_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Weekday_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_Weekday_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.monday", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.tuesday", 24, 24); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.wednesday", 26, 26); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.thursday", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.friday", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.saturday", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.sunday", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12_(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +case 2: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1024u); +x_16 = lean_nat_dec_le(x_15, x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18; +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20; +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +} +case 3: +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(1024u); +x_22 = lean_nat_dec_le(x_21, x_2); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24; +x_24 = l_Repr_addAppParen(x_23, x_2); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26; +x_26 = l_Repr_addAppParen(x_25, x_2); +return x_26; +} +} +case 4: +{ +lean_object* x_27; uint8_t x_28; +x_27 = lean_unsigned_to_nat(1024u); +x_28 = lean_nat_dec_le(x_27, x_2); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30; +x_30 = l_Repr_addAppParen(x_29, x_2); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32; +x_32 = l_Repr_addAppParen(x_31, x_2); +return x_32; +} +} +case 5: +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_unsigned_to_nat(1024u); +x_34 = lean_nat_dec_le(x_33, x_2); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; +x_35 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36; +x_36 = l_Repr_addAppParen(x_35, x_2); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38; +x_38 = l_Repr_addAppParen(x_37, x_2); +return x_38; +} +} +default: +{ +lean_object* x_39; uint8_t x_40; +x_39 = lean_unsigned_to_nat(1024u); +x_40 = lean_nat_dec_le(x_39, x_2); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42; +x_42 = l_Repr_addAppParen(x_41, x_2); +return x_42; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44; +x_44 = l_Repr_addAppParen(x_43, x_2); +return x_44; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprWeekday___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprWeekday() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprWeekday___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedWeekday() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265_(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Std_Time_Weekday_toCtorIdx(x_1); +x_4 = l_Std_Time_Weekday_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_instBEqWeekday___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_beqWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_265____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instBEqWeekday() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instBEqWeekday___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_unsigned_to_nat(6u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Weekday_ofOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_nat_sub(x_9, x_4); +lean_dec(x_9); +x_12 = lean_nat_dec_eq(x_11, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_nat_sub(x_11, x_4); +lean_dec(x_11); +x_14 = lean_nat_dec_eq(x_13, x_3); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_nat_sub(x_13, x_4); +lean_dec(x_13); +x_16 = lean_nat_dec_eq(x_15, x_3); +lean_dec(x_15); +if (x_16 == 0) +{ +uint8_t x_17; +x_17 = 6; +return x_17; +} +else +{ +uint8_t x_18; +x_18 = 5; +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_13); +x_19 = 4; +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_11); +x_20 = 3; +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_9); +x_21 = 2; +return x_21; +} +} +else +{ +uint8_t x_22; +lean_dec(x_7); +x_22 = 1; +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_5); +x_23 = 0; +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_Weekday_ofOrdinal(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__2; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__5; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__6; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__7; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__8; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__10; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__11; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__12; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__13; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__15; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__16; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__17; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__18; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__19; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__21; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__22; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__23; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__24; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__25; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__27; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__28; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__29; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__30; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__31; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__1; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__33; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__34; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__35; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__36; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__38; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__39; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__40; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__42() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__41; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_toOrdinal___closed__43() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Weekday_toOrdinal___closed__42; +x_2 = l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toOrdinal(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l_Std_Time_Weekday_toOrdinal___closed__9; +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = l_Std_Time_Weekday_toOrdinal___closed__14; +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = l_Std_Time_Weekday_toOrdinal___closed__20; +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = l_Std_Time_Weekday_toOrdinal___closed__26; +return x_5; +} +case 4: +{ +lean_object* x_6; +x_6 = l_Std_Time_Weekday_toOrdinal___closed__32; +return x_6; +} +case 5: +{ +lean_object* x_7; +x_7 = l_Std_Time_Weekday_toOrdinal___closed__37; +return x_7; +} +default: +{ +lean_object* x_8; +x_8 = l_Std_Time_Weekday_toOrdinal___closed__43; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Weekday_toOrdinal(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toNat(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(1u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(2u); +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(3u); +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(4u); +return x_5; +} +case 4: +{ +lean_object* x_6; +x_6 = lean_unsigned_to_nat(5u); +return x_6; +} +case 5: +{ +lean_object* x_7; +x_7 = lean_unsigned_to_nat(6u); +return x_7; +} +default: +{ +lean_object* x_8; +x_8 = lean_unsigned_to_nat(7u); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_toNat___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Weekday_toNat(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 6; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 5; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 4; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__4() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 3; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 2; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__6() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x3f___closed__7() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(3u); +x_7 = lean_nat_dec_eq(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(4u); +x_9 = lean_nat_dec_eq(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(5u); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_unsigned_to_nat(6u); +x_13 = lean_nat_dec_eq(x_1, x_12); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(7u); +x_15 = lean_nat_dec_eq(x_1, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_box(0); +return x_16; +} +else +{ +lean_object* x_17; +x_17 = l_Std_Time_Weekday_ofNat_x3f___closed__1; +return x_17; +} +} +else +{ +lean_object* x_18; +x_18 = l_Std_Time_Weekday_ofNat_x3f___closed__2; +return x_18; +} +} +else +{ +lean_object* x_19; +x_19 = l_Std_Time_Weekday_ofNat_x3f___closed__3; +return x_19; +} +} +else +{ +lean_object* x_20; +x_20 = l_Std_Time_Weekday_ofNat_x3f___closed__4; +return x_20; +} +} +else +{ +lean_object* x_21; +x_21 = l_Std_Time_Weekday_ofNat_x3f___closed__5; +return x_21; +} +} +else +{ +lean_object* x_22; +x_22 = l_Std_Time_Weekday_ofNat_x3f___closed__6; +return x_22; +} +} +else +{ +lean_object* x_23; +x_23 = l_Std_Time_Weekday_ofNat_x3f___closed__7; +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Weekday_ofNat_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_Weekday_ofNat_x21___spec__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_instInhabitedWeekday; +x_3 = lean_box(x_2); +x_4 = lean_panic_fn(x_3, x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x21___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Date.Unit.Weekday", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x21___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Weekday.ofNat!", 23, 23); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x21___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid weekday", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Weekday_ofNat_x21___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Std_Time_Weekday_ofNat_x21___closed__1; +x_2 = l_Std_Time_Weekday_ofNat_x21___closed__2; +x_3 = lean_unsigned_to_nat(110u); +x_4 = lean_unsigned_to_nat(12u); +x_5 = l_Std_Time_Weekday_ofNat_x21___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x21(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Weekday_ofNat_x3f(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_Weekday_ofNat_x21___closed__4; +x_4 = l_panic___at_Std_Time_Weekday_ofNat_x21___spec__1(x_3); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_ofNat_x21___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Weekday_ofNat_x21(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Weekday_next(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +case 1: +{ +uint8_t x_3; +x_3 = 2; +return x_3; +} +case 2: +{ +uint8_t x_4; +x_4 = 3; +return x_4; +} +case 3: +{ +uint8_t x_5; +x_5 = 4; +return x_5; +} +case 4: +{ +uint8_t x_6; +x_6 = 5; +return x_6; +} +case 5: +{ +uint8_t x_7; +x_7 = 6; +return x_7; +} +default: +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_next___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Weekday_next(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Weekday_isWeekend(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(x_1); +switch (lean_obj_tag(x_2)) { +case 5: +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +case 6: +{ +uint8_t x_4; +x_4 = 1; +return x_4; +} +default: +{ +uint8_t x_5; +lean_dec(x_2); +x_5 = 0; +return x_5; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Weekday_isWeekend___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Weekday_isWeekend(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Weekday(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Weekday_noConfusion___rarg___closed__1 = _init_l_Std_Time_Weekday_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Weekday_noConfusion___rarg___closed__1); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__1); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__2); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__3); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__4); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__5); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__6); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__7); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__8); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__9); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__10); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__11); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__12); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__13); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__14); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__15); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__16); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__17); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__18); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__19); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__20); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__21); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__22); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__23); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__24); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__25); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__26); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__27); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__28); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__29); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__30); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__31); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__32); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__33); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__34); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__35); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__36); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__37); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__38); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__39); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__40); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__41); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__42); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__43); +l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44 = _init_l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Weekday_0__Std_Time_reprWeekday____x40_Std_Time_Date_Unit_Weekday___hyg_12____closed__44); +l_Std_Time_instReprWeekday___closed__1 = _init_l_Std_Time_instReprWeekday___closed__1(); +lean_mark_persistent(l_Std_Time_instReprWeekday___closed__1); +l_Std_Time_instReprWeekday = _init_l_Std_Time_instReprWeekday(); +lean_mark_persistent(l_Std_Time_instReprWeekday); +l_Std_Time_instInhabitedWeekday = _init_l_Std_Time_instInhabitedWeekday(); +l_Std_Time_instBEqWeekday___closed__1 = _init_l_Std_Time_instBEqWeekday___closed__1(); +lean_mark_persistent(l_Std_Time_instBEqWeekday___closed__1); +l_Std_Time_instBEqWeekday = _init_l_Std_Time_instBEqWeekday(); +lean_mark_persistent(l_Std_Time_instBEqWeekday); +l_Std_Time_Weekday_ofOrdinal___closed__1 = _init_l_Std_Time_Weekday_ofOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Weekday_ofOrdinal___closed__1); +l_Std_Time_Weekday_toOrdinal___closed__1 = _init_l_Std_Time_Weekday_toOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__1); +l_Std_Time_Weekday_toOrdinal___closed__2 = _init_l_Std_Time_Weekday_toOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__2); +l_Std_Time_Weekday_toOrdinal___closed__3 = _init_l_Std_Time_Weekday_toOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__3); +l_Std_Time_Weekday_toOrdinal___closed__4 = _init_l_Std_Time_Weekday_toOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__4); +l_Std_Time_Weekday_toOrdinal___closed__5 = _init_l_Std_Time_Weekday_toOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__5); +l_Std_Time_Weekday_toOrdinal___closed__6 = _init_l_Std_Time_Weekday_toOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__6); +l_Std_Time_Weekday_toOrdinal___closed__7 = _init_l_Std_Time_Weekday_toOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__7); +l_Std_Time_Weekday_toOrdinal___closed__8 = _init_l_Std_Time_Weekday_toOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__8); +l_Std_Time_Weekday_toOrdinal___closed__9 = _init_l_Std_Time_Weekday_toOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__9); +l_Std_Time_Weekday_toOrdinal___closed__10 = _init_l_Std_Time_Weekday_toOrdinal___closed__10(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__10); +l_Std_Time_Weekday_toOrdinal___closed__11 = _init_l_Std_Time_Weekday_toOrdinal___closed__11(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__11); +l_Std_Time_Weekday_toOrdinal___closed__12 = _init_l_Std_Time_Weekday_toOrdinal___closed__12(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__12); +l_Std_Time_Weekday_toOrdinal___closed__13 = _init_l_Std_Time_Weekday_toOrdinal___closed__13(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__13); +l_Std_Time_Weekday_toOrdinal___closed__14 = _init_l_Std_Time_Weekday_toOrdinal___closed__14(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__14); +l_Std_Time_Weekday_toOrdinal___closed__15 = _init_l_Std_Time_Weekday_toOrdinal___closed__15(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__15); +l_Std_Time_Weekday_toOrdinal___closed__16 = _init_l_Std_Time_Weekday_toOrdinal___closed__16(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__16); +l_Std_Time_Weekday_toOrdinal___closed__17 = _init_l_Std_Time_Weekday_toOrdinal___closed__17(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__17); +l_Std_Time_Weekday_toOrdinal___closed__18 = _init_l_Std_Time_Weekday_toOrdinal___closed__18(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__18); +l_Std_Time_Weekday_toOrdinal___closed__19 = _init_l_Std_Time_Weekday_toOrdinal___closed__19(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__19); +l_Std_Time_Weekday_toOrdinal___closed__20 = _init_l_Std_Time_Weekday_toOrdinal___closed__20(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__20); +l_Std_Time_Weekday_toOrdinal___closed__21 = _init_l_Std_Time_Weekday_toOrdinal___closed__21(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__21); +l_Std_Time_Weekday_toOrdinal___closed__22 = _init_l_Std_Time_Weekday_toOrdinal___closed__22(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__22); +l_Std_Time_Weekday_toOrdinal___closed__23 = _init_l_Std_Time_Weekday_toOrdinal___closed__23(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__23); +l_Std_Time_Weekday_toOrdinal___closed__24 = _init_l_Std_Time_Weekday_toOrdinal___closed__24(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__24); +l_Std_Time_Weekday_toOrdinal___closed__25 = _init_l_Std_Time_Weekday_toOrdinal___closed__25(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__25); +l_Std_Time_Weekday_toOrdinal___closed__26 = _init_l_Std_Time_Weekday_toOrdinal___closed__26(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__26); +l_Std_Time_Weekday_toOrdinal___closed__27 = _init_l_Std_Time_Weekday_toOrdinal___closed__27(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__27); +l_Std_Time_Weekday_toOrdinal___closed__28 = _init_l_Std_Time_Weekday_toOrdinal___closed__28(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__28); +l_Std_Time_Weekday_toOrdinal___closed__29 = _init_l_Std_Time_Weekday_toOrdinal___closed__29(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__29); +l_Std_Time_Weekday_toOrdinal___closed__30 = _init_l_Std_Time_Weekday_toOrdinal___closed__30(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__30); +l_Std_Time_Weekday_toOrdinal___closed__31 = _init_l_Std_Time_Weekday_toOrdinal___closed__31(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__31); +l_Std_Time_Weekday_toOrdinal___closed__32 = _init_l_Std_Time_Weekday_toOrdinal___closed__32(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__32); +l_Std_Time_Weekday_toOrdinal___closed__33 = _init_l_Std_Time_Weekday_toOrdinal___closed__33(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__33); +l_Std_Time_Weekday_toOrdinal___closed__34 = _init_l_Std_Time_Weekday_toOrdinal___closed__34(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__34); +l_Std_Time_Weekday_toOrdinal___closed__35 = _init_l_Std_Time_Weekday_toOrdinal___closed__35(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__35); +l_Std_Time_Weekday_toOrdinal___closed__36 = _init_l_Std_Time_Weekday_toOrdinal___closed__36(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__36); +l_Std_Time_Weekday_toOrdinal___closed__37 = _init_l_Std_Time_Weekday_toOrdinal___closed__37(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__37); +l_Std_Time_Weekday_toOrdinal___closed__38 = _init_l_Std_Time_Weekday_toOrdinal___closed__38(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__38); +l_Std_Time_Weekday_toOrdinal___closed__39 = _init_l_Std_Time_Weekday_toOrdinal___closed__39(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__39); +l_Std_Time_Weekday_toOrdinal___closed__40 = _init_l_Std_Time_Weekday_toOrdinal___closed__40(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__40); +l_Std_Time_Weekday_toOrdinal___closed__41 = _init_l_Std_Time_Weekday_toOrdinal___closed__41(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__41); +l_Std_Time_Weekday_toOrdinal___closed__42 = _init_l_Std_Time_Weekday_toOrdinal___closed__42(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__42); +l_Std_Time_Weekday_toOrdinal___closed__43 = _init_l_Std_Time_Weekday_toOrdinal___closed__43(); +lean_mark_persistent(l_Std_Time_Weekday_toOrdinal___closed__43); +l_Std_Time_Weekday_ofNat_x3f___closed__1 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__1(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__1); +l_Std_Time_Weekday_ofNat_x3f___closed__2 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__2(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__2); +l_Std_Time_Weekday_ofNat_x3f___closed__3 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__3(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__3); +l_Std_Time_Weekday_ofNat_x3f___closed__4 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__4(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__4); +l_Std_Time_Weekday_ofNat_x3f___closed__5 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__5(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__5); +l_Std_Time_Weekday_ofNat_x3f___closed__6 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__6(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__6); +l_Std_Time_Weekday_ofNat_x3f___closed__7 = _init_l_Std_Time_Weekday_ofNat_x3f___closed__7(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x3f___closed__7); +l_Std_Time_Weekday_ofNat_x21___closed__1 = _init_l_Std_Time_Weekday_ofNat_x21___closed__1(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x21___closed__1); +l_Std_Time_Weekday_ofNat_x21___closed__2 = _init_l_Std_Time_Weekday_ofNat_x21___closed__2(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x21___closed__2); +l_Std_Time_Weekday_ofNat_x21___closed__3 = _init_l_Std_Time_Weekday_ofNat_x21___closed__3(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x21___closed__3); +l_Std_Time_Weekday_ofNat_x21___closed__4 = _init_l_Std_Time_Weekday_ofNat_x21___closed__4(); +lean_mark_persistent(l_Std_Time_Weekday_ofNat_x21___closed__4); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/Unit/Year.c b/stage0/stdlib/Std/Time/Date/Unit/Year.c new file mode 100644 index 000000000000..cf3a887366dc --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/Unit/Year.c @@ -0,0 +1,1273 @@ +// Lean compiler output +// Module: Std.Time.Date.Unit.Year +// Imports: Std.Time.Internal Std.Internal.Rat Std.Time.Date.Unit.Day Std.Time.Date.Unit.Month +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* lean_int_mod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_instOffsetBEq___closed__1; +lean_object* l_instBEqOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_Offset_weeks___closed__5; +extern lean_object* l_Int_instInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion(lean_object*); +static lean_object* l_Std_Time_Year_Era_noConfusion___rarg___closed__1; +static lean_object* l_Std_Time_Year_Offset_days___closed__11; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetLT; +static lean_object* l_Std_Time_Year_Offset_days___closed__7; +static lean_object* l_Std_Time_Year_Offset_weeks___closed__4; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_era___boxed(lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_Offset_weeks___closed__1; +static lean_object* l_Std_Time_Year_Offset_days___closed__13; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_toCtorIdx___boxed(lean_object*); +static lean_object* l_Std_Time_Year_instToStringEra___closed__2; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13_(uint8_t, lean_object*); +static lean_object* l_Std_Time_Year_Offset_weeks___closed__2; +static lean_object* l_Std_Time_Year_Offset_days___closed__6; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_Year_instToStringEra___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetSub; +static lean_object* l_Std_Time_Year_Offset_days___closed__9; +LEAN_EXPORT uint8_t l_Std_Time_Year_Offset_isLeap(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_isLeap___boxed(lean_object*); +static lean_object* l_Std_Time_Year_Offset_days___closed__3; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +static lean_object* l_Std_Time_Year_Offset_days___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___lambda__1___boxed(lean_object*); +uint8_t l_instDecidableNot___rarg(uint8_t); +extern lean_object* l_Int_instNegInt; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Std_Time_Year_Offset_days___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetLE; +static lean_object* l_Std_Time_Year_instReprEra___closed__1; +static lean_object* l_Std_Time_Year_Offset_days___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instDecidableLeOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Year_instDecidableLeOffset(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_Offset_days___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetBEq; +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____boxed(lean_object*, lean_object*); +lean_object* l_instReprInt___boxed(lean_object*, lean_object*); +extern lean_object* l_Int_instLTInt; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instReprEra; +static lean_object* l_Std_Time_Year_Offset_weeks___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_weeks___boxed(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetNeg; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofInt___boxed(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_Offset_isLeap___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_weeks(lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2; +static lean_object* l_Std_Time_Year_Offset_days___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toInt___boxed(lean_object*); +static lean_object* l_Std_Time_Year_Offset_days___closed__14; +static lean_object* l_Std_Time_Year_Offset_isLeap___closed__3; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_instOffsetToString___closed__1; +lean_object* l_Int_instDecidableEq___boxed(lean_object*, lean_object*); +lean_object* l_instToStringInt___boxed(lean_object*); +static lean_object* l_Std_Time_Year_instOffsetRepr___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_Year_instInhabitedEra; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13; +extern lean_object* l_Int_instAdd; +static lean_object* l_Std_Time_Year_Offset_toMonths___closed__1; +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5; +LEAN_EXPORT uint8_t l_Std_Time_Year_Offset_era(lean_object*); +extern lean_object* l_Int_instSub; +static lean_object* l_Std_Time_Year_Offset_isLeap___closed__1; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_Offset_days___closed__1; +static lean_object* l_Std_Time_Year_Offset_days___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetToString; +LEAN_EXPORT uint8_t l_Std_Time_Year_instDecidableLtOffset(lean_object*, lean_object*); +static lean_object* l_Std_Time_Year_instOffsetBEq___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toMonths___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instOfNatOffset(lean_object*); +static lean_object* l_Std_Time_Year_Offset_isLeap___closed__2; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetAdd; +LEAN_EXPORT lean_object* l_Std_Time_Year_instDecidableLtOffset___boxed(lean_object*, lean_object*); +extern lean_object* l_Int_instLEInt; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toMonths(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_instOffsetRepr; +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_days___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_days(lean_object*); +static lean_object* l_Std_Time_Year_instToStringEra___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Year_instToStringEra(uint8_t); +static lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Year_Era_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_Era_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Year_Era_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Year_Era_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Year_Era_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Era_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Era_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_Year_Era_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.Era.bce", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.Era.ce", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Year_instReprEra___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instReprEra() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Year_instReprEra___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_Year_instInhabitedEra() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instToStringEra___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("BCE", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instToStringEra___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("CE", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_instToStringEra(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_instToStringEra___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l_Std_Time_Year_instToStringEra___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_instToStringEra___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Year_instToStringEra(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instReprInt___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Year_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Int_instDecidableEq___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetBEq___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Year_instOffsetBEq___closed__1; +x_2 = lean_alloc_closure((void*)(l_instBEqOfDecidableEq___rarg), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Year_instOffsetBEq___closed__2; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instInhabited; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instAdd; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instSub; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instNegInt; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instLEInt; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = l_Int_instLTInt; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_instToStringInt___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Year_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Year_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Year_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Year_instDecidableLtOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_instDecidableLtOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Year_instDecidableLtOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Offset_toInt(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_toMonths___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toMonths(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Year_Offset_toMonths___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_toMonths___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Offset_toMonths(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_isLeap___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_isLeap___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_isLeap___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_isLeap___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Year_Offset_isLeap(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = l_Std_Time_Year_Offset_isLeap___closed__1; +x_3 = lean_int_mod(x_1, x_2); +x_4 = l_Std_Time_Year_Offset_isLeap___closed__2; +x_5 = lean_int_dec_eq(x_3, x_4); +lean_dec(x_3); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; +x_7 = l_Std_Time_Year_Offset_isLeap___closed__3; +x_8 = lean_int_mod(x_1, x_7); +x_9 = lean_int_dec_eq(x_8, x_4); +lean_dec(x_8); +x_10 = l_instDecidableNot___rarg(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Std_Time_Year_Offset_isLeap___closed__4; +x_12 = lean_int_mod(x_1, x_11); +x_13 = lean_int_dec_eq(x_12, x_4); +lean_dec(x_12); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = 1; +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_isLeap___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_Year_Offset_isLeap(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Year_Offset_era(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_3 = lean_int_dec_le(x_2, x_1); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +else +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_era___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_Year_Offset_era(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(365u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(366u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(355u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__2; +x_2 = l_Std_Time_Year_Offset_days___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__4; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__3; +x_2 = l_Std_Time_Year_Offset_days___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__6; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__7; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__8; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__9; +x_2 = l_Std_Time_Year_Offset_days___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__4; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__11; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__12; +x_2 = l_Std_Time_Year_Offset_days___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_days___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_days___closed__13; +x_2 = l_Std_Time_Year_Offset_days___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_days(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = l_Std_Time_Year_Offset_isLeap___closed__1; +x_3 = lean_int_mod(x_1, x_2); +x_4 = l_Std_Time_Year_Offset_isLeap___closed__2; +x_5 = lean_int_dec_eq(x_3, x_4); +lean_dec(x_3); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = l_Std_Time_Year_Offset_days___closed__10; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; +x_7 = l_Std_Time_Year_Offset_isLeap___closed__3; +x_8 = lean_int_mod(x_1, x_7); +x_9 = lean_int_dec_eq(x_8, x_4); +lean_dec(x_8); +x_10 = l_instDecidableNot___rarg(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Std_Time_Year_Offset_isLeap___closed__4; +x_12 = lean_int_mod(x_1, x_11); +x_13 = lean_int_dec_eq(x_12, x_4); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +x_14 = l_Std_Time_Year_Offset_days___closed__10; +return x_14; +} +else +{ +lean_object* x_15; +x_15 = l_Std_Time_Year_Offset_days___closed__14; +return x_15; +} +} +else +{ +lean_object* x_16; +x_16 = l_Std_Time_Year_Offset_days___closed__14; +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_days___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Offset_days(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_weeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_weeks___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(52u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_weeks___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_weeks___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_weeks___closed__2; +x_2 = l_Std_Time_Year_Offset_isLeap___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Year_Offset_weeks___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Year_Offset_weeks___closed__2; +x_2 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_weeks(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_2 = l_Std_Time_Year_Offset_isLeap___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +x_4 = lean_int_add(x_1, x_3); +lean_dec(x_3); +x_5 = l_Std_Time_Year_Offset_isLeap___closed__3; +x_6 = lean_int_ediv(x_1, x_5); +x_7 = lean_int_sub(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = l_Std_Time_Year_Offset_isLeap___closed__4; +x_9 = lean_int_ediv(x_1, x_8); +x_10 = lean_int_add(x_7, x_9); +lean_dec(x_9); +lean_dec(x_7); +x_11 = l_Std_Time_Year_Offset_weeks___closed__1; +x_12 = lean_int_emod(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_dec_eq(x_12, x_2); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_14 = l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6; +x_15 = lean_int_sub(x_1, x_14); +x_16 = lean_int_ediv(x_15, x_2); +x_17 = lean_int_add(x_15, x_16); +lean_dec(x_16); +x_18 = lean_int_ediv(x_15, x_5); +x_19 = lean_int_sub(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +x_20 = lean_int_ediv(x_15, x_8); +lean_dec(x_15); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +x_22 = lean_int_emod(x_21, x_11); +lean_dec(x_21); +x_23 = l_Std_Time_Year_Offset_weeks___closed__3; +x_24 = lean_int_dec_eq(x_22, x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = l_Std_Time_Year_Offset_weeks___closed__4; +return x_25; +} +else +{ +lean_object* x_26; +x_26 = l_Std_Time_Year_Offset_weeks___closed__5; +return x_26; +} +} +else +{ +lean_object* x_27; +x_27 = l_Std_Time_Year_Offset_weeks___closed__5; +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_Offset_weeks___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Year_Offset_weeks(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Month(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_Unit_Year(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Month(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Year_Era_noConfusion___rarg___closed__1 = _init_l_Std_Time_Year_Era_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Year_Era_noConfusion___rarg___closed__1); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__1); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__2); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__3); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__4); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__5); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__6); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__7); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__8); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__9); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__10); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__11); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__12); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__13); +l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14 = _init_l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14(); +lean_mark_persistent(l___private_Std_Time_Date_Unit_Year_0__Std_Time_Year_reprEra____x40_Std_Time_Date_Unit_Year___hyg_13____closed__14); +l_Std_Time_Year_instReprEra___closed__1 = _init_l_Std_Time_Year_instReprEra___closed__1(); +lean_mark_persistent(l_Std_Time_Year_instReprEra___closed__1); +l_Std_Time_Year_instReprEra = _init_l_Std_Time_Year_instReprEra(); +lean_mark_persistent(l_Std_Time_Year_instReprEra); +l_Std_Time_Year_instInhabitedEra = _init_l_Std_Time_Year_instInhabitedEra(); +l_Std_Time_Year_instToStringEra___closed__1 = _init_l_Std_Time_Year_instToStringEra___closed__1(); +lean_mark_persistent(l_Std_Time_Year_instToStringEra___closed__1); +l_Std_Time_Year_instToStringEra___closed__2 = _init_l_Std_Time_Year_instToStringEra___closed__2(); +lean_mark_persistent(l_Std_Time_Year_instToStringEra___closed__2); +l_Std_Time_Year_instOffsetRepr___closed__1 = _init_l_Std_Time_Year_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Year_instOffsetRepr___closed__1); +l_Std_Time_Year_instOffsetRepr = _init_l_Std_Time_Year_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Year_instOffsetRepr); +l_Std_Time_Year_instOffsetBEq___closed__1 = _init_l_Std_Time_Year_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Year_instOffsetBEq___closed__1); +l_Std_Time_Year_instOffsetBEq___closed__2 = _init_l_Std_Time_Year_instOffsetBEq___closed__2(); +lean_mark_persistent(l_Std_Time_Year_instOffsetBEq___closed__2); +l_Std_Time_Year_instOffsetBEq = _init_l_Std_Time_Year_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Year_instOffsetBEq); +l_Std_Time_Year_instOffsetInhabited = _init_l_Std_Time_Year_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Year_instOffsetInhabited); +l_Std_Time_Year_instOffsetAdd = _init_l_Std_Time_Year_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Year_instOffsetAdd); +l_Std_Time_Year_instOffsetSub = _init_l_Std_Time_Year_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Year_instOffsetSub); +l_Std_Time_Year_instOffsetNeg = _init_l_Std_Time_Year_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Year_instOffsetNeg); +l_Std_Time_Year_instOffsetLE = _init_l_Std_Time_Year_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Year_instOffsetLE); +l_Std_Time_Year_instOffsetLT = _init_l_Std_Time_Year_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Year_instOffsetLT); +l_Std_Time_Year_instOffsetToString___closed__1 = _init_l_Std_Time_Year_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Year_instOffsetToString___closed__1); +l_Std_Time_Year_instOffsetToString = _init_l_Std_Time_Year_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Year_instOffsetToString); +l_Std_Time_Year_Offset_toMonths___closed__1 = _init_l_Std_Time_Year_Offset_toMonths___closed__1(); +lean_mark_persistent(l_Std_Time_Year_Offset_toMonths___closed__1); +l_Std_Time_Year_Offset_isLeap___closed__1 = _init_l_Std_Time_Year_Offset_isLeap___closed__1(); +lean_mark_persistent(l_Std_Time_Year_Offset_isLeap___closed__1); +l_Std_Time_Year_Offset_isLeap___closed__2 = _init_l_Std_Time_Year_Offset_isLeap___closed__2(); +lean_mark_persistent(l_Std_Time_Year_Offset_isLeap___closed__2); +l_Std_Time_Year_Offset_isLeap___closed__3 = _init_l_Std_Time_Year_Offset_isLeap___closed__3(); +lean_mark_persistent(l_Std_Time_Year_Offset_isLeap___closed__3); +l_Std_Time_Year_Offset_isLeap___closed__4 = _init_l_Std_Time_Year_Offset_isLeap___closed__4(); +lean_mark_persistent(l_Std_Time_Year_Offset_isLeap___closed__4); +l_Std_Time_Year_Offset_days___closed__1 = _init_l_Std_Time_Year_Offset_days___closed__1(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__1); +l_Std_Time_Year_Offset_days___closed__2 = _init_l_Std_Time_Year_Offset_days___closed__2(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__2); +l_Std_Time_Year_Offset_days___closed__3 = _init_l_Std_Time_Year_Offset_days___closed__3(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__3); +l_Std_Time_Year_Offset_days___closed__4 = _init_l_Std_Time_Year_Offset_days___closed__4(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__4); +l_Std_Time_Year_Offset_days___closed__5 = _init_l_Std_Time_Year_Offset_days___closed__5(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__5); +l_Std_Time_Year_Offset_days___closed__6 = _init_l_Std_Time_Year_Offset_days___closed__6(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__6); +l_Std_Time_Year_Offset_days___closed__7 = _init_l_Std_Time_Year_Offset_days___closed__7(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__7); +l_Std_Time_Year_Offset_days___closed__8 = _init_l_Std_Time_Year_Offset_days___closed__8(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__8); +l_Std_Time_Year_Offset_days___closed__9 = _init_l_Std_Time_Year_Offset_days___closed__9(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__9); +l_Std_Time_Year_Offset_days___closed__10 = _init_l_Std_Time_Year_Offset_days___closed__10(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__10); +l_Std_Time_Year_Offset_days___closed__11 = _init_l_Std_Time_Year_Offset_days___closed__11(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__11); +l_Std_Time_Year_Offset_days___closed__12 = _init_l_Std_Time_Year_Offset_days___closed__12(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__12); +l_Std_Time_Year_Offset_days___closed__13 = _init_l_Std_Time_Year_Offset_days___closed__13(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__13); +l_Std_Time_Year_Offset_days___closed__14 = _init_l_Std_Time_Year_Offset_days___closed__14(); +lean_mark_persistent(l_Std_Time_Year_Offset_days___closed__14); +l_Std_Time_Year_Offset_weeks___closed__1 = _init_l_Std_Time_Year_Offset_weeks___closed__1(); +lean_mark_persistent(l_Std_Time_Year_Offset_weeks___closed__1); +l_Std_Time_Year_Offset_weeks___closed__2 = _init_l_Std_Time_Year_Offset_weeks___closed__2(); +lean_mark_persistent(l_Std_Time_Year_Offset_weeks___closed__2); +l_Std_Time_Year_Offset_weeks___closed__3 = _init_l_Std_Time_Year_Offset_weeks___closed__3(); +lean_mark_persistent(l_Std_Time_Year_Offset_weeks___closed__3); +l_Std_Time_Year_Offset_weeks___closed__4 = _init_l_Std_Time_Year_Offset_weeks___closed__4(); +lean_mark_persistent(l_Std_Time_Year_Offset_weeks___closed__4); +l_Std_Time_Year_Offset_weeks___closed__5 = _init_l_Std_Time_Year_Offset_weeks___closed__5(); +lean_mark_persistent(l_Std_Time_Year_Offset_weeks___closed__5); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Date/ValidDate.c b/stage0/stdlib/Std/Time/Date/ValidDate.c new file mode 100644 index 000000000000..59e2c935825f --- /dev/null +++ b/stage0/stdlib/Std/Time/Date/ValidDate.c @@ -0,0 +1,419 @@ +// Lean compiler output +// Module: Std.Time.Date.ValidDate +// Imports: Std.Internal.Rat Std.Time.Date.Unit.Day Std.Time.Date.Unit.Month +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_instInhabitedValidDate___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__18; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__8; +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__4; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__5; +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__9; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_dayOfYear___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__1; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__13; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__15; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedValidDate___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal_go(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__16; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__19; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__11; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__12; +lean_object* l_Std_Time_Month_Ordinal_cumulativeDays(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__14; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__7; +static lean_object* l_Std_Time_instInhabitedValidDate___closed__6; +lean_object* lean_int_add(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +static lean_object* l_Std_Time_instInhabitedValidDate___closed__17; +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedValidDate(uint8_t); +static lean_object* l_Std_Time_ValidDate_ofOrdinal___closed__1; +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__1; +x_2 = l_Std_Time_instInhabitedValidDate___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__3; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__4; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedValidDate___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__6; +x_2 = l_Std_Time_instInhabitedValidDate___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__7; +x_2 = l_Std_Time_instInhabitedValidDate___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__8; +x_2 = l_Std_Time_instInhabitedValidDate___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__9; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__1; +x_2 = l_Std_Time_instInhabitedValidDate___closed__11; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__12; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__13; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__6; +x_2 = l_Std_Time_instInhabitedValidDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__15; +x_2 = l_Std_Time_instInhabitedValidDate___closed__14; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__16; +x_2 = l_Std_Time_instInhabitedValidDate___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__17; +x_2 = l_Std_Time_instInhabitedValidDate___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedValidDate___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedValidDate___closed__10; +x_2 = l_Std_Time_instInhabitedValidDate___closed__18; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedValidDate(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instInhabitedValidDate___closed__19; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedValidDate___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_instInhabitedValidDate(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_Month_Ordinal_cumulativeDays(x_1, x_3); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_dayOfYear___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_ValidDate_dayOfYear(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal_go(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Std_Time_Month_Ordinal_days(x_1, x_3); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +x_9 = lean_int_dec_le(x_2, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_4); +x_10 = l_Std_Time_instInhabitedValidDate___closed__1; +x_11 = lean_int_add(x_3, x_10); +lean_dec(x_3); +x_3 = x_11; +x_4 = x_8; +x_5 = lean_box(0); +x_6 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_8); +x_13 = lean_int_neg(x_4); +lean_dec(x_4); +x_14 = lean_int_add(x_2, x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_3); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_1); +lean_dec(x_1); +x_8 = l_Std_Time_ValidDate_ofOrdinal_go(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_2); +return x_8; +} +} +static lean_object* _init_l_Std_Time_ValidDate_ofOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal(uint8_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_instInhabitedValidDate___closed__10; +x_4 = l_Std_Time_ValidDate_ofOrdinal___closed__1; +x_5 = l_Std_Time_ValidDate_ofOrdinal_go(x_1, x_2, x_3, x_4, lean_box(0), lean_box(0)); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ValidDate_ofOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_ValidDate_ofOrdinal(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Day(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date_Unit_Month(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Date_ValidDate(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Day(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date_Unit_Month(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_instInhabitedValidDate___closed__1 = _init_l_Std_Time_instInhabitedValidDate___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__1); +l_Std_Time_instInhabitedValidDate___closed__2 = _init_l_Std_Time_instInhabitedValidDate___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__2); +l_Std_Time_instInhabitedValidDate___closed__3 = _init_l_Std_Time_instInhabitedValidDate___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__3); +l_Std_Time_instInhabitedValidDate___closed__4 = _init_l_Std_Time_instInhabitedValidDate___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__4); +l_Std_Time_instInhabitedValidDate___closed__5 = _init_l_Std_Time_instInhabitedValidDate___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__5); +l_Std_Time_instInhabitedValidDate___closed__6 = _init_l_Std_Time_instInhabitedValidDate___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__6); +l_Std_Time_instInhabitedValidDate___closed__7 = _init_l_Std_Time_instInhabitedValidDate___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__7); +l_Std_Time_instInhabitedValidDate___closed__8 = _init_l_Std_Time_instInhabitedValidDate___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__8); +l_Std_Time_instInhabitedValidDate___closed__9 = _init_l_Std_Time_instInhabitedValidDate___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__9); +l_Std_Time_instInhabitedValidDate___closed__10 = _init_l_Std_Time_instInhabitedValidDate___closed__10(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__10); +l_Std_Time_instInhabitedValidDate___closed__11 = _init_l_Std_Time_instInhabitedValidDate___closed__11(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__11); +l_Std_Time_instInhabitedValidDate___closed__12 = _init_l_Std_Time_instInhabitedValidDate___closed__12(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__12); +l_Std_Time_instInhabitedValidDate___closed__13 = _init_l_Std_Time_instInhabitedValidDate___closed__13(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__13); +l_Std_Time_instInhabitedValidDate___closed__14 = _init_l_Std_Time_instInhabitedValidDate___closed__14(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__14); +l_Std_Time_instInhabitedValidDate___closed__15 = _init_l_Std_Time_instInhabitedValidDate___closed__15(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__15); +l_Std_Time_instInhabitedValidDate___closed__16 = _init_l_Std_Time_instInhabitedValidDate___closed__16(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__16); +l_Std_Time_instInhabitedValidDate___closed__17 = _init_l_Std_Time_instInhabitedValidDate___closed__17(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__17); +l_Std_Time_instInhabitedValidDate___closed__18 = _init_l_Std_Time_instInhabitedValidDate___closed__18(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__18); +l_Std_Time_instInhabitedValidDate___closed__19 = _init_l_Std_Time_instInhabitedValidDate___closed__19(); +lean_mark_persistent(l_Std_Time_instInhabitedValidDate___closed__19); +l_Std_Time_ValidDate_ofOrdinal___closed__1 = _init_l_Std_Time_ValidDate_ofOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_ValidDate_ofOrdinal___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/DateTime.c b/stage0/stdlib/Std/Time/DateTime.c new file mode 100644 index 000000000000..3ce389286a99 --- /dev/null +++ b/stage0/stdlib/Std/Time/DateTime.c @@ -0,0 +1,588 @@ +// Lean compiler output +// Module: Std.Time.DateTime +// Imports: Std.Time.DateTime.Timestamp Std.Time.DateTime.PlainDateTime +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1; +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHSubDuration(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofPlainDateTimeAssumingUTC(lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainDate___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainTime(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__16; +lean_object* lean_nat_to_int(lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +extern lean_object* l_Std_Time_PlainTime_midnight; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__13; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__18; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofPlainTime(lean_object*); +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__2; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__10; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__12; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__3; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__9; +static lean_object* l_Std_Time_PlainDate_instHSubDuration___closed__1; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_getTimeAssumingUTC___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainDate(lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_getTimeAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofPlainDateAssumingUTC(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateTimeAssumingUTC___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofPlainDate(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateAssumingUTC___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainTime___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__17; +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__4; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1; +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubDuration(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateTimeAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateAssumingUTC(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofPlainTime___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofPlainDateTimeAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateTimeAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateTimeAssumingUTC___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toPlainDateTimeAssumingUTC(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofPlainDateAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_3 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1; +x_4 = lean_int_mul(x_2, x_3); +lean_dec(x_2); +x_5 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1; +x_4 = lean_int_ediv(x_2, x_3); +x_5 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toPlainDateAssumingUTC___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toPlainDateAssumingUTC(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_getTimeAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_PlainTime_ofNanoseconds(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_getTimeAssumingUTC___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_getTimeAssumingUTC(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_3 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1; +x_4 = lean_int_mul(x_2, x_3); +lean_dec(x_2); +x_5 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instHSubDuration___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instHSubDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_4 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1; +x_5 = lean_int_mul(x_3, x_4); +lean_dec(x_3); +x_6 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_2); +x_7 = lean_int_mul(x_6, x_4); +lean_dec(x_6); +x_8 = lean_int_neg(x_7); +lean_dec(x_7); +x_9 = l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1; +x_10 = lean_int_mul(x_5, x_9); +lean_dec(x_5); +x_11 = l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_mul(x_8, x_9); +lean_dec(x_8); +x_14 = l_Std_Time_PlainDate_instHSubDuration___closed__1; +x_15 = lean_int_add(x_13, x_14); +lean_dec(x_13); +x_16 = lean_int_add(x_12, x_15); +lean_dec(x_15); +lean_dec(x_12); +x_17 = l_Std_Time_Duration_ofNanoseconds(x_16); +lean_dec(x_16); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofPlainDate(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainTime_midnight; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainDate(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainDate___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_toPlainDate(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__3; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__4; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__6; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__7; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__8; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__9; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__11; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__12; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__13; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__6; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__15; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__14; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__16; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__17; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_PlainDateTime_ofPlainTime___closed__1; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__10; +x_3 = l_Std_Time_PlainDateTime_ofPlainTime___closed__18; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofPlainTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDateTime_ofPlainTime___closed__19; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toPlainTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_toPlainTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_4 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_int_neg(x_5); +lean_dec(x_5); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_int_neg(x_7); +lean_dec(x_7); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1; +x_11 = lean_int_mul(x_9, x_10); +lean_dec(x_9); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_dec(x_3); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +x_14 = lean_int_mul(x_6, x_10); +lean_dec(x_6); +x_15 = lean_int_add(x_14, x_8); +lean_dec(x_8); +lean_dec(x_14); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_Duration_ofNanoseconds(x_16); +lean_dec(x_16); +return x_17; +} +} +lean_object* initialize_Std_Time_DateTime_Timestamp(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime_PlainDateTime(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_DateTime_Timestamp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime_PlainDateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1 = _init_l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__1); +l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2 = _init_l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2(); +lean_mark_persistent(l_Std_Time_Timestamp_ofPlainDateAssumingUTC___closed__2); +l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1 = _init_l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_getTimeAssumingUTC___closed__1); +l_Std_Time_PlainDate_instHSubDuration___closed__1 = _init_l_Std_Time_PlainDate_instHSubDuration___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instHSubDuration___closed__1); +l_Std_Time_PlainDateTime_ofPlainTime___closed__1 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__1); +l_Std_Time_PlainDateTime_ofPlainTime___closed__2 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__2); +l_Std_Time_PlainDateTime_ofPlainTime___closed__3 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__3); +l_Std_Time_PlainDateTime_ofPlainTime___closed__4 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__4); +l_Std_Time_PlainDateTime_ofPlainTime___closed__5 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__5(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__5); +l_Std_Time_PlainDateTime_ofPlainTime___closed__6 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__6(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__6); +l_Std_Time_PlainDateTime_ofPlainTime___closed__7 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__7(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__7); +l_Std_Time_PlainDateTime_ofPlainTime___closed__8 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__8(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__8); +l_Std_Time_PlainDateTime_ofPlainTime___closed__9 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__9(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__9); +l_Std_Time_PlainDateTime_ofPlainTime___closed__10 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__10(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__10); +l_Std_Time_PlainDateTime_ofPlainTime___closed__11 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__11(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__11); +l_Std_Time_PlainDateTime_ofPlainTime___closed__12 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__12(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__12); +l_Std_Time_PlainDateTime_ofPlainTime___closed__13 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__13(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__13); +l_Std_Time_PlainDateTime_ofPlainTime___closed__14 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__14(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__14); +l_Std_Time_PlainDateTime_ofPlainTime___closed__15 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__15(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__15); +l_Std_Time_PlainDateTime_ofPlainTime___closed__16 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__16(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__16); +l_Std_Time_PlainDateTime_ofPlainTime___closed__17 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__17(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__17); +l_Std_Time_PlainDateTime_ofPlainTime___closed__18 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__18(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__18); +l_Std_Time_PlainDateTime_ofPlainTime___closed__19 = _init_l_Std_Time_PlainDateTime_ofPlainTime___closed__19(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofPlainTime___closed__19); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/DateTime/PlainDateTime.c b/stage0/stdlib/Std/Time/DateTime/PlainDateTime.c new file mode 100644 index 000000000000..1b6142537361 --- /dev/null +++ b/stage0/stdlib/Std/Time/DateTime/PlainDateTime.c @@ -0,0 +1,6439 @@ +// Lean compiler output +// Module: Std.Time.DateTime.PlainDateTime +// Imports: Std.Time.Date Std.Time.Time Std.Time.Internal Std.Time.DateTime.Timestamp +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__21; +lean_object* lean_int_mod(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_year___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addWeeks___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_atTime(lean_object*, lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_era___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___boxed(lean_object*); +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_day(lean_object*); +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsClip___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_millisecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_day___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__36; +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__31; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +static lean_object* l_Std_Time_instReprPlainDateTime___closed__1; +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__9; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsRollOver___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_month(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withWeekday___boxed(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_atTime(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_second(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__30; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__26; +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10; +static lean_object* l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_millisecond___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__34; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subSeconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__22; +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withYearRollOver(lean_object*, lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Fin_add(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__17; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addDays___boxed(lean_object*, lean_object*); +lean_object* l_Fin_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__5; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subSeconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMonthRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_inLeapYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toDaysSinceUNIXEpoch(lean_object*); +lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_weekday(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__24; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofDaysSinceUNIXEpoch(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_dayOfYear(lean_object*); +uint8_t l_instDecidableNot___rarg(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMonthClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfMonth(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_year(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedPlainDateTime; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsClip___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__18; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subDays(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subNanoseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_nanosecond___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__7; +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_nanosecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_quarter(lean_object*); +lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_hour___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddDuration___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__14; +static lean_object* l_Std_Time_PlainDateTime_instHAddOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekday___boxed(lean_object*); +lean_object* l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withHours(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__12; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_month___boxed(lean_object*); +lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_atDate(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysRollOver___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__15; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__27; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__23; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprPlainDateTime; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_quarter___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withSeconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withYearClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_minute(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addWeeks(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__6; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subNanoseconds(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqPlainDateTime; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofDaysSinceUNIXEpoch___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_alignedWeekOfMonth(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__10; +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMilliseconds(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_rollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withWeekday(lean_object*, uint8_t); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddDuration(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__33; +lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__4; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addSeconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__29; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfMonth___boxed(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subHours___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_dayOfYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__13; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2; +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_hour(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__32; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addHours(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__35; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__39; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsRollOver___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13; +lean_object* l_Std_Time_PlainTime_addMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__2; +uint8_t l_Std_Time_Year_Offset_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsRollOver(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__2; +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4; +lean_object* l_Int_toNat(lean_object*); +uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__2; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset__1; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__38; +lean_object* l_Std_Time_PlainDate_withWeekday(lean_object*, uint8_t); +static lean_object* l_Std_Time_PlainDateTime_withMilliseconds___closed__2; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__20; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset; +static lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_minute___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__37; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__25; +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysRollOver(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_atDate(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l_Std_Time_instBEqPlainDateTime___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_inLeapYear(lean_object*); +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver___boxed(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addDays(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subWeeks(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHSubOffset__5; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__40; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1; +static lean_object* l_Std_Time_PlainDateTime_addMilliseconds___closed__1; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addSeconds(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysClip(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_weekOfMonth(lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddOffset; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__28; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_second___boxed(lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addHours___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_withMilliseconds___closed__1; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3; +static lean_object* l_Std_Time_PlainDateTime_addWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4; +static lean_object* l_Std_Time_instInhabitedPlainDateTime___closed__1; +LEAN_EXPORT uint8_t l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37_(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1; +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15; +static lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6; +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__4; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__5; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__7; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__8; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__9; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__10; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__12; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__13; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__14; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__7; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__15; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__16; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__15; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__17; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__15; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__18; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__11; +x_3 = l_Std_Time_instInhabitedPlainDateTime___closed__19; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__21; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__22; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__23; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__25; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__24; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__26; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__24; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__27; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__24; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__28; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__30; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__31; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__32; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__25; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__33; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__34; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__33; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__35; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__33; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__36; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__38() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__37; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__29; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__37; +x_3 = l_Std_Time_instInhabitedPlainDateTime___closed__38; +x_4 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__20; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__39; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainDateTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__40; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_3, 2); +x_8 = lean_ctor_get(x_5, 2); +x_9 = lean_int_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_5, 1); +x_13 = lean_int_dec_eq(x_11, x_12); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_3, 0); +x_16 = lean_ctor_get(x_5, 0); +x_17 = lean_int_dec_eq(x_15, x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_4, 0); +x_20 = lean_ctor_get(x_6, 0); +x_21 = lean_int_dec_eq(x_19, x_20); +if (x_21 == 0) +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_4, 1); +x_24 = lean_ctor_get(x_6, 1); +x_25 = lean_int_dec_eq(x_23, x_24); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = 0; +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_27 = lean_ctor_get(x_4, 2); +x_28 = lean_ctor_get(x_27, 1); +x_29 = lean_ctor_get(x_6, 2); +x_30 = lean_ctor_get(x_29, 1); +x_31 = lean_int_dec_eq(x_28, x_30); +if (x_31 == 0) +{ +uint8_t x_32; +x_32 = 0; +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_33 = lean_ctor_get(x_4, 3); +x_34 = lean_ctor_get(x_6, 3); +x_35 = lean_int_dec_eq(x_33, x_34); +return x_35; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instBEqPlainDateTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_beqPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_37____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instBEqPlainDateTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instBEqPlainDateTime___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("date", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3; +x_2 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("time", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Date_PlainDate_0__Std_Time_reprPlainDate____x40_Std_Time_Date_PlainDate___hyg_40_(x_3, x_4); +lean_dec(x_3); +x_6 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43_(x_20, x_4); +x_22 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set_uint8(x_23, sizeof(void*)*1, x_8); +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_23); +x_25 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15; +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17; +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14; +x_30 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_8); +return x_31; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprPlainDateTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprPlainDateTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprPlainDateTime___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_2); +x_4 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_5 = lean_int_mul(x_3, x_4); +lean_dec(x_3); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec(x_1); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +x_9 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_10 = lean_int_mul(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_6, 3); +lean_inc(x_11); +lean_dec(x_6); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = l_Std_Time_Duration_ofNanoseconds(x_12); +lean_dec(x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_int_sub(x_3, x_1); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +if (lean_obj_tag(x_4) == 0) +{ +return x_5; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_5, 1); +x_12 = lean_unsigned_to_nat(13u); +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Fin_add(x_12, x_10, x_13); +lean_dec(x_10); +x_15 = lean_int_dec_lt(x_11, x_8); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_free_object(x_5); +x_16 = lean_box(0); +x_17 = l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1(x_8, x_14, x_11, x_16); +lean_dec(x_11); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_4 = x_9; +x_5 = x_18; +x_6 = lean_box(0); +goto _start; +} +else +{ +lean_ctor_set(x_5, 0, x_14); +return x_5; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_20 = lean_ctor_get(x_4, 0); +x_21 = lean_ctor_get(x_4, 1); +x_22 = lean_ctor_get(x_5, 0); +x_23 = lean_ctor_get(x_5, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_5); +x_24 = lean_unsigned_to_nat(13u); +x_25 = lean_unsigned_to_nat(1u); +x_26 = l_Fin_add(x_24, x_22, x_25); +lean_dec(x_22); +x_27 = lean_int_dec_lt(x_23, x_20); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_box(0); +x_29 = l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1(x_20, x_26, x_23, x_28); +lean_dec(x_23); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +x_4 = x_21; +x_5 = x_30; +x_6 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_32; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_26); +lean_ctor_set(x_32, 1, x_23); +return x_32; +} +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_11 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1; +x_12 = lean_int_emod(x_1, x_11); +x_13 = lean_int_ediv(x_1, x_11); +x_14 = lean_int_emod(x_13, x_11); +lean_dec(x_13); +x_15 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2; +x_16 = lean_int_ediv(x_1, x_15); +x_17 = lean_int_emod(x_2, x_3); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_4, x_18); +x_20 = lean_nat_dec_le(x_18, x_19); +x_21 = lean_int_mod(x_9, x_5); +x_22 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_23 = lean_int_dec_eq(x_21, x_22); +lean_dec(x_21); +x_24 = 0; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_12); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_16); +lean_ctor_set(x_27, 1, x_14); +lean_ctor_set(x_27, 2, x_26); +lean_ctor_set(x_27, 3, x_17); +if (x_20 == 0) +{ +lean_dec(x_19); +x_28 = x_8; +goto block_60; +} +else +{ +lean_object* x_61; +lean_dec(x_8); +x_61 = lean_nat_to_int(x_19); +x_28 = x_61; +goto block_60; +} +block_60: +{ +if (x_23 == 0) +{ +lean_object* x_29; uint8_t x_30; +x_29 = l_Std_Time_Month_Ordinal_days(x_24, x_10); +x_30 = lean_int_dec_lt(x_29, x_28); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_29); +x_31 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_31, 0, x_9); +lean_ctor_set(x_31, 1, x_10); +lean_ctor_set(x_31, 2, x_28); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_27); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_28); +x_33 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_33, 0, x_9); +lean_ctor_set(x_33, 1, x_10); +lean_ctor_set(x_33, 2, x_29); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_27); +return x_34; +} +} +else +{ +lean_object* x_35; uint8_t x_36; uint8_t x_37; +x_35 = lean_int_mod(x_9, x_6); +x_36 = lean_int_dec_eq(x_35, x_22); +lean_dec(x_35); +x_37 = l_instDecidableNot___rarg(x_36); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_int_mod(x_9, x_7); +x_39 = lean_int_dec_eq(x_38, x_22); +lean_dec(x_38); +if (x_39 == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = l_Std_Time_Month_Ordinal_days(x_24, x_10); +x_41 = lean_int_dec_lt(x_40, x_28); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_40); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_9); +lean_ctor_set(x_42, 1, x_10); +lean_ctor_set(x_42, 2, x_28); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_27); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_28); +x_44 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_44, 0, x_9); +lean_ctor_set(x_44, 1, x_10); +lean_ctor_set(x_44, 2, x_40); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_27); +return x_45; +} +} +else +{ +uint8_t x_46; lean_object* x_47; uint8_t x_48; +x_46 = 1; +x_47 = l_Std_Time_Month_Ordinal_days(x_46, x_10); +x_48 = lean_int_dec_lt(x_47, x_28); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_49, 0, x_9); +lean_ctor_set(x_49, 1, x_10); +lean_ctor_set(x_49, 2, x_28); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_27); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; +lean_dec(x_28); +x_51 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_51, 0, x_9); +lean_ctor_set(x_51, 1, x_10); +lean_ctor_set(x_51, 2, x_47); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_27); +return x_52; +} +} +} +else +{ +uint8_t x_53; lean_object* x_54; uint8_t x_55; +x_53 = 1; +x_54 = l_Std_Time_Month_Ordinal_days(x_53, x_10); +x_55 = lean_int_dec_lt(x_54, x_28); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_56, 0, x_9); +lean_ctor_set(x_56, 1, x_10); +lean_ctor_set(x_56, 2, x_28); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_27); +return x_57; +} +else +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_28); +x_58 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_58, 0, x_9); +lean_ctor_set(x_58, 1, x_10); +lean_ctor_set(x_58, 2, x_54); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_27); +return x_59; +} +} +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(31u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(29u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__12; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__12; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__12; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__12; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_15 = lean_int_mul(x_13, x_1); +x_16 = lean_int_sub(x_12, x_15); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1; +x_18 = lean_int_add(x_17, x_13); +x_19 = lean_int_mul(x_2, x_3); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +x_21 = lean_int_mul(x_4, x_5); +x_22 = lean_int_add(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +x_23 = lean_int_mul(x_6, x_7); +x_24 = lean_int_add(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +x_25 = lean_box(0); +x_26 = lean_unsigned_to_nat(0u); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_16); +x_28 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15; +x_29 = l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1(x_28, x_25, x_28, x_28, x_27, lean_box(0)); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Int_toNat(x_31); +lean_dec(x_31); +x_33 = lean_unsigned_to_nat(30u); +x_34 = l_Fin_ofNat(x_33, x_32); +lean_dec(x_32); +x_35 = lean_unsigned_to_nat(10u); +x_36 = lean_nat_dec_lt(x_35, x_30); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_unsigned_to_nat(2u); +x_38 = lean_nat_add(x_30, x_37); +lean_dec(x_30); +x_39 = lean_nat_to_int(x_38); +x_40 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1(x_8, x_9, x_10, x_34, x_2, x_4, x_6, x_11, x_24, x_39); +lean_dec(x_34); +return x_40; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_int_add(x_24, x_11); +lean_dec(x_24); +x_42 = lean_nat_sub(x_30, x_35); +lean_dec(x_30); +x_43 = lean_nat_to_int(x_42); +x_44 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1(x_8, x_9, x_10, x_34, x_2, x_4, x_6, x_11, x_41, x_43); +lean_dec(x_34); +return x_44; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_int_mul(x_12, x_1); +x_16 = lean_int_sub(x_13, x_15); +lean_dec(x_15); +x_17 = lean_int_ediv(x_16, x_2); +x_18 = lean_int_dec_eq(x_17, x_3); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2(x_2, x_3, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16, x_17, x_19); +lean_dec(x_17); +lean_dec(x_16); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_int_sub(x_17, x_11); +lean_dec(x_17); +x_22 = lean_box(0); +x_23 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2(x_2, x_3, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16, x_21, x_22); +lean_dec(x_21); +lean_dec(x_16); +return x_23; +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(25u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_int_mul(x_12, x_1); +x_16 = lean_int_sub(x_13, x_15); +lean_dec(x_15); +x_17 = lean_int_ediv(x_16, x_2); +x_18 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1; +x_19 = lean_int_dec_eq(x_17, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_box(0); +x_21 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3(x_2, x_3, x_4, x_5, x_12, x_6, x_7, x_8, x_9, x_10, x_11, x_17, x_16, x_20); +lean_dec(x_16); +lean_dec(x_17); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_int_sub(x_17, x_11); +lean_dec(x_17); +x_23 = lean_box(0); +x_24 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3(x_2, x_3, x_4, x_5, x_12, x_6, x_7, x_8, x_9, x_10, x_11, x_22, x_16, x_23); +lean_dec(x_16); +lean_dec(x_22); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_int_ediv(x_12, x_1); +x_15 = lean_int_dec_eq(x_14, x_4); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_box(0); +x_17 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_7, x_8, x_9, x_10, x_14, x_12, x_16); +lean_dec(x_14); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_int_sub(x_14, x_10); +lean_dec(x_14); +x_19 = lean_box(0); +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_11, x_7, x_8, x_9, x_10, x_18, x_12, x_19); +lean_dec(x_18); +return x_20; +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11017u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(365u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(97u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12; +x_2 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_41; uint8_t x_42; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +x_7 = lean_int_ediv(x_6, x_3); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +x_10 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1; +x_11 = lean_int_sub(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mod(x_7, x_8); +lean_dec(x_7); +x_41 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14; +x_42 = lean_int_dec_le(x_12, x_41); +if (x_42 == 0) +{ +x_13 = x_12; +x_14 = x_11; +goto block_40; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_int_add(x_12, x_8); +lean_dec(x_12); +x_44 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_45 = lean_int_sub(x_11, x_44); +lean_dec(x_11); +x_13 = x_43; +x_14 = x_45; +goto block_40; +} +block_40: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6; +x_16 = lean_int_ediv(x_14, x_15); +x_17 = lean_int_emod(x_14, x_15); +lean_dec(x_14); +x_18 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_19 = lean_int_dec_lt(x_17, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10; +x_21 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13; +x_22 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +x_23 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_24 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_25 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_26 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_27 = lean_box(0); +x_28 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5(x_20, x_21, x_22, x_23, x_24, x_25, x_13, x_6, x_3, x_26, x_16, x_17, x_27); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_6); +lean_dec(x_13); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = lean_int_add(x_17, x_15); +lean_dec(x_17); +x_30 = l_Std_Time_instInhabitedPlainDateTime___closed__2; +x_31 = lean_int_sub(x_16, x_30); +lean_dec(x_16); +x_32 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10; +x_33 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13; +x_34 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2; +x_35 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_36 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_37 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_38 = lean_box(0); +x_39 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5(x_32, x_33, x_34, x_35, x_36, x_37, x_13, x_6, x_3, x_30, x_31, x_29, x_38); +lean_dec(x_29); +lean_dec(x_31); +lean_dec(x_6); +lean_dec(x_13); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_List_forIn_x27_loop___at_Std_Time_PlainDateTime_ofTimestampAssumingUTC___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; +x_14 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toDaysSinceUNIXEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofDaysSinceUNIXEpoch(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_ofDaysSinceUNIXEpoch___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_ofDaysSinceUNIXEpoch(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withWeekday(lean_object* x_1, uint8_t x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDate_withWeekday(x_4, x_2); +lean_ctor_set(x_1, 0, x_5); +return x_1; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_8 = l_Std_Time_PlainDate_withWeekday(x_6, x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withWeekday___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Std_Time_PlainDateTime_withWeekday(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_ctor_get(x_4, 2); +lean_dec(x_8); +x_9 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_10 = lean_int_mod(x_6, x_9); +x_11 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_12 = lean_int_dec_eq(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +uint8_t x_13; lean_object* x_14; uint8_t x_15; +x_13 = 0; +x_14 = l_Std_Time_Month_Ordinal_days(x_13, x_7); +x_15 = lean_int_dec_lt(x_14, x_2); +if (x_15 == 0) +{ +lean_dec(x_14); +lean_ctor_set(x_4, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_4, 2, x_14); +return x_1; +} +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; +x_16 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_17 = lean_int_mod(x_6, x_16); +x_18 = lean_int_dec_eq(x_17, x_11); +lean_dec(x_17); +x_19 = l_instDecidableNot___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_21 = lean_int_mod(x_6, x_20); +x_22 = lean_int_dec_eq(x_21, x_11); +lean_dec(x_21); +if (x_22 == 0) +{ +uint8_t x_23; lean_object* x_24; uint8_t x_25; +x_23 = 0; +x_24 = l_Std_Time_Month_Ordinal_days(x_23, x_7); +x_25 = lean_int_dec_lt(x_24, x_2); +if (x_25 == 0) +{ +lean_dec(x_24); +lean_ctor_set(x_4, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_4, 2, x_24); +return x_1; +} +} +else +{ +uint8_t x_26; lean_object* x_27; uint8_t x_28; +x_26 = 1; +x_27 = l_Std_Time_Month_Ordinal_days(x_26, x_7); +x_28 = lean_int_dec_lt(x_27, x_2); +if (x_28 == 0) +{ +lean_dec(x_27); +lean_ctor_set(x_4, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_4, 2, x_27); +return x_1; +} +} +} +else +{ +uint8_t x_29; lean_object* x_30; uint8_t x_31; +x_29 = 1; +x_30 = l_Std_Time_Month_Ordinal_days(x_29, x_7); +x_31 = lean_int_dec_lt(x_30, x_2); +if (x_31 == 0) +{ +lean_dec(x_30); +lean_ctor_set(x_4, 2, x_2); +return x_1; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_4, 2, x_30); +return x_1; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_32 = lean_ctor_get(x_4, 0); +x_33 = lean_ctor_get(x_4, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_4); +x_34 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_35 = lean_int_mod(x_32, x_34); +x_36 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_37 = lean_int_dec_eq(x_35, x_36); +lean_dec(x_35); +if (x_37 == 0) +{ +uint8_t x_38; lean_object* x_39; uint8_t x_40; +x_38 = 0; +x_39 = l_Std_Time_Month_Ordinal_days(x_38, x_33); +x_40 = lean_int_dec_lt(x_39, x_2); +if (x_40 == 0) +{ +lean_object* x_41; +lean_dec(x_39); +x_41 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 1, x_33); +lean_ctor_set(x_41, 2, x_2); +lean_ctor_set(x_1, 0, x_41); +return x_1; +} +else +{ +lean_object* x_42; +lean_dec(x_2); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_32); +lean_ctor_set(x_42, 1, x_33); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_1, 0, x_42); +return x_1; +} +} +else +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; uint8_t x_46; +x_43 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_44 = lean_int_mod(x_32, x_43); +x_45 = lean_int_dec_eq(x_44, x_36); +lean_dec(x_44); +x_46 = l_instDecidableNot___rarg(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_48 = lean_int_mod(x_32, x_47); +x_49 = lean_int_dec_eq(x_48, x_36); +lean_dec(x_48); +if (x_49 == 0) +{ +uint8_t x_50; lean_object* x_51; uint8_t x_52; +x_50 = 0; +x_51 = l_Std_Time_Month_Ordinal_days(x_50, x_33); +x_52 = lean_int_dec_lt(x_51, x_2); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_32); +lean_ctor_set(x_53, 1, x_33); +lean_ctor_set(x_53, 2, x_2); +lean_ctor_set(x_1, 0, x_53); +return x_1; +} +else +{ +lean_object* x_54; +lean_dec(x_2); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_32); +lean_ctor_set(x_54, 1, x_33); +lean_ctor_set(x_54, 2, x_51); +lean_ctor_set(x_1, 0, x_54); +return x_1; +} +} +else +{ +uint8_t x_55; lean_object* x_56; uint8_t x_57; +x_55 = 1; +x_56 = l_Std_Time_Month_Ordinal_days(x_55, x_33); +x_57 = lean_int_dec_lt(x_56, x_2); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_56); +x_58 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_58, 0, x_32); +lean_ctor_set(x_58, 1, x_33); +lean_ctor_set(x_58, 2, x_2); +lean_ctor_set(x_1, 0, x_58); +return x_1; +} +else +{ +lean_object* x_59; +lean_dec(x_2); +x_59 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_59, 0, x_32); +lean_ctor_set(x_59, 1, x_33); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_1, 0, x_59); +return x_1; +} +} +} +else +{ +uint8_t x_60; lean_object* x_61; uint8_t x_62; +x_60 = 1; +x_61 = l_Std_Time_Month_Ordinal_days(x_60, x_33); +x_62 = lean_int_dec_lt(x_61, x_2); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_61); +x_63 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_63, 0, x_32); +lean_ctor_set(x_63, 1, x_33); +lean_ctor_set(x_63, 2, x_2); +lean_ctor_set(x_1, 0, x_63); +return x_1; +} +else +{ +lean_object* x_64; +lean_dec(x_2); +x_64 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_64, 0, x_32); +lean_ctor_set(x_64, 1, x_33); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_1, 0, x_64); +return x_1; +} +} +} +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_65 = lean_ctor_get(x_1, 0); +x_66 = lean_ctor_get(x_1, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_1); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + x_69 = x_65; +} else { + lean_dec_ref(x_65); + x_69 = lean_box(0); +} +x_70 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_71 = lean_int_mod(x_67, x_70); +x_72 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_73 = lean_int_dec_eq(x_71, x_72); +lean_dec(x_71); +if (x_73 == 0) +{ +uint8_t x_74; lean_object* x_75; uint8_t x_76; +x_74 = 0; +x_75 = l_Std_Time_Month_Ordinal_days(x_74, x_68); +x_76 = lean_int_dec_lt(x_75, x_2); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_75); +if (lean_is_scalar(x_69)) { + x_77 = lean_alloc_ctor(0, 3, 0); +} else { + x_77 = x_69; +} +lean_ctor_set(x_77, 0, x_67); +lean_ctor_set(x_77, 1, x_68); +lean_ctor_set(x_77, 2, x_2); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_66); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_2); +if (lean_is_scalar(x_69)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_69; +} +lean_ctor_set(x_79, 0, x_67); +lean_ctor_set(x_79, 1, x_68); +lean_ctor_set(x_79, 2, x_75); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_66); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; uint8_t x_83; uint8_t x_84; +x_81 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_82 = lean_int_mod(x_67, x_81); +x_83 = lean_int_dec_eq(x_82, x_72); +lean_dec(x_82); +x_84 = l_instDecidableNot___rarg(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_85 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_86 = lean_int_mod(x_67, x_85); +x_87 = lean_int_dec_eq(x_86, x_72); +lean_dec(x_86); +if (x_87 == 0) +{ +uint8_t x_88; lean_object* x_89; uint8_t x_90; +x_88 = 0; +x_89 = l_Std_Time_Month_Ordinal_days(x_88, x_68); +x_90 = lean_int_dec_lt(x_89, x_2); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_89); +if (lean_is_scalar(x_69)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_69; +} +lean_ctor_set(x_91, 0, x_67); +lean_ctor_set(x_91, 1, x_68); +lean_ctor_set(x_91, 2, x_2); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_66); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_2); +if (lean_is_scalar(x_69)) { + x_93 = lean_alloc_ctor(0, 3, 0); +} else { + x_93 = x_69; +} +lean_ctor_set(x_93, 0, x_67); +lean_ctor_set(x_93, 1, x_68); +lean_ctor_set(x_93, 2, x_89); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_66); +return x_94; +} +} +else +{ +uint8_t x_95; lean_object* x_96; uint8_t x_97; +x_95 = 1; +x_96 = l_Std_Time_Month_Ordinal_days(x_95, x_68); +x_97 = lean_int_dec_lt(x_96, x_2); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; +lean_dec(x_96); +if (lean_is_scalar(x_69)) { + x_98 = lean_alloc_ctor(0, 3, 0); +} else { + x_98 = x_69; +} +lean_ctor_set(x_98, 0, x_67); +lean_ctor_set(x_98, 1, x_68); +lean_ctor_set(x_98, 2, x_2); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_66); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_2); +if (lean_is_scalar(x_69)) { + x_100 = lean_alloc_ctor(0, 3, 0); +} else { + x_100 = x_69; +} +lean_ctor_set(x_100, 0, x_67); +lean_ctor_set(x_100, 1, x_68); +lean_ctor_set(x_100, 2, x_96); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_66); +return x_101; +} +} +} +else +{ +uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_102 = 1; +x_103 = l_Std_Time_Month_Ordinal_days(x_102, x_68); +x_104 = lean_int_dec_lt(x_103, x_2); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_103); +if (lean_is_scalar(x_69)) { + x_105 = lean_alloc_ctor(0, 3, 0); +} else { + x_105 = x_69; +} +lean_ctor_set(x_105, 0, x_67); +lean_ctor_set(x_105, 1, x_68); +lean_ctor_set(x_105, 2, x_2); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_66); +return x_106; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_2); +if (lean_is_scalar(x_69)) { + x_107 = lean_alloc_ctor(0, 3, 0); +} else { + x_107 = x_69; +} +lean_ctor_set(x_107, 0, x_67); +lean_ctor_set(x_107, 1, x_68); +lean_ctor_set(x_107, 2, x_103); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_66); +return x_108; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_rollOver(x_5, x_6, x_2); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = l_Std_Time_PlainDate_rollOver(x_10, x_11, x_2); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withDaysRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_withDaysRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMonthClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 2); +x_8 = lean_ctor_get(x_4, 1); +lean_dec(x_8); +x_9 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_10 = lean_int_mod(x_6, x_9); +x_11 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_12 = lean_int_dec_eq(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +uint8_t x_13; lean_object* x_14; uint8_t x_15; +x_13 = 0; +x_14 = l_Std_Time_Month_Ordinal_days(x_13, x_2); +x_15 = lean_int_dec_lt(x_14, x_7); +if (x_15 == 0) +{ +lean_dec(x_14); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_14); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; +x_16 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_17 = lean_int_mod(x_6, x_16); +x_18 = lean_int_dec_eq(x_17, x_11); +lean_dec(x_17); +x_19 = l_instDecidableNot___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_21 = lean_int_mod(x_6, x_20); +x_22 = lean_int_dec_eq(x_21, x_11); +lean_dec(x_21); +if (x_22 == 0) +{ +uint8_t x_23; lean_object* x_24; uint8_t x_25; +x_23 = 0; +x_24 = l_Std_Time_Month_Ordinal_days(x_23, x_2); +x_25 = lean_int_dec_lt(x_24, x_7); +if (x_25 == 0) +{ +lean_dec(x_24); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_24); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +} +else +{ +uint8_t x_26; lean_object* x_27; uint8_t x_28; +x_26 = 1; +x_27 = l_Std_Time_Month_Ordinal_days(x_26, x_2); +x_28 = lean_int_dec_lt(x_27, x_7); +if (x_28 == 0) +{ +lean_dec(x_27); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_27); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +} +} +else +{ +uint8_t x_29; lean_object* x_30; uint8_t x_31; +x_29 = 1; +x_30 = l_Std_Time_Month_Ordinal_days(x_29, x_2); +x_31 = lean_int_dec_lt(x_30, x_7); +if (x_31 == 0) +{ +lean_dec(x_30); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_30); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_32 = lean_ctor_get(x_4, 0); +x_33 = lean_ctor_get(x_4, 2); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_4); +x_34 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_35 = lean_int_mod(x_32, x_34); +x_36 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_37 = lean_int_dec_eq(x_35, x_36); +lean_dec(x_35); +if (x_37 == 0) +{ +uint8_t x_38; lean_object* x_39; uint8_t x_40; +x_38 = 0; +x_39 = l_Std_Time_Month_Ordinal_days(x_38, x_2); +x_40 = lean_int_dec_lt(x_39, x_33); +if (x_40 == 0) +{ +lean_object* x_41; +lean_dec(x_39); +x_41 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 1, x_2); +lean_ctor_set(x_41, 2, x_33); +lean_ctor_set(x_1, 0, x_41); +return x_1; +} +else +{ +lean_object* x_42; +lean_dec(x_33); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_32); +lean_ctor_set(x_42, 1, x_2); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_1, 0, x_42); +return x_1; +} +} +else +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; uint8_t x_46; +x_43 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_44 = lean_int_mod(x_32, x_43); +x_45 = lean_int_dec_eq(x_44, x_36); +lean_dec(x_44); +x_46 = l_instDecidableNot___rarg(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_48 = lean_int_mod(x_32, x_47); +x_49 = lean_int_dec_eq(x_48, x_36); +lean_dec(x_48); +if (x_49 == 0) +{ +uint8_t x_50; lean_object* x_51; uint8_t x_52; +x_50 = 0; +x_51 = l_Std_Time_Month_Ordinal_days(x_50, x_2); +x_52 = lean_int_dec_lt(x_51, x_33); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_32); +lean_ctor_set(x_53, 1, x_2); +lean_ctor_set(x_53, 2, x_33); +lean_ctor_set(x_1, 0, x_53); +return x_1; +} +else +{ +lean_object* x_54; +lean_dec(x_33); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_32); +lean_ctor_set(x_54, 1, x_2); +lean_ctor_set(x_54, 2, x_51); +lean_ctor_set(x_1, 0, x_54); +return x_1; +} +} +else +{ +uint8_t x_55; lean_object* x_56; uint8_t x_57; +x_55 = 1; +x_56 = l_Std_Time_Month_Ordinal_days(x_55, x_2); +x_57 = lean_int_dec_lt(x_56, x_33); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_56); +x_58 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_58, 0, x_32); +lean_ctor_set(x_58, 1, x_2); +lean_ctor_set(x_58, 2, x_33); +lean_ctor_set(x_1, 0, x_58); +return x_1; +} +else +{ +lean_object* x_59; +lean_dec(x_33); +x_59 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_59, 0, x_32); +lean_ctor_set(x_59, 1, x_2); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_1, 0, x_59); +return x_1; +} +} +} +else +{ +uint8_t x_60; lean_object* x_61; uint8_t x_62; +x_60 = 1; +x_61 = l_Std_Time_Month_Ordinal_days(x_60, x_2); +x_62 = lean_int_dec_lt(x_61, x_33); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_61); +x_63 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_63, 0, x_32); +lean_ctor_set(x_63, 1, x_2); +lean_ctor_set(x_63, 2, x_33); +lean_ctor_set(x_1, 0, x_63); +return x_1; +} +else +{ +lean_object* x_64; +lean_dec(x_33); +x_64 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_64, 0, x_32); +lean_ctor_set(x_64, 1, x_2); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_1, 0, x_64); +return x_1; +} +} +} +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_65 = lean_ctor_get(x_1, 0); +x_66 = lean_ctor_get(x_1, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_1); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_65, 2); +lean_inc(x_68); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + x_69 = x_65; +} else { + lean_dec_ref(x_65); + x_69 = lean_box(0); +} +x_70 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_71 = lean_int_mod(x_67, x_70); +x_72 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_73 = lean_int_dec_eq(x_71, x_72); +lean_dec(x_71); +if (x_73 == 0) +{ +uint8_t x_74; lean_object* x_75; uint8_t x_76; +x_74 = 0; +x_75 = l_Std_Time_Month_Ordinal_days(x_74, x_2); +x_76 = lean_int_dec_lt(x_75, x_68); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_75); +if (lean_is_scalar(x_69)) { + x_77 = lean_alloc_ctor(0, 3, 0); +} else { + x_77 = x_69; +} +lean_ctor_set(x_77, 0, x_67); +lean_ctor_set(x_77, 1, x_2); +lean_ctor_set(x_77, 2, x_68); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_66); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_69; +} +lean_ctor_set(x_79, 0, x_67); +lean_ctor_set(x_79, 1, x_2); +lean_ctor_set(x_79, 2, x_75); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_66); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; uint8_t x_83; uint8_t x_84; +x_81 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_82 = lean_int_mod(x_67, x_81); +x_83 = lean_int_dec_eq(x_82, x_72); +lean_dec(x_82); +x_84 = l_instDecidableNot___rarg(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_85 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_86 = lean_int_mod(x_67, x_85); +x_87 = lean_int_dec_eq(x_86, x_72); +lean_dec(x_86); +if (x_87 == 0) +{ +uint8_t x_88; lean_object* x_89; uint8_t x_90; +x_88 = 0; +x_89 = l_Std_Time_Month_Ordinal_days(x_88, x_2); +x_90 = lean_int_dec_lt(x_89, x_68); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_89); +if (lean_is_scalar(x_69)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_69; +} +lean_ctor_set(x_91, 0, x_67); +lean_ctor_set(x_91, 1, x_2); +lean_ctor_set(x_91, 2, x_68); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_66); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_93 = lean_alloc_ctor(0, 3, 0); +} else { + x_93 = x_69; +} +lean_ctor_set(x_93, 0, x_67); +lean_ctor_set(x_93, 1, x_2); +lean_ctor_set(x_93, 2, x_89); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_66); +return x_94; +} +} +else +{ +uint8_t x_95; lean_object* x_96; uint8_t x_97; +x_95 = 1; +x_96 = l_Std_Time_Month_Ordinal_days(x_95, x_2); +x_97 = lean_int_dec_lt(x_96, x_68); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; +lean_dec(x_96); +if (lean_is_scalar(x_69)) { + x_98 = lean_alloc_ctor(0, 3, 0); +} else { + x_98 = x_69; +} +lean_ctor_set(x_98, 0, x_67); +lean_ctor_set(x_98, 1, x_2); +lean_ctor_set(x_98, 2, x_68); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_66); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_100 = lean_alloc_ctor(0, 3, 0); +} else { + x_100 = x_69; +} +lean_ctor_set(x_100, 0, x_67); +lean_ctor_set(x_100, 1, x_2); +lean_ctor_set(x_100, 2, x_96); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_66); +return x_101; +} +} +} +else +{ +uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_102 = 1; +x_103 = l_Std_Time_Month_Ordinal_days(x_102, x_2); +x_104 = lean_int_dec_lt(x_103, x_68); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_103); +if (lean_is_scalar(x_69)) { + x_105 = lean_alloc_ctor(0, 3, 0); +} else { + x_105 = x_69; +} +lean_ctor_set(x_105, 0, x_67); +lean_ctor_set(x_105, 1, x_2); +lean_ctor_set(x_105, 2, x_68); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_66); +return x_106; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_107 = lean_alloc_ctor(0, 3, 0); +} else { + x_107 = x_69; +} +lean_ctor_set(x_107, 0, x_67); +lean_ctor_set(x_107, 1, x_2); +lean_ctor_set(x_107, 2, x_103); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_66); +return x_108; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMonthRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 2); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_rollOver(x_5, x_2, x_6); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 2); +lean_inc(x_11); +lean_dec(x_8); +x_12 = l_Std_Time_PlainDate_rollOver(x_10, x_2, x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withYearClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_6 = lean_ctor_get(x_4, 1); +x_7 = lean_ctor_get(x_4, 2); +x_8 = lean_ctor_get(x_4, 0); +lean_dec(x_8); +x_9 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_10 = lean_int_mod(x_2, x_9); +x_11 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_12 = lean_int_dec_eq(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) +{ +uint8_t x_13; lean_object* x_14; uint8_t x_15; +x_13 = 0; +x_14 = l_Std_Time_Month_Ordinal_days(x_13, x_6); +x_15 = lean_int_dec_lt(x_14, x_7); +if (x_15 == 0) +{ +lean_dec(x_14); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_14); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; +x_16 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_17 = lean_int_mod(x_2, x_16); +x_18 = lean_int_dec_eq(x_17, x_11); +lean_dec(x_17); +x_19 = l_instDecidableNot___rarg(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_21 = lean_int_mod(x_2, x_20); +x_22 = lean_int_dec_eq(x_21, x_11); +lean_dec(x_21); +if (x_22 == 0) +{ +uint8_t x_23; lean_object* x_24; uint8_t x_25; +x_23 = 0; +x_24 = l_Std_Time_Month_Ordinal_days(x_23, x_6); +x_25 = lean_int_dec_lt(x_24, x_7); +if (x_25 == 0) +{ +lean_dec(x_24); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_24); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +} +else +{ +uint8_t x_26; lean_object* x_27; uint8_t x_28; +x_26 = 1; +x_27 = l_Std_Time_Month_Ordinal_days(x_26, x_6); +x_28 = lean_int_dec_lt(x_27, x_7); +if (x_28 == 0) +{ +lean_dec(x_27); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_27); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +} +} +else +{ +uint8_t x_29; lean_object* x_30; uint8_t x_31; +x_29 = 1; +x_30 = l_Std_Time_Month_Ordinal_days(x_29, x_6); +x_31 = lean_int_dec_lt(x_30, x_7); +if (x_31 == 0) +{ +lean_dec(x_30); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +else +{ +lean_dec(x_7); +lean_ctor_set(x_4, 2, x_30); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_32 = lean_ctor_get(x_4, 1); +x_33 = lean_ctor_get(x_4, 2); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_4); +x_34 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_35 = lean_int_mod(x_2, x_34); +x_36 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_37 = lean_int_dec_eq(x_35, x_36); +lean_dec(x_35); +if (x_37 == 0) +{ +uint8_t x_38; lean_object* x_39; uint8_t x_40; +x_38 = 0; +x_39 = l_Std_Time_Month_Ordinal_days(x_38, x_32); +x_40 = lean_int_dec_lt(x_39, x_33); +if (x_40 == 0) +{ +lean_object* x_41; +lean_dec(x_39); +x_41 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_41, 0, x_2); +lean_ctor_set(x_41, 1, x_32); +lean_ctor_set(x_41, 2, x_33); +lean_ctor_set(x_1, 0, x_41); +return x_1; +} +else +{ +lean_object* x_42; +lean_dec(x_33); +x_42 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_42, 0, x_2); +lean_ctor_set(x_42, 1, x_32); +lean_ctor_set(x_42, 2, x_39); +lean_ctor_set(x_1, 0, x_42); +return x_1; +} +} +else +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; uint8_t x_46; +x_43 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_44 = lean_int_mod(x_2, x_43); +x_45 = lean_int_dec_eq(x_44, x_36); +lean_dec(x_44); +x_46 = l_instDecidableNot___rarg(x_45); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_47 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_48 = lean_int_mod(x_2, x_47); +x_49 = lean_int_dec_eq(x_48, x_36); +lean_dec(x_48); +if (x_49 == 0) +{ +uint8_t x_50; lean_object* x_51; uint8_t x_52; +x_50 = 0; +x_51 = l_Std_Time_Month_Ordinal_days(x_50, x_32); +x_52 = lean_int_dec_lt(x_51, x_33); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_53, 0, x_2); +lean_ctor_set(x_53, 1, x_32); +lean_ctor_set(x_53, 2, x_33); +lean_ctor_set(x_1, 0, x_53); +return x_1; +} +else +{ +lean_object* x_54; +lean_dec(x_33); +x_54 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_54, 0, x_2); +lean_ctor_set(x_54, 1, x_32); +lean_ctor_set(x_54, 2, x_51); +lean_ctor_set(x_1, 0, x_54); +return x_1; +} +} +else +{ +uint8_t x_55; lean_object* x_56; uint8_t x_57; +x_55 = 1; +x_56 = l_Std_Time_Month_Ordinal_days(x_55, x_32); +x_57 = lean_int_dec_lt(x_56, x_33); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_56); +x_58 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_58, 0, x_2); +lean_ctor_set(x_58, 1, x_32); +lean_ctor_set(x_58, 2, x_33); +lean_ctor_set(x_1, 0, x_58); +return x_1; +} +else +{ +lean_object* x_59; +lean_dec(x_33); +x_59 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_59, 0, x_2); +lean_ctor_set(x_59, 1, x_32); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_1, 0, x_59); +return x_1; +} +} +} +else +{ +uint8_t x_60; lean_object* x_61; uint8_t x_62; +x_60 = 1; +x_61 = l_Std_Time_Month_Ordinal_days(x_60, x_32); +x_62 = lean_int_dec_lt(x_61, x_33); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_61); +x_63 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_63, 0, x_2); +lean_ctor_set(x_63, 1, x_32); +lean_ctor_set(x_63, 2, x_33); +lean_ctor_set(x_1, 0, x_63); +return x_1; +} +else +{ +lean_object* x_64; +lean_dec(x_33); +x_64 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_64, 0, x_2); +lean_ctor_set(x_64, 1, x_32); +lean_ctor_set(x_64, 2, x_61); +lean_ctor_set(x_1, 0, x_64); +return x_1; +} +} +} +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_65 = lean_ctor_get(x_1, 0); +x_66 = lean_ctor_get(x_1, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_1); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +x_68 = lean_ctor_get(x_65, 2); +lean_inc(x_68); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + x_69 = x_65; +} else { + lean_dec_ref(x_65); + x_69 = lean_box(0); +} +x_70 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_71 = lean_int_mod(x_2, x_70); +x_72 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_73 = lean_int_dec_eq(x_71, x_72); +lean_dec(x_71); +if (x_73 == 0) +{ +uint8_t x_74; lean_object* x_75; uint8_t x_76; +x_74 = 0; +x_75 = l_Std_Time_Month_Ordinal_days(x_74, x_67); +x_76 = lean_int_dec_lt(x_75, x_68); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_75); +if (lean_is_scalar(x_69)) { + x_77 = lean_alloc_ctor(0, 3, 0); +} else { + x_77 = x_69; +} +lean_ctor_set(x_77, 0, x_2); +lean_ctor_set(x_77, 1, x_67); +lean_ctor_set(x_77, 2, x_68); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_66); +return x_78; +} +else +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_69; +} +lean_ctor_set(x_79, 0, x_2); +lean_ctor_set(x_79, 1, x_67); +lean_ctor_set(x_79, 2, x_75); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_66); +return x_80; +} +} +else +{ +lean_object* x_81; lean_object* x_82; uint8_t x_83; uint8_t x_84; +x_81 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_82 = lean_int_mod(x_2, x_81); +x_83 = lean_int_dec_eq(x_82, x_72); +lean_dec(x_82); +x_84 = l_instDecidableNot___rarg(x_83); +if (x_84 == 0) +{ +lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_85 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_86 = lean_int_mod(x_2, x_85); +x_87 = lean_int_dec_eq(x_86, x_72); +lean_dec(x_86); +if (x_87 == 0) +{ +uint8_t x_88; lean_object* x_89; uint8_t x_90; +x_88 = 0; +x_89 = l_Std_Time_Month_Ordinal_days(x_88, x_67); +x_90 = lean_int_dec_lt(x_89, x_68); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_89); +if (lean_is_scalar(x_69)) { + x_91 = lean_alloc_ctor(0, 3, 0); +} else { + x_91 = x_69; +} +lean_ctor_set(x_91, 0, x_2); +lean_ctor_set(x_91, 1, x_67); +lean_ctor_set(x_91, 2, x_68); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_66); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_93 = lean_alloc_ctor(0, 3, 0); +} else { + x_93 = x_69; +} +lean_ctor_set(x_93, 0, x_2); +lean_ctor_set(x_93, 1, x_67); +lean_ctor_set(x_93, 2, x_89); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_66); +return x_94; +} +} +else +{ +uint8_t x_95; lean_object* x_96; uint8_t x_97; +x_95 = 1; +x_96 = l_Std_Time_Month_Ordinal_days(x_95, x_67); +x_97 = lean_int_dec_lt(x_96, x_68); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; +lean_dec(x_96); +if (lean_is_scalar(x_69)) { + x_98 = lean_alloc_ctor(0, 3, 0); +} else { + x_98 = x_69; +} +lean_ctor_set(x_98, 0, x_2); +lean_ctor_set(x_98, 1, x_67); +lean_ctor_set(x_98, 2, x_68); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_66); +return x_99; +} +else +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_100 = lean_alloc_ctor(0, 3, 0); +} else { + x_100 = x_69; +} +lean_ctor_set(x_100, 0, x_2); +lean_ctor_set(x_100, 1, x_67); +lean_ctor_set(x_100, 2, x_96); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_66); +return x_101; +} +} +} +else +{ +uint8_t x_102; lean_object* x_103; uint8_t x_104; +x_102 = 1; +x_103 = l_Std_Time_Month_Ordinal_days(x_102, x_67); +x_104 = lean_int_dec_lt(x_103, x_68); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_103); +if (lean_is_scalar(x_69)) { + x_105 = lean_alloc_ctor(0, 3, 0); +} else { + x_105 = x_69; +} +lean_ctor_set(x_105, 0, x_2); +lean_ctor_set(x_105, 1, x_67); +lean_ctor_set(x_105, 2, x_68); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_66); +return x_106; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_dec(x_68); +if (lean_is_scalar(x_69)) { + x_107 = lean_alloc_ctor(0, 3, 0); +} else { + x_107 = x_69; +} +lean_ctor_set(x_107, 0, x_2); +lean_ctor_set(x_107, 1, x_67); +lean_ctor_set(x_107, 2, x_103); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_66); +return x_108; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withYearRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 2); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_rollOver(x_2, x_5, x_6); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 2); +lean_inc(x_11); +lean_dec(x_8); +x_12 = l_Std_Time_PlainDate_rollOver(x_2, x_10, x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_4, 0); +lean_dec(x_6); +lean_ctor_set(x_4, 0, x_2); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_ctor_get(x_4, 2); +x_9 = lean_ctor_get(x_4, 3); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_10, 1, x_7); +lean_ctor_set(x_10, 2, x_8); +lean_ctor_set(x_10, 3, x_9); +lean_ctor_set(x_1, 1, x_10); +return x_1; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 3); +lean_inc(x_15); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + lean_ctor_release(x_11, 1); + lean_ctor_release(x_11, 2); + lean_ctor_release(x_11, 3); + x_16 = x_11; +} else { + lean_dec_ref(x_11); + x_16 = lean_box(0); +} +if (lean_is_scalar(x_16)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_16; +} +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_13); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_4, 1); +lean_dec(x_6); +lean_ctor_set(x_4, 1, x_2); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 2); +x_9 = lean_ctor_get(x_4, 3); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_2); +lean_ctor_set(x_10, 2, x_8); +lean_ctor_set(x_10, 3, x_9); +lean_ctor_set(x_1, 1, x_10); +return x_1; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 3); +lean_inc(x_15); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + lean_ctor_release(x_11, 1); + lean_ctor_release(x_11, 2); + lean_ctor_release(x_11, 3); + x_16 = x_11; +} else { + lean_dec_ref(x_11); + x_16 = lean_box(0); +} +if (lean_is_scalar(x_16)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_16; +} +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_4, 2); +lean_dec(x_6); +lean_ctor_set(x_4, 2, x_2); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 1); +x_9 = lean_ctor_get(x_4, 3); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_2); +lean_ctor_set(x_10, 3, x_9); +lean_ctor_set(x_1, 1, x_10); +return x_1; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 3); +lean_inc(x_15); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + lean_ctor_release(x_11, 1); + lean_ctor_release(x_11, 2); + lean_ctor_release(x_11, 3); + x_16 = x_11; +} else { + lean_dec_ref(x_11); + x_16 = lean_box(0); +} +if (lean_is_scalar(x_16)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_16; +} +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_14); +lean_ctor_set(x_17, 2, x_2); +lean_ctor_set(x_17, 3, x_15); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_withMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_withMilliseconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_ctor_get(x_4, 3); +x_7 = l_Std_Time_PlainDateTime_withMilliseconds___closed__1; +x_8 = lean_int_emod(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_PlainDateTime_withMilliseconds___closed__2; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_add(x_10, x_8); +lean_dec(x_8); +lean_dec(x_10); +lean_ctor_set(x_4, 3, x_11); +return x_1; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_12 = lean_ctor_get(x_4, 0); +x_13 = lean_ctor_get(x_4, 1); +x_14 = lean_ctor_get(x_4, 2); +x_15 = lean_ctor_get(x_4, 3); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_4); +x_16 = l_Std_Time_PlainDateTime_withMilliseconds___closed__1; +x_17 = lean_int_emod(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainDateTime_withMilliseconds___closed__2; +x_19 = lean_int_mul(x_2, x_18); +x_20 = lean_int_add(x_19, x_17); +lean_dec(x_17); +lean_dec(x_19); +x_21 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_21, 0, x_12); +lean_ctor_set(x_21, 1, x_13); +lean_ctor_set(x_21, 2, x_14); +lean_ctor_set(x_21, 3, x_20); +lean_ctor_set(x_1, 1, x_21); +return x_1; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_22 = lean_ctor_get(x_1, 1); +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_inc(x_23); +lean_dec(x_1); +x_24 = lean_ctor_get(x_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +x_26 = lean_ctor_get(x_22, 2); +lean_inc(x_26); +x_27 = lean_ctor_get(x_22, 3); +lean_inc(x_27); +if (lean_is_exclusive(x_22)) { + lean_ctor_release(x_22, 0); + lean_ctor_release(x_22, 1); + lean_ctor_release(x_22, 2); + lean_ctor_release(x_22, 3); + x_28 = x_22; +} else { + lean_dec_ref(x_22); + x_28 = lean_box(0); +} +x_29 = l_Std_Time_PlainDateTime_withMilliseconds___closed__1; +x_30 = lean_int_emod(x_27, x_29); +lean_dec(x_27); +x_31 = l_Std_Time_PlainDateTime_withMilliseconds___closed__2; +x_32 = lean_int_mul(x_2, x_31); +x_33 = lean_int_add(x_32, x_30); +lean_dec(x_30); +lean_dec(x_32); +if (lean_is_scalar(x_28)) { + x_34 = lean_alloc_ctor(0, 4, 0); +} else { + x_34 = x_28; +} +lean_ctor_set(x_34, 0, x_24); +lean_ctor_set(x_34, 1, x_25); +lean_ctor_set(x_34, 2, x_26); +lean_ctor_set(x_34, 3, x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_23); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_withMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_withNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_1, 1); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_4, 3); +lean_dec(x_6); +lean_ctor_set(x_4, 3, x_2); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_ctor_get(x_4, 1); +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_4); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_9); +lean_ctor_set(x_10, 3, x_2); +lean_ctor_set(x_1, 1, x_10); +return x_1; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_11, 2); +lean_inc(x_15); +if (lean_is_exclusive(x_11)) { + lean_ctor_release(x_11, 0); + lean_ctor_release(x_11, 1); + lean_ctor_release(x_11, 2); + lean_ctor_release(x_11, 3); + x_16 = x_11; +} else { + lean_dec_ref(x_11); + x_16 = lean_box(0); +} +if (lean_is_scalar(x_16)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_16; +} +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_14); +lean_ctor_set(x_17, 2, x_15); +lean_ctor_set(x_17, 3, x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +x_6 = lean_int_add(x_5, x_2); +lean_dec(x_5); +x_7 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_6); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_11 = lean_int_add(x_10, x_2); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_int_neg(x_2); +x_6 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +x_7 = lean_int_add(x_6, x_5); +lean_dec(x_5); +lean_dec(x_6); +x_8 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_7); +lean_dec(x_7); +lean_ctor_set(x_1, 0, x_8); +return x_1; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = lean_int_neg(x_2); +x_12 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_13 = lean_int_add(x_12, x_11); +lean_dec(x_11); +lean_dec(x_12); +x_14 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_addWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +x_6 = l_Std_Time_PlainDateTime_addWeeks___closed__1; +x_7 = lean_int_mul(x_2, x_6); +x_8 = lean_int_add(x_5, x_7); +lean_dec(x_7); +lean_dec(x_5); +x_9 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_8); +lean_dec(x_8); +lean_ctor_set(x_1, 0, x_9); +return x_1; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_1); +x_12 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_10); +x_13 = l_Std_Time_PlainDateTime_addWeeks___closed__1; +x_14 = lean_int_mul(x_2, x_13); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_15); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_11); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_int_neg(x_2); +x_6 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +x_7 = l_Std_Time_PlainDateTime_addWeeks___closed__1; +x_8 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_9 = lean_int_add(x_6, x_8); +lean_dec(x_8); +lean_dec(x_6); +x_10 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_9); +lean_dec(x_9); +lean_ctor_set(x_1, 0, x_10); +return x_1; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_ctor_get(x_1, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_1); +x_13 = lean_int_neg(x_2); +x_14 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_15 = l_Std_Time_PlainDateTime_addWeeks___closed__1; +x_16 = lean_int_mul(x_13, x_15); +lean_dec(x_13); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_12); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDate_addMonthsClip(x_4, x_2); +lean_ctor_set(x_1, 0, x_5); +return x_1; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_8 = l_Std_Time_PlainDate_addMonthsClip(x_6, x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_int_neg(x_2); +x_6 = l_Std_Time_PlainDate_addMonthsClip(x_4, x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 0, x_6); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_addMonthsClip(x_7, x_9); +lean_dec(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDate_addMonthsRollOver(x_4, x_2); +lean_ctor_set(x_1, 0, x_5); +return x_1; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_1); +x_8 = l_Std_Time_PlainDate_addMonthsRollOver(x_6, x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_int_neg(x_2); +x_6 = l_Std_Time_PlainDate_addMonthsRollOver(x_4, x_5); +lean_dec(x_5); +lean_ctor_set(x_1, 0, x_6); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_addMonthsRollOver(x_7, x_9); +lean_dec(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_addYearsRollOver___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_6 = lean_int_mul(x_2, x_5); +x_7 = l_Std_Time_PlainDate_addMonthsRollOver(x_4, x_6); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_11 = lean_int_mul(x_2, x_10); +x_12 = l_Std_Time_PlainDate_addMonthsRollOver(x_8, x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_6 = lean_int_mul(x_2, x_5); +x_7 = l_Std_Time_PlainDate_addMonthsClip(x_4, x_6); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_10 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_11 = lean_int_mul(x_2, x_10); +x_12 = l_Std_Time_PlainDate_addMonthsClip(x_8, x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_9); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_6 = lean_int_mul(x_2, x_5); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = l_Std_Time_PlainDate_addMonthsRollOver(x_4, x_7); +lean_dec(x_7); +lean_ctor_set(x_1, 0, x_8); +return x_1; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_12 = lean_int_mul(x_2, x_11); +x_13 = lean_int_neg(x_12); +lean_dec(x_12); +x_14 = l_Std_Time_PlainDate_addMonthsRollOver(x_9, x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_6 = lean_int_mul(x_2, x_5); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = l_Std_Time_PlainDate_addMonthsClip(x_4, x_7); +lean_dec(x_7); +lean_ctor_set(x_1, 0, x_8); +return x_1; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Std_Time_PlainDateTime_addYearsRollOver___closed__1; +x_12 = lean_int_mul(x_2, x_11); +x_13 = lean_int_neg(x_12); +lean_dec(x_12); +x_14 = l_Std_Time_PlainDate_addMonthsClip(x_9, x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = l_Std_Time_PlainTime_toSeconds(x_3); +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2; +x_6 = lean_int_mul(x_2, x_5); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_3); +lean_dec(x_3); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_12 = lean_int_mul(x_6, x_11); +lean_dec(x_6); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_9); +lean_dec(x_9); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addHours(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = l_Std_Time_PlainTime_toSeconds(x_4); +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2; +x_7 = lean_int_mul(x_3, x_6); +lean_dec(x_3); +x_8 = lean_int_add(x_5, x_7); +lean_dec(x_5); +x_9 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_4); +lean_dec(x_4); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_13 = lean_int_mul(x_7, x_12); +lean_dec(x_7); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subHours(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = l_Std_Time_PlainTime_toSeconds(x_3); +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1; +x_6 = lean_int_mul(x_2, x_5); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_3); +lean_dec(x_3); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_12 = lean_int_mul(x_6, x_11); +lean_dec(x_6); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_9); +lean_dec(x_9); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addMinutes(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = l_Std_Time_PlainTime_toSeconds(x_4); +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1; +x_7 = lean_int_mul(x_3, x_6); +lean_dec(x_3); +x_8 = lean_int_add(x_5, x_7); +lean_dec(x_5); +x_9 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_4); +lean_dec(x_4); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_13 = lean_int_mul(x_7, x_12); +lean_dec(x_7); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subMinutes(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = l_Std_Time_PlainTime_toSeconds(x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_7 = lean_int_ediv(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_PlainTime_toNanoseconds(x_3); +lean_dec(x_3); +x_9 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_10); +lean_dec(x_8); +x_12 = l_Std_Time_PlainTime_ofNanoseconds(x_11); +lean_dec(x_11); +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_1); +x_14 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_15 = lean_int_add(x_14, x_7); +lean_dec(x_7); +lean_dec(x_14); +x_16 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_15); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_12); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addSeconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = l_Std_Time_PlainTime_toSeconds(x_4); +x_6 = lean_int_add(x_5, x_3); +lean_dec(x_5); +x_7 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1; +x_8 = lean_int_ediv(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_PlainTime_toNanoseconds(x_4); +lean_dec(x_4); +x_10 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_11 = lean_int_mul(x_3, x_10); +lean_dec(x_3); +x_12 = lean_int_add(x_9, x_11); +lean_dec(x_11); +lean_dec(x_9); +x_13 = l_Std_Time_PlainTime_ofNanoseconds(x_12); +lean_dec(x_12); +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +lean_dec(x_1); +x_15 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +x_16 = lean_int_add(x_15, x_8); +lean_dec(x_8); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_16); +lean_dec(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subSeconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_addMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = l_Std_Time_PlainTime_toMilliseconds(x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDateTime_addMilliseconds___closed__1; +x_7 = lean_int_ediv(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_PlainTime_addMilliseconds(x_3, x_2); +lean_dec(x_3); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_11 = lean_int_add(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_8); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = l_Std_Time_PlainTime_toMilliseconds(x_4); +x_6 = lean_int_add(x_5, x_3); +lean_dec(x_5); +x_7 = l_Std_Time_PlainDateTime_addMilliseconds___closed__1; +x_8 = lean_int_ediv(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_PlainTime_addMilliseconds(x_4, x_3); +lean_dec(x_3); +lean_dec(x_4); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_10); +x_12 = lean_int_add(x_11, x_8); +lean_dec(x_8); +lean_dec(x_11); +x_13 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_12); +lean_dec(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_9); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_4, 3); +lean_inc(x_5); +x_6 = lean_int_add(x_5, x_2); +lean_dec(x_5); +x_7 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_8 = lean_int_ediv(x_6, x_7); +x_9 = lean_int_emod(x_6, x_7); +lean_dec(x_6); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_4); +lean_dec(x_4); +x_11 = lean_int_mul(x_8, x_7); +lean_dec(x_8); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_ofNanoseconds(x_12); +lean_dec(x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 3); +lean_dec(x_15); +lean_ctor_set(x_13, 3, x_9); +lean_ctor_set(x_1, 1, x_13); +return x_1; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_13, 0); +x_17 = lean_ctor_get(x_13, 1); +x_18 = lean_ctor_get(x_13, 2); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_13); +x_19 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_17); +lean_ctor_set(x_19, 2, x_18); +lean_ctor_set(x_19, 3, x_9); +lean_ctor_set(x_1, 1, x_19); +return x_1; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_ctor_get(x_1, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_1); +x_22 = lean_ctor_get(x_21, 3); +lean_inc(x_22); +x_23 = lean_int_add(x_22, x_2); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_25 = lean_int_ediv(x_23, x_24); +x_26 = lean_int_emod(x_23, x_24); +lean_dec(x_23); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_21); +lean_dec(x_21); +x_28 = lean_int_mul(x_25, x_24); +lean_dec(x_25); +x_29 = lean_int_add(x_27, x_28); +lean_dec(x_28); +lean_dec(x_27); +x_30 = l_Std_Time_PlainTime_ofNanoseconds(x_29); +lean_dec(x_29); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +x_33 = lean_ctor_get(x_30, 2); +lean_inc(x_33); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + lean_ctor_release(x_30, 2); + lean_ctor_release(x_30, 3); + x_34 = x_30; +} else { + lean_dec_ref(x_30); + x_34 = lean_box(0); +} +if (lean_is_scalar(x_34)) { + x_35 = lean_alloc_ctor(0, 4, 0); +} else { + x_35 = x_34; +} +lean_ctor_set(x_35, 0, x_31); +lean_ctor_set(x_35, 1, x_32); +lean_ctor_set(x_35, 2, x_33); +lean_ctor_set(x_35, 3, x_26); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_20); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_addNanoseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_int_neg(x_2); +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_5, 3); +lean_inc(x_6); +x_7 = lean_int_add(x_6, x_3); +lean_dec(x_3); +lean_dec(x_6); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_9 = lean_int_ediv(x_7, x_8); +x_10 = lean_int_emod(x_7, x_8); +lean_dec(x_7); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_5); +lean_dec(x_5); +x_12 = lean_int_mul(x_9, x_8); +lean_dec(x_9); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_14, 3); +lean_dec(x_16); +lean_ctor_set(x_14, 3, x_10); +lean_ctor_set(x_1, 1, x_14); +return x_1; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_14, 0); +x_18 = lean_ctor_get(x_14, 1); +x_19 = lean_ctor_get(x_14, 2); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_14); +x_20 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_20, 0, x_17); +lean_ctor_set(x_20, 1, x_18); +lean_ctor_set(x_20, 2, x_19); +lean_ctor_set(x_20, 3, x_10); +lean_ctor_set(x_1, 1, x_20); +return x_1; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_ctor_get(x_1, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_1); +x_23 = lean_ctor_get(x_22, 3); +lean_inc(x_23); +x_24 = lean_int_add(x_23, x_3); +lean_dec(x_3); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_26 = lean_int_ediv(x_24, x_25); +x_27 = lean_int_emod(x_24, x_25); +lean_dec(x_24); +x_28 = l_Std_Time_PlainTime_toNanoseconds(x_22); +lean_dec(x_22); +x_29 = lean_int_mul(x_26, x_25); +lean_dec(x_26); +x_30 = lean_int_add(x_28, x_29); +lean_dec(x_29); +lean_dec(x_28); +x_31 = l_Std_Time_PlainTime_ofNanoseconds(x_30); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 2); +lean_inc(x_34); +if (lean_is_exclusive(x_31)) { + lean_ctor_release(x_31, 0); + lean_ctor_release(x_31, 1); + lean_ctor_release(x_31, 2); + lean_ctor_release(x_31, 3); + x_35 = x_31; +} else { + lean_dec_ref(x_31); + x_35 = lean_box(0); +} +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(0, 4, 0); +} else { + x_36 = x_35; +} +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_33); +lean_ctor_set(x_36, 2, x_34); +lean_ctor_set(x_36, 3, x_27); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_21); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_subNanoseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_year(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_year___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_year(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_month(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_month___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_month(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_day(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 2); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_day___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_day(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_weekday(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l_Std_Time_PlainDate_weekday(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekday___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDateTime_weekday(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_hour___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_hour(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_minute(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 1); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_minute___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_minute(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_millisecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 3); +x_4 = l_Std_Time_PlainDateTime_withMilliseconds___closed__2; +x_5 = lean_int_ediv(x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_millisecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_millisecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_second(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_second___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_second(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_nanosecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 3); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_nanosecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_nanosecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_era(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_Year_Offset_era(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_era___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDateTime_era(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_PlainDateTime_inLeapYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_5 = lean_int_mod(x_3, x_4); +x_6 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; +x_9 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_10 = lean_int_mod(x_3, x_9); +x_11 = lean_int_dec_eq(x_10, x_6); +lean_dec(x_10); +x_12 = l_instDecidableNot___rarg(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_14 = lean_int_mod(x_3, x_13); +x_15 = lean_int_dec_eq(x_14, x_6); +lean_dec(x_14); +return x_15; +} +else +{ +uint8_t x_16; +x_16 = 1; +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_inLeapYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_PlainDateTime_inLeapYear(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l_Std_Time_PlainDate_weekOfYear(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainDate_weekOfMonth(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_weekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_weekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_alignedWeekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +lean_dec(x_1); +x_3 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_dayOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11; +x_5 = lean_int_mod(x_3, x_4); +x_6 = l_Std_Time_instInhabitedPlainDateTime___closed__1; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_ctor_get(x_2, 2); +lean_inc(x_9); +lean_inc(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +if (x_7 == 0) +{ +uint8_t x_11; lean_object* x_12; +x_11 = 0; +x_12 = l_Std_Time_ValidDate_dayOfYear(x_11, x_10); +lean_dec(x_10); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; +x_13 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7; +x_14 = lean_int_mod(x_3, x_13); +x_15 = lean_int_dec_eq(x_14, x_6); +lean_dec(x_14); +x_16 = l_instDecidableNot___rarg(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3; +x_18 = lean_int_mod(x_3, x_17); +x_19 = lean_int_dec_eq(x_18, x_6); +lean_dec(x_18); +if (x_19 == 0) +{ +uint8_t x_20; lean_object* x_21; +x_20 = 0; +x_21 = l_Std_Time_ValidDate_dayOfYear(x_20, x_10); +lean_dec(x_10); +return x_21; +} +else +{ +uint8_t x_22; lean_object* x_23; +x_22 = 1; +x_23 = l_Std_Time_ValidDate_dayOfYear(x_22, x_10); +lean_dec(x_10); +return x_23; +} +} +else +{ +uint8_t x_24; lean_object* x_25; +x_24 = 1; +x_25 = l_Std_Time_ValidDate_dayOfYear(x_24, x_10); +lean_dec(x_10); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_dayOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_dayOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_quarter(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainDate_quarter(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_quarter___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainDateTime_quarter(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_atTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_atDate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_addNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHAddOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_subNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instHSubOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_ctor_get(x_9, 3); +lean_inc(x_10); +x_11 = lean_int_add(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = lean_int_ediv(x_11, x_4); +x_13 = lean_int_emod(x_11, x_4); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_15 = lean_int_mul(x_12, x_4); +lean_dec(x_12); +x_16 = lean_int_add(x_14, x_15); +lean_dec(x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 3); +lean_dec(x_19); +lean_ctor_set(x_17, 3, x_13); +lean_ctor_set(x_1, 1, x_17); +return x_1; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_17, 0); +x_21 = lean_ctor_get(x_17, 1); +x_22 = lean_ctor_get(x_17, 2); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_17); +x_23 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_22); +lean_ctor_set(x_23, 3, x_13); +lean_ctor_set(x_1, 1, x_23); +return x_1; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_ctor_get(x_1, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_1); +x_26 = lean_ctor_get(x_25, 3); +lean_inc(x_26); +x_27 = lean_int_add(x_26, x_7); +lean_dec(x_7); +lean_dec(x_26); +x_28 = lean_int_ediv(x_27, x_4); +x_29 = lean_int_emod(x_27, x_4); +lean_dec(x_27); +x_30 = l_Std_Time_PlainTime_toNanoseconds(x_25); +lean_dec(x_25); +x_31 = lean_int_mul(x_28, x_4); +lean_dec(x_28); +x_32 = lean_int_add(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +x_33 = l_Std_Time_PlainTime_ofNanoseconds(x_32); +lean_dec(x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +x_36 = lean_ctor_get(x_33, 2); +lean_inc(x_36); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + lean_ctor_release(x_33, 2); + lean_ctor_release(x_33, 3); + x_37 = x_33; +} else { + lean_dec_ref(x_33); + x_37 = lean_box(0); +} +if (lean_is_scalar(x_37)) { + x_38 = lean_alloc_ctor(0, 4, 0); +} else { + x_38 = x_37; +} +lean_ctor_set(x_38, 0, x_34); +lean_ctor_set(x_38, 1, x_35); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_29); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_24); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instHAddDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_instHAddDuration(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_atTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_atDate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime_Timestamp(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_DateTime_PlainDateTime(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime_Timestamp(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_instInhabitedPlainDateTime___closed__1 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__1); +l_Std_Time_instInhabitedPlainDateTime___closed__2 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__2); +l_Std_Time_instInhabitedPlainDateTime___closed__3 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__3); +l_Std_Time_instInhabitedPlainDateTime___closed__4 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__4); +l_Std_Time_instInhabitedPlainDateTime___closed__5 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__5); +l_Std_Time_instInhabitedPlainDateTime___closed__6 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__6); +l_Std_Time_instInhabitedPlainDateTime___closed__7 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__7); +l_Std_Time_instInhabitedPlainDateTime___closed__8 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__8); +l_Std_Time_instInhabitedPlainDateTime___closed__9 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__9); +l_Std_Time_instInhabitedPlainDateTime___closed__10 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__10(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__10); +l_Std_Time_instInhabitedPlainDateTime___closed__11 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__11(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__11); +l_Std_Time_instInhabitedPlainDateTime___closed__12 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__12(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__12); +l_Std_Time_instInhabitedPlainDateTime___closed__13 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__13(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__13); +l_Std_Time_instInhabitedPlainDateTime___closed__14 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__14(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__14); +l_Std_Time_instInhabitedPlainDateTime___closed__15 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__15(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__15); +l_Std_Time_instInhabitedPlainDateTime___closed__16 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__16(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__16); +l_Std_Time_instInhabitedPlainDateTime___closed__17 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__17(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__17); +l_Std_Time_instInhabitedPlainDateTime___closed__18 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__18(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__18); +l_Std_Time_instInhabitedPlainDateTime___closed__19 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__19(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__19); +l_Std_Time_instInhabitedPlainDateTime___closed__20 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__20(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__20); +l_Std_Time_instInhabitedPlainDateTime___closed__21 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__21(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__21); +l_Std_Time_instInhabitedPlainDateTime___closed__22 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__22(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__22); +l_Std_Time_instInhabitedPlainDateTime___closed__23 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__23(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__23); +l_Std_Time_instInhabitedPlainDateTime___closed__24 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__24(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__24); +l_Std_Time_instInhabitedPlainDateTime___closed__25 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__25(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__25); +l_Std_Time_instInhabitedPlainDateTime___closed__26 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__26(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__26); +l_Std_Time_instInhabitedPlainDateTime___closed__27 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__27(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__27); +l_Std_Time_instInhabitedPlainDateTime___closed__28 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__28(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__28); +l_Std_Time_instInhabitedPlainDateTime___closed__29 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__29(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__29); +l_Std_Time_instInhabitedPlainDateTime___closed__30 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__30(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__30); +l_Std_Time_instInhabitedPlainDateTime___closed__31 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__31(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__31); +l_Std_Time_instInhabitedPlainDateTime___closed__32 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__32(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__32); +l_Std_Time_instInhabitedPlainDateTime___closed__33 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__33(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__33); +l_Std_Time_instInhabitedPlainDateTime___closed__34 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__34(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__34); +l_Std_Time_instInhabitedPlainDateTime___closed__35 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__35(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__35); +l_Std_Time_instInhabitedPlainDateTime___closed__36 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__36(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__36); +l_Std_Time_instInhabitedPlainDateTime___closed__37 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__37(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__37); +l_Std_Time_instInhabitedPlainDateTime___closed__38 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__38(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__38); +l_Std_Time_instInhabitedPlainDateTime___closed__39 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__39(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__39); +l_Std_Time_instInhabitedPlainDateTime___closed__40 = _init_l_Std_Time_instInhabitedPlainDateTime___closed__40(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime___closed__40); +l_Std_Time_instInhabitedPlainDateTime = _init_l_Std_Time_instInhabitedPlainDateTime(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainDateTime); +l_Std_Time_instBEqPlainDateTime___closed__1 = _init_l_Std_Time_instBEqPlainDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_instBEqPlainDateTime___closed__1); +l_Std_Time_instBEqPlainDateTime = _init_l_Std_Time_instBEqPlainDateTime(); +lean_mark_persistent(l_Std_Time_instBEqPlainDateTime); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__1); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__2); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__3); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__4); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__5); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__6); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__7); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__8); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__9); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__10); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__11); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__12); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__13); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__14); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__15); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__16); +l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17 = _init_l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17(); +lean_mark_persistent(l___private_Std_Time_DateTime_PlainDateTime_0__Std_Time_reprPlainDateTime____x40_Std_Time_DateTime_PlainDateTime___hyg_111____closed__17); +l_Std_Time_instReprPlainDateTime___closed__1 = _init_l_Std_Time_instReprPlainDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_instReprPlainDateTime___closed__1); +l_Std_Time_instReprPlainDateTime = _init_l_Std_Time_instReprPlainDateTime(); +lean_mark_persistent(l_Std_Time_instReprPlainDateTime); +l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1 = _init_l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__1); +l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2 = _init_l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_toTimestampAssumingUTC___closed__2); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__1); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__1___closed__2); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__1); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__2); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__3); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__4); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__5); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__6); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__7); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__8); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__9); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__10); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__11); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__12); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__13); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__14); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__2___closed__15); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___lambda__4___closed__1); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__1); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__2); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__3); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__4); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__5); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__6); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__7); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__8); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__9); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__10); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__11); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__12); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__13); +l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14 = _init_l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14(); +lean_mark_persistent(l_Std_Time_PlainDateTime_ofTimestampAssumingUTC___closed__14); +l_Std_Time_PlainDateTime_withMilliseconds___closed__1 = _init_l_Std_Time_PlainDateTime_withMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_withMilliseconds___closed__1); +l_Std_Time_PlainDateTime_withMilliseconds___closed__2 = _init_l_Std_Time_PlainDateTime_withMilliseconds___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_withMilliseconds___closed__2); +l_Std_Time_PlainDateTime_addWeeks___closed__1 = _init_l_Std_Time_PlainDateTime_addWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_addWeeks___closed__1); +l_Std_Time_PlainDateTime_addYearsRollOver___closed__1 = _init_l_Std_Time_PlainDateTime_addYearsRollOver___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_addYearsRollOver___closed__1); +l_Std_Time_PlainDateTime_addMilliseconds___closed__1 = _init_l_Std_Time_PlainDateTime_addMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_addMilliseconds___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset = _init_l_Std_Time_PlainDateTime_instHAddOffset(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset); +l_Std_Time_PlainDateTime_instHSubOffset___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset = _init_l_Std_Time_PlainDateTime_instHSubOffset(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset); +l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__1___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__1); +l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__1___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__1); +l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__2___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__2 = _init_l_Std_Time_PlainDateTime_instHAddOffset__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__2); +l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__2___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset__2 = _init_l_Std_Time_PlainDateTime_instHSubOffset__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__2); +l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__3___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__3 = _init_l_Std_Time_PlainDateTime_instHAddOffset__3(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__3); +l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__3___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset__3 = _init_l_Std_Time_PlainDateTime_instHSubOffset__3(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__3); +l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__4___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__4 = _init_l_Std_Time_PlainDateTime_instHAddOffset__4(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__4); +l_Std_Time_PlainDateTime_instHSubOffset__4 = _init_l_Std_Time_PlainDateTime_instHSubOffset__4(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__4); +l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__5___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__5 = _init_l_Std_Time_PlainDateTime_instHAddOffset__5(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__5); +l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__5___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset__5 = _init_l_Std_Time_PlainDateTime_instHSubOffset__5(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__5); +l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1 = _init_l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__6___closed__1); +l_Std_Time_PlainDateTime_instHAddOffset__6 = _init_l_Std_Time_PlainDateTime_instHAddOffset__6(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHAddOffset__6); +l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1 = _init_l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__6___closed__1); +l_Std_Time_PlainDateTime_instHSubOffset__6 = _init_l_Std_Time_PlainDateTime_instHSubOffset__6(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instHSubOffset__6); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/DateTime/Timestamp.c b/stage0/stdlib/Std/Time/DateTime/Timestamp.c new file mode 100644 index 000000000000..1ee5561d7dcd --- /dev/null +++ b/stage0/stdlib/Std/Time/DateTime/Timestamp.c @@ -0,0 +1,2260 @@ +// Lean compiler output +// Module: Std.Time.DateTime.Timestamp +// Imports: Std.Time.Internal Init.Data.Int Std.Time.Time Std.Time.Date Std.Time.Duration +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMinutes___boxed(lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subSeconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addHours(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHSubOffset__5___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMilliseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHSubOffset__1___closed__1; +static lean_object* l_Std_Time_Timestamp_toMinutes___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDurationSinceUnixEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofNanosecondsSinceUnixEpoch___boxed(lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMillisecondsSinceUnixEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqTimestamp; +lean_object* l_String_quote(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addSeconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprTimestamp__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset__6___closed__1; +static lean_object* l_Std_Time_Timestamp_instHSubOffset__4___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instReprTimestamp__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subNanoseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedTimestamp; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubDuration; +static lean_object* l_Std_Time_instInhabitedTimestamp___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +static lean_object* l_Std_Time_instBEqTimestamp___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDuration(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instDecidableLeTimestamp___boxed(lean_object*, lean_object*); +uint8_t l_Std_Time_Duration_instDecidableLe(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addSeconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofMillisecondsSinceUnixEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubDuration__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_addHours___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_instDecidableLeTimestamp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toSecondsSinceUnixEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofSecondsSinceUnixEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDays___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addWeeks(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subSeconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset___closed__1; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Std_Time_instToStringTimestamp___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instToStringTimestamp___boxed(lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDuration___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instOfNatTimestamp(lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddDuration___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDuration___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__6; +static lean_object* l_Std_Time_Timestamp_instHSubOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDays(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__5; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subWeeks___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11; +static lean_object* l_Std_Time_Timestamp_instHSubDuration___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofNanosecondsSinceUnixEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMillisecondsSinceUnixEpoch(lean_object*); +lean_object* lean_int_div(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDays___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toSecondsSinceUnixEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddDuration; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDays(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset__4___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_now___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_since(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instLETimestamp; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofMillisecondsSinceUnixEpoch(lean_object*); +lean_object* lean_nat_abs(lean_object*); +static lean_object* l_Std_Time_Timestamp_instHSubOffset__2___closed__1; +lean_object* lean_int_mul(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___boxed(lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__6; +static lean_object* l_Std_Time_Timestamp_instHSubOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_since___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_toDays___closed__1; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__4; +LEAN_EXPORT uint8_t l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__1; +LEAN_EXPORT lean_object* l_Std_Time_instReprTimestamp; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subHours___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_instToStringDuration_leftPad(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMilliseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMilliseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subWeeks(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6; +static lean_object* l_Std_Time_Timestamp_instHAddOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset__2; +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13; +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDays(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__2; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDurationSinceUnixEpoch(lean_object*); +lean_object* lean_get_current_time(lean_object*); +static lean_object* l_Std_Time_instToStringTimestamp___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addNanoseconds(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9; +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubDuration__1(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instToStringTimestamp(lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addHours___boxed(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMinutes(lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__4; +static lean_object* l_Std_Time_Timestamp_instHSubOffset__6___closed__1; +static lean_object* l_Std_Time_instReprTimestamp___closed__1; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubOffset__5; +static lean_object* l_Std_Time_Timestamp_subSeconds___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18_(lean_object*, lean_object*); +static lean_object* l_Std_Time_Timestamp_instHAddOffset__5___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDuration(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHAddOffset; +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("val", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3; +x_2 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("s", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_22; lean_object* x_23; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_6 = lean_int_dec_lt(x_5, x_3); +if (x_6 == 0) +{ +uint8_t x_107; +x_107 = lean_int_dec_lt(x_3, x_5); +if (x_107 == 0) +{ +uint8_t x_108; +x_108 = lean_int_dec_lt(x_4, x_5); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; +lean_inc(x_4); +lean_inc(x_3); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_3); +lean_ctor_set(x_109, 1, x_4); +x_110 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18; +x_22 = x_110; +x_23 = x_109; +goto block_106; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_111 = lean_int_neg(x_3); +x_112 = lean_int_neg(x_4); +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_22 = x_114; +x_23 = x_113; +goto block_106; +} +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_115 = lean_int_neg(x_3); +x_116 = lean_int_neg(x_4); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +x_118 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_22 = x_118; +x_23 = x_117; +goto block_106; +} +} +else +{ +lean_object* x_119; lean_object* x_120; +lean_inc(x_4); +lean_inc(x_3); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_3); +lean_ctor_set(x_119, 1, x_4); +x_120 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18; +x_22 = x_120; +x_23 = x_119; +goto block_106; +} +block_21: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_8 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7; +x_9 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +x_10 = 0; +x_11 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*1, x_10); +x_12 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11; +x_19 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_10); +return x_20; +} +block_106: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_1, 1); +x_27 = lean_int_dec_eq(x_26, x_5); +x_28 = lean_int_dec_lt(x_24, x_5); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_nat_abs(x_24); +lean_dec(x_24); +x_30 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +x_31 = lean_string_append(x_22, x_30); +lean_dec(x_30); +if (x_27 == 0) +{ +uint8_t x_32; +x_32 = lean_int_dec_lt(x_25, x_5); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_33 = lean_nat_abs(x_25); +lean_dec(x_25); +x_34 = l___private_Init_Data_Repr_0__Nat_reprFast(x_33); +x_35 = lean_unsigned_to_nat(9u); +x_36 = l_Std_Time_instToStringDuration_leftPad(x_35, x_34); +lean_dec(x_34); +x_37 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15; +x_38 = lean_string_append(x_37, x_36); +lean_dec(x_36); +x_39 = lean_string_append(x_31, x_38); +lean_dec(x_38); +x_40 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_41 = lean_string_append(x_39, x_40); +x_42 = l_String_quote(x_41); +lean_dec(x_41); +x_43 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_7 = x_43; +goto block_21; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_44 = lean_nat_abs(x_25); +lean_dec(x_25); +x_45 = lean_unsigned_to_nat(1u); +x_46 = lean_nat_sub(x_44, x_45); +lean_dec(x_44); +x_47 = lean_nat_add(x_46, x_45); +lean_dec(x_46); +x_48 = l___private_Init_Data_Repr_0__Nat_reprFast(x_47); +x_49 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_50 = lean_string_append(x_49, x_48); +lean_dec(x_48); +x_51 = lean_unsigned_to_nat(9u); +x_52 = l_Std_Time_instToStringDuration_leftPad(x_51, x_50); +lean_dec(x_50); +x_53 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15; +x_54 = lean_string_append(x_53, x_52); +lean_dec(x_52); +x_55 = lean_string_append(x_31, x_54); +lean_dec(x_54); +x_56 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_57 = lean_string_append(x_55, x_56); +x_58 = l_String_quote(x_57); +lean_dec(x_57); +x_59 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_7 = x_59; +goto block_21; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_25); +x_60 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18; +x_61 = lean_string_append(x_31, x_60); +x_62 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_63 = lean_string_append(x_61, x_62); +x_64 = l_String_quote(x_63); +lean_dec(x_63); +x_65 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_7 = x_65; +goto block_21; +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_66 = lean_nat_abs(x_24); +lean_dec(x_24); +x_67 = lean_unsigned_to_nat(1u); +x_68 = lean_nat_sub(x_66, x_67); +lean_dec(x_66); +x_69 = lean_nat_add(x_68, x_67); +lean_dec(x_68); +x_70 = l___private_Init_Data_Repr_0__Nat_reprFast(x_69); +x_71 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_72 = lean_string_append(x_71, x_70); +lean_dec(x_70); +x_73 = lean_string_append(x_22, x_72); +lean_dec(x_72); +if (x_27 == 0) +{ +uint8_t x_74; +x_74 = lean_int_dec_lt(x_25, x_5); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_75 = lean_nat_abs(x_25); +lean_dec(x_25); +x_76 = l___private_Init_Data_Repr_0__Nat_reprFast(x_75); +x_77 = lean_unsigned_to_nat(9u); +x_78 = l_Std_Time_instToStringDuration_leftPad(x_77, x_76); +lean_dec(x_76); +x_79 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15; +x_80 = lean_string_append(x_79, x_78); +lean_dec(x_78); +x_81 = lean_string_append(x_73, x_80); +lean_dec(x_80); +x_82 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_83 = lean_string_append(x_81, x_82); +x_84 = l_String_quote(x_83); +lean_dec(x_83); +x_85 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_7 = x_85; +goto block_21; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_86 = lean_nat_abs(x_25); +lean_dec(x_25); +x_87 = lean_nat_sub(x_86, x_67); +lean_dec(x_86); +x_88 = lean_nat_add(x_87, x_67); +lean_dec(x_87); +x_89 = l___private_Init_Data_Repr_0__Nat_reprFast(x_88); +x_90 = lean_string_append(x_71, x_89); +lean_dec(x_89); +x_91 = lean_unsigned_to_nat(9u); +x_92 = l_Std_Time_instToStringDuration_leftPad(x_91, x_90); +lean_dec(x_90); +x_93 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15; +x_94 = lean_string_append(x_93, x_92); +lean_dec(x_92); +x_95 = lean_string_append(x_73, x_94); +lean_dec(x_94); +x_96 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_97 = lean_string_append(x_95, x_96); +x_98 = l_String_quote(x_97); +lean_dec(x_97); +x_99 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_99, 0, x_98); +x_7 = x_99; +goto block_21; +} +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_25); +x_100 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18; +x_101 = lean_string_append(x_73, x_100); +x_102 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16; +x_103 = lean_string_append(x_101, x_102); +x_104 = l_String_quote(x_103); +lean_dec(x_103); +x_105 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_105, 0, x_104); +x_7 = x_105; +goto block_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprTimestamp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprTimestamp() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprTimestamp___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_int_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_dec_eq(x_7, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instBEqTimestamp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instBEqTimestamp() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instBEqTimestamp___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedTimestamp___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedTimestamp() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedTimestamp___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instLETimestamp() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_instDecidableLeTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_Std_Time_Duration_instDecidableLe(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instDecidableLeTimestamp___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_instDecidableLeTimestamp(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instOfNatTimestamp(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_nat_to_int(x_1); +x_3 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instToStringTimestamp___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instToStringTimestamp___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringTimestamp(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_instToStringTimestamp___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Std_Time_instToStringTimestamp___closed__2; +x_7 = lean_int_ediv(x_5, x_6); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +lean_dec(x_4); +x_9 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_10 = lean_int_dec_lt(x_8, x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_nat_abs(x_8); +lean_dec(x_8); +x_12 = l___private_Init_Data_Repr_0__Nat_reprFast(x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_13 = lean_nat_abs(x_8); +lean_dec(x_8); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_13, x_14); +lean_dec(x_13); +x_16 = lean_nat_add(x_15, x_14); +lean_dec(x_15); +x_17 = l___private_Init_Data_Repr_0__Nat_reprFast(x_16); +x_18 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringTimestamp___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instToStringTimestamp(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprTimestamp__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_instToStringTimestamp___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Std_Time_instToStringTimestamp___closed__2; +x_8 = lean_int_ediv(x_6, x_7); +x_9 = lean_int_add(x_5, x_8); +lean_dec(x_8); +lean_dec(x_5); +x_10 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_11 = lean_int_dec_lt(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_nat_abs(x_9); +lean_dec(x_9); +x_13 = l___private_Init_Data_Repr_0__Nat_reprFast(x_12); +x_14 = l_String_quote(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_16 = lean_nat_abs(x_9); +lean_dec(x_9); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_16, x_17); +lean_dec(x_16); +x_19 = lean_nat_add(x_18, x_17); +lean_dec(x_18); +x_20 = l___private_Init_Data_Repr_0__Nat_reprFast(x_19); +x_21 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17; +x_22 = lean_string_append(x_21, x_20); +lean_dec(x_20); +x_23 = l_String_quote(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprTimestamp__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instReprTimestamp__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_now___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_get_current_time(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_toMinutes___closed__1; +x_4 = lean_int_ediv(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_toDays___closed__1; +x_4 = lean_int_ediv(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofSecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofNanosecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_ofNanoseconds(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofNanosecondsSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_ofNanosecondsSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofMillisecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_instToStringTimestamp___closed__2; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_Duration_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_ofMillisecondsSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_ofMillisecondsSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toSecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toSecondsSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toSecondsSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMillisecondsSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_instToStringTimestamp___closed__2; +x_8 = lean_int_div(x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toMillisecondsSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toMillisecondsSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_since(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_get_current_time(x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_int_neg(x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_neg(x_8); +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +x_11 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_12 = lean_int_mul(x_10, x_11); +lean_dec(x_10); +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_15 = lean_int_mul(x_7, x_11); +lean_dec(x_7); +x_16 = lean_int_add(x_15, x_9); +lean_dec(x_9); +lean_dec(x_15); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_Duration_ofNanoseconds(x_17); +lean_dec(x_17); +lean_ctor_set(x_3, 0, x_18); +return x_3; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_19 = lean_ctor_get(x_3, 0); +x_20 = lean_ctor_get(x_3, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_3); +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_int_neg(x_21); +x_23 = lean_ctor_get(x_1, 1); +x_24 = lean_int_neg(x_23); +x_25 = lean_ctor_get(x_19, 0); +lean_inc(x_25); +x_26 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_27 = lean_int_mul(x_25, x_26); +lean_dec(x_25); +x_28 = lean_ctor_get(x_19, 1); +lean_inc(x_28); +lean_dec(x_19); +x_29 = lean_int_add(x_27, x_28); +lean_dec(x_28); +lean_dec(x_27); +x_30 = lean_int_mul(x_22, x_26); +lean_dec(x_22); +x_31 = lean_int_add(x_30, x_24); +lean_dec(x_24); +lean_dec(x_30); +x_32 = lean_int_add(x_29, x_31); +lean_dec(x_31); +lean_dec(x_29); +x_33 = l_Std_Time_Duration_ofNanoseconds(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_20); +return x_34; +} +} +else +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_3); +if (x_35 == 0) +{ +return x_3; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_3, 0); +x_37 = lean_ctor_get(x_3, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_3); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_since___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_since(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDurationSinceUnixEpoch(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_toDurationSinceUnixEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Timestamp_toDurationSinceUnixEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_3 = l_Std_Time_instToStringTimestamp___closed__2; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Duration_ofNanoseconds(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_5, 0); +lean_inc(x_11); +x_12 = lean_int_mul(x_11, x_7); +lean_dec(x_11); +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_15 = lean_int_add(x_10, x_14); +lean_dec(x_14); +lean_dec(x_10); +x_16 = l_Std_Time_Duration_ofNanoseconds(x_15); +lean_dec(x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_3 = l_Std_Time_instToStringTimestamp___closed__2; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Duration_ofNanoseconds(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_int_neg(x_8); +lean_dec(x_8); +x_10 = lean_ctor_get(x_1, 0); +x_11 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_12 = lean_int_mul(x_10, x_11); +x_13 = lean_ctor_get(x_1, 1); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_12); +x_15 = lean_int_mul(x_7, x_11); +lean_dec(x_7); +x_16 = lean_int_add(x_15, x_9); +lean_dec(x_9); +lean_dec(x_15); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_Duration_ofNanoseconds(x_17); +lean_dec(x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_ofNanoseconds(x_2); +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_int_mul(x_9, x_5); +lean_dec(x_9); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_dec(x_3); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_8, x_12); +lean_dec(x_12); +lean_dec(x_8); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_3 = l_Std_Time_Duration_ofNanoseconds(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_dec(x_3); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 0); +x_9 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_10 = lean_int_mul(x_8, x_9); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_mul(x_5, x_9); +lean_dec(x_5); +x_14 = lean_int_add(x_13, x_7); +lean_dec(x_7); +lean_dec(x_13); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_Duration_ofNanoseconds(x_15); +lean_dec(x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_mul(x_2, x_4); +x_9 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_add(x_7, x_10); +lean_dec(x_10); +lean_dec(x_7); +x_12 = l_Std_Time_Duration_ofNanoseconds(x_11); +lean_dec(x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Timestamp_subSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_10 = l_Std_Time_Timestamp_subSeconds___closed__1; +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_add(x_8, x_11); +lean_dec(x_11); +lean_dec(x_8); +x_13 = l_Std_Time_Duration_ofNanoseconds(x_12); +lean_dec(x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Timestamp_toMinutes___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Timestamp_toMinutes___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Timestamp_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Timestamp_addHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Timestamp_addHours___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Timestamp_addHours___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Timestamp_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Timestamp_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addDays(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Timestamp_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Timestamp_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subDays(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_3 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Timestamp_toDays___closed__1; +x_6 = lean_int_mul(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 0); +x_8 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_9 = lean_int_mul(x_7, x_8); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mul(x_6, x_8); +lean_dec(x_6); +x_13 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8; +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_12); +x_15 = lean_int_add(x_11, x_14); +lean_dec(x_14); +lean_dec(x_11); +x_16 = l_Std_Time_Duration_ofNanoseconds(x_15); +lean_dec(x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addWeeks(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Timestamp_toDays___closed__1; +x_6 = lean_int_mul(x_4, x_5); +lean_dec(x_4); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 0); +x_9 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_10 = lean_int_mul(x_8, x_9); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_mul(x_7, x_9); +lean_dec(x_7); +x_14 = l_Std_Time_Timestamp_subSeconds___closed__1; +x_15 = lean_int_add(x_13, x_14); +lean_dec(x_13); +x_16 = lean_int_add(x_12, x_15); +lean_dec(x_15); +lean_dec(x_12); +x_17 = l_Std_Time_Duration_ofNanoseconds(x_16); +lean_dec(x_16); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subWeeks(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_int_mul(x_8, x_4); +x_10 = lean_ctor_get(x_2, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_add(x_7, x_11); +lean_dec(x_11); +lean_dec(x_7); +x_13 = l_Std_Time_Duration_ofNanoseconds(x_12); +lean_dec(x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_addDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_addDuration(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_neg(x_3); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_int_neg(x_5); +x_7 = lean_ctor_get(x_1, 0); +x_8 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_9 = lean_int_mul(x_7, x_8); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mul(x_4, x_8); +lean_dec(x_4); +x_13 = lean_int_add(x_12, x_6); +lean_dec(x_6); +lean_dec(x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_subDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_subDuration(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddDuration___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addDuration___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddDuration() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddDuration___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubDuration___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subDuration___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubDuration() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubDuration___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_addNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHAddOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHAddOffset__6___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Timestamp_subNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Timestamp_instHSubOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Timestamp_instHSubOffset__6___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubDuration__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_neg(x_3); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_int_neg(x_5); +x_7 = lean_ctor_get(x_1, 0); +x_8 = l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1; +x_9 = lean_int_mul(x_7, x_8); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mul(x_4, x_8); +lean_dec(x_4); +x_13 = lean_int_add(x_12, x_6); +lean_dec(x_6); +lean_dec(x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Timestamp_instHSubDuration__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Timestamp_instHSubDuration__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Duration(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_DateTime_Timestamp(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Duration(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__1); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__2); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__3); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__4); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__5); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__6); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__7); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__8); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__9); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__10); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__11); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__12); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__13); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__14); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__15); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__16); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__17); +l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18 = _init_l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18(); +lean_mark_persistent(l___private_Std_Time_DateTime_Timestamp_0__Std_Time_reprTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_18____closed__18); +l_Std_Time_instReprTimestamp___closed__1 = _init_l_Std_Time_instReprTimestamp___closed__1(); +lean_mark_persistent(l_Std_Time_instReprTimestamp___closed__1); +l_Std_Time_instReprTimestamp = _init_l_Std_Time_instReprTimestamp(); +lean_mark_persistent(l_Std_Time_instReprTimestamp); +l_Std_Time_instBEqTimestamp___closed__1 = _init_l_Std_Time_instBEqTimestamp___closed__1(); +lean_mark_persistent(l_Std_Time_instBEqTimestamp___closed__1); +l_Std_Time_instBEqTimestamp = _init_l_Std_Time_instBEqTimestamp(); +lean_mark_persistent(l_Std_Time_instBEqTimestamp); +l_Std_Time_instInhabitedTimestamp___closed__1 = _init_l_Std_Time_instInhabitedTimestamp___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedTimestamp___closed__1); +l_Std_Time_instInhabitedTimestamp = _init_l_Std_Time_instInhabitedTimestamp(); +lean_mark_persistent(l_Std_Time_instInhabitedTimestamp); +l_Std_Time_instLETimestamp = _init_l_Std_Time_instLETimestamp(); +lean_mark_persistent(l_Std_Time_instLETimestamp); +l_Std_Time_instToStringTimestamp___closed__1 = _init_l_Std_Time_instToStringTimestamp___closed__1(); +lean_mark_persistent(l_Std_Time_instToStringTimestamp___closed__1); +l_Std_Time_instToStringTimestamp___closed__2 = _init_l_Std_Time_instToStringTimestamp___closed__2(); +lean_mark_persistent(l_Std_Time_instToStringTimestamp___closed__2); +l_Std_Time_Timestamp_toMinutes___closed__1 = _init_l_Std_Time_Timestamp_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_toMinutes___closed__1); +l_Std_Time_Timestamp_toDays___closed__1 = _init_l_Std_Time_Timestamp_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_toDays___closed__1); +l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1 = _init_l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_toNanosecondsSinceUnixEpoch___closed__1); +l_Std_Time_Timestamp_subSeconds___closed__1 = _init_l_Std_Time_Timestamp_subSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_subSeconds___closed__1); +l_Std_Time_Timestamp_addHours___closed__1 = _init_l_Std_Time_Timestamp_addHours___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_addHours___closed__1); +l_Std_Time_Timestamp_instHAddDuration___closed__1 = _init_l_Std_Time_Timestamp_instHAddDuration___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddDuration___closed__1); +l_Std_Time_Timestamp_instHAddDuration = _init_l_Std_Time_Timestamp_instHAddDuration(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddDuration); +l_Std_Time_Timestamp_instHSubDuration___closed__1 = _init_l_Std_Time_Timestamp_instHSubDuration___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubDuration___closed__1); +l_Std_Time_Timestamp_instHSubDuration = _init_l_Std_Time_Timestamp_instHSubDuration(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubDuration); +l_Std_Time_Timestamp_instHAddOffset___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset___closed__1); +l_Std_Time_Timestamp_instHAddOffset = _init_l_Std_Time_Timestamp_instHAddOffset(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset); +l_Std_Time_Timestamp_instHSubOffset___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset___closed__1); +l_Std_Time_Timestamp_instHSubOffset = _init_l_Std_Time_Timestamp_instHSubOffset(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset); +l_Std_Time_Timestamp_instHAddOffset__1___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__1___closed__1); +l_Std_Time_Timestamp_instHAddOffset__1 = _init_l_Std_Time_Timestamp_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__1); +l_Std_Time_Timestamp_instHSubOffset__1___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__1___closed__1); +l_Std_Time_Timestamp_instHSubOffset__1 = _init_l_Std_Time_Timestamp_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__1); +l_Std_Time_Timestamp_instHAddOffset__2___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__2___closed__1); +l_Std_Time_Timestamp_instHAddOffset__2 = _init_l_Std_Time_Timestamp_instHAddOffset__2(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__2); +l_Std_Time_Timestamp_instHSubOffset__2___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__2___closed__1); +l_Std_Time_Timestamp_instHSubOffset__2 = _init_l_Std_Time_Timestamp_instHSubOffset__2(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__2); +l_Std_Time_Timestamp_instHAddOffset__3___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__3___closed__1); +l_Std_Time_Timestamp_instHAddOffset__3 = _init_l_Std_Time_Timestamp_instHAddOffset__3(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__3); +l_Std_Time_Timestamp_instHSubOffset__3___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__3___closed__1); +l_Std_Time_Timestamp_instHSubOffset__3 = _init_l_Std_Time_Timestamp_instHSubOffset__3(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__3); +l_Std_Time_Timestamp_instHAddOffset__4___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__4___closed__1); +l_Std_Time_Timestamp_instHAddOffset__4 = _init_l_Std_Time_Timestamp_instHAddOffset__4(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__4); +l_Std_Time_Timestamp_instHSubOffset__4___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__4___closed__1); +l_Std_Time_Timestamp_instHSubOffset__4 = _init_l_Std_Time_Timestamp_instHSubOffset__4(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__4); +l_Std_Time_Timestamp_instHAddOffset__5___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__5___closed__1); +l_Std_Time_Timestamp_instHAddOffset__5 = _init_l_Std_Time_Timestamp_instHAddOffset__5(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__5); +l_Std_Time_Timestamp_instHSubOffset__5___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__5___closed__1); +l_Std_Time_Timestamp_instHSubOffset__5 = _init_l_Std_Time_Timestamp_instHSubOffset__5(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__5); +l_Std_Time_Timestamp_instHAddOffset__6___closed__1 = _init_l_Std_Time_Timestamp_instHAddOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__6___closed__1); +l_Std_Time_Timestamp_instHAddOffset__6 = _init_l_Std_Time_Timestamp_instHAddOffset__6(); +lean_mark_persistent(l_Std_Time_Timestamp_instHAddOffset__6); +l_Std_Time_Timestamp_instHSubOffset__6___closed__1 = _init_l_Std_Time_Timestamp_instHSubOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__6___closed__1); +l_Std_Time_Timestamp_instHSubOffset__6 = _init_l_Std_Time_Timestamp_instHSubOffset__6(); +lean_mark_persistent(l_Std_Time_Timestamp_instHSubOffset__6); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Duration.c b/stage0/stdlib/Std/Time/Duration.c new file mode 100644 index 000000000000..7d5b91e64d2a --- /dev/null +++ b/stage0/stdlib/Std/Time/Duration.c @@ -0,0 +1,2924 @@ +// Lean compiler output +// Module: Std.Time.Duration +// Imports: Std.Time.Date Std.Time.Time +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +static lean_object* l_Std_Time_Duration_instHSubOffset__1___closed__1; +static lean_object* l_Std_Time_Duration_addWeeks___closed__1; +lean_object* lean_int_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset__5; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_Duration_subSeconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset; +static lean_object* l_Std_Time_Duration_instHAddOffset__4___closed__1; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__4; +static lean_object* l_Std_Time_Duration_toMinutes___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_toNanoseconds___boxed(lean_object*); +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_sub___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1; +static lean_object* l_Std_Time_Duration_instCoeOffset__4___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_addHours___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_subSeconds(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7; +static lean_object* l_Std_Time_Duration_instCoeOffset__5___closed__1; +static lean_object* l_Std_Time_instToStringDuration___closed__3; +static lean_object* l_Std_Time_Duration_instCoeOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_fromComponents(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset__3; +static lean_object* l_Std_Time_Duration_instHSubOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofSeconds(lean_object*); +lean_object* l_String_quote(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddPlainTime___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22; +static lean_object* l_Std_Time_instToStringDuration_leftPad___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__6; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_Duration_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_subNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprDuration___closed__1; +static lean_object* l_Std_Time_Duration_instHSubOffset__2___closed__1; +static lean_object* l_Std_Time_Duration_instHAddOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__2; +static lean_object* l_Std_Time_Duration_instHAddOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_addDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHSubOffset__4___closed__1; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19; +static lean_object* l_Std_Time_Duration_instHAddOffset__6___closed__1; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_subDays(lean_object*, lean_object*); +lean_object* l_Std_Time_Day_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subWeeks(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1(uint32_t, lean_object*, lean_object*); +lean_object* lean_string_push(lean_object*, uint32_t); +LEAN_EXPORT uint8_t l_Std_Time_Duration_instDecidableLe(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAdd; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instDecidableLe___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMilliseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__4; +LEAN_EXPORT lean_object* l_Std_Time_Duration_isZero___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset__4; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddPlainTime(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__5; +static lean_object* l_Std_Time_Duration_instCoeOffset__4___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration___boxed(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration_leftPad___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13; +LEAN_EXPORT lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__6; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMilliseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addWeeks(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addNanoseconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9; +lean_object* l_Std_Time_Minute_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addDays(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_ofNanoseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration(lean_object*); +lean_object* l_Int_repr(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17; +lean_object* l_Std_Time_Hour_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprDuration__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofMillisecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_toSeconds(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_instBEqDuration___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instBEqDuration(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_toMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubPlainTime(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_ofMillisecond___closed__1; +static lean_object* l_Std_Time_Duration_addHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_toDays(lean_object*); +static lean_object* l_Std_Time_Duration_instHSubOffset__5___closed__1; +static lean_object* l_Std_Time_Duration_instCoeOffset__2___closed__1; +static lean_object* l_Std_Time_instToStringDuration___closed__1; +static lean_object* l_Std_Time_Duration_instCoeOffset__4___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Duration_neg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofNanoseconds___boxed(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20; +LEAN_EXPORT lean_object* l_Std_Time_instReprDuration__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHAddOffset__5___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instOfNatDuration(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMilliseconds(lean_object*); +lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subSeconds___boxed(lean_object*, lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_fromComponents___boxed(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59_(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHAddOffset___closed__1; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23; +static lean_object* l_Std_Time_Duration_toDays___closed__1; +lean_object* l_Std_Time_Week_Offset_toDays___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_toDays___boxed(lean_object*); +static lean_object* l_Std_Time_Duration_instHAddOffset__2___closed__1; +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__1; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedDuration___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMinutes(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_instReprDuration; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__5; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration_leftPad(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_add___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instCoeOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subDays___boxed(lean_object*, lean_object*); +lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Duration_isZero(lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6; +static lean_object* l_Std_Time_Duration_instCoeOffset__2___closed__2; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instCoeOffset__3___closed__1; +static lean_object* l_Std_Time_Duration_instCoeOffset__4___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Duration_addNanoseconds(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Std_Time_instToStringDuration___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset__2; +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset__1; +static lean_object* l_Std_Time_Duration_instHSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubPlainTime___boxed(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMinutes___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instCoeOffset__3___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Duration_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHSubOffset__6___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instLE; +LEAN_EXPORT lean_object* l_Std_Time_Duration_subNanoseconds___boxed(lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_Duration_instHSubOffset__3___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofMillisecond___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSub; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddOffset__2; +LEAN_EXPORT lean_object* l_Std_Time_Duration_add(lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_subHours___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addSeconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDuration; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instCoeOffset; +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubOffset; +LEAN_EXPORT lean_object* l_Std_Time_Duration_addSeconds(lean_object*, lean_object*); +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("second", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3; +x_2 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nano", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("proof", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_dec_lt(x_6, x_4); +if (x_5 == 0) +{ +lean_object* x_66; lean_object* x_67; +x_66 = l_Int_repr(x_3); +x_67 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_67, 0, x_66); +x_8 = x_67; +goto block_65; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_68 = l_Int_repr(x_3); +x_69 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_69, 0, x_68); +x_70 = lean_unsigned_to_nat(0u); +x_71 = l_Repr_addAppParen(x_69, x_70); +x_8 = x_71; +goto block_65; +} +block_65: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_9 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7; +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10; +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_box(1); +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5; +x_22 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_23 = l_Int_repr(x_6); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_11); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_15); +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15; +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_21); +x_34 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17; +x_35 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21; +x_37 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23; +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20; +x_41 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set_uint8(x_42, sizeof(void*)*1, x_11); +return x_42; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_43 = l_Int_repr(x_6); +x_44 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = lean_unsigned_to_nat(0u); +x_46 = l_Repr_addAppParen(x_44, x_45); +x_47 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13; +x_48 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_11); +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_22); +lean_ctor_set(x_50, 1, x_49); +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_15); +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_17); +x_53 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15; +x_54 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_21); +x_56 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17; +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21; +x_59 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +x_60 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23; +x_61 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20; +x_63 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set_uint8(x_64, sizeof(void*)*1, x_11); +return x_64; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprDuration___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprDuration() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprDuration___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_2, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_2, x_6); +lean_dec(x_2); +x_8 = lean_string_push(x_3, x_1); +x_2 = x_7; +x_3 = x_8; +goto _start; +} +else +{ +lean_dec(x_2); +return x_3; +} +} +} +static lean_object* _init_l_Std_Time_instToStringDuration_leftPad___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration_leftPad(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = 48; +x_4 = lean_string_length(x_2); +x_5 = lean_nat_sub(x_1, x_4); +lean_dec(x_4); +x_6 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_7 = l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1(x_3, x_5, x_6); +x_8 = lean_string_append(x_7, x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Nat_repeatTR_loop___at_Std_Time_instToStringDuration_leftPad___spec__1(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration_leftPad___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instToStringDuration_leftPad(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instToStringDuration___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instToStringDuration___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("s", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instToStringDuration___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = lean_ctor_get(x_1, 0); +x_77 = lean_ctor_get(x_1, 1); +x_78 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_79 = lean_int_dec_lt(x_78, x_76); +if (x_79 == 0) +{ +uint8_t x_80; +x_80 = lean_int_dec_lt(x_76, x_78); +if (x_80 == 0) +{ +uint8_t x_81; +x_81 = lean_int_dec_lt(x_77, x_78); +if (x_81 == 0) +{ +lean_object* x_82; lean_object* x_83; +lean_inc(x_77); +lean_inc(x_76); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_76); +lean_ctor_set(x_82, 1, x_77); +x_83 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_2 = x_83; +x_3 = x_82; +goto block_75; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_int_neg(x_76); +x_85 = lean_int_neg(x_77); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Std_Time_instToStringDuration___closed__3; +x_2 = x_87; +x_3 = x_86; +goto block_75; +} +} +else +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_int_neg(x_76); +x_89 = lean_int_neg(x_77); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Std_Time_instToStringDuration___closed__3; +x_2 = x_91; +x_3 = x_90; +goto block_75; +} +} +else +{ +lean_object* x_92; lean_object* x_93; +lean_inc(x_77); +lean_inc(x_76); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_76); +lean_ctor_set(x_92, 1, x_77); +x_93 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_2 = x_93; +x_3 = x_92; +goto block_75; +} +block_75: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_8 = lean_int_dec_eq(x_6, x_7); +x_9 = lean_int_dec_lt(x_4, x_7); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_nat_abs(x_4); +lean_dec(x_4); +x_11 = l___private_Init_Data_Repr_0__Nat_reprFast(x_10); +x_12 = lean_string_append(x_2, x_11); +lean_dec(x_11); +if (x_8 == 0) +{ +uint8_t x_13; +x_13 = lean_int_dec_lt(x_5, x_7); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_14 = lean_nat_abs(x_5); +lean_dec(x_5); +x_15 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); +x_16 = lean_unsigned_to_nat(9u); +x_17 = l_Std_Time_instToStringDuration_leftPad(x_16, x_15); +lean_dec(x_15); +x_18 = l_Std_Time_instToStringDuration___closed__1; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = lean_string_append(x_12, x_19); +lean_dec(x_19); +x_21 = l_Std_Time_instToStringDuration___closed__2; +x_22 = lean_string_append(x_20, x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_23 = lean_nat_abs(x_5); +lean_dec(x_5); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = lean_nat_add(x_25, x_24); +lean_dec(x_25); +x_27 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_28 = l_Std_Time_instToStringDuration___closed__3; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +x_30 = lean_unsigned_to_nat(9u); +x_31 = l_Std_Time_instToStringDuration_leftPad(x_30, x_29); +lean_dec(x_29); +x_32 = l_Std_Time_instToStringDuration___closed__1; +x_33 = lean_string_append(x_32, x_31); +lean_dec(x_31); +x_34 = lean_string_append(x_12, x_33); +lean_dec(x_33); +x_35 = l_Std_Time_instToStringDuration___closed__2; +x_36 = lean_string_append(x_34, x_35); +return x_36; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_5); +x_37 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_38 = lean_string_append(x_12, x_37); +x_39 = l_Std_Time_instToStringDuration___closed__2; +x_40 = lean_string_append(x_38, x_39); +return x_40; +} +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_41 = lean_nat_abs(x_4); +lean_dec(x_4); +x_42 = lean_unsigned_to_nat(1u); +x_43 = lean_nat_sub(x_41, x_42); +lean_dec(x_41); +x_44 = lean_nat_add(x_43, x_42); +lean_dec(x_43); +x_45 = l___private_Init_Data_Repr_0__Nat_reprFast(x_44); +x_46 = l_Std_Time_instToStringDuration___closed__3; +x_47 = lean_string_append(x_46, x_45); +lean_dec(x_45); +x_48 = lean_string_append(x_2, x_47); +lean_dec(x_47); +if (x_8 == 0) +{ +uint8_t x_49; +x_49 = lean_int_dec_lt(x_5, x_7); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_50 = lean_nat_abs(x_5); +lean_dec(x_5); +x_51 = l___private_Init_Data_Repr_0__Nat_reprFast(x_50); +x_52 = lean_unsigned_to_nat(9u); +x_53 = l_Std_Time_instToStringDuration_leftPad(x_52, x_51); +lean_dec(x_51); +x_54 = l_Std_Time_instToStringDuration___closed__1; +x_55 = lean_string_append(x_54, x_53); +lean_dec(x_53); +x_56 = lean_string_append(x_48, x_55); +lean_dec(x_55); +x_57 = l_Std_Time_instToStringDuration___closed__2; +x_58 = lean_string_append(x_56, x_57); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_59 = lean_nat_abs(x_5); +lean_dec(x_5); +x_60 = lean_nat_sub(x_59, x_42); +lean_dec(x_59); +x_61 = lean_nat_add(x_60, x_42); +lean_dec(x_60); +x_62 = l___private_Init_Data_Repr_0__Nat_reprFast(x_61); +x_63 = lean_string_append(x_46, x_62); +lean_dec(x_62); +x_64 = lean_unsigned_to_nat(9u); +x_65 = l_Std_Time_instToStringDuration_leftPad(x_64, x_63); +lean_dec(x_63); +x_66 = l_Std_Time_instToStringDuration___closed__1; +x_67 = lean_string_append(x_66, x_65); +lean_dec(x_65); +x_68 = lean_string_append(x_48, x_67); +lean_dec(x_67); +x_69 = l_Std_Time_instToStringDuration___closed__2; +x_70 = lean_string_append(x_68, x_69); +return x_70; +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_5); +x_71 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_72 = lean_string_append(x_48, x_71); +x_73 = l_Std_Time_instToStringDuration___closed__2; +x_74 = lean_string_append(x_72, x_73); +return x_74; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instToStringDuration___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instToStringDuration(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprDuration__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_1, 0); +x_90 = lean_ctor_get(x_1, 1); +x_91 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_92 = lean_int_dec_lt(x_91, x_89); +if (x_92 == 0) +{ +uint8_t x_93; +x_93 = lean_int_dec_lt(x_89, x_91); +if (x_93 == 0) +{ +uint8_t x_94; +x_94 = lean_int_dec_lt(x_90, x_91); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +lean_inc(x_90); +lean_inc(x_89); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_89); +lean_ctor_set(x_95, 1, x_90); +x_96 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_3 = x_96; +x_4 = x_95; +goto block_88; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_int_neg(x_89); +x_98 = lean_int_neg(x_90); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = l_Std_Time_instToStringDuration___closed__3; +x_3 = x_100; +x_4 = x_99; +goto block_88; +} +} +else +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_101 = lean_int_neg(x_89); +x_102 = lean_int_neg(x_90); +x_103 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +x_104 = l_Std_Time_instToStringDuration___closed__3; +x_3 = x_104; +x_4 = x_103; +goto block_88; +} +} +else +{ +lean_object* x_105; lean_object* x_106; +lean_inc(x_90); +lean_inc(x_89); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_89); +lean_ctor_set(x_105, 1, x_90); +x_106 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_3 = x_106; +x_4 = x_105; +goto block_88; +} +block_88: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint8_t x_10; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +x_8 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_9 = lean_int_dec_eq(x_7, x_8); +x_10 = lean_int_dec_lt(x_5, x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_nat_abs(x_5); +lean_dec(x_5); +x_12 = l___private_Init_Data_Repr_0__Nat_reprFast(x_11); +x_13 = lean_string_append(x_3, x_12); +lean_dec(x_12); +if (x_9 == 0) +{ +uint8_t x_14; +x_14 = lean_int_dec_lt(x_6, x_8); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_nat_abs(x_6); +lean_dec(x_6); +x_16 = l___private_Init_Data_Repr_0__Nat_reprFast(x_15); +x_17 = lean_unsigned_to_nat(9u); +x_18 = l_Std_Time_instToStringDuration_leftPad(x_17, x_16); +lean_dec(x_16); +x_19 = l_Std_Time_instToStringDuration___closed__1; +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = lean_string_append(x_13, x_20); +lean_dec(x_20); +x_22 = l_Std_Time_instToStringDuration___closed__2; +x_23 = lean_string_append(x_21, x_22); +x_24 = l_String_quote(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_25, 0, x_24); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_26 = lean_nat_abs(x_6); +lean_dec(x_6); +x_27 = lean_unsigned_to_nat(1u); +x_28 = lean_nat_sub(x_26, x_27); +lean_dec(x_26); +x_29 = lean_nat_add(x_28, x_27); +lean_dec(x_28); +x_30 = l___private_Init_Data_Repr_0__Nat_reprFast(x_29); +x_31 = l_Std_Time_instToStringDuration___closed__3; +x_32 = lean_string_append(x_31, x_30); +lean_dec(x_30); +x_33 = lean_unsigned_to_nat(9u); +x_34 = l_Std_Time_instToStringDuration_leftPad(x_33, x_32); +lean_dec(x_32); +x_35 = l_Std_Time_instToStringDuration___closed__1; +x_36 = lean_string_append(x_35, x_34); +lean_dec(x_34); +x_37 = lean_string_append(x_13, x_36); +lean_dec(x_36); +x_38 = l_Std_Time_instToStringDuration___closed__2; +x_39 = lean_string_append(x_37, x_38); +x_40 = l_String_quote(x_39); +lean_dec(x_39); +x_41 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_6); +x_42 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_43 = lean_string_append(x_13, x_42); +x_44 = l_Std_Time_instToStringDuration___closed__2; +x_45 = lean_string_append(x_43, x_44); +x_46 = l_String_quote(x_45); +lean_dec(x_45); +x_47 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_47, 0, x_46); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_48 = lean_nat_abs(x_5); +lean_dec(x_5); +x_49 = lean_unsigned_to_nat(1u); +x_50 = lean_nat_sub(x_48, x_49); +lean_dec(x_48); +x_51 = lean_nat_add(x_50, x_49); +lean_dec(x_50); +x_52 = l___private_Init_Data_Repr_0__Nat_reprFast(x_51); +x_53 = l_Std_Time_instToStringDuration___closed__3; +x_54 = lean_string_append(x_53, x_52); +lean_dec(x_52); +x_55 = lean_string_append(x_3, x_54); +lean_dec(x_54); +if (x_9 == 0) +{ +uint8_t x_56; +x_56 = lean_int_dec_lt(x_6, x_8); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_57 = lean_nat_abs(x_6); +lean_dec(x_6); +x_58 = l___private_Init_Data_Repr_0__Nat_reprFast(x_57); +x_59 = lean_unsigned_to_nat(9u); +x_60 = l_Std_Time_instToStringDuration_leftPad(x_59, x_58); +lean_dec(x_58); +x_61 = l_Std_Time_instToStringDuration___closed__1; +x_62 = lean_string_append(x_61, x_60); +lean_dec(x_60); +x_63 = lean_string_append(x_55, x_62); +lean_dec(x_62); +x_64 = l_Std_Time_instToStringDuration___closed__2; +x_65 = lean_string_append(x_63, x_64); +x_66 = l_String_quote(x_65); +lean_dec(x_65); +x_67 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_68 = lean_nat_abs(x_6); +lean_dec(x_6); +x_69 = lean_nat_sub(x_68, x_49); +lean_dec(x_68); +x_70 = lean_nat_add(x_69, x_49); +lean_dec(x_69); +x_71 = l___private_Init_Data_Repr_0__Nat_reprFast(x_70); +x_72 = lean_string_append(x_53, x_71); +lean_dec(x_71); +x_73 = lean_unsigned_to_nat(9u); +x_74 = l_Std_Time_instToStringDuration_leftPad(x_73, x_72); +lean_dec(x_72); +x_75 = l_Std_Time_instToStringDuration___closed__1; +x_76 = lean_string_append(x_75, x_74); +lean_dec(x_74); +x_77 = lean_string_append(x_55, x_76); +lean_dec(x_76); +x_78 = l_Std_Time_instToStringDuration___closed__2; +x_79 = lean_string_append(x_77, x_78); +x_80 = l_String_quote(x_79); +lean_dec(x_79); +x_81 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_6); +x_82 = l_Std_Time_instToStringDuration_leftPad___closed__1; +x_83 = lean_string_append(x_55, x_82); +x_84 = l_Std_Time_instToStringDuration___closed__2; +x_85 = lean_string_append(x_83, x_84); +x_86 = l_String_quote(x_85); +lean_dec(x_85); +x_87 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_87, 0, x_86); +return x_87; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprDuration__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_instReprDuration__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_instBEqDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_int_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_dec_eq(x_7, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_instBEqDuration(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDuration___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDuration() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedDuration___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instOfNatDuration(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_nat_to_int(x_1); +x_3 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_neg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_int_neg(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_neg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_neg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_ofNanoseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_3 = lean_int_ediv(x_1, x_2); +x_4 = lean_int_mod(x_1, x_2); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Duration_ofMillisecond___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofMillisecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Duration_ofMillisecond___closed__1; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_Duration_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_ofMillisecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_ofMillisecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Duration_isZero(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_4 = lean_int_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_dec_eq(x_6, x_3); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_isZero___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_Duration_isZero(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Duration_toMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Duration_toMilliseconds___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Std_Time_Duration_ofMillisecond___closed__1; +x_7 = lean_int_ediv(x_5, x_6); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +lean_dec(x_4); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Duration_instLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Duration_instDecidableLe(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_int_mul(x_8, x_4); +x_10 = lean_ctor_get(x_2, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_dec_le(x_7, x_11); +lean_dec(x_11); +lean_dec(x_7); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instDecidableLe___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Duration_instDecidableLe(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Duration_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Duration_toMinutes___closed__1; +x_4 = lean_int_ediv(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Duration_toDays___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toDays(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_Duration_toDays___closed__1; +x_4 = lean_int_ediv(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_toDays___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Duration_toDays(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_fromComponents(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = lean_int_add(x_4, x_2); +lean_dec(x_4); +x_6 = l_Std_Time_Duration_ofNanoseconds(x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_fromComponents___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_fromComponents(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_add(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 0); +x_9 = lean_int_mul(x_8, x_4); +x_10 = lean_ctor_get(x_2, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_add(x_7, x_11); +lean_dec(x_11); +lean_dec(x_7); +x_13 = l_Std_Time_Duration_ofNanoseconds(x_12); +lean_dec(x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_add___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_add(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_sub(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_neg(x_3); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_int_neg(x_5); +x_7 = lean_ctor_get(x_1, 0); +x_8 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_9 = lean_int_mul(x_7, x_8); +x_10 = lean_ctor_get(x_1, 1); +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_mul(x_4, x_8); +lean_dec(x_4); +x_13 = lean_int_add(x_12, x_6); +lean_dec(x_6); +lean_dec(x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_sub___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_sub(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_ofNanoseconds(x_2); +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_int_mul(x_9, x_5); +lean_dec(x_9); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_dec(x_3); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_8, x_12); +lean_dec(x_12); +lean_dec(x_8); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_3 = l_Std_Time_Duration_ofMillisecond___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Duration_ofNanoseconds(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_5, 0); +lean_inc(x_11); +x_12 = lean_int_mul(x_11, x_7); +lean_dec(x_11); +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_15 = lean_int_add(x_10, x_14); +lean_dec(x_14); +lean_dec(x_10); +x_16 = l_Std_Time_Duration_ofNanoseconds(x_15); +lean_dec(x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_3 = l_Std_Time_Duration_ofMillisecond___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = l_Std_Time_Duration_ofNanoseconds(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_int_neg(x_8); +lean_dec(x_8); +x_10 = lean_ctor_get(x_1, 0); +x_11 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_12 = lean_int_mul(x_10, x_11); +x_13 = lean_ctor_get(x_1, 1); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_12); +x_15 = lean_int_mul(x_7, x_11); +lean_dec(x_7); +x_16 = lean_int_add(x_15, x_9); +lean_dec(x_9); +lean_dec(x_15); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_Duration_ofNanoseconds(x_17); +lean_dec(x_17); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_3 = l_Std_Time_Duration_ofNanoseconds(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_dec(x_3); +x_7 = lean_int_neg(x_6); +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 0); +x_9 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_10 = lean_int_mul(x_8, x_9); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_mul(x_5, x_9); +lean_dec(x_5); +x_14 = lean_int_add(x_13, x_7); +lean_dec(x_7); +lean_dec(x_13); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_Duration_ofNanoseconds(x_15); +lean_dec(x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_mul(x_2, x_4); +x_9 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_add(x_7, x_10); +lean_dec(x_10); +lean_dec(x_7); +x_12 = l_Std_Time_Duration_ofNanoseconds(x_11); +lean_dec(x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_subSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = lean_int_neg(x_2); +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_10 = l_Std_Time_Duration_subSeconds___closed__1; +x_11 = lean_int_add(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_add(x_8, x_11); +lean_dec(x_11); +lean_dec(x_8); +x_13 = l_Std_Time_Duration_ofNanoseconds(x_12); +lean_dec(x_12); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_toMinutes___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Duration_toMinutes___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Duration_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_addHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_addHours___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Duration_addHours___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Duration_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addDays(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Duration_toDays___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Duration_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subDays(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_addWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(604800u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_3 = l_Std_Time_Duration_addWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_mul(x_4, x_6); +lean_dec(x_4); +x_11 = l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = lean_int_add(x_9, x_12); +lean_dec(x_12); +lean_dec(x_9); +x_14 = l_Std_Time_Duration_ofNanoseconds(x_13); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_addWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_addWeeks(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_Duration_addWeeks___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_Duration_subSeconds___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_subWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_subWeeks(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_addMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAddOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAddOffset__6___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_subMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSubOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSubOffset__6___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_sub___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_add___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instHAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instHAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_ofNanoseconds___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Duration_ofSeconds), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Minute_Offset_toSeconds___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Duration_instCoeOffset__1___closed__1; +x_2 = l_Std_Time_Duration_instCoeOffset__2___closed__1; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset__2___closed__2; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Hour_Offset_toSeconds___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Duration_instCoeOffset__1___closed__1; +x_2 = l_Std_Time_Duration_instCoeOffset__3___closed__1; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset__3___closed__2; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Day_Offset_toSeconds___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Week_Offset_toDays___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__4___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Duration_instCoeOffset__4___closed__1; +x_2 = l_Std_Time_Duration_instCoeOffset__4___closed__2; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__4___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Duration_instCoeOffset__1___closed__1; +x_2 = l_Std_Time_Duration_instCoeOffset__4___closed__3; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset__4___closed__4; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__5___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Duration_instCoeOffset__1___closed__1; +x_2 = l_Std_Time_Duration_instCoeOffset__4___closed__1; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Duration_instCoeOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Duration_instCoeOffset__5___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_mul(x_7, x_1); +lean_dec(x_7); +x_9 = l_Std_Time_Duration_ofNanoseconds(x_8); +lean_dec(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_instHMulInt(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_mul(x_7, x_2); +lean_dec(x_7); +x_9 = l_Std_Time_Duration_ofNanoseconds(x_8); +lean_dec(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHMulInt__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_instHMulInt__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddPlainTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_ofNanoseconds(x_9); +lean_dec(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHAddPlainTime___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_instHAddPlainTime(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubPlainTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_Duration_ofNanoseconds___closed__1; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_9 = lean_int_sub(x_7, x_8); +lean_dec(x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_ofNanoseconds(x_9); +lean_dec(x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Duration_instHSubPlainTime___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Duration_instHSubPlainTime(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Duration(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__1); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__2); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__3); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__4); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__5); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__6); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__7); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__8); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__9); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__10); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__11); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__12); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__13); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__14); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__15); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__16); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__17); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__18); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__19); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__20); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__21); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__22); +l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23 = _init_l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23(); +lean_mark_persistent(l___private_Std_Time_Duration_0__Std_Time_reprDuration____x40_Std_Time_Duration___hyg_59____closed__23); +l_Std_Time_instReprDuration___closed__1 = _init_l_Std_Time_instReprDuration___closed__1(); +lean_mark_persistent(l_Std_Time_instReprDuration___closed__1); +l_Std_Time_instReprDuration = _init_l_Std_Time_instReprDuration(); +lean_mark_persistent(l_Std_Time_instReprDuration); +l_Std_Time_instToStringDuration_leftPad___closed__1 = _init_l_Std_Time_instToStringDuration_leftPad___closed__1(); +lean_mark_persistent(l_Std_Time_instToStringDuration_leftPad___closed__1); +l_Std_Time_instToStringDuration___closed__1 = _init_l_Std_Time_instToStringDuration___closed__1(); +lean_mark_persistent(l_Std_Time_instToStringDuration___closed__1); +l_Std_Time_instToStringDuration___closed__2 = _init_l_Std_Time_instToStringDuration___closed__2(); +lean_mark_persistent(l_Std_Time_instToStringDuration___closed__2); +l_Std_Time_instToStringDuration___closed__3 = _init_l_Std_Time_instToStringDuration___closed__3(); +lean_mark_persistent(l_Std_Time_instToStringDuration___closed__3); +l_Std_Time_instInhabitedDuration___closed__1 = _init_l_Std_Time_instInhabitedDuration___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedDuration___closed__1); +l_Std_Time_instInhabitedDuration = _init_l_Std_Time_instInhabitedDuration(); +lean_mark_persistent(l_Std_Time_instInhabitedDuration); +l_Std_Time_Duration_ofNanoseconds___closed__1 = _init_l_Std_Time_Duration_ofNanoseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_ofNanoseconds___closed__1); +l_Std_Time_Duration_ofMillisecond___closed__1 = _init_l_Std_Time_Duration_ofMillisecond___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_ofMillisecond___closed__1); +l_Std_Time_Duration_toMilliseconds___closed__1 = _init_l_Std_Time_Duration_toMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_toMilliseconds___closed__1); +l_Std_Time_Duration_instLE = _init_l_Std_Time_Duration_instLE(); +lean_mark_persistent(l_Std_Time_Duration_instLE); +l_Std_Time_Duration_toMinutes___closed__1 = _init_l_Std_Time_Duration_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_toMinutes___closed__1); +l_Std_Time_Duration_toDays___closed__1 = _init_l_Std_Time_Duration_toDays___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_toDays___closed__1); +l_Std_Time_Duration_subSeconds___closed__1 = _init_l_Std_Time_Duration_subSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_subSeconds___closed__1); +l_Std_Time_Duration_addHours___closed__1 = _init_l_Std_Time_Duration_addHours___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_addHours___closed__1); +l_Std_Time_Duration_addWeeks___closed__1 = _init_l_Std_Time_Duration_addWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_addWeeks___closed__1); +l_Std_Time_Duration_instHAddOffset___closed__1 = _init_l_Std_Time_Duration_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset___closed__1); +l_Std_Time_Duration_instHAddOffset = _init_l_Std_Time_Duration_instHAddOffset(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset); +l_Std_Time_Duration_instHSubOffset___closed__1 = _init_l_Std_Time_Duration_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset___closed__1); +l_Std_Time_Duration_instHSubOffset = _init_l_Std_Time_Duration_instHSubOffset(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset); +l_Std_Time_Duration_instHAddOffset__1___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__1___closed__1); +l_Std_Time_Duration_instHAddOffset__1 = _init_l_Std_Time_Duration_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__1); +l_Std_Time_Duration_instHSubOffset__1___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__1___closed__1); +l_Std_Time_Duration_instHSubOffset__1 = _init_l_Std_Time_Duration_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__1); +l_Std_Time_Duration_instHAddOffset__2___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__2___closed__1); +l_Std_Time_Duration_instHAddOffset__2 = _init_l_Std_Time_Duration_instHAddOffset__2(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__2); +l_Std_Time_Duration_instHSubOffset__2___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__2___closed__1); +l_Std_Time_Duration_instHSubOffset__2 = _init_l_Std_Time_Duration_instHSubOffset__2(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__2); +l_Std_Time_Duration_instHAddOffset__3___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__3___closed__1); +l_Std_Time_Duration_instHAddOffset__3 = _init_l_Std_Time_Duration_instHAddOffset__3(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__3); +l_Std_Time_Duration_instHSubOffset__3___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__3___closed__1); +l_Std_Time_Duration_instHSubOffset__3 = _init_l_Std_Time_Duration_instHSubOffset__3(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__3); +l_Std_Time_Duration_instHAddOffset__4___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__4___closed__1); +l_Std_Time_Duration_instHAddOffset__4 = _init_l_Std_Time_Duration_instHAddOffset__4(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__4); +l_Std_Time_Duration_instHSubOffset__4___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__4___closed__1); +l_Std_Time_Duration_instHSubOffset__4 = _init_l_Std_Time_Duration_instHSubOffset__4(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__4); +l_Std_Time_Duration_instHAddOffset__5___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__5___closed__1); +l_Std_Time_Duration_instHAddOffset__5 = _init_l_Std_Time_Duration_instHAddOffset__5(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__5); +l_Std_Time_Duration_instHSubOffset__5___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__5___closed__1); +l_Std_Time_Duration_instHSubOffset__5 = _init_l_Std_Time_Duration_instHSubOffset__5(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__5); +l_Std_Time_Duration_instHAddOffset__6___closed__1 = _init_l_Std_Time_Duration_instHAddOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__6___closed__1); +l_Std_Time_Duration_instHAddOffset__6 = _init_l_Std_Time_Duration_instHAddOffset__6(); +lean_mark_persistent(l_Std_Time_Duration_instHAddOffset__6); +l_Std_Time_Duration_instHSubOffset__6___closed__1 = _init_l_Std_Time_Duration_instHSubOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__6___closed__1); +l_Std_Time_Duration_instHSubOffset__6 = _init_l_Std_Time_Duration_instHSubOffset__6(); +lean_mark_persistent(l_Std_Time_Duration_instHSubOffset__6); +l_Std_Time_Duration_instHSub___closed__1 = _init_l_Std_Time_Duration_instHSub___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHSub___closed__1); +l_Std_Time_Duration_instHSub = _init_l_Std_Time_Duration_instHSub(); +lean_mark_persistent(l_Std_Time_Duration_instHSub); +l_Std_Time_Duration_instHAdd___closed__1 = _init_l_Std_Time_Duration_instHAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instHAdd___closed__1); +l_Std_Time_Duration_instHAdd = _init_l_Std_Time_Duration_instHAdd(); +lean_mark_persistent(l_Std_Time_Duration_instHAdd); +l_Std_Time_Duration_instCoeOffset___closed__1 = _init_l_Std_Time_Duration_instCoeOffset___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset___closed__1); +l_Std_Time_Duration_instCoeOffset = _init_l_Std_Time_Duration_instCoeOffset(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset); +l_Std_Time_Duration_instCoeOffset__1___closed__1 = _init_l_Std_Time_Duration_instCoeOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__1___closed__1); +l_Std_Time_Duration_instCoeOffset__1 = _init_l_Std_Time_Duration_instCoeOffset__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__1); +l_Std_Time_Duration_instCoeOffset__2___closed__1 = _init_l_Std_Time_Duration_instCoeOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__2___closed__1); +l_Std_Time_Duration_instCoeOffset__2___closed__2 = _init_l_Std_Time_Duration_instCoeOffset__2___closed__2(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__2___closed__2); +l_Std_Time_Duration_instCoeOffset__2 = _init_l_Std_Time_Duration_instCoeOffset__2(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__2); +l_Std_Time_Duration_instCoeOffset__3___closed__1 = _init_l_Std_Time_Duration_instCoeOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__3___closed__1); +l_Std_Time_Duration_instCoeOffset__3___closed__2 = _init_l_Std_Time_Duration_instCoeOffset__3___closed__2(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__3___closed__2); +l_Std_Time_Duration_instCoeOffset__3 = _init_l_Std_Time_Duration_instCoeOffset__3(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__3); +l_Std_Time_Duration_instCoeOffset__4___closed__1 = _init_l_Std_Time_Duration_instCoeOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__4___closed__1); +l_Std_Time_Duration_instCoeOffset__4___closed__2 = _init_l_Std_Time_Duration_instCoeOffset__4___closed__2(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__4___closed__2); +l_Std_Time_Duration_instCoeOffset__4___closed__3 = _init_l_Std_Time_Duration_instCoeOffset__4___closed__3(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__4___closed__3); +l_Std_Time_Duration_instCoeOffset__4___closed__4 = _init_l_Std_Time_Duration_instCoeOffset__4___closed__4(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__4___closed__4); +l_Std_Time_Duration_instCoeOffset__4 = _init_l_Std_Time_Duration_instCoeOffset__4(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__4); +l_Std_Time_Duration_instCoeOffset__5___closed__1 = _init_l_Std_Time_Duration_instCoeOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__5___closed__1); +l_Std_Time_Duration_instCoeOffset__5 = _init_l_Std_Time_Duration_instCoeOffset__5(); +lean_mark_persistent(l_Std_Time_Duration_instCoeOffset__5); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Format.c b/stage0/stdlib/Std/Time/Format.c new file mode 100644 index 000000000000..6ba68d6fe67d --- /dev/null +++ b/stage0/stdlib/Std/Time/Format.c @@ -0,0 +1,7956 @@ +// Lean compiler output +// Module: Std.Time.Format +// Imports: Std.Time.Notation.Spec Std.Time.Format.Basic Std.Time.Internal.Bounded +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Formats_rfc850; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instRepr___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDate___closed__4; +lean_object* lean_int_mod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_time24Hour___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromDateTimeString(lean_object*); +static lean_object* l_Std_Time_PlainDate_format___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos; +static lean_object* l_Std_Time_Formats_americanDate___closed__1; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromRFC822String(lean_object*); +static lean_object* l_Std_Time_PlainDate_instRepr___closed__3; +static lean_object* l_Std_Time_Formats_iso8601___closed__25; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instRepr___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1; +static lean_object* l_Std_Time_TimeZone_fromTimeZone___closed__2; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDateTimeWithZoneString(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__24; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toLeanDateTimeWithZoneString(lean_object*); +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__4; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7; +static lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1; +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__2; +static lean_object* l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Formats_ascTime; +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +static lean_object* l_Std_Time_TimeZone_fromTimeZone___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_format(lean_object*, lean_object*); +extern lean_object* l_Std_Time_TimeZone_GMT; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__9; +static lean_object* l_Std_Time_Formats_ascTime___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromLeanDateTimeString(lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instToString; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTimeWithZone; +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithIdentifierString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDate; +static lean_object* l_Std_Time_PlainDate_format___lambda__1___closed__3; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__10; +static lean_object* l_Std_Time_Formats_iso8601___closed__28; +static lean_object* l_Std_Time_Formats_iso8601___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Formats_iso8601; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__6; +static lean_object* l_Std_Time_Formats_rfc822___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromSQLDateString(lean_object*); +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__14; +static lean_object* l_Std_Time_Formats_iso8601___closed__15; +static lean_object* l_Std_Time_ZonedDateTime_instToString___closed__1; +static lean_object* l_Std_Time_Formats_leanDate___closed__3; +static lean_object* l_Std_Time_Formats_iso8601___closed__4; +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_parse(lean_object*); +static lean_object* l_Std_Time_Formats_leanDate___closed__5; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithZoneString(lean_object*); +lean_object* l_Std_Time_PlainDate_dayOfYear(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__5; +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__3; +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__6; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6; +lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_fromAscTimeString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__9; +static lean_object* l_Std_Time_Formats_iso8601___closed__8; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__10; +static lean_object* l_Std_Time_Formats_time12Hour___closed__13; +static lean_object* l_Std_Time_TimeZone_Offset_fromOffset___closed__1; +static lean_object* l_Std_Time_PlainTime_fromTime12Hour___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toTime12Hour(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_format(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__23; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_parse(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__3; +lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object*, uint8_t); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__3; +static lean_object* l_Std_Time_Formats_iso8601___closed__19; +static lean_object* l_Std_Time_Formats_iso8601___closed__3; +static lean_object* l_Std_Time_Formats_americanDate___closed__3; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instToString; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainTime_instRepr___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__5; +static lean_object* l_Std_Time_Formats_iso8601___closed__29; +static lean_object* l_Std_Time_Formats_rfc822___closed__15; +static lean_object* l_Std_Time_Formats_ascTime___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime24Hour(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toRFC822String(lean_object*); +lean_object* l_Fin_ofNat(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc850___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_fromLongDateFormatString(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__8; +lean_object* l_Std_Time_HourMarker_toAbsolute(uint8_t, lean_object*); +static lean_object* l_Std_Time_PlainDate_fromSQLDateString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC850String___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromLeanDateString(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20; +lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_format___lambda__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__14; +static lean_object* l_Std_Time_PlainDate_format___lambda__1___closed__4; +static lean_object* l_Std_Time_Formats_time24Hour___closed__4; +static lean_object* l_Std_Time_Formats_iso8601___closed__30; +static lean_object* l_Std_Time_Formats_rfc822___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Formats_dateTime24Hour; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__11; +static lean_object* l_Std_Time_Formats_europeanDate___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__11; +static lean_object* l_Std_Time_Formats_ascTime___closed__15; +uint8_t l_instDecidableNot___rarg(uint8_t); +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__10; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__12; +static lean_object* l_Std_Time_Formats_time24Hour___closed__5; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5; +lean_object* l_Std_Time_PlainDateTime_weekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toISO8601String___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toAscTimeString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_parse(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instRepr(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Formats_sqlDate; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__2; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__6; +static lean_object* l_Std_Time_Formats_europeanDate___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toLongDateFormatString(lean_object*); +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__7; +static lean_object* l_Std_Time_Formats_ascTime___closed__12; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Std_Time_Formats_ascTime___closed__3; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16; +static lean_object* l_Std_Time_Formats_rfc822___closed__9; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toTime24Hour(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__12; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__31; +static lean_object* l_Std_Time_Formats_americanDate___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour(lean_object*); +static lean_object* l_Std_Time_Formats_rfc850___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_format(lean_object*, lean_object*); +lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__4; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__6; +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__22; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_Formats_time12Hour; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5; +static lean_object* l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__13; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__7; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__2; +static lean_object* l_Std_Time_Formats_iso8601___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toISO8601String(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLeanDateTimeWithZoneString___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTime24Hour; +static lean_object* l_Std_Time_Formats_iso8601___closed__17; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__15; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__1; +static lean_object* l_Std_Time_PlainDate_instToString___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLeanDateTimeWithZoneString(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_ascTime___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDateTimeWithZoneString___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instRepr___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__9; +static lean_object* l_Std_Time_Formats_leanDate___closed__1; +lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(lean_object*); +static lean_object* l_Std_Time_PlainDate_format___closed__2; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__8; +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14; +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC822String___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__32; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toAmericanDateString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Formats_longDateFormat; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toLeanDateString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_fromOffset(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17; +static lean_object* l_Std_Time_PlainDate_format___closed__3; +lean_object* lean_thunk_get_own(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10; +static lean_object* l_Std_Time_ZonedDateTime_instRepr___closed__1; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__16; +static lean_object* l_Std_Time_Formats_time12Hour___closed__5; +static lean_object* l_Std_Time_Formats_time12Hour___closed__8; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__15; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toISO8601String(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_fromOffset___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_format___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instRepr___closed__2; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_parse(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromISO8601String(lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_parse(lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toSQLDateString(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__8; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9; +static lean_object* l_Std_Time_Formats_iso8601___closed__33; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__14; +lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +static lean_object* l_Std_Time_PlainTime_fromTime24Hour___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__16; +static lean_object* l_Std_Time_Formats_ascTime___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDateTimeWithZoneString(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanTime24HourNoNanos; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toRFC850String(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime24Hour___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_ascTime___closed__5; +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toDateTimeString(lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__13; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3; +lean_object* lean_mk_thunk(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromAmericanDateString(lean_object*); +lean_object* l_Std_Time_GenericFormat_parse(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__17; +static lean_object* l_Std_Time_Formats_ascTime___closed__10; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1; +static lean_object* l_Std_Time_Formats_rfc822___closed__1; +lean_object* l_Std_Time_PlainDate_ofYearMonthDay_x3f(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString___lambda__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_time24Hour___closed__3; +static lean_object* l_Std_Time_Formats_ascTime___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC822String(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_longDateFormat___closed__15; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__3; +lean_object* l_Std_Time_GenericFormat_formatBuilder_go(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_time12Hour___closed__9; +lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_fromOffset___closed__4; +static lean_object* l_Std_Time_Formats_americanDate___closed__4; +static lean_object* l_Std_Time_Formats_rfc822___closed__5; +lean_object* l_Std_Time_GenericFormat_parseBuilder___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc850___closed__4; +static lean_object* l_Std_Time_PlainDate_instRepr___closed__2; +static lean_object* l_Std_Time_Formats_time12Hour___closed__7; +lean_object* l_Std_Time_GenericFormat_spec___rarg(lean_object*); +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__7; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__4; +static lean_object* l_Std_Time_Formats_ascTime___closed__13; +static lean_object* l_Std_Time_Formats_rfc850___closed__1; +static lean_object* l_Std_Time_ZonedDateTime_format___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_fromTimeZone(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9; +static lean_object* l_Std_Time_TimeZone_fromTimeZone___closed__1; +static lean_object* l_Std_Time_ZonedDateTime_format___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromDateTimeWithZoneString(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanTime24Hour; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__4; +static lean_object* l_Std_Time_Formats_time12Hour___closed__3; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__12; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_instRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instRepr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toLeanDateTimeWithIdentifierString(lean_object*); +static lean_object* l_Std_Time_PlainDate_instRepr___closed__1; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__6; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__12; +static lean_object* l_Std_Time_Formats_rfc822___closed__4; +static lean_object* l_Std_Time_Formats_rfc822___closed__10; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3; +static lean_object* l_Std_Time_Formats_dateTime24Hour___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromLongDateFormatString(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_Formats_dateTimeWithZone; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instToString; +uint8_t l_Std_Time_Year_Offset_era(lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__11; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__1; +lean_object* l_Std_Time_Hour_Ordinal_toRelative(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instRepr(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_ascTime___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instRepr(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_americanDate___closed__2; +static lean_object* l_Std_Time_Formats_iso8601___closed__5; +static lean_object* l_Std_Time_Formats_ascTime___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__1; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__20; +LEAN_EXPORT lean_object* l_Std_Time_Formats_americanDate; +static lean_object* l_Std_Time_PlainDateTime_instRepr___closed__2; +static lean_object* l_Std_Time_PlainDate_format___closed__1; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4; +extern lean_object* l_Std_Time_TimeZone_UTC; +static lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2; +static lean_object* l_Std_Time_Formats_iso8601___closed__21; +static lean_object* l_Std_Time_Formats_time12Hour___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLongDateFormatString(lean_object*); +static lean_object* l_Std_Time_Formats_time12Hour___closed__6; +static lean_object* l_Std_Time_Formats_time12Hour___closed__2; +static lean_object* l_Std_Time_Formats_europeanDate___closed__2; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__7; +static lean_object* l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5; +static lean_object* l_Std_Time_TimeZone_fromTimeZone___closed__3; +static lean_object* l_Std_Time_Formats_ascTime___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instToString(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7; +static lean_object* l_Std_Time_Formats_time24Hour___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instToString; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instRepr(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC850String(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_time12Hour___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_format___lambda__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZone___closed__14; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instRepr(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9; +static lean_object* l_Std_Time_PlainTime_instToString___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__10; +static lean_object* l_Std_Time_Formats_iso8601___closed__27; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18; +uint8_t l_Std_Time_HourMarker_ofOrdinal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +static lean_object* l_Std_Time_Formats_ascTime___closed__18; +LEAN_EXPORT lean_object* l_Std_Time_Formats_europeanDate; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDate_instRepr___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_fromOffset___lambda__1(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6; +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc822___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Formats_rfc822; +static lean_object* l_Std_Time_Formats_time12Hour___closed__10; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12; +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainTime_format___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_fromTimeZone___lambda__1(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromAscTimeString(lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_fromOffset___closed__2; +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc850___closed__2; +static lean_object* l_Std_Time_Formats_iso8601___closed__18; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instRepr___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromAmericanDateString___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_iso8601___closed__2; +static lean_object* l_Std_Time_Formats_ascTime___closed__8; +static lean_object* l_Std_Time_PlainDate_format___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toLeanTime24Hour(lean_object*); +static lean_object* l_Std_Time_Formats_time12Hour___closed__1; +static lean_object* l_Std_Time_Formats_rfc822___closed__12; +static lean_object* l_Std_Time_Formats_time12Hour___closed__11; +lean_object* l_Std_Time_GenericFormat_formatGeneric_go(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTime24HourNoNanos; +LEAN_EXPORT lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos; +static lean_object* l_Std_Time_Formats_rfc822___closed__13; +static lean_object* l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainTime_instRepr___closed__2; +static lean_object* l_Std_Time_Formats_ascTime___closed__2; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toLeanDateTimeString(lean_object*); +lean_object* l_Std_Time_PlainDate_weekOfMonth(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12; +static lean_object* l_Std_Time_Formats_iso8601___closed__35; +static lean_object* l_Std_Time_Formats_dateTimeWithZone___closed__10; +lean_object* l_Std_Time_GenericFormat_format(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Formats_rfc850___closed__3; +static lean_object* l_Std_Time_Formats_iso8601___closed__34; +static lean_object* l_Std_Time_Formats_leanDateTime24Hour___closed__3; +static lean_object* l_Std_Time_Formats_ascTime___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromRFC850String(lean_object*); +static lean_object* l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1; +static lean_object* l_Std_Time_PlainDate_fromAmericanDateString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Formats_time24Hour; +static lean_object* l_Std_Time_Formats_longDateFormat___closed__12; +static lean_object* l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8; +static lean_object* l_Std_Time_Formats_ascTime___closed__17; +static lean_object* l_Std_Time_Formats_leanDate___closed__2; +static lean_object* l_Std_Time_PlainDateTime_instToString___closed__1; +static lean_object* l_Std_Time_Formats_iso8601___closed__26; +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(1); +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__5; +x_2 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__6; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__8; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("T", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(16, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__12; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__14; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(17, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__16; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__18; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__20; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__22() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(28, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_iso8601___closed__22; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_iso8601___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_iso8601___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__19; +x_2 = l_Std_Time_Formats_iso8601___closed__25; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_iso8601___closed__26; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_iso8601___closed__27; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_iso8601___closed__28; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_iso8601___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_iso8601___closed__30; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_iso8601___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_iso8601___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_iso8601___closed__33; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_iso8601___closed__34; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_iso8601() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_iso8601___closed__35; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_iso8601___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_americanDate___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_americanDate___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_americanDate___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_americanDate___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_americanDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_americanDate___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_europeanDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_americanDate___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_europeanDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_europeanDate___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_europeanDate___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_europeanDate___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_europeanDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_europeanDate___closed__3; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_ctor(13, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_time12Hour___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" ", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_time12Hour___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(12, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_time12Hour___closed__5; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_time12Hour___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_time12Hour___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_time12Hour___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_time12Hour___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_time12Hour___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_time12Hour___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__2; +x_2 = l_Std_Time_Formats_time12Hour___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time12Hour() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_time12Hour___closed__13; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_iso8601___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_time24Hour___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_time24Hour___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_time24Hour___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_time24Hour___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_time24Hour() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_time24Hour___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(19, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__19; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTime24Hour() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__15; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__2; +x_2 = l_Std_Time_Formats_iso8601___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__19; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_dateTimeWithZone___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_dateTimeWithZone() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_dateTimeWithZone___closed__13; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanTime24Hour() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__9; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanTime24HourNoNanos() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_time24Hour___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_dateTime24Hour___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTime24Hour___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTime24Hour___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTime24Hour___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTime24Hour___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTime24Hour___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24Hour() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTime24Hour___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_time24Hour___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTime24HourNoNanos() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 2; +x_2 = lean_alloc_ctor(28, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_leanDateTimeWithZone___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__19; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZone() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTimeWithZone___closed__16; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(24, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifier() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_dateTime24Hour___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__19; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__11; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_Formats_iso8601___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDate___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_leanDate___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_leanDate___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_leanDate___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_leanDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDate___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_sqlDate() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_leanDate___closed__5; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(9, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_longDateFormat___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", ", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_longDateFormat___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_longDateFormat___closed__5; +x_2 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_longDateFormat___closed__6; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_longDateFormat___closed__8; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_time24Hour___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_longDateFormat___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__4; +x_2 = l_Std_Time_Formats_longDateFormat___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__9; +x_2 = l_Std_Time_Formats_longDateFormat___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_longDateFormat___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__7; +x_2 = l_Std_Time_Formats_longDateFormat___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__4; +x_2 = l_Std_Time_Formats_longDateFormat___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__2; +x_2 = l_Std_Time_Formats_longDateFormat___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_longDateFormat() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_longDateFormat___closed__17; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(9, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_ascTime___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_ascTime___closed__3; +x_2 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_ascTime___closed__4; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_ascTime___closed__6; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_americanDate___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_ascTime___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_ascTime___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_ascTime___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_ascTime___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_ascTime___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_ascTime___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_ascTime___closed__7; +x_2 = l_Std_Time_Formats_ascTime___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_ascTime___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_ascTime___closed__5; +x_2 = l_Std_Time_Formats_ascTime___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_ascTime___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_ascTime___closed__2; +x_2 = l_Std_Time_Formats_ascTime___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_ascTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_ascTime___closed__19; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_ascTime___closed__3; +x_2 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Formats_rfc822___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_iso8601___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__21; +x_2 = l_Std_Time_Formats_rfc822___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_rfc822___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__17; +x_2 = l_Std_Time_Formats_rfc822___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__15; +x_2 = l_Std_Time_Formats_rfc822___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__13; +x_2 = l_Std_Time_Formats_rfc822___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_rfc822___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__2; +x_2 = l_Std_Time_Formats_rfc822___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_rfc822___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_ascTime___closed__5; +x_2 = l_Std_Time_Formats_rfc822___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_rfc822___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_rfc822___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__4; +x_2 = l_Std_Time_Formats_rfc822___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_rfc822___closed__2; +x_2 = l_Std_Time_Formats_rfc822___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc822() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_rfc822___closed__16; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_rfc822___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__7; +x_2 = l_Std_Time_Formats_rfc850___closed__1; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__4; +x_2 = l_Std_Time_Formats_rfc850___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_iso8601___closed__9; +x_2 = l_Std_Time_Formats_rfc850___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_longDateFormat___closed__4; +x_2 = l_Std_Time_Formats_rfc850___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_rfc822___closed__2; +x_2 = l_Std_Time_Formats_rfc850___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Formats_rfc850() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Formats_rfc850___closed__6; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_fromTimeZone___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; +x_3 = 1; +lean_inc(x_2); +x_4 = l_Std_Time_TimeZone_Offset_toIsoString(x_2, x_3); +x_5 = 0; +x_6 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +static lean_object* _init_l_Std_Time_TimeZone_fromTimeZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(23); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_fromTimeZone___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Formats_time12Hour___closed__4; +x_2 = l_Std_Time_Formats_leanDateTimeWithZone___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_fromTimeZone___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_TimeZone_fromTimeZone___closed__1; +x_2 = l_Std_Time_TimeZone_fromTimeZone___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_fromTimeZone___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_fromTimeZone___lambda__1), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_fromTimeZone(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_TimeZone_fromTimeZone___closed__3; +x_3 = l_Std_Time_TimeZone_fromTimeZone___closed__4; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_fromOffset___lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 2; +x_2 = lean_alloc_ctor(27, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_TimeZone_Offset_fromOffset___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_TimeZone_Offset_fromOffset___closed__2; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_Offset_fromOffset___lambda__1), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_fromOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_TimeZone_Offset_fromOffset___closed__3; +x_3 = l_Std_Time_TimeZone_Offset_fromOffset___closed__4; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_format___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Std_Time_Year_Offset_era(x_3); +lean_dec(x_3); +x_5 = lean_box(x_4); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +case 1: +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_2); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 0); +lean_dec(x_8); +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_9); +return x_2; +} +else +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_2); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +case 2: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_2); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_2, 0); +lean_dec(x_13); +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_14); +return x_2; +} +else +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_2); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; +} +} +case 3: +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_2); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_18 = lean_ctor_get(x_2, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +x_20 = l_Std_Time_PlainDate_format___lambda__1___closed__1; +x_21 = lean_int_mod(x_19, x_20); +x_22 = l_Std_Time_PlainDate_format___lambda__1___closed__2; +x_23 = lean_int_dec_eq(x_21, x_22); +lean_dec(x_21); +x_24 = l_Std_Time_PlainDate_dayOfYear(x_1); +lean_dec(x_1); +if (x_23 == 0) +{ +uint8_t x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_19); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_27); +return x_2; +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; +x_28 = l_Std_Time_PlainDate_format___lambda__1___closed__3; +x_29 = lean_int_mod(x_19, x_28); +x_30 = lean_int_dec_eq(x_29, x_22); +lean_dec(x_29); +x_31 = l_instDecidableNot___rarg(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = l_Std_Time_PlainDate_format___lambda__1___closed__4; +x_33 = lean_int_mod(x_19, x_32); +lean_dec(x_19); +x_34 = lean_int_dec_eq(x_33, x_22); +lean_dec(x_33); +if (x_34 == 0) +{ +uint8_t x_35; lean_object* x_36; lean_object* x_37; +x_35 = 0; +x_36 = lean_box(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_24); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_37); +return x_2; +} +else +{ +uint8_t x_38; lean_object* x_39; lean_object* x_40; +x_38 = 1; +x_39 = lean_box(x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_24); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_40); +return x_2; +} +} +else +{ +uint8_t x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_19); +x_41 = 1; +x_42 = lean_box(x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_43); +return x_2; +} +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; +lean_dec(x_2); +x_44 = lean_ctor_get(x_1, 0); +lean_inc(x_44); +x_45 = l_Std_Time_PlainDate_format___lambda__1___closed__1; +x_46 = lean_int_mod(x_44, x_45); +x_47 = l_Std_Time_PlainDate_format___lambda__1___closed__2; +x_48 = lean_int_dec_eq(x_46, x_47); +lean_dec(x_46); +x_49 = l_Std_Time_PlainDate_dayOfYear(x_1); +lean_dec(x_1); +if (x_48 == 0) +{ +uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_dec(x_44); +x_50 = 0; +x_51 = lean_box(x_50); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_49); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_52); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; uint8_t x_56; uint8_t x_57; +x_54 = l_Std_Time_PlainDate_format___lambda__1___closed__3; +x_55 = lean_int_mod(x_44, x_54); +x_56 = lean_int_dec_eq(x_55, x_47); +lean_dec(x_55); +x_57 = l_instDecidableNot___rarg(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = l_Std_Time_PlainDate_format___lambda__1___closed__4; +x_59 = lean_int_mod(x_44, x_58); +lean_dec(x_44); +x_60 = lean_int_dec_eq(x_59, x_47); +lean_dec(x_59); +if (x_60 == 0) +{ +uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = 0; +x_62 = lean_box(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_49); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; +} +else +{ +uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = 1; +x_66 = lean_box(x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_49); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_67); +return x_68; +} +} +else +{ +uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_44); +x_69 = 1; +x_70 = lean_box(x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_49); +x_72 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_72, 0, x_71); +return x_72; +} +} +} +} +case 4: +{ +uint8_t x_73; +x_73 = !lean_is_exclusive(x_2); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_2, 0); +lean_dec(x_74); +x_75 = lean_ctor_get(x_1, 1); +lean_inc(x_75); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_75); +return x_2; +} +else +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_2); +x_76 = lean_ctor_get(x_1, 1); +lean_inc(x_76); +lean_dec(x_1); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_76); +return x_77; +} +} +case 5: +{ +uint8_t x_78; +x_78 = !lean_is_exclusive(x_2); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_2, 0); +lean_dec(x_79); +x_80 = lean_ctor_get(x_1, 2); +lean_inc(x_80); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_80); +return x_2; +} +else +{ +lean_object* x_81; lean_object* x_82; +lean_dec(x_2); +x_81 = lean_ctor_get(x_1, 2); +lean_inc(x_81); +lean_dec(x_1); +x_82 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_82, 0, x_81); +return x_82; +} +} +case 6: +{ +uint8_t x_83; +x_83 = !lean_is_exclusive(x_2); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_2, 0); +lean_dec(x_84); +x_85 = l_Std_Time_PlainDate_quarter(x_1); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_85); +return x_2; +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_2); +x_86 = l_Std_Time_PlainDate_quarter(x_1); +lean_dec(x_1); +x_87 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_87, 0, x_86); +return x_87; +} +} +case 7: +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_2); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; +x_89 = lean_ctor_get(x_2, 0); +lean_dec(x_89); +x_90 = l_Std_Time_PlainDate_weekOfYear(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_90); +return x_2; +} +else +{ +lean_object* x_91; lean_object* x_92; +lean_dec(x_2); +x_91 = l_Std_Time_PlainDate_weekOfYear(x_1); +x_92 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_92, 0, x_91); +return x_92; +} +} +case 8: +{ +uint8_t x_93; +x_93 = !lean_is_exclusive(x_2); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_2, 0); +lean_dec(x_94); +x_95 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_95); +return x_2; +} +else +{ +lean_object* x_96; lean_object* x_97; +lean_dec(x_2); +x_96 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_1); +x_97 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_97, 0, x_96); +return x_97; +} +} +case 9: +{ +uint8_t x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_2); +x_98 = l_Std_Time_PlainDate_weekday(x_1); +x_99 = lean_box(x_98); +x_100 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_100, 0, x_99); +return x_100; +} +case 10: +{ +uint8_t x_101; +x_101 = !lean_is_exclusive(x_2); +if (x_101 == 0) +{ +lean_object* x_102; uint8_t x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_2, 0); +lean_dec(x_102); +x_103 = l_Std_Time_PlainDate_weekday(x_1); +x_104 = lean_box(x_103); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_104); +return x_2; +} +else +{ +uint8_t x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_2); +x_105 = l_Std_Time_PlainDate_weekday(x_1); +x_106 = lean_box(x_105); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_106); +return x_107; +} +} +case 11: +{ +uint8_t x_108; +x_108 = !lean_is_exclusive(x_2); +if (x_108 == 0) +{ +lean_object* x_109; lean_object* x_110; +x_109 = lean_ctor_get(x_2, 0); +lean_dec(x_109); +x_110 = l_Std_Time_PlainDate_weekOfMonth(x_1); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_110); +return x_2; +} +else +{ +lean_object* x_111; lean_object* x_112; +lean_dec(x_2); +x_111 = l_Std_Time_PlainDate_weekOfMonth(x_1); +lean_dec(x_1); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +return x_112; +} +} +default: +{ +lean_object* x_113; +lean_dec(x_2); +lean_dec(x_1); +x_113 = lean_box(0); +return x_113; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("error: ", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_format___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid time", 12, 12); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_spec___rarg(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_format___closed__1; +x_6 = lean_string_append(x_5, x_4); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_format___closed__2; +x_8 = lean_string_append(x_6, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_format___lambda__1), 2, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Std_Time_PlainDate_format___closed__2; +x_12 = l_Std_Time_GenericFormat_formatGeneric_go(x_10, x_11, x_9); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = l_Std_Time_PlainDate_format___closed__3; +return x_13; +} +else +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromAmericanDateString___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = l_Std_Time_PlainDate_format___lambda__1___closed__1; +x_5 = lean_int_mod(x_3, x_4); +x_6 = l_Std_Time_PlainDate_format___lambda__1___closed__2; +x_7 = lean_int_dec_eq(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; uint8_t x_10; +x_8 = 0; +x_9 = l_Std_Time_Month_Ordinal_days(x_8, x_1); +x_10 = lean_int_dec_le(x_2, x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_11 = lean_box(0); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_1); +lean_ctor_set(x_12, 2, x_2); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; +x_14 = l_Std_Time_PlainDate_format___lambda__1___closed__3; +x_15 = lean_int_mod(x_3, x_14); +x_16 = lean_int_dec_eq(x_15, x_6); +lean_dec(x_15); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = l_Std_Time_PlainDate_format___lambda__1___closed__4; +x_19 = lean_int_mod(x_3, x_18); +x_20 = lean_int_dec_eq(x_19, x_6); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; lean_object* x_22; uint8_t x_23; +x_21 = 0; +x_22 = l_Std_Time_Month_Ordinal_days(x_21, x_1); +x_23 = lean_int_dec_le(x_2, x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +lean_object* x_24; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_box(0); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_3); +lean_ctor_set(x_25, 1, x_1); +lean_ctor_set(x_25, 2, x_2); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +} +else +{ +uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_27 = 1; +x_28 = l_Std_Time_Month_Ordinal_days(x_27, x_1); +x_29 = lean_int_dec_le(x_2, x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +lean_object* x_30; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_30 = lean_box(0); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_31, 0, x_3); +lean_ctor_set(x_31, 1, x_1); +lean_ctor_set(x_31, 2, x_2); +x_32 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_31); +return x_32; +} +} +} +else +{ +uint8_t x_33; lean_object* x_34; uint8_t x_35; +x_33 = 1; +x_34 = l_Std_Time_Month_Ordinal_days(x_33, x_1); +x_35 = lean_int_dec_le(x_2, x_34); +lean_dec(x_34); +if (x_35 == 0) +{ +lean_object* x_36; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_box(0); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_37, 0, x_3); +lean_ctor_set(x_37, 1, x_1); +lean_ctor_set(x_37, 2, x_2); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +return x_38; +} +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDate_fromAmericanDateString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_fromAmericanDateString___lambda__1), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromAmericanDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_americanDate; +x_3 = l_Std_Time_PlainDate_fromAmericanDateString___closed__1; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toAmericanDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 2); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_format___closed__2; +x_6 = l_Std_Time_Formats_americanDate; +x_7 = l_Std_Time_GenericFormat_formatBuilder_go(x_5, x_6); +x_8 = lean_apply_3(x_7, x_2, x_3, x_4); +return x_8; +} +} +static lean_object* _init_l_Std_Time_PlainDate_fromSQLDateString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_ofYearMonthDay_x3f), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromSQLDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_sqlDate; +x_3 = l_Std_Time_PlainDate_fromSQLDateString___closed__1; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toSQLDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_format___closed__2; +x_6 = l_Std_Time_Formats_sqlDate; +x_7 = l_Std_Time_GenericFormat_formatBuilder_go(x_5, x_6); +x_8 = lean_apply_3(x_7, x_2, x_3, x_4); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_fromLeanDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_leanDate; +x_3 = l_Std_Time_PlainDate_fromSQLDateString___closed__1; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toLeanDateString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_format___closed__2; +x_6 = l_Std_Time_Formats_leanDate; +x_7 = l_Std_Time_GenericFormat_formatBuilder_go(x_5, x_6); +x_8 = lean_apply_3(x_7, x_2, x_3, x_4); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_americanDate; +x_3 = l_Std_Time_PlainDate_fromAmericanDateString___closed__1; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_sqlDate; +x_6 = l_Std_Time_PlainDate_fromSQLDateString___closed__1; +x_7 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_5, x_6, x_1); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +lean_dec(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDate_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_toLeanDateString), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDate_instToString___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("date(\"", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instRepr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDate_instRepr___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instRepr___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("\")", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDate_instRepr___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDate_instRepr___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instRepr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = l_Std_Time_PlainDate_toLeanDateString(x_1); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = l_Std_Time_PlainDate_instRepr___closed__2; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDate_instRepr___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_instRepr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_instRepr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainTime_format___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 12: +{ +lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_HourMarker_ofOrdinal(x_3); +x_5 = lean_box(x_4); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +case 13: +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_2); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_2, 0); +lean_dec(x_8); +x_9 = lean_ctor_get(x_1, 0); +x_10 = l_Std_Time_Hour_Ordinal_toRelative(x_9); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_10); +return x_2; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_2); +x_11 = lean_ctor_get(x_1, 0); +x_12 = l_Std_Time_Hour_Ordinal_toRelative(x_11); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +case 14: +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_2, 0); +lean_dec(x_15); +x_16 = lean_ctor_get(x_1, 0); +x_17 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_18 = lean_int_emod(x_16, x_17); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_18); +return x_2; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_2); +x_19 = lean_ctor_get(x_1, 0); +x_20 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_21 = lean_int_emod(x_19, x_20); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +case 15: +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_2, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_1, 0); +x_26 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_25); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_26); +return x_2; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_2); +x_27 = lean_ctor_get(x_1, 0); +x_28 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_27); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +return x_29; +} +} +case 16: +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_2); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_2, 0); +lean_dec(x_31); +x_32 = lean_ctor_get(x_1, 0); +lean_inc(x_32); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_32); +return x_2; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_2); +x_33 = lean_ctor_get(x_1, 0); +lean_inc(x_33); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; +} +} +case 17: +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_2); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_2, 0); +lean_dec(x_36); +x_37 = lean_ctor_get(x_1, 1); +lean_inc(x_37); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_37); +return x_2; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_2); +x_38 = lean_ctor_get(x_1, 1); +lean_inc(x_38); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_38); +return x_39; +} +} +case 18: +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_2); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_2, 0); +lean_dec(x_41); +x_42 = lean_ctor_get(x_1, 2); +lean_inc(x_42); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_42); +return x_2; +} +else +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_2); +x_43 = lean_ctor_get(x_1, 2); +lean_inc(x_43); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_43); +return x_44; +} +} +case 19: +{ +uint8_t x_45; +x_45 = !lean_is_exclusive(x_2); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_2, 0); +lean_dec(x_46); +x_47 = lean_ctor_get(x_1, 3); +lean_inc(x_47); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_47); +return x_2; +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_2); +x_48 = lean_ctor_get(x_1, 3); +lean_inc(x_48); +x_49 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_49, 0, x_48); +return x_49; +} +} +case 20: +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_2); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_2, 0); +lean_dec(x_51); +x_52 = l_Std_Time_PlainTime_toMilliseconds(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_52); +return x_2; +} +else +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_2); +x_53 = l_Std_Time_PlainTime_toMilliseconds(x_1); +x_54 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_54, 0, x_53); +return x_54; +} +} +case 21: +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_2); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_2, 0); +lean_dec(x_56); +x_57 = lean_ctor_get(x_1, 3); +lean_inc(x_57); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_57); +return x_2; +} +else +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_2); +x_58 = lean_ctor_get(x_1, 3); +lean_inc(x_58); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_58); +return x_59; +} +} +case 22: +{ +uint8_t x_60; +x_60 = !lean_is_exclusive(x_2); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_2, 0); +lean_dec(x_61); +x_62 = l_Std_Time_PlainTime_toNanoseconds(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_62); +return x_2; +} +else +{ +lean_object* x_63; lean_object* x_64; +lean_dec(x_2); +x_63 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_64 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_64, 0, x_63); +return x_64; +} +} +default: +{ +lean_object* x_65; +lean_dec(x_2); +x_65 = lean_box(0); +return x_65; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_spec___rarg(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_format___closed__1; +x_6 = lean_string_append(x_5, x_4); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_format___closed__2; +x_8 = lean_string_append(x_6, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_format___lambda__1___boxed), 2, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Std_Time_PlainDate_format___closed__2; +x_12 = l_Std_Time_GenericFormat_formatGeneric_go(x_10, x_11, x_9); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = l_Std_Time_PlainDate_format___closed__3; +return x_13; +} +else +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_format___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_format___lambda__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(999999999u); +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Fin_ofNat(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime24Hour___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2; +x_6 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_3, 0); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_3); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +x_11 = l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2; +x_12 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_2); +lean_ctor_set(x_12, 2, x_10); +lean_ctor_set(x_12, 3, x_11); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromTime24Hour___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_fromTime24Hour___lambda__1), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime24Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_time24Hour; +x_3 = l_Std_Time_PlainTime_fromTime24Hour___closed__1; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toTime24Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_PlainDate_format___closed__2; +x_6 = l_Std_Time_Formats_time24Hour; +x_7 = l_Std_Time_GenericFormat_formatBuilder_go(x_5, x_6); +x_8 = lean_apply_3(x_7, x_2, x_3, x_4); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_4); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_3, 0); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_3); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +x_11 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_2); +lean_ctor_set(x_11, 2, x_10); +lean_ctor_set(x_11, 3, x_4); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +return x_12; +} +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_fromLeanTime24Hour___lambda__1), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_leanTime24Hour; +x_3 = l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_leanTime24HourNoNanos; +x_6 = l_Std_Time_PlainTime_fromTime24Hour___closed__1; +x_7 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_5, x_6, x_1); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +lean_dec(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toLeanTime24Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 3); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l_Std_Time_PlainDate_format___closed__2; +x_7 = l_Std_Time_Formats_leanTime24Hour; +x_8 = l_Std_Time_GenericFormat_formatBuilder_go(x_6, x_7); +x_9 = lean_apply_4(x_8, x_2, x_3, x_4, x_5); +return x_9; +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1; +x_6 = lean_int_dec_le(x_5, x_1); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_3); +lean_dec(x_2); +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; +x_8 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_9 = lean_int_dec_le(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_3); +lean_dec(x_2); +x_10 = lean_box(0); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_3); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_Std_Time_HourMarker_toAbsolute(x_4, x_1); +x_13 = l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2; +x_14 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_2); +lean_ctor_set(x_14, 2, x_3); +lean_ctor_set(x_14, 3, x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_3); +x_18 = l_Std_Time_HourMarker_toAbsolute(x_4, x_1); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_16); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2; +x_21 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_2); +lean_ctor_set(x_21, 2, x_19); +lean_ctor_set(x_21, 3, x_20); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +} +} +} +static lean_object* _init_l_Std_Time_PlainTime_fromTime12Hour___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_fromTime12Hour___lambda__1___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_time12Hour; +x_3 = l_Std_Time_PlainTime_fromTime12Hour___closed__1; +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_fromTime12Hour___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_4); +lean_dec(x_4); +x_6 = l_Std_Time_PlainTime_fromTime12Hour___lambda__1(x_1, x_2, x_3, x_5); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toTime12Hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_4 = lean_int_emod(x_2, x_3); +x_5 = l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1; +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_int_dec_le(x_3, x_2); +lean_dec(x_2); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = l_Std_Time_PlainDate_format___closed__2; +x_11 = l_Std_Time_Formats_time12Hour; +x_12 = 0; +x_13 = l_Std_Time_GenericFormat_formatBuilder_go(x_10, x_11); +x_14 = lean_box(x_12); +x_15 = lean_apply_4(x_13, x_6, x_7, x_8, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = l_Std_Time_PlainDate_format___closed__2; +x_17 = l_Std_Time_Formats_time12Hour; +x_18 = 1; +x_19 = l_Std_Time_GenericFormat_formatBuilder_go(x_16, x_17); +x_20 = lean_box(x_18); +x_21 = lean_apply_4(x_19, x_6, x_7, x_8, x_20); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Formats_time12Hour; +x_3 = l_Std_Time_PlainTime_fromTime12Hour___closed__1; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_time24Hour; +x_6 = l_Std_Time_PlainTime_fromTime24Hour___closed__1; +x_7 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_5, x_6, x_1); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +lean_dec(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainTime_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_toLeanTime24Hour), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instToString___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("time(\"", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instRepr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainTime_instRepr___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instRepr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = l_Std_Time_PlainTime_toLeanTime24Hour(x_1); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = l_Std_Time_PlainTime_instRepr___closed__2; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDate_instRepr___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instRepr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_instRepr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_format___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_format___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_5); +lean_dec(x_7); +x_9 = l_Std_Time_ZonedDateTime_format___lambda__1___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l_Std_Time_ZonedDateTime_format___lambda__1___closed__2; +x_13 = lean_int_mul(x_5, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_spec___rarg(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_format___closed__1; +x_6 = lean_string_append(x_5, x_4); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_format___closed__2; +x_8 = lean_string_append(x_6, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_10); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +lean_dec(x_1); +lean_inc(x_10); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_format___lambda__1___boxed), 3, 2); +lean_closure_set(x_12, 0, x_11); +lean_closure_set(x_12, 1, x_10); +x_13 = lean_mk_thunk(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_11); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_box(1); +x_16 = l_Std_Time_GenericFormat_format(x_15, x_10, x_9, x_14); +lean_dec(x_10); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_format___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_ZonedDateTime_format___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromISO8601String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_iso8601; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toISO8601String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +lean_inc(x_2); +lean_inc(x_3); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_format___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_3); +lean_closure_set(x_4, 1, x_2); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_box(1); +x_8 = l_Std_Time_Formats_iso8601; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_2, x_8, x_6); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromRFC822String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_rfc822; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toRFC822String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +lean_inc(x_2); +lean_inc(x_3); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_format___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_3); +lean_closure_set(x_4, 1, x_2); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_box(1); +x_8 = l_Std_Time_Formats_rfc822; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_2, x_8, x_6); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromRFC850String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_rfc850; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toRFC850String(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +lean_inc(x_2); +lean_inc(x_3); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_format___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_3); +lean_closure_set(x_4, 1, x_2); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_box(1); +x_8 = l_Std_Time_Formats_rfc850; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_2, x_8, x_6); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromDateTimeWithZoneString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_dateTimeWithZone; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDateTimeWithZoneString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_2 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +lean_inc(x_2); +lean_inc(x_3); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_format___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_3); +lean_closure_set(x_4, 1, x_2); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_box(1); +x_8 = l_Std_Time_Formats_dateTimeWithZone; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_2, x_8, x_6); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithZoneString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_leanDateTimeWithZone; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_leanDateTimeWithZoneNoNanos; +x_6 = l_Std_Time_GenericFormat_parse(x_2, x_5, x_1); +return x_6; +} +else +{ +uint8_t x_7; +lean_dec(x_1); +x_7 = !lean_is_exclusive(x_4); +if (x_7 == 0) +{ +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +lean_dec(x_4); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithIdentifierString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_leanDateTimeWithIdentifier; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos; +x_6 = l_Std_Time_GenericFormat_parse(x_2, x_5, x_1); +return x_6; +} +else +{ +uint8_t x_7; +lean_dec(x_1); +x_7 = !lean_is_exclusive(x_4); +if (x_7 == 0) +{ +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +lean_dec(x_4); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toLeanDateTimeWithZoneString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_thunk_get_own(x_2); +lean_dec(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_4, 2); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_dec(x_3); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_8, 3); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_ctor_get(x_1, 3); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Std_Time_PlainDate_format___closed__2; +x_16 = l_Std_Time_Formats_leanDateTimeWithZone; +x_17 = l_Std_Time_GenericFormat_formatBuilder_go(x_15, x_16); +x_18 = lean_apply_8(x_17, x_5, x_6, x_7, x_9, x_10, x_11, x_12, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toLeanDateTimeWithIdentifierString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_thunk_get_own(x_2); +lean_dec(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = lean_ctor_get(x_4, 2); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_dec(x_3); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_8, 3); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_ctor_get(x_1, 3); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_Std_Time_PlainDate_format___closed__2; +x_16 = l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos; +x_17 = l_Std_Time_GenericFormat_formatBuilder_go(x_15, x_16); +x_18 = lean_apply_8(x_17, x_5, x_6, x_7, x_9, x_10, x_11, x_12, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_box(1); +x_3 = l_Std_Time_Formats_iso8601; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_rfc822; +lean_inc(x_1); +x_6 = l_Std_Time_GenericFormat_parse(x_2, x_5, x_1); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_6); +x_7 = l_Std_Time_Formats_rfc850; +lean_inc(x_1); +x_8 = l_Std_Time_GenericFormat_parse(x_2, x_7, x_1); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; +lean_dec(x_8); +x_9 = l_Std_Time_Formats_dateTimeWithZone; +lean_inc(x_1); +x_10 = l_Std_Time_GenericFormat_parse(x_2, x_9, x_1); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +lean_dec(x_10); +x_11 = l_Std_Time_ZonedDateTime_fromLeanDateTimeWithIdentifierString(x_1); +return x_11; +} +else +{ +uint8_t x_12; +lean_dec(x_1); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +return x_10; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) +{ +return x_8; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_8, 0); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_6); +if (x_18 == 0) +{ +return x_6; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +lean_dec(x_6); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_19); +return x_20; +} +} +} +else +{ +uint8_t x_21; +lean_dec(x_1); +x_21 = !lean_is_exclusive(x_4); +if (x_21 == 0) +{ +return x_4; +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_4, 0); +lean_inc(x_22); +lean_dec(x_4); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +return x_23; +} +} +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_toLeanDateTimeWithIdentifierString), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instToString___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("zoned(\"", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instRepr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_ZonedDateTime_instRepr___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instRepr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = l_Std_Time_ZonedDateTime_toLeanDateTimeWithZoneString(x_1); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = l_Std_Time_ZonedDateTime_instRepr___closed__2; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDate_instRepr___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instRepr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_instRepr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_format___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_Year_Offset_era(x_4); +lean_dec(x_4); +x_6 = lean_box(x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +case 1: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_2); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_2, 0); +lean_dec(x_9); +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +lean_ctor_set(x_2, 0, x_11); +return x_2; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_2); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +case 2: +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_2, 0); +lean_dec(x_16); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_18); +return x_2; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_2); +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_20); +return x_21; +} +} +case 3: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + x_22 = x_2; +} else { + lean_dec_ref(x_2); + x_22 = lean_box(0); +} +x_23 = lean_ctor_get(x_1, 0); +lean_inc(x_23); +lean_dec(x_1); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = l_Std_Time_PlainDate_format___lambda__1___closed__1; +x_26 = lean_int_mod(x_24, x_25); +x_27 = l_Std_Time_PlainDate_format___lambda__1___closed__2; +x_28 = lean_int_dec_eq(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +x_30 = lean_ctor_get(x_23, 2); +lean_inc(x_30); +lean_dec(x_23); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +if (x_28 == 0) +{ +uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_24); +lean_dec(x_22); +x_56 = 0; +x_57 = l_Std_Time_ValidDate_dayOfYear(x_56, x_31); +lean_dec(x_31); +x_58 = lean_box(x_56); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_59); +return x_60; +} +else +{ +lean_object* x_61; lean_object* x_62; uint8_t x_63; uint8_t x_64; +x_61 = l_Std_Time_PlainDate_format___lambda__1___closed__3; +x_62 = lean_int_mod(x_24, x_61); +x_63 = lean_int_dec_eq(x_62, x_27); +lean_dec(x_62); +x_64 = l_instDecidableNot___rarg(x_63); +if (x_64 == 0) +{ +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = l_Std_Time_PlainDate_format___lambda__1___closed__4; +x_66 = lean_int_mod(x_24, x_65); +x_67 = lean_int_dec_eq(x_66, x_27); +lean_dec(x_66); +if (x_67 == 0) +{ +uint8_t x_68; +x_68 = 0; +x_32 = x_68; +goto block_55; +} +else +{ +uint8_t x_69; +x_69 = 1; +x_32 = x_69; +goto block_55; +} +} +else +{ +uint8_t x_70; +x_70 = 1; +x_32 = x_70; +goto block_55; +} +} +block_55: +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; +x_33 = l_Std_Time_PlainDate_format___lambda__1___closed__3; +x_34 = lean_int_mod(x_24, x_33); +x_35 = lean_int_dec_eq(x_34, x_27); +lean_dec(x_34); +x_36 = l_instDecidableNot___rarg(x_35); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = l_Std_Time_PlainDate_format___lambda__1___closed__4; +x_38 = lean_int_mod(x_24, x_37); +lean_dec(x_24); +x_39 = lean_int_dec_eq(x_38, x_27); +lean_dec(x_38); +if (x_39 == 0) +{ +uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_40 = 0; +x_41 = l_Std_Time_ValidDate_dayOfYear(x_40, x_31); +lean_dec(x_31); +x_42 = lean_box(x_32); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +if (lean_is_scalar(x_22)) { + x_44 = lean_alloc_ctor(1, 1, 0); +} else { + x_44 = x_22; + lean_ctor_set_tag(x_44, 1); +} +lean_ctor_set(x_44, 0, x_43); +return x_44; +} +else +{ +uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_45 = 1; +x_46 = l_Std_Time_ValidDate_dayOfYear(x_45, x_31); +lean_dec(x_31); +x_47 = lean_box(x_32); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +if (lean_is_scalar(x_22)) { + x_49 = lean_alloc_ctor(1, 1, 0); +} else { + x_49 = x_22; + lean_ctor_set_tag(x_49, 1); +} +lean_ctor_set(x_49, 0, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_dec(x_24); +x_50 = 1; +x_51 = l_Std_Time_ValidDate_dayOfYear(x_50, x_31); +lean_dec(x_31); +x_52 = lean_box(x_32); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +if (lean_is_scalar(x_22)) { + x_54 = lean_alloc_ctor(1, 1, 0); +} else { + x_54 = x_22; + lean_ctor_set_tag(x_54, 1); +} +lean_ctor_set(x_54, 0, x_53); +return x_54; +} +} +} +case 4: +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_2); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_2, 0); +lean_dec(x_72); +x_73 = lean_ctor_get(x_1, 0); +lean_inc(x_73); +lean_dec(x_1); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +lean_dec(x_73); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_74); +return x_2; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +lean_dec(x_2); +x_75 = lean_ctor_get(x_1, 0); +lean_inc(x_75); +lean_dec(x_1); +x_76 = lean_ctor_get(x_75, 1); +lean_inc(x_76); +lean_dec(x_75); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_76); +return x_77; +} +} +case 5: +{ +uint8_t x_78; +x_78 = !lean_is_exclusive(x_2); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_2, 0); +lean_dec(x_79); +x_80 = lean_ctor_get(x_1, 0); +lean_inc(x_80); +lean_dec(x_1); +x_81 = lean_ctor_get(x_80, 2); +lean_inc(x_81); +lean_dec(x_80); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_81); +return x_2; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +lean_dec(x_2); +x_82 = lean_ctor_get(x_1, 0); +lean_inc(x_82); +lean_dec(x_1); +x_83 = lean_ctor_get(x_82, 2); +lean_inc(x_83); +lean_dec(x_82); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_83); +return x_84; +} +} +case 6: +{ +uint8_t x_85; +x_85 = !lean_is_exclusive(x_2); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_2, 0); +lean_dec(x_86); +x_87 = lean_ctor_get(x_1, 0); +lean_inc(x_87); +lean_dec(x_1); +x_88 = l_Std_Time_PlainDate_quarter(x_87); +lean_dec(x_87); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_88); +return x_2; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_2); +x_89 = lean_ctor_get(x_1, 0); +lean_inc(x_89); +lean_dec(x_1); +x_90 = l_Std_Time_PlainDate_quarter(x_89); +lean_dec(x_89); +x_91 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_90); +return x_91; +} +} +case 7: +{ +uint8_t x_92; +x_92 = !lean_is_exclusive(x_2); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_2, 0); +lean_dec(x_93); +x_94 = lean_ctor_get(x_1, 0); +lean_inc(x_94); +lean_dec(x_1); +x_95 = l_Std_Time_PlainDate_weekOfYear(x_94); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_95); +return x_2; +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_2); +x_96 = lean_ctor_get(x_1, 0); +lean_inc(x_96); +lean_dec(x_1); +x_97 = l_Std_Time_PlainDate_weekOfYear(x_96); +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_97); +return x_98; +} +} +case 8: +{ +uint8_t x_99; +x_99 = !lean_is_exclusive(x_2); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_2, 0); +lean_dec(x_100); +x_101 = lean_ctor_get(x_1, 0); +lean_inc(x_101); +lean_dec(x_1); +x_102 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_101); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_102); +return x_2; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_2); +x_103 = lean_ctor_get(x_1, 0); +lean_inc(x_103); +lean_dec(x_1); +x_104 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_103); +x_105 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_105, 0, x_104); +return x_105; +} +} +case 9: +{ +lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_2); +x_106 = lean_ctor_get(x_1, 0); +lean_inc(x_106); +lean_dec(x_1); +x_107 = l_Std_Time_PlainDate_weekday(x_106); +x_108 = lean_box(x_107); +x_109 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_109, 0, x_108); +return x_109; +} +case 10: +{ +uint8_t x_110; +x_110 = !lean_is_exclusive(x_2); +if (x_110 == 0) +{ +lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; +x_111 = lean_ctor_get(x_2, 0); +lean_dec(x_111); +x_112 = lean_ctor_get(x_1, 0); +lean_inc(x_112); +lean_dec(x_1); +x_113 = l_Std_Time_PlainDate_weekday(x_112); +x_114 = lean_box(x_113); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_114); +return x_2; +} +else +{ +lean_object* x_115; uint8_t x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_2); +x_115 = lean_ctor_get(x_1, 0); +lean_inc(x_115); +lean_dec(x_1); +x_116 = l_Std_Time_PlainDate_weekday(x_115); +x_117 = lean_box(x_116); +x_118 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_118, 0, x_117); +return x_118; +} +} +case 11: +{ +uint8_t x_119; +x_119 = !lean_is_exclusive(x_2); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +x_120 = lean_ctor_get(x_2, 0); +lean_dec(x_120); +x_121 = l_Std_Time_PlainDateTime_weekOfMonth(x_1); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_121); +return x_2; +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_2); +x_122 = l_Std_Time_PlainDateTime_weekOfMonth(x_1); +lean_dec(x_1); +x_123 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_123, 0, x_122); +return x_123; +} +} +case 12: +{ +lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_2); +x_124 = lean_ctor_get(x_1, 1); +lean_inc(x_124); +lean_dec(x_1); +x_125 = lean_ctor_get(x_124, 0); +lean_inc(x_125); +lean_dec(x_124); +x_126 = l_Std_Time_HourMarker_ofOrdinal(x_125); +lean_dec(x_125); +x_127 = lean_box(x_126); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_127); +return x_128; +} +case 13: +{ +uint8_t x_129; +x_129 = !lean_is_exclusive(x_2); +if (x_129 == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_130 = lean_ctor_get(x_2, 0); +lean_dec(x_130); +x_131 = lean_ctor_get(x_1, 1); +lean_inc(x_131); +lean_dec(x_1); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_dec(x_131); +x_133 = l_Std_Time_Hour_Ordinal_toRelative(x_132); +lean_dec(x_132); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_133); +return x_2; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_dec(x_2); +x_134 = lean_ctor_get(x_1, 1); +lean_inc(x_134); +lean_dec(x_1); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +lean_dec(x_134); +x_136 = l_Std_Time_Hour_Ordinal_toRelative(x_135); +lean_dec(x_135); +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_136); +return x_137; +} +} +case 14: +{ +uint8_t x_138; +x_138 = !lean_is_exclusive(x_2); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_139 = lean_ctor_get(x_2, 0); +lean_dec(x_139); +x_140 = lean_ctor_get(x_1, 1); +lean_inc(x_140); +lean_dec(x_1); +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +lean_dec(x_140); +x_142 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_143 = lean_int_emod(x_141, x_142); +lean_dec(x_141); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_143); +return x_2; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_dec(x_2); +x_144 = lean_ctor_get(x_1, 1); +lean_inc(x_144); +lean_dec(x_1); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +lean_dec(x_144); +x_146 = l_Std_Time_PlainTime_format___lambda__1___closed__1; +x_147 = lean_int_emod(x_145, x_146); +lean_dec(x_145); +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_147); +return x_148; +} +} +case 15: +{ +uint8_t x_149; +x_149 = !lean_is_exclusive(x_2); +if (x_149 == 0) +{ +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_150 = lean_ctor_get(x_2, 0); +lean_dec(x_150); +x_151 = lean_ctor_get(x_1, 1); +lean_inc(x_151); +lean_dec(x_1); +x_152 = lean_ctor_get(x_151, 0); +lean_inc(x_152); +lean_dec(x_151); +x_153 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_152); +lean_dec(x_152); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_153); +return x_2; +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_dec(x_2); +x_154 = lean_ctor_get(x_1, 1); +lean_inc(x_154); +lean_dec(x_1); +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +lean_dec(x_154); +x_156 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_155); +lean_dec(x_155); +x_157 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_157, 0, x_156); +return x_157; +} +} +case 16: +{ +uint8_t x_158; +x_158 = !lean_is_exclusive(x_2); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_2, 0); +lean_dec(x_159); +x_160 = lean_ctor_get(x_1, 1); +lean_inc(x_160); +lean_dec(x_1); +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +lean_dec(x_160); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_161); +return x_2; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +lean_dec(x_2); +x_162 = lean_ctor_get(x_1, 1); +lean_inc(x_162); +lean_dec(x_1); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +lean_dec(x_162); +x_164 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_164, 0, x_163); +return x_164; +} +} +case 17: +{ +uint8_t x_165; +x_165 = !lean_is_exclusive(x_2); +if (x_165 == 0) +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_2, 0); +lean_dec(x_166); +x_167 = lean_ctor_get(x_1, 1); +lean_inc(x_167); +lean_dec(x_1); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +lean_dec(x_167); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_168); +return x_2; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_2); +x_169 = lean_ctor_get(x_1, 1); +lean_inc(x_169); +lean_dec(x_1); +x_170 = lean_ctor_get(x_169, 1); +lean_inc(x_170); +lean_dec(x_169); +x_171 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_171, 0, x_170); +return x_171; +} +} +case 18: +{ +uint8_t x_172; +x_172 = !lean_is_exclusive(x_2); +if (x_172 == 0) +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_2, 0); +lean_dec(x_173); +x_174 = lean_ctor_get(x_1, 1); +lean_inc(x_174); +lean_dec(x_1); +x_175 = lean_ctor_get(x_174, 2); +lean_inc(x_175); +lean_dec(x_174); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_175); +return x_2; +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; +lean_dec(x_2); +x_176 = lean_ctor_get(x_1, 1); +lean_inc(x_176); +lean_dec(x_1); +x_177 = lean_ctor_get(x_176, 2); +lean_inc(x_177); +lean_dec(x_176); +x_178 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_178, 0, x_177); +return x_178; +} +} +case 19: +{ +uint8_t x_179; +x_179 = !lean_is_exclusive(x_2); +if (x_179 == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_2, 0); +lean_dec(x_180); +x_181 = lean_ctor_get(x_1, 1); +lean_inc(x_181); +lean_dec(x_1); +x_182 = lean_ctor_get(x_181, 3); +lean_inc(x_182); +lean_dec(x_181); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_182); +return x_2; +} +else +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_dec(x_2); +x_183 = lean_ctor_get(x_1, 1); +lean_inc(x_183); +lean_dec(x_1); +x_184 = lean_ctor_get(x_183, 3); +lean_inc(x_184); +lean_dec(x_183); +x_185 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_185, 0, x_184); +return x_185; +} +} +case 20: +{ +uint8_t x_186; +x_186 = !lean_is_exclusive(x_2); +if (x_186 == 0) +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_2, 0); +lean_dec(x_187); +x_188 = lean_ctor_get(x_1, 1); +lean_inc(x_188); +lean_dec(x_1); +x_189 = l_Std_Time_PlainTime_toMilliseconds(x_188); +lean_dec(x_188); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_189); +return x_2; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; +lean_dec(x_2); +x_190 = lean_ctor_get(x_1, 1); +lean_inc(x_190); +lean_dec(x_1); +x_191 = l_Std_Time_PlainTime_toMilliseconds(x_190); +lean_dec(x_190); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, x_191); +return x_192; +} +} +case 21: +{ +uint8_t x_193; +x_193 = !lean_is_exclusive(x_2); +if (x_193 == 0) +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; +x_194 = lean_ctor_get(x_2, 0); +lean_dec(x_194); +x_195 = lean_ctor_get(x_1, 1); +lean_inc(x_195); +lean_dec(x_1); +x_196 = lean_ctor_get(x_195, 3); +lean_inc(x_196); +lean_dec(x_195); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_196); +return x_2; +} +else +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_2); +x_197 = lean_ctor_get(x_1, 1); +lean_inc(x_197); +lean_dec(x_1); +x_198 = lean_ctor_get(x_197, 3); +lean_inc(x_198); +lean_dec(x_197); +x_199 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_199, 0, x_198); +return x_199; +} +} +case 22: +{ +uint8_t x_200; +x_200 = !lean_is_exclusive(x_2); +if (x_200 == 0) +{ +lean_object* x_201; lean_object* x_202; lean_object* x_203; +x_201 = lean_ctor_get(x_2, 0); +lean_dec(x_201); +x_202 = lean_ctor_get(x_1, 1); +lean_inc(x_202); +lean_dec(x_1); +x_203 = l_Std_Time_PlainTime_toNanoseconds(x_202); +lean_dec(x_202); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_203); +return x_2; +} +else +{ +lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_2); +x_204 = lean_ctor_get(x_1, 1); +lean_inc(x_204); +lean_dec(x_1); +x_205 = l_Std_Time_PlainTime_toNanoseconds(x_204); +lean_dec(x_204); +x_206 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_206, 0, x_205); +return x_206; +} +} +default: +{ +lean_object* x_207; +lean_dec(x_2); +lean_dec(x_1); +x_207 = lean_box(0); +return x_207; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_format(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_spec___rarg(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_format___closed__1; +x_6 = lean_string_append(x_5, x_4); +lean_dec(x_4); +x_7 = l_Std_Time_PlainDate_format___closed__2; +x_8 = lean_string_append(x_6, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_format___lambda__1), 2, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Std_Time_PlainDate_format___closed__2; +x_12 = l_Std_Time_GenericFormat_formatGeneric_go(x_10, x_11, x_9); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +x_13 = l_Std_Time_PlainDate_format___closed__3; +return x_13; +} +else +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +return x_14; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_fromAscTimeString___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_TimeZone_GMT; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromAscTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_ascTime; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_thunk_get_own(x_10); +lean_dec(x_10); +lean_ctor_set(x_4, 0, x_11); +return x_4; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_dec(x_4); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_thunk_get_own(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_TimeZone_UTC; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = l_Std_Time_ZonedDateTime_format___lambda__1___closed__2; +x_4 = lean_int_mul(x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_3 = l_Std_Time_TimeZone_UTC; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = l_Std_Time_PlainTime_toSeconds(x_5); +x_7 = lean_int_add(x_6, x_4); +lean_dec(x_4); +lean_dec(x_6); +x_8 = l_Std_Time_ZonedDateTime_format___lambda__1___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_5); +lean_dec(x_5); +x_11 = l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1; +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_ofNanoseconds(x_12); +lean_dec(x_12); +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +lean_dec(x_1); +x_15 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +x_16 = lean_int_add(x_15, x_9); +lean_dec(x_9); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_16); +lean_dec(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_2 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_3 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___boxed), 2, 1); +lean_closure_set(x_3, 0, x_1); +x_4 = lean_mk_thunk(x_3); +x_5 = l_Std_Time_TimeZone_UTC; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_8 = l_Std_Time_Formats_ascTime; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_5, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_toAscTimeString___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromLongDateFormatString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_longDateFormat; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_thunk_get_own(x_10); +lean_dec(x_10); +lean_ctor_set(x_4, 0, x_11); +return x_4; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_dec(x_4); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_thunk_get_own(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toLongDateFormatString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_inc(x_1); +x_2 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_3 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___boxed), 2, 1); +lean_closure_set(x_3, 0, x_1); +x_4 = lean_mk_thunk(x_3); +x_5 = l_Std_Time_TimeZone_UTC; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_8 = l_Std_Time_Formats_longDateFormat; +x_9 = l_Std_Time_GenericFormat_format(x_7, x_5, x_8, x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromDateTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_dateTime24Hour; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_thunk_get_own(x_10); +lean_dec(x_10); +lean_ctor_set(x_4, 0, x_11); +return x_4; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_dec(x_4); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_thunk_get_own(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toDateTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_6, 3); +lean_inc(x_10); +lean_dec(x_6); +x_11 = l_Std_Time_PlainDate_format___closed__2; +x_12 = l_Std_Time_Formats_dateTime24Hour; +x_13 = l_Std_Time_GenericFormat_formatBuilder_go(x_11, x_12); +x_14 = lean_apply_7(x_13, x_3, x_4, x_5, x_7, x_8, x_9, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_fromLeanDateTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_leanDateTime24Hour; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = l_Std_Time_Formats_leanDateTime24HourNoNanos; +x_6 = l_Std_Time_GenericFormat_parse(x_2, x_5, x_1); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +return x_6; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_6); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_6, 0); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_thunk_get_own(x_12); +lean_dec(x_12); +lean_ctor_set(x_6, 0, x_13); +return x_6; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_6, 0); +lean_inc(x_14); +lean_dec(x_6); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_thunk_get_own(x_15); +lean_dec(x_15); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +return x_17; +} +} +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_4); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_4, 0); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_thunk_get_own(x_20); +lean_dec(x_20); +lean_ctor_set(x_4, 0, x_21); +return x_4; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_4, 0); +lean_inc(x_22); +lean_dec(x_4); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_thunk_get_own(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toLeanDateTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_6, 3); +lean_inc(x_10); +lean_dec(x_6); +x_11 = l_Std_Time_PlainDate_format___closed__2; +x_12 = l_Std_Time_Formats_leanDateTime24Hour; +x_13 = l_Std_Time_GenericFormat_formatBuilder_go(x_11, x_12); +x_14 = lean_apply_7(x_13, x_3, x_4, x_5, x_7, x_8, x_9, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; +lean_inc(x_1); +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +lean_inc(x_1); +x_3 = l_Std_Time_PlainDateTime_fromLongDateFormatString(x_1); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_3); +lean_inc(x_1); +x_4 = l_Std_Time_PlainDateTime_fromDateTimeString(x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +lean_dec(x_4); +x_5 = l_Std_Time_PlainDateTime_fromLeanDateTimeString(x_1); +return x_5; +} +else +{ +uint8_t x_6; +lean_dec(x_1); +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) +{ +return x_4; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +} +else +{ +uint8_t x_9; +lean_dec(x_1); +x_9 = !lean_is_exclusive(x_3); +if (x_9 == 0) +{ +return x_3; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_3, 0); +lean_inc(x_10); +lean_dec(x_3); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +} +else +{ +uint8_t x_12; +lean_dec(x_1); +x_12 = !lean_is_exclusive(x_2); +if (x_12 == 0) +{ +return x_2; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +lean_dec(x_2); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDateTime_toLeanDateTimeString), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainDateTime_instToString___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("datetime(\"", 10, 10); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_instRepr___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainDateTime_instRepr___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instRepr(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_3 = l_Std_Time_PlainDateTime_toLeanDateTimeString(x_1); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = l_Std_Time_PlainDateTime_instRepr___closed__2; +x_6 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainDate_instRepr___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_instRepr___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_instRepr(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_format(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_GenericFormat_spec___rarg(x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_PlainDate_format___closed__1; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +x_8 = l_Std_Time_PlainDate_format___closed__2; +x_9 = lean_string_append(x_7, x_8); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_4, 0); +lean_inc(x_10); +lean_dec(x_4); +x_11 = lean_box(1); +x_12 = l_Std_Time_GenericFormat_format(x_11, x_1, x_10, x_2); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_format___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_format(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_fromAscTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_ascTime; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toAscTimeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_TimeZone_GMT; +x_4 = l_Std_Time_Formats_ascTime; +x_5 = l_Std_Time_GenericFormat_format(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_fromLongDateFormatString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_Formats_longDateFormat; +x_4 = l_Std_Time_GenericFormat_parse(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLongDateFormatString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Std_Time_PlainDateTime_fromAscTimeString___closed__1; +x_3 = l_Std_Time_TimeZone_GMT; +x_4 = l_Std_Time_Formats_longDateFormat; +x_5 = l_Std_Time_GenericFormat_format(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toISO8601String(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(1); +x_4 = l_Std_Time_Formats_iso8601; +x_5 = l_Std_Time_GenericFormat_format(x_3, x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toISO8601String___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_toISO8601String(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC822String(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(1); +x_4 = l_Std_Time_Formats_rfc822; +x_5 = l_Std_Time_GenericFormat_format(x_3, x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC822String___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_toRFC822String(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC850String(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(1); +x_4 = l_Std_Time_Formats_rfc850; +x_5 = l_Std_Time_GenericFormat_format(x_3, x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toRFC850String___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_toRFC850String(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDateTimeWithZoneString(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(1); +x_4 = l_Std_Time_Formats_dateTimeWithZone; +x_5 = l_Std_Time_GenericFormat_format(x_3, x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDateTimeWithZoneString___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_toDateTimeWithZoneString(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLeanDateTimeWithZoneString(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_box(1); +x_4 = l_Std_Time_Formats_leanDateTimeWithZone; +x_5 = l_Std_Time_GenericFormat_format(x_3, x_1, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toLeanDateTimeWithZoneString___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_toLeanDateTimeWithZoneString(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; +lean_inc(x_1); +x_2 = l_Std_Time_DateTime_fromAscTimeString(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = l_Std_Time_DateTime_fromLongDateFormatString(x_1); +return x_3; +} +else +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +return x_2; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_2, 0); +lean_inc(x_5); +lean_dec(x_2); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instRepr(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_box(1); +x_5 = l_Std_Time_Formats_leanDateTimeWithZone; +x_6 = l_Std_Time_GenericFormat_format(x_4, x_1, x_5, x_2); +x_7 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = l_Repr_addAppParen(x_7, x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instRepr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_instRepr(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instToString(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toLeanDateTimeWithZoneString___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Notation_Spec(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Format_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal_Bounded(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Format(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Notation_Spec(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Format_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal_Bounded(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Formats_iso8601___closed__1 = _init_l_Std_Time_Formats_iso8601___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__1); +l_Std_Time_Formats_iso8601___closed__2 = _init_l_Std_Time_Formats_iso8601___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__2); +l_Std_Time_Formats_iso8601___closed__3 = _init_l_Std_Time_Formats_iso8601___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__3); +l_Std_Time_Formats_iso8601___closed__4 = _init_l_Std_Time_Formats_iso8601___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__4); +l_Std_Time_Formats_iso8601___closed__5 = _init_l_Std_Time_Formats_iso8601___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__5); +l_Std_Time_Formats_iso8601___closed__6 = _init_l_Std_Time_Formats_iso8601___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__6); +l_Std_Time_Formats_iso8601___closed__7 = _init_l_Std_Time_Formats_iso8601___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__7); +l_Std_Time_Formats_iso8601___closed__8 = _init_l_Std_Time_Formats_iso8601___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__8); +l_Std_Time_Formats_iso8601___closed__9 = _init_l_Std_Time_Formats_iso8601___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__9); +l_Std_Time_Formats_iso8601___closed__10 = _init_l_Std_Time_Formats_iso8601___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__10); +l_Std_Time_Formats_iso8601___closed__11 = _init_l_Std_Time_Formats_iso8601___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__11); +l_Std_Time_Formats_iso8601___closed__12 = _init_l_Std_Time_Formats_iso8601___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__12); +l_Std_Time_Formats_iso8601___closed__13 = _init_l_Std_Time_Formats_iso8601___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__13); +l_Std_Time_Formats_iso8601___closed__14 = _init_l_Std_Time_Formats_iso8601___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__14); +l_Std_Time_Formats_iso8601___closed__15 = _init_l_Std_Time_Formats_iso8601___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__15); +l_Std_Time_Formats_iso8601___closed__16 = _init_l_Std_Time_Formats_iso8601___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__16); +l_Std_Time_Formats_iso8601___closed__17 = _init_l_Std_Time_Formats_iso8601___closed__17(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__17); +l_Std_Time_Formats_iso8601___closed__18 = _init_l_Std_Time_Formats_iso8601___closed__18(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__18); +l_Std_Time_Formats_iso8601___closed__19 = _init_l_Std_Time_Formats_iso8601___closed__19(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__19); +l_Std_Time_Formats_iso8601___closed__20 = _init_l_Std_Time_Formats_iso8601___closed__20(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__20); +l_Std_Time_Formats_iso8601___closed__21 = _init_l_Std_Time_Formats_iso8601___closed__21(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__21); +l_Std_Time_Formats_iso8601___closed__22 = _init_l_Std_Time_Formats_iso8601___closed__22(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__22); +l_Std_Time_Formats_iso8601___closed__23 = _init_l_Std_Time_Formats_iso8601___closed__23(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__23); +l_Std_Time_Formats_iso8601___closed__24 = _init_l_Std_Time_Formats_iso8601___closed__24(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__24); +l_Std_Time_Formats_iso8601___closed__25 = _init_l_Std_Time_Formats_iso8601___closed__25(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__25); +l_Std_Time_Formats_iso8601___closed__26 = _init_l_Std_Time_Formats_iso8601___closed__26(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__26); +l_Std_Time_Formats_iso8601___closed__27 = _init_l_Std_Time_Formats_iso8601___closed__27(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__27); +l_Std_Time_Formats_iso8601___closed__28 = _init_l_Std_Time_Formats_iso8601___closed__28(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__28); +l_Std_Time_Formats_iso8601___closed__29 = _init_l_Std_Time_Formats_iso8601___closed__29(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__29); +l_Std_Time_Formats_iso8601___closed__30 = _init_l_Std_Time_Formats_iso8601___closed__30(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__30); +l_Std_Time_Formats_iso8601___closed__31 = _init_l_Std_Time_Formats_iso8601___closed__31(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__31); +l_Std_Time_Formats_iso8601___closed__32 = _init_l_Std_Time_Formats_iso8601___closed__32(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__32); +l_Std_Time_Formats_iso8601___closed__33 = _init_l_Std_Time_Formats_iso8601___closed__33(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__33); +l_Std_Time_Formats_iso8601___closed__34 = _init_l_Std_Time_Formats_iso8601___closed__34(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__34); +l_Std_Time_Formats_iso8601___closed__35 = _init_l_Std_Time_Formats_iso8601___closed__35(); +lean_mark_persistent(l_Std_Time_Formats_iso8601___closed__35); +l_Std_Time_Formats_iso8601 = _init_l_Std_Time_Formats_iso8601(); +lean_mark_persistent(l_Std_Time_Formats_iso8601); +l_Std_Time_Formats_americanDate___closed__1 = _init_l_Std_Time_Formats_americanDate___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_americanDate___closed__1); +l_Std_Time_Formats_americanDate___closed__2 = _init_l_Std_Time_Formats_americanDate___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_americanDate___closed__2); +l_Std_Time_Formats_americanDate___closed__3 = _init_l_Std_Time_Formats_americanDate___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_americanDate___closed__3); +l_Std_Time_Formats_americanDate___closed__4 = _init_l_Std_Time_Formats_americanDate___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_americanDate___closed__4); +l_Std_Time_Formats_americanDate___closed__5 = _init_l_Std_Time_Formats_americanDate___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_americanDate___closed__5); +l_Std_Time_Formats_americanDate = _init_l_Std_Time_Formats_americanDate(); +lean_mark_persistent(l_Std_Time_Formats_americanDate); +l_Std_Time_Formats_europeanDate___closed__1 = _init_l_Std_Time_Formats_europeanDate___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_europeanDate___closed__1); +l_Std_Time_Formats_europeanDate___closed__2 = _init_l_Std_Time_Formats_europeanDate___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_europeanDate___closed__2); +l_Std_Time_Formats_europeanDate___closed__3 = _init_l_Std_Time_Formats_europeanDate___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_europeanDate___closed__3); +l_Std_Time_Formats_europeanDate = _init_l_Std_Time_Formats_europeanDate(); +lean_mark_persistent(l_Std_Time_Formats_europeanDate); +l_Std_Time_Formats_time12Hour___closed__1 = _init_l_Std_Time_Formats_time12Hour___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__1); +l_Std_Time_Formats_time12Hour___closed__2 = _init_l_Std_Time_Formats_time12Hour___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__2); +l_Std_Time_Formats_time12Hour___closed__3 = _init_l_Std_Time_Formats_time12Hour___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__3); +l_Std_Time_Formats_time12Hour___closed__4 = _init_l_Std_Time_Formats_time12Hour___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__4); +l_Std_Time_Formats_time12Hour___closed__5 = _init_l_Std_Time_Formats_time12Hour___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__5); +l_Std_Time_Formats_time12Hour___closed__6 = _init_l_Std_Time_Formats_time12Hour___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__6); +l_Std_Time_Formats_time12Hour___closed__7 = _init_l_Std_Time_Formats_time12Hour___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__7); +l_Std_Time_Formats_time12Hour___closed__8 = _init_l_Std_Time_Formats_time12Hour___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__8); +l_Std_Time_Formats_time12Hour___closed__9 = _init_l_Std_Time_Formats_time12Hour___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__9); +l_Std_Time_Formats_time12Hour___closed__10 = _init_l_Std_Time_Formats_time12Hour___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__10); +l_Std_Time_Formats_time12Hour___closed__11 = _init_l_Std_Time_Formats_time12Hour___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__11); +l_Std_Time_Formats_time12Hour___closed__12 = _init_l_Std_Time_Formats_time12Hour___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__12); +l_Std_Time_Formats_time12Hour___closed__13 = _init_l_Std_Time_Formats_time12Hour___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour___closed__13); +l_Std_Time_Formats_time12Hour = _init_l_Std_Time_Formats_time12Hour(); +lean_mark_persistent(l_Std_Time_Formats_time12Hour); +l_Std_Time_Formats_time24Hour___closed__1 = _init_l_Std_Time_Formats_time24Hour___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour___closed__1); +l_Std_Time_Formats_time24Hour___closed__2 = _init_l_Std_Time_Formats_time24Hour___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour___closed__2); +l_Std_Time_Formats_time24Hour___closed__3 = _init_l_Std_Time_Formats_time24Hour___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour___closed__3); +l_Std_Time_Formats_time24Hour___closed__4 = _init_l_Std_Time_Formats_time24Hour___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour___closed__4); +l_Std_Time_Formats_time24Hour___closed__5 = _init_l_Std_Time_Formats_time24Hour___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour___closed__5); +l_Std_Time_Formats_time24Hour = _init_l_Std_Time_Formats_time24Hour(); +lean_mark_persistent(l_Std_Time_Formats_time24Hour); +l_Std_Time_Formats_dateTime24Hour___closed__1 = _init_l_Std_Time_Formats_dateTime24Hour___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__1); +l_Std_Time_Formats_dateTime24Hour___closed__2 = _init_l_Std_Time_Formats_dateTime24Hour___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__2); +l_Std_Time_Formats_dateTime24Hour___closed__3 = _init_l_Std_Time_Formats_dateTime24Hour___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__3); +l_Std_Time_Formats_dateTime24Hour___closed__4 = _init_l_Std_Time_Formats_dateTime24Hour___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__4); +l_Std_Time_Formats_dateTime24Hour___closed__5 = _init_l_Std_Time_Formats_dateTime24Hour___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__5); +l_Std_Time_Formats_dateTime24Hour___closed__6 = _init_l_Std_Time_Formats_dateTime24Hour___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__6); +l_Std_Time_Formats_dateTime24Hour___closed__7 = _init_l_Std_Time_Formats_dateTime24Hour___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__7); +l_Std_Time_Formats_dateTime24Hour___closed__8 = _init_l_Std_Time_Formats_dateTime24Hour___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__8); +l_Std_Time_Formats_dateTime24Hour___closed__9 = _init_l_Std_Time_Formats_dateTime24Hour___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__9); +l_Std_Time_Formats_dateTime24Hour___closed__10 = _init_l_Std_Time_Formats_dateTime24Hour___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__10); +l_Std_Time_Formats_dateTime24Hour___closed__11 = _init_l_Std_Time_Formats_dateTime24Hour___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__11); +l_Std_Time_Formats_dateTime24Hour___closed__12 = _init_l_Std_Time_Formats_dateTime24Hour___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__12); +l_Std_Time_Formats_dateTime24Hour___closed__13 = _init_l_Std_Time_Formats_dateTime24Hour___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__13); +l_Std_Time_Formats_dateTime24Hour___closed__14 = _init_l_Std_Time_Formats_dateTime24Hour___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__14); +l_Std_Time_Formats_dateTime24Hour___closed__15 = _init_l_Std_Time_Formats_dateTime24Hour___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour___closed__15); +l_Std_Time_Formats_dateTime24Hour = _init_l_Std_Time_Formats_dateTime24Hour(); +lean_mark_persistent(l_Std_Time_Formats_dateTime24Hour); +l_Std_Time_Formats_dateTimeWithZone___closed__1 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__1); +l_Std_Time_Formats_dateTimeWithZone___closed__2 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__2); +l_Std_Time_Formats_dateTimeWithZone___closed__3 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__3); +l_Std_Time_Formats_dateTimeWithZone___closed__4 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__4); +l_Std_Time_Formats_dateTimeWithZone___closed__5 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__5); +l_Std_Time_Formats_dateTimeWithZone___closed__6 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__6); +l_Std_Time_Formats_dateTimeWithZone___closed__7 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__7); +l_Std_Time_Formats_dateTimeWithZone___closed__8 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__8); +l_Std_Time_Formats_dateTimeWithZone___closed__9 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__9); +l_Std_Time_Formats_dateTimeWithZone___closed__10 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__10); +l_Std_Time_Formats_dateTimeWithZone___closed__11 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__11); +l_Std_Time_Formats_dateTimeWithZone___closed__12 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__12); +l_Std_Time_Formats_dateTimeWithZone___closed__13 = _init_l_Std_Time_Formats_dateTimeWithZone___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone___closed__13); +l_Std_Time_Formats_dateTimeWithZone = _init_l_Std_Time_Formats_dateTimeWithZone(); +lean_mark_persistent(l_Std_Time_Formats_dateTimeWithZone); +l_Std_Time_Formats_leanTime24Hour = _init_l_Std_Time_Formats_leanTime24Hour(); +lean_mark_persistent(l_Std_Time_Formats_leanTime24Hour); +l_Std_Time_Formats_leanTime24HourNoNanos = _init_l_Std_Time_Formats_leanTime24HourNoNanos(); +lean_mark_persistent(l_Std_Time_Formats_leanTime24HourNoNanos); +l_Std_Time_Formats_leanDateTime24Hour___closed__1 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__1); +l_Std_Time_Formats_leanDateTime24Hour___closed__2 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__2); +l_Std_Time_Formats_leanDateTime24Hour___closed__3 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__3); +l_Std_Time_Formats_leanDateTime24Hour___closed__4 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__4); +l_Std_Time_Formats_leanDateTime24Hour___closed__5 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__5); +l_Std_Time_Formats_leanDateTime24Hour___closed__6 = _init_l_Std_Time_Formats_leanDateTime24Hour___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour___closed__6); +l_Std_Time_Formats_leanDateTime24Hour = _init_l_Std_Time_Formats_leanDateTime24Hour(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24Hour); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__1); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__2); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__3); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__4); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__5); +l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6 = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos___closed__6); +l_Std_Time_Formats_leanDateTime24HourNoNanos = _init_l_Std_Time_Formats_leanDateTime24HourNoNanos(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTime24HourNoNanos); +l_Std_Time_Formats_leanDateTimeWithZone___closed__1 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__1); +l_Std_Time_Formats_leanDateTimeWithZone___closed__2 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__2); +l_Std_Time_Formats_leanDateTimeWithZone___closed__3 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__3); +l_Std_Time_Formats_leanDateTimeWithZone___closed__4 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__4); +l_Std_Time_Formats_leanDateTimeWithZone___closed__5 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__5); +l_Std_Time_Formats_leanDateTimeWithZone___closed__6 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__6); +l_Std_Time_Formats_leanDateTimeWithZone___closed__7 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__7); +l_Std_Time_Formats_leanDateTimeWithZone___closed__8 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__8); +l_Std_Time_Formats_leanDateTimeWithZone___closed__9 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__9); +l_Std_Time_Formats_leanDateTimeWithZone___closed__10 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__10); +l_Std_Time_Formats_leanDateTimeWithZone___closed__11 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__11); +l_Std_Time_Formats_leanDateTimeWithZone___closed__12 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__12); +l_Std_Time_Formats_leanDateTimeWithZone___closed__13 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__13); +l_Std_Time_Formats_leanDateTimeWithZone___closed__14 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__14); +l_Std_Time_Formats_leanDateTimeWithZone___closed__15 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__15); +l_Std_Time_Formats_leanDateTimeWithZone___closed__16 = _init_l_Std_Time_Formats_leanDateTimeWithZone___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone___closed__16); +l_Std_Time_Formats_leanDateTimeWithZone = _init_l_Std_Time_Formats_leanDateTimeWithZone(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZone); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__1); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__2); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__3); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__4); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__5); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__6); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__7); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__8); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__9); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__10); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11 = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos___closed__11); +l_Std_Time_Formats_leanDateTimeWithZoneNoNanos = _init_l_Std_Time_Formats_leanDateTimeWithZoneNoNanos(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithZoneNoNanos); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__1); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__2); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__3); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__4); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__5); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__6); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__7); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__8); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__9); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__10); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__11); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__12); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__13); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__14); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__15); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__16); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__17); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__18); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__19); +l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier___closed__20); +l_Std_Time_Formats_leanDateTimeWithIdentifier = _init_l_Std_Time_Formats_leanDateTimeWithIdentifier(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifier); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__1); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__2); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__3); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__4); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__5); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__6); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__7); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__8); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__9); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__10); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__11); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__12); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13 = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos___closed__13); +l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos = _init_l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos(); +lean_mark_persistent(l_Std_Time_Formats_leanDateTimeWithIdentifierAndNanos); +l_Std_Time_Formats_leanDate___closed__1 = _init_l_Std_Time_Formats_leanDate___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_leanDate___closed__1); +l_Std_Time_Formats_leanDate___closed__2 = _init_l_Std_Time_Formats_leanDate___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_leanDate___closed__2); +l_Std_Time_Formats_leanDate___closed__3 = _init_l_Std_Time_Formats_leanDate___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_leanDate___closed__3); +l_Std_Time_Formats_leanDate___closed__4 = _init_l_Std_Time_Formats_leanDate___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_leanDate___closed__4); +l_Std_Time_Formats_leanDate___closed__5 = _init_l_Std_Time_Formats_leanDate___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_leanDate___closed__5); +l_Std_Time_Formats_leanDate = _init_l_Std_Time_Formats_leanDate(); +lean_mark_persistent(l_Std_Time_Formats_leanDate); +l_Std_Time_Formats_sqlDate = _init_l_Std_Time_Formats_sqlDate(); +lean_mark_persistent(l_Std_Time_Formats_sqlDate); +l_Std_Time_Formats_longDateFormat___closed__1 = _init_l_Std_Time_Formats_longDateFormat___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__1); +l_Std_Time_Formats_longDateFormat___closed__2 = _init_l_Std_Time_Formats_longDateFormat___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__2); +l_Std_Time_Formats_longDateFormat___closed__3 = _init_l_Std_Time_Formats_longDateFormat___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__3); +l_Std_Time_Formats_longDateFormat___closed__4 = _init_l_Std_Time_Formats_longDateFormat___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__4); +l_Std_Time_Formats_longDateFormat___closed__5 = _init_l_Std_Time_Formats_longDateFormat___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__5); +l_Std_Time_Formats_longDateFormat___closed__6 = _init_l_Std_Time_Formats_longDateFormat___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__6); +l_Std_Time_Formats_longDateFormat___closed__7 = _init_l_Std_Time_Formats_longDateFormat___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__7); +l_Std_Time_Formats_longDateFormat___closed__8 = _init_l_Std_Time_Formats_longDateFormat___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__8); +l_Std_Time_Formats_longDateFormat___closed__9 = _init_l_Std_Time_Formats_longDateFormat___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__9); +l_Std_Time_Formats_longDateFormat___closed__10 = _init_l_Std_Time_Formats_longDateFormat___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__10); +l_Std_Time_Formats_longDateFormat___closed__11 = _init_l_Std_Time_Formats_longDateFormat___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__11); +l_Std_Time_Formats_longDateFormat___closed__12 = _init_l_Std_Time_Formats_longDateFormat___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__12); +l_Std_Time_Formats_longDateFormat___closed__13 = _init_l_Std_Time_Formats_longDateFormat___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__13); +l_Std_Time_Formats_longDateFormat___closed__14 = _init_l_Std_Time_Formats_longDateFormat___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__14); +l_Std_Time_Formats_longDateFormat___closed__15 = _init_l_Std_Time_Formats_longDateFormat___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__15); +l_Std_Time_Formats_longDateFormat___closed__16 = _init_l_Std_Time_Formats_longDateFormat___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__16); +l_Std_Time_Formats_longDateFormat___closed__17 = _init_l_Std_Time_Formats_longDateFormat___closed__17(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat___closed__17); +l_Std_Time_Formats_longDateFormat = _init_l_Std_Time_Formats_longDateFormat(); +lean_mark_persistent(l_Std_Time_Formats_longDateFormat); +l_Std_Time_Formats_ascTime___closed__1 = _init_l_Std_Time_Formats_ascTime___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__1); +l_Std_Time_Formats_ascTime___closed__2 = _init_l_Std_Time_Formats_ascTime___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__2); +l_Std_Time_Formats_ascTime___closed__3 = _init_l_Std_Time_Formats_ascTime___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__3); +l_Std_Time_Formats_ascTime___closed__4 = _init_l_Std_Time_Formats_ascTime___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__4); +l_Std_Time_Formats_ascTime___closed__5 = _init_l_Std_Time_Formats_ascTime___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__5); +l_Std_Time_Formats_ascTime___closed__6 = _init_l_Std_Time_Formats_ascTime___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__6); +l_Std_Time_Formats_ascTime___closed__7 = _init_l_Std_Time_Formats_ascTime___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__7); +l_Std_Time_Formats_ascTime___closed__8 = _init_l_Std_Time_Formats_ascTime___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__8); +l_Std_Time_Formats_ascTime___closed__9 = _init_l_Std_Time_Formats_ascTime___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__9); +l_Std_Time_Formats_ascTime___closed__10 = _init_l_Std_Time_Formats_ascTime___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__10); +l_Std_Time_Formats_ascTime___closed__11 = _init_l_Std_Time_Formats_ascTime___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__11); +l_Std_Time_Formats_ascTime___closed__12 = _init_l_Std_Time_Formats_ascTime___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__12); +l_Std_Time_Formats_ascTime___closed__13 = _init_l_Std_Time_Formats_ascTime___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__13); +l_Std_Time_Formats_ascTime___closed__14 = _init_l_Std_Time_Formats_ascTime___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__14); +l_Std_Time_Formats_ascTime___closed__15 = _init_l_Std_Time_Formats_ascTime___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__15); +l_Std_Time_Formats_ascTime___closed__16 = _init_l_Std_Time_Formats_ascTime___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__16); +l_Std_Time_Formats_ascTime___closed__17 = _init_l_Std_Time_Formats_ascTime___closed__17(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__17); +l_Std_Time_Formats_ascTime___closed__18 = _init_l_Std_Time_Formats_ascTime___closed__18(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__18); +l_Std_Time_Formats_ascTime___closed__19 = _init_l_Std_Time_Formats_ascTime___closed__19(); +lean_mark_persistent(l_Std_Time_Formats_ascTime___closed__19); +l_Std_Time_Formats_ascTime = _init_l_Std_Time_Formats_ascTime(); +lean_mark_persistent(l_Std_Time_Formats_ascTime); +l_Std_Time_Formats_rfc822___closed__1 = _init_l_Std_Time_Formats_rfc822___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__1); +l_Std_Time_Formats_rfc822___closed__2 = _init_l_Std_Time_Formats_rfc822___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__2); +l_Std_Time_Formats_rfc822___closed__3 = _init_l_Std_Time_Formats_rfc822___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__3); +l_Std_Time_Formats_rfc822___closed__4 = _init_l_Std_Time_Formats_rfc822___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__4); +l_Std_Time_Formats_rfc822___closed__5 = _init_l_Std_Time_Formats_rfc822___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__5); +l_Std_Time_Formats_rfc822___closed__6 = _init_l_Std_Time_Formats_rfc822___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__6); +l_Std_Time_Formats_rfc822___closed__7 = _init_l_Std_Time_Formats_rfc822___closed__7(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__7); +l_Std_Time_Formats_rfc822___closed__8 = _init_l_Std_Time_Formats_rfc822___closed__8(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__8); +l_Std_Time_Formats_rfc822___closed__9 = _init_l_Std_Time_Formats_rfc822___closed__9(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__9); +l_Std_Time_Formats_rfc822___closed__10 = _init_l_Std_Time_Formats_rfc822___closed__10(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__10); +l_Std_Time_Formats_rfc822___closed__11 = _init_l_Std_Time_Formats_rfc822___closed__11(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__11); +l_Std_Time_Formats_rfc822___closed__12 = _init_l_Std_Time_Formats_rfc822___closed__12(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__12); +l_Std_Time_Formats_rfc822___closed__13 = _init_l_Std_Time_Formats_rfc822___closed__13(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__13); +l_Std_Time_Formats_rfc822___closed__14 = _init_l_Std_Time_Formats_rfc822___closed__14(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__14); +l_Std_Time_Formats_rfc822___closed__15 = _init_l_Std_Time_Formats_rfc822___closed__15(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__15); +l_Std_Time_Formats_rfc822___closed__16 = _init_l_Std_Time_Formats_rfc822___closed__16(); +lean_mark_persistent(l_Std_Time_Formats_rfc822___closed__16); +l_Std_Time_Formats_rfc822 = _init_l_Std_Time_Formats_rfc822(); +lean_mark_persistent(l_Std_Time_Formats_rfc822); +l_Std_Time_Formats_rfc850___closed__1 = _init_l_Std_Time_Formats_rfc850___closed__1(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__1); +l_Std_Time_Formats_rfc850___closed__2 = _init_l_Std_Time_Formats_rfc850___closed__2(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__2); +l_Std_Time_Formats_rfc850___closed__3 = _init_l_Std_Time_Formats_rfc850___closed__3(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__3); +l_Std_Time_Formats_rfc850___closed__4 = _init_l_Std_Time_Formats_rfc850___closed__4(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__4); +l_Std_Time_Formats_rfc850___closed__5 = _init_l_Std_Time_Formats_rfc850___closed__5(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__5); +l_Std_Time_Formats_rfc850___closed__6 = _init_l_Std_Time_Formats_rfc850___closed__6(); +lean_mark_persistent(l_Std_Time_Formats_rfc850___closed__6); +l_Std_Time_Formats_rfc850 = _init_l_Std_Time_Formats_rfc850(); +lean_mark_persistent(l_Std_Time_Formats_rfc850); +l_Std_Time_TimeZone_fromTimeZone___closed__1 = _init_l_Std_Time_TimeZone_fromTimeZone___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_fromTimeZone___closed__1); +l_Std_Time_TimeZone_fromTimeZone___closed__2 = _init_l_Std_Time_TimeZone_fromTimeZone___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_fromTimeZone___closed__2); +l_Std_Time_TimeZone_fromTimeZone___closed__3 = _init_l_Std_Time_TimeZone_fromTimeZone___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_fromTimeZone___closed__3); +l_Std_Time_TimeZone_fromTimeZone___closed__4 = _init_l_Std_Time_TimeZone_fromTimeZone___closed__4(); +lean_mark_persistent(l_Std_Time_TimeZone_fromTimeZone___closed__4); +l_Std_Time_TimeZone_Offset_fromOffset___closed__1 = _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_fromOffset___closed__1); +l_Std_Time_TimeZone_Offset_fromOffset___closed__2 = _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_fromOffset___closed__2); +l_Std_Time_TimeZone_Offset_fromOffset___closed__3 = _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_fromOffset___closed__3); +l_Std_Time_TimeZone_Offset_fromOffset___closed__4 = _init_l_Std_Time_TimeZone_Offset_fromOffset___closed__4(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_fromOffset___closed__4); +l_Std_Time_PlainDate_format___lambda__1___closed__1 = _init_l_Std_Time_PlainDate_format___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_format___lambda__1___closed__1); +l_Std_Time_PlainDate_format___lambda__1___closed__2 = _init_l_Std_Time_PlainDate_format___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_format___lambda__1___closed__2); +l_Std_Time_PlainDate_format___lambda__1___closed__3 = _init_l_Std_Time_PlainDate_format___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_format___lambda__1___closed__3); +l_Std_Time_PlainDate_format___lambda__1___closed__4 = _init_l_Std_Time_PlainDate_format___lambda__1___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDate_format___lambda__1___closed__4); +l_Std_Time_PlainDate_format___closed__1 = _init_l_Std_Time_PlainDate_format___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_format___closed__1); +l_Std_Time_PlainDate_format___closed__2 = _init_l_Std_Time_PlainDate_format___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_format___closed__2); +l_Std_Time_PlainDate_format___closed__3 = _init_l_Std_Time_PlainDate_format___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_format___closed__3); +l_Std_Time_PlainDate_fromAmericanDateString___closed__1 = _init_l_Std_Time_PlainDate_fromAmericanDateString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_fromAmericanDateString___closed__1); +l_Std_Time_PlainDate_fromSQLDateString___closed__1 = _init_l_Std_Time_PlainDate_fromSQLDateString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_fromSQLDateString___closed__1); +l_Std_Time_PlainDate_instToString___closed__1 = _init_l_Std_Time_PlainDate_instToString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instToString___closed__1); +l_Std_Time_PlainDate_instToString = _init_l_Std_Time_PlainDate_instToString(); +lean_mark_persistent(l_Std_Time_PlainDate_instToString); +l_Std_Time_PlainDate_instRepr___closed__1 = _init_l_Std_Time_PlainDate_instRepr___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDate_instRepr___closed__1); +l_Std_Time_PlainDate_instRepr___closed__2 = _init_l_Std_Time_PlainDate_instRepr___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDate_instRepr___closed__2); +l_Std_Time_PlainDate_instRepr___closed__3 = _init_l_Std_Time_PlainDate_instRepr___closed__3(); +lean_mark_persistent(l_Std_Time_PlainDate_instRepr___closed__3); +l_Std_Time_PlainDate_instRepr___closed__4 = _init_l_Std_Time_PlainDate_instRepr___closed__4(); +lean_mark_persistent(l_Std_Time_PlainDate_instRepr___closed__4); +l_Std_Time_PlainTime_format___lambda__1___closed__1 = _init_l_Std_Time_PlainTime_format___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_format___lambda__1___closed__1); +l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1 = _init_l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__1); +l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2 = _init_l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_fromTime24Hour___lambda__1___closed__2); +l_Std_Time_PlainTime_fromTime24Hour___closed__1 = _init_l_Std_Time_PlainTime_fromTime24Hour___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_fromTime24Hour___closed__1); +l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1 = _init_l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_fromLeanTime24Hour___closed__1); +l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1 = _init_l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_fromTime12Hour___lambda__1___closed__1); +l_Std_Time_PlainTime_fromTime12Hour___closed__1 = _init_l_Std_Time_PlainTime_fromTime12Hour___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_fromTime12Hour___closed__1); +l_Std_Time_PlainTime_instToString___closed__1 = _init_l_Std_Time_PlainTime_instToString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instToString___closed__1); +l_Std_Time_PlainTime_instToString = _init_l_Std_Time_PlainTime_instToString(); +lean_mark_persistent(l_Std_Time_PlainTime_instToString); +l_Std_Time_PlainTime_instRepr___closed__1 = _init_l_Std_Time_PlainTime_instRepr___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instRepr___closed__1); +l_Std_Time_PlainTime_instRepr___closed__2 = _init_l_Std_Time_PlainTime_instRepr___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_instRepr___closed__2); +l_Std_Time_ZonedDateTime_format___lambda__1___closed__1 = _init_l_Std_Time_ZonedDateTime_format___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_format___lambda__1___closed__1); +l_Std_Time_ZonedDateTime_format___lambda__1___closed__2 = _init_l_Std_Time_ZonedDateTime_format___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_format___lambda__1___closed__2); +l_Std_Time_ZonedDateTime_instToString___closed__1 = _init_l_Std_Time_ZonedDateTime_instToString___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instToString___closed__1); +l_Std_Time_ZonedDateTime_instToString = _init_l_Std_Time_ZonedDateTime_instToString(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instToString); +l_Std_Time_ZonedDateTime_instRepr___closed__1 = _init_l_Std_Time_ZonedDateTime_instRepr___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instRepr___closed__1); +l_Std_Time_ZonedDateTime_instRepr___closed__2 = _init_l_Std_Time_ZonedDateTime_instRepr___closed__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instRepr___closed__2); +l_Std_Time_PlainDateTime_fromAscTimeString___closed__1 = _init_l_Std_Time_PlainDateTime_fromAscTimeString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_fromAscTimeString___closed__1); +l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1 = _init_l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_toAscTimeString___lambda__1___closed__1); +l_Std_Time_PlainDateTime_instToString___closed__1 = _init_l_Std_Time_PlainDateTime_instToString___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instToString___closed__1); +l_Std_Time_PlainDateTime_instToString = _init_l_Std_Time_PlainDateTime_instToString(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instToString); +l_Std_Time_PlainDateTime_instRepr___closed__1 = _init_l_Std_Time_PlainDateTime_instRepr___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instRepr___closed__1); +l_Std_Time_PlainDateTime_instRepr___closed__2 = _init_l_Std_Time_PlainDateTime_instRepr___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_instRepr___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Format/Basic.c b/stage0/stdlib/Std/Time/Format/Basic.c new file mode 100644 index 000000000000..608dbf8bdcf3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Format/Basic.c @@ -0,0 +1,89589 @@ +// Lean compiler output +// Module: Std.Time.Format.Basic +// Imports: Std.Internal.Parsec Std.Time.Date Std.Time.Time Std.Time.Zoned Std.Time.DateTime +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprText; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1; +lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +lean_object* l_Char_isDigit___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1; +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__7; +lean_object* lean_int_mod(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11; +LEAN_EXPORT lean_object* l_Std_Time_instCoeStringFormatPart(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22; +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraShort(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instCoeTimeZone(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398_(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5; +LEAN_EXPORT lean_object* l_Std_Time_parseMonthShort(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +uint32_t lean_string_utf8_get(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__2(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong(uint8_t); +static lean_object* l_Std_Time_OffsetX_classify___closed__3; +uint8_t l_Std_Time_Weekday_ofOrdinal(lean_object*); +lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParser(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29; +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__3; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Text_classify___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instReprGenericFormat___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21; +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40; +lean_object* l_String_toNat_x21(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___lambda__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18; +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7; +static lean_object* l_Std_Time_OffsetX_classify___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_GenericFormat_spec_x21___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_instCoeModifierFormatPart(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_instReprOffsetX; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24___boxed(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1; +static lean_object* l_Std_Time_OffsetO_classify___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZoneId_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2(uint32_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedNumber; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60; +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_toCtorIdx(uint8_t); +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26(uint8_t); +lean_object* l_Std_Time_DateTime_dayOfYear___rarg(lean_object*); +lean_object* l_String_quote(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder___rarg(lean_object*); +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerNarrow(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31; +static lean_object* l_Std_Time_ZoneId_classify___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterNumber(lean_object*); +lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61; +LEAN_EXPORT lean_object* l_Std_Time_instReprZoneName; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2; +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthNarrow(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedFraction; +extern lean_object* l_Std_Internal_Parsec_expectedEndOfInput; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_toCtorIdx(uint8_t); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15; +static lean_object* l_Std_Time_OffsetZ_classify___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4(uint32_t, lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_toCtorIdx(uint8_t); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26___boxed(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4; +static uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49; +static lean_object* l_Std_Time_Text_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayShort(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7; +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50; +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(lean_object*, uint8_t, uint8_t, uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec___boxed(lean_object*); +lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseSigned(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Text_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprOffsetZ; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayLong(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort(uint8_t); +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25; +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerShort(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272_(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Text_classify(lean_object*); +lean_object* l_Std_Time_DateTime_weekOfMonth___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseYear(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3(uint32_t, lean_object*); +uint8_t l_Std_Time_DateTime_era___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier(lean_object*); +static uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52; +lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object*, uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10; +static lean_object* l_Std_Time_GenericFormat_spec_x21___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toSigned(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1; +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong(uint8_t); +static lean_object* l_Std_Time_instReprModifier___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___boxed(lean_object*); +lean_object* lean_string_utf8_byte_size(lean_object*); +static lean_object* l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28; +lean_object* lean_string_push(lean_object*, uint32_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___boxed(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___rarg(lean_object*); +lean_object* l_Fin_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedModifier; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Awareness_getD___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41; +static lean_object* l_Std_Time_OffsetX_classify___closed__5; +static lean_object* l_Std_Time_OffsetZ_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5; +lean_object* l_Std_Time_HourMarker_toAbsolute(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84; +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprNumber; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +static lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Awareness_getD(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__3(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12; +lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13; +static lean_object* l_Std_Time_Year_classify___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_classifyNumberText(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___boxed(lean_object*); +lean_object* l_panic___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1(uint32_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Format_Basic_0__Std_Time_specParser___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11; +uint8_t l_instDecidableNot___rarg(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4; +static lean_object* l_Std_Time_OffsetX_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instReprOffsetO; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprFormatPart; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedOffsetX; +uint8_t lean_uint32_dec_le(uint32_t, uint32_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38; +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser_go(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraNarrow(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87; +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedGenericFormat(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46; +LEAN_EXPORT lean_object* l_Std_Time_instReprFraction; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25___boxed(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterShort(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_classify(lean_object*); +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayNarrow(lean_object*); +static lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28___boxed(lean_object*); +extern lean_object* l_Std_Internal_Parsec_unexpectedEndOfInput; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21(lean_object*, lean_object*); +lean_object* l_Array_back_x3f___rarg(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterLong(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11; +static lean_object* l_Std_Time_instReprZoneName___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instReprYear; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort(uint8_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__14(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1(uint32_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__18(lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow(uint8_t); +lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45; +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_classify___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(lean_object*, lean_object*); +static lean_object* l_Std_Time_OffsetZ_classify___closed__3; +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8; +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedOffsetO; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraLong(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16; +lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__21(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14; +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction(lean_object*, lean_object*, lean_object*); +static lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__8(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(lean_object*, uint32_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Text_classify___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__2(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +lean_object* l_List_foldl___at_String_join___spec__1(lean_object*, lean_object*); +lean_object* lean_thunk_get_own(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprZoneId; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19; +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43; +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11; +static lean_object* l_Std_Time_GenericFormat_parse_x21___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15; +lean_object* lean_int_div(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1; +uint32_t lean_string_utf8_get_fast(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__16(lean_object*); +lean_object* l_Std_Time_Weekday_toOrdinal(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71; +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_insert(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1(uint32_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprModifier; +lean_object* l_Std_Internal_Parsec_String_pstring(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4(uint32_t, uint32_t, uint32_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser_go___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74; +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedOffsetZ; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +static lean_object* l_Std_Time_instReprFormatPart___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2; +lean_object* l_Std_Internal_Parsec_String_pchar___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18; +static lean_object* l_Std_Time_instReprOffsetZ___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse_x21(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +lean_object* lean_mk_thunk(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec_x21(lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprOffsetX___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1(uint8_t); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_toCtorIdx(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12; +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedZoneName; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18; +static lean_object* l_Std_Time_instReprGenericFormat___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1; +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric___rarg(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1; +lean_object* l_String_push___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2; +lean_object* l_Substring_nextn(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder_go(lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprText___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedText; +lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__23(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprFraction___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Text_toCtorIdx(uint8_t); +static lean_object* l_Std_Time_ZoneName_classify___closed__2; +static lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec___rarg(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995_(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13; +lean_object* lean_nat_abs(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow(uint8_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerLong(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parseWithDate(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_classify(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2; +LEAN_EXPORT uint8_t l_Std_Time_instInhabitedZoneId; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_pad___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13; +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__3(lean_object*, lean_object*); +lean_object* l_Std_Internal_Parsec_String_Parser_run___rarg(lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__5(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15; +LEAN_EXPORT lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_classify___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec_x21___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1; +uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__12(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_any___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__6(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49; +extern lean_object* l_Std_Time_TimeZone_Offset_zero; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParse(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_format___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__7(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__9(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__6(lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__2(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55; +static lean_object* l_Std_Time_OffsetX_classify___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_classify(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1; +static lean_object* l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1; +lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9; +static lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_instReprGenericFormat(lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24; +static lean_object* l_Std_Time_instReprOffsetO___closed__1; +static lean_object* l_Std_Time_Year_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseText(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Fraction_classify(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedGenericFormat___boxed(lean_object*); +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__4(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Year_classify(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +lean_object* l_List_reverse___rarg(lean_object*); +lean_object* lean_string_utf8_next_fast(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__11(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(uint8_t, lean_object*); +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_classify(uint32_t, lean_object*); +static lean_object* l_Std_Time_instReprYear___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed(lean_object*, lean_object*); +lean_object* l_Function_comp___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_classify___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_classify___boxed(lean_object*); +uint8_t l_Std_Time_HourMarker_ofOrdinal(lean_object*); +static lean_object* l_Std_Time_instInhabitedModifier___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3; +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___spec__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_OffsetO_classify___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1; +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4; +static lean_object* l_Std_Time_Fraction_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(lean_object*, lean_object*); +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90; +static lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1(lean_object*, lean_object*, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedYear; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956_(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_toCtorIdx___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1; +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7; +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_classify___boxed(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26; +lean_object* lean_int_neg(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20; +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1(uint32_t, uint32_t); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__19(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseText___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric_go(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(uint8_t, uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2; +static lean_object* l_Std_Time_Text_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(lean_object*, uint32_t, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__22(lean_object*); +static lean_object* l_Std_Time_ZoneName_classify___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27; +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_format(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3; +LEAN_EXPORT lean_object* l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_leftPad___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong(uint8_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__20(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly(lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7; +lean_object* l_Std_Time_HourMarker_toRelative(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_pad(lean_object*, lean_object*, uint8_t); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprZoneId___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2; +static lean_object* l_Std_Time_Text_classify___closed__3; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1; +static lean_object* l_Std_Time_instReprNumber___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__15(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1(uint32_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20; +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_classify(lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__17(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightPad___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_instFormatGenericFormatFormatTypeString(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10; +static lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1; +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Text_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Text_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Text_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Text_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Text_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Text_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_Text_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.short", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.full", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.narrow", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +default: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1024u); +x_16 = lean_nat_dec_le(x_15, x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18; +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20; +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprText___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprText() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprText___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedText() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Text_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 2; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Text_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Text_classify___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(4u); +x_3 = lean_nat_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = lean_nat_dec_eq(x_1, x_2); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; +x_5 = lean_unsigned_to_nat(5u); +x_6 = lean_nat_dec_eq(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = l_Std_Time_Text_classify___closed__1; +return x_8; +} +} +else +{ +lean_object* x_9; +x_9 = l_Std_Time_Text_classify___closed__2; +return x_9; +} +} +else +{ +lean_object* x_10; +x_10 = l_Std_Time_Text_classify___closed__3; +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Text_classify___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Text_classify(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("padding", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_6 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = 0; +x_8 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*1, x_7); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6; +x_10 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10; +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_7); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprNumber___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprNumber() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprNumber___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedNumber() { +_start: +{ +lean_object* x_1; +x_1 = lean_unsigned_to_nat(0u); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_classifyNumberText(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(3u); +x_3 = lean_nat_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_classify(x_1); +lean_dec(x_1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_4, 0); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_4, 0, x_8); +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +lean_dec(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.nano", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.truncated", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313_(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_unsigned_to_nat(1024u); +x_12 = lean_nat_dec_le(x_11, x_2); +x_13 = l___private_Init_Data_Repr_0__Nat_reprFast(x_10); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_13); +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9; +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_1); +if (x_12 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_17 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = 0; +x_19 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_18); +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_21 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_22 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_15); +x_23 = 0; +x_24 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_23); +x_25 = l_Repr_addAppParen(x_24, x_2); +return x_25; +} +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = lean_unsigned_to_nat(1024u); +x_28 = lean_nat_dec_le(x_27, x_2); +x_29 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_30 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_30, 0, x_29); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9; +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +if (x_28 == 0) +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_34 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = 0; +x_36 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_35); +x_37 = l_Repr_addAppParen(x_36, x_2); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_39 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_32); +x_40 = 0; +x_41 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set_uint8(x_41, sizeof(void*)*1, x_40); +x_42 = l_Repr_addAppParen(x_41, x_2); +return x_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprFraction___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprFraction() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprFraction___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedFraction() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Fraction_classify___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Fraction_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(9u); +x_3 = lean_nat_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = lean_nat_dec_eq(x_1, x_2); +lean_dec(x_1); +if (x_4 == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +else +{ +lean_object* x_6; +x_6 = l_Std_Time_Fraction_classify___closed__1; +return x_6; +} +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.twoDigit", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.fourDigit", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.extended", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466_(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +default: +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_1); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_1, 0); +x_17 = lean_unsigned_to_nat(1024u); +x_18 = lean_nat_dec_le(x_17, x_2); +x_19 = l___private_Init_Data_Repr_0__Nat_reprFast(x_16); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15; +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_1); +if (x_18 == 0) +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = 0; +x_25 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set_uint8(x_25, sizeof(void*)*1, x_24); +x_26 = l_Repr_addAppParen(x_25, x_2); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_28 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_21); +x_29 = 0; +x_30 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set_uint8(x_30, sizeof(void*)*1, x_29); +x_31 = l_Repr_addAppParen(x_30, x_2); +return x_31; +} +} +else +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_32 = lean_ctor_get(x_1, 0); +lean_inc(x_32); +lean_dec(x_1); +x_33 = lean_unsigned_to_nat(1024u); +x_34 = lean_nat_dec_le(x_33, x_2); +x_35 = l___private_Init_Data_Repr_0__Nat_reprFast(x_32); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15; +x_38 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +if (x_34 == 0) +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; +x_39 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_40 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = 0; +x_42 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set_uint8(x_42, sizeof(void*)*1, x_41); +x_43 = l_Repr_addAppParen(x_42, x_2); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_44 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_45 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_38); +x_46 = 0; +x_47 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set_uint8(x_47, sizeof(void*)*1, x_46); +x_48 = l_Repr_addAppParen(x_47, x_2); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprYear___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprYear() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprYear___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedYear() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Year_classify___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(1); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Year_classify___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Year_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(2u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(4u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = lean_nat_dec_lt(x_4, x_1); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(3u); +x_8 = lean_nat_dec_eq(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_1); +x_9 = lean_box(0); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_10, 0, x_1); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_10); +return x_11; +} +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_12, 0, x_1); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; +lean_dec(x_1); +x_14 = l_Std_Time_Year_classify___closed__1; +return x_14; +} +} +else +{ +lean_object* x_15; +lean_dec(x_1); +x_15 = l_Std_Time_Year_classify___closed__2; +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_ZoneId_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_ZoneId_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_ZoneId_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneId.short", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneId.full", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprZoneId___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprZoneId() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprZoneId___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedZoneId() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZoneId_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZoneId_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(2u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(4u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = l_Std_Time_ZoneId_classify___closed__1; +return x_7; +} +} +else +{ +lean_object* x_8; +x_8 = l_Std_Time_ZoneId_classify___closed__2; +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneId_classify___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZoneId_classify(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_ZoneName_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_ZoneName_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_ZoneName_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.short", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.full", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprZoneName___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprZoneName() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprZoneName___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedZoneName() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZoneName_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZoneName_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_classify(uint32_t x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; uint8_t x_4; +x_3 = 122; +x_4 = lean_uint32_dec_eq(x_1, x_3); +if (x_4 == 0) +{ +uint32_t x_5; uint8_t x_6; +x_5 = 118; +x_6 = lean_uint32_dec_eq(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_dec_eq(x_2, x_8); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(4u); +x_11 = lean_nat_dec_eq(x_2, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; +x_13 = l_Std_Time_ZoneName_classify___closed__1; +return x_13; +} +} +else +{ +lean_object* x_14; +x_14 = l_Std_Time_ZoneName_classify___closed__2; +return x_14; +} +} +} +else +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(4u); +x_16 = lean_nat_dec_lt(x_2, x_15); +if (x_16 == 0) +{ +uint8_t x_17; +x_17 = lean_nat_dec_eq(x_2, x_15); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_box(0); +return x_18; +} +else +{ +lean_object* x_19; +x_19 = l_Std_Time_ZoneName_classify___closed__1; +return x_19; +} +} +else +{ +lean_object* x_20; +x_20 = l_Std_Time_ZoneName_classify___closed__2; +return x_20; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZoneName_classify___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_ZoneName_classify(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = lean_unsigned_to_nat(3u); +return x_5; +} +default: +{ +lean_object* x_6; +x_6 = lean_unsigned_to_nat(4u); +return x_6; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_OffsetX_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_OffsetX_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_OffsetX_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hour", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteColon", 32, 32); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecond", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995_(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +case 2: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1024u); +x_16 = lean_nat_dec_le(x_15, x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16; +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18; +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +} +case 3: +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(1024u); +x_22 = lean_nat_dec_le(x_21, x_2); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22; +x_24 = l_Repr_addAppParen(x_23, x_2); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24; +x_26 = l_Repr_addAppParen(x_25, x_2); +return x_26; +} +} +default: +{ +lean_object* x_27; uint8_t x_28; +x_27 = lean_unsigned_to_nat(1024u); +x_28 = lean_nat_dec_le(x_27, x_2); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28; +x_30 = l_Repr_addAppParen(x_29, x_2); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30; +x_32 = l_Repr_addAppParen(x_31, x_2); +return x_32; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetX___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetX() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprOffsetX___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedOffsetX() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_OffsetX_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 4; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetX_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 3; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetX_classify___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 2; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetX_classify___closed__4() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetX_classify___closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(3u); +x_7 = lean_nat_dec_eq(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(4u); +x_9 = lean_nat_dec_eq(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(5u); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; +x_13 = l_Std_Time_OffsetX_classify___closed__1; +return x_13; +} +} +else +{ +lean_object* x_14; +x_14 = l_Std_Time_OffsetX_classify___closed__2; +return x_14; +} +} +else +{ +lean_object* x_15; +x_15 = l_Std_Time_OffsetX_classify___closed__3; +return x_15; +} +} +else +{ +lean_object* x_16; +x_16 = l_Std_Time_OffsetX_classify___closed__4; +return x_16; +} +} +else +{ +lean_object* x_17; +x_17 = l_Std_Time_OffsetX_classify___closed__5; +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetX_classify___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_OffsetX_classify(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_OffsetO_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_OffsetO_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_OffsetO_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.short", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetO___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetO() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprOffsetO___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedOffsetO() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_OffsetO_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetO_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(4u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_box(0); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = l_Std_Time_OffsetO_classify___closed__1; +return x_7; +} +} +else +{ +lean_object* x_8; +x_8 = l_Std_Time_OffsetO_classify___closed__2; +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetO_classify___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_OffsetO_classify(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_OffsetZ_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_OffsetZ_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_OffsetZ_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398_(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +default: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1024u); +x_16 = lean_nat_dec_le(x_15, x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16; +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18; +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetZ___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprOffsetZ() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprOffsetZ___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_instInhabitedOffsetZ() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l_Std_Time_OffsetZ_classify___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 2; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetZ_classify___closed__2() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_OffsetZ_classify___closed__3() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 0; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_classify(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_dec_eq(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_unsigned_to_nat(3u); +x_7 = lean_nat_dec_eq(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_unsigned_to_nat(4u); +x_9 = lean_nat_dec_eq(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_unsigned_to_nat(5u); +x_11 = lean_nat_dec_eq(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_box(0); +return x_12; +} +else +{ +lean_object* x_13; +x_13 = l_Std_Time_OffsetZ_classify___closed__1; +return x_13; +} +} +else +{ +lean_object* x_14; +x_14 = l_Std_Time_OffsetZ_classify___closed__2; +return x_14; +} +} +else +{ +lean_object* x_15; +x_15 = l_Std_Time_OffsetZ_classify___closed__3; +return x_15; +} +} +else +{ +lean_object* x_16; +x_16 = l_Std_Time_OffsetZ_classify___closed__3; +return x_16; +} +} +else +{ +lean_object* x_17; +x_17 = l_Std_Time_OffsetZ_classify___closed__3; +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_OffsetZ_classify___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_OffsetZ_classify(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sum.inl ", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sum.inr ", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_unsigned_to_nat(1024u); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_3, x_4); +x_6 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2; +x_7 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_unsigned_to_nat(1024u); +x_11 = lean_unbox(x_9); +lean_dec(x_9); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(x_11, x_10); +x_13 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l_Repr_addAppParen(x_14, x_2); +return x_15; +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.G", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.y", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.u", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.D", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.MorL", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.d", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Qorq", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.w", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.W", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.E", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.eorc", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.F", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.a", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.h", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.K", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.k", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.H", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.m", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.s", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.S", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.A", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.n", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.N", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.V", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.O", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.X", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.x", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836_(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_4 = lean_unsigned_to_nat(1024u); +x_5 = lean_nat_dec_le(x_4, x_2); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(x_3, x_4); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +if (x_5 == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l_Repr_addAppParen(x_12, x_2); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_15 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_8); +x_16 = 0; +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_16); +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +} +case 1: +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_19 = lean_ctor_get(x_1, 0); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_unsigned_to_nat(1024u); +x_21 = lean_nat_dec_le(x_20, x_2); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466_(x_19, x_20); +x_23 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6; +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +if (x_21 == 0) +{ +lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = 0; +x_28 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set_uint8(x_28, sizeof(void*)*1, x_27); +x_29 = l_Repr_addAppParen(x_28, x_2); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; +x_30 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +x_32 = 0; +x_33 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); +x_34 = l_Repr_addAppParen(x_33, x_2); +return x_34; +} +} +case 2: +{ +lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_35 = lean_ctor_get(x_1, 0); +lean_inc(x_35); +lean_dec(x_1); +x_36 = lean_unsigned_to_nat(1024u); +x_37 = lean_nat_dec_le(x_36, x_2); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466_(x_35, x_36); +x_39 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9; +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +if (x_37 == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; +x_41 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_42 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = 0; +x_44 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set_uint8(x_44, sizeof(void*)*1, x_43); +x_45 = l_Repr_addAppParen(x_44, x_2); +return x_45; +} +else +{ +lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_47 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_40); +x_48 = 0; +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_48); +x_50 = l_Repr_addAppParen(x_49, x_2); +return x_50; +} +} +case 3: +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_51 = lean_ctor_get(x_1, 0); +lean_inc(x_51); +lean_dec(x_1); +x_52 = lean_unsigned_to_nat(1024u); +x_53 = lean_nat_dec_le(x_52, x_2); +x_54 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_51, x_52); +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12; +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +if (x_53 == 0) +{ +lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; +x_57 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_58 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +x_59 = 0; +x_60 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set_uint8(x_60, sizeof(void*)*1, x_59); +x_61 = l_Repr_addAppParen(x_60, x_2); +return x_61; +} +else +{ +lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; +x_62 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_63 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_56); +x_64 = 0; +x_65 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set_uint8(x_65, sizeof(void*)*1, x_64); +x_66 = l_Repr_addAppParen(x_65, x_2); +return x_66; +} +} +case 4: +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_67 = lean_ctor_get(x_1, 0); +lean_inc(x_67); +lean_dec(x_1); +x_68 = lean_unsigned_to_nat(1024u); +x_69 = lean_nat_dec_le(x_68, x_2); +x_70 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(x_67, x_68); +x_71 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15; +x_72 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_70); +if (x_69 == 0) +{ +lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; +x_73 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_74 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_72); +x_75 = 0; +x_76 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set_uint8(x_76, sizeof(void*)*1, x_75); +x_77 = l_Repr_addAppParen(x_76, x_2); +return x_77; +} +else +{ +lean_object* x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_79 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_72); +x_80 = 0; +x_81 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_81, 0, x_79); +lean_ctor_set_uint8(x_81, sizeof(void*)*1, x_80); +x_82 = l_Repr_addAppParen(x_81, x_2); +return x_82; +} +} +case 5: +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_83 = lean_ctor_get(x_1, 0); +lean_inc(x_83); +lean_dec(x_1); +x_84 = lean_unsigned_to_nat(1024u); +x_85 = lean_nat_dec_le(x_84, x_2); +x_86 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_83, x_84); +x_87 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18; +x_88 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_86); +if (x_85 == 0) +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; +x_89 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_90 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_88); +x_91 = 0; +x_92 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set_uint8(x_92, sizeof(void*)*1, x_91); +x_93 = l_Repr_addAppParen(x_92, x_2); +return x_93; +} +else +{ +lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; +x_94 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_95 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_88); +x_96 = 0; +x_97 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set_uint8(x_97, sizeof(void*)*1, x_96); +x_98 = l_Repr_addAppParen(x_97, x_2); +return x_98; +} +} +case 6: +{ +lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_99 = lean_ctor_get(x_1, 0); +lean_inc(x_99); +lean_dec(x_1); +x_100 = lean_unsigned_to_nat(1024u); +x_101 = lean_nat_dec_le(x_100, x_2); +x_102 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(x_99, x_100); +x_103 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21; +x_104 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_102); +if (x_101 == 0) +{ +lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; +x_105 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_106 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_104); +x_107 = 0; +x_108 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set_uint8(x_108, sizeof(void*)*1, x_107); +x_109 = l_Repr_addAppParen(x_108, x_2); +return x_109; +} +else +{ +lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; lean_object* x_114; +x_110 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_111 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_104); +x_112 = 0; +x_113 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set_uint8(x_113, sizeof(void*)*1, x_112); +x_114 = l_Repr_addAppParen(x_113, x_2); +return x_114; +} +} +case 7: +{ +lean_object* x_115; lean_object* x_116; uint8_t x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_115 = lean_ctor_get(x_1, 0); +lean_inc(x_115); +lean_dec(x_1); +x_116 = lean_unsigned_to_nat(1024u); +x_117 = lean_nat_dec_le(x_116, x_2); +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_115, x_116); +x_119 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24; +x_120 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_118); +if (x_117 == 0) +{ +lean_object* x_121; lean_object* x_122; uint8_t x_123; lean_object* x_124; lean_object* x_125; +x_121 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_122 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_120); +x_123 = 0; +x_124 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set_uint8(x_124, sizeof(void*)*1, x_123); +x_125 = l_Repr_addAppParen(x_124, x_2); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_127 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_120); +x_128 = 0; +x_129 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_129, 0, x_127); +lean_ctor_set_uint8(x_129, sizeof(void*)*1, x_128); +x_130 = l_Repr_addAppParen(x_129, x_2); +return x_130; +} +} +case 8: +{ +lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_131 = lean_ctor_get(x_1, 0); +lean_inc(x_131); +lean_dec(x_1); +x_132 = lean_unsigned_to_nat(1024u); +x_133 = lean_nat_dec_le(x_132, x_2); +x_134 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_131, x_132); +x_135 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27; +x_136 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_134); +if (x_133 == 0) +{ +lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; +x_137 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_138 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_136); +x_139 = 0; +x_140 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set_uint8(x_140, sizeof(void*)*1, x_139); +x_141 = l_Repr_addAppParen(x_140, x_2); +return x_141; +} +else +{ +lean_object* x_142; lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; +x_142 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_143 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_136); +x_144 = 0; +x_145 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set_uint8(x_145, sizeof(void*)*1, x_144); +x_146 = l_Repr_addAppParen(x_145, x_2); +return x_146; +} +} +case 9: +{ +uint8_t x_147; lean_object* x_148; uint8_t x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_147 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_148 = lean_unsigned_to_nat(1024u); +x_149 = lean_nat_dec_le(x_148, x_2); +x_150 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(x_147, x_148); +x_151 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30; +x_152 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_150); +if (x_149 == 0) +{ +lean_object* x_153; lean_object* x_154; uint8_t x_155; lean_object* x_156; lean_object* x_157; +x_153 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_154 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_152); +x_155 = 0; +x_156 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set_uint8(x_156, sizeof(void*)*1, x_155); +x_157 = l_Repr_addAppParen(x_156, x_2); +return x_157; +} +else +{ +lean_object* x_158; lean_object* x_159; uint8_t x_160; lean_object* x_161; lean_object* x_162; +x_158 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_159 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_152); +x_160 = 0; +x_161 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_161, 0, x_159); +lean_ctor_set_uint8(x_161, sizeof(void*)*1, x_160); +x_162 = l_Repr_addAppParen(x_161, x_2); +return x_162; +} +} +case 10: +{ +lean_object* x_163; lean_object* x_164; uint8_t x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_163 = lean_ctor_get(x_1, 0); +lean_inc(x_163); +lean_dec(x_1); +x_164 = lean_unsigned_to_nat(1024u); +x_165 = lean_nat_dec_le(x_164, x_2); +x_166 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(x_163, x_164); +x_167 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33; +x_168 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_166); +if (x_165 == 0) +{ +lean_object* x_169; lean_object* x_170; uint8_t x_171; lean_object* x_172; lean_object* x_173; +x_169 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_170 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_168); +x_171 = 0; +x_172 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set_uint8(x_172, sizeof(void*)*1, x_171); +x_173 = l_Repr_addAppParen(x_172, x_2); +return x_173; +} +else +{ +lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; lean_object* x_178; +x_174 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_175 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_168); +x_176 = 0; +x_177 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set_uint8(x_177, sizeof(void*)*1, x_176); +x_178 = l_Repr_addAppParen(x_177, x_2); +return x_178; +} +} +case 11: +{ +lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_179 = lean_ctor_get(x_1, 0); +lean_inc(x_179); +lean_dec(x_1); +x_180 = lean_unsigned_to_nat(1024u); +x_181 = lean_nat_dec_le(x_180, x_2); +x_182 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_179, x_180); +x_183 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36; +x_184 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_182); +if (x_181 == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +x_185 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_186 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_184); +x_187 = 0; +x_188 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set_uint8(x_188, sizeof(void*)*1, x_187); +x_189 = l_Repr_addAppParen(x_188, x_2); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; uint8_t x_192; lean_object* x_193; lean_object* x_194; +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_191 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_184); +x_192 = 0; +x_193 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set_uint8(x_193, sizeof(void*)*1, x_192); +x_194 = l_Repr_addAppParen(x_193, x_2); +return x_194; +} +} +case 12: +{ +uint8_t x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_195 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_196 = lean_unsigned_to_nat(1024u); +x_197 = lean_nat_dec_le(x_196, x_2); +x_198 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14_(x_195, x_196); +x_199 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39; +x_200 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_198); +if (x_197 == 0) +{ +lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; lean_object* x_205; +x_201 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_202 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_200); +x_203 = 0; +x_204 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_204, 0, x_202); +lean_ctor_set_uint8(x_204, sizeof(void*)*1, x_203); +x_205 = l_Repr_addAppParen(x_204, x_2); +return x_205; +} +else +{ +lean_object* x_206; lean_object* x_207; uint8_t x_208; lean_object* x_209; lean_object* x_210; +x_206 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_207 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_200); +x_208 = 0; +x_209 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set_uint8(x_209, sizeof(void*)*1, x_208); +x_210 = l_Repr_addAppParen(x_209, x_2); +return x_210; +} +} +case 13: +{ +lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_211 = lean_ctor_get(x_1, 0); +lean_inc(x_211); +lean_dec(x_1); +x_212 = lean_unsigned_to_nat(1024u); +x_213 = lean_nat_dec_le(x_212, x_2); +x_214 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_211, x_212); +x_215 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42; +x_216 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_214); +if (x_213 == 0) +{ +lean_object* x_217; lean_object* x_218; uint8_t x_219; lean_object* x_220; lean_object* x_221; +x_217 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_218 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_218, 0, x_217); +lean_ctor_set(x_218, 1, x_216); +x_219 = 0; +x_220 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_220, 0, x_218); +lean_ctor_set_uint8(x_220, sizeof(void*)*1, x_219); +x_221 = l_Repr_addAppParen(x_220, x_2); +return x_221; +} +else +{ +lean_object* x_222; lean_object* x_223; uint8_t x_224; lean_object* x_225; lean_object* x_226; +x_222 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_223 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_216); +x_224 = 0; +x_225 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_225, 0, x_223); +lean_ctor_set_uint8(x_225, sizeof(void*)*1, x_224); +x_226 = l_Repr_addAppParen(x_225, x_2); +return x_226; +} +} +case 14: +{ +lean_object* x_227; lean_object* x_228; uint8_t x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +x_227 = lean_ctor_get(x_1, 0); +lean_inc(x_227); +lean_dec(x_1); +x_228 = lean_unsigned_to_nat(1024u); +x_229 = lean_nat_dec_le(x_228, x_2); +x_230 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_227, x_228); +x_231 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45; +x_232 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_232, 0, x_231); +lean_ctor_set(x_232, 1, x_230); +if (x_229 == 0) +{ +lean_object* x_233; lean_object* x_234; uint8_t x_235; lean_object* x_236; lean_object* x_237; +x_233 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_234 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_234, 0, x_233); +lean_ctor_set(x_234, 1, x_232); +x_235 = 0; +x_236 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_236, 0, x_234); +lean_ctor_set_uint8(x_236, sizeof(void*)*1, x_235); +x_237 = l_Repr_addAppParen(x_236, x_2); +return x_237; +} +else +{ +lean_object* x_238; lean_object* x_239; uint8_t x_240; lean_object* x_241; lean_object* x_242; +x_238 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_239 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_239, 0, x_238); +lean_ctor_set(x_239, 1, x_232); +x_240 = 0; +x_241 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set_uint8(x_241, sizeof(void*)*1, x_240); +x_242 = l_Repr_addAppParen(x_241, x_2); +return x_242; +} +} +case 15: +{ +lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_243 = lean_ctor_get(x_1, 0); +lean_inc(x_243); +lean_dec(x_1); +x_244 = lean_unsigned_to_nat(1024u); +x_245 = lean_nat_dec_le(x_244, x_2); +x_246 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_243, x_244); +x_247 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48; +x_248 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_248, 0, x_247); +lean_ctor_set(x_248, 1, x_246); +if (x_245 == 0) +{ +lean_object* x_249; lean_object* x_250; uint8_t x_251; lean_object* x_252; lean_object* x_253; +x_249 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_250 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_248); +x_251 = 0; +x_252 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_252, 0, x_250); +lean_ctor_set_uint8(x_252, sizeof(void*)*1, x_251); +x_253 = l_Repr_addAppParen(x_252, x_2); +return x_253; +} +else +{ +lean_object* x_254; lean_object* x_255; uint8_t x_256; lean_object* x_257; lean_object* x_258; +x_254 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_255 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_248); +x_256 = 0; +x_257 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set_uint8(x_257, sizeof(void*)*1, x_256); +x_258 = l_Repr_addAppParen(x_257, x_2); +return x_258; +} +} +case 16: +{ +lean_object* x_259; lean_object* x_260; uint8_t x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; +x_259 = lean_ctor_get(x_1, 0); +lean_inc(x_259); +lean_dec(x_1); +x_260 = lean_unsigned_to_nat(1024u); +x_261 = lean_nat_dec_le(x_260, x_2); +x_262 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_259, x_260); +x_263 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51; +x_264 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_262); +if (x_261 == 0) +{ +lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; lean_object* x_269; +x_265 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_266 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_264); +x_267 = 0; +x_268 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_268, 0, x_266); +lean_ctor_set_uint8(x_268, sizeof(void*)*1, x_267); +x_269 = l_Repr_addAppParen(x_268, x_2); +return x_269; +} +else +{ +lean_object* x_270; lean_object* x_271; uint8_t x_272; lean_object* x_273; lean_object* x_274; +x_270 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_271 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_264); +x_272 = 0; +x_273 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_273, 0, x_271); +lean_ctor_set_uint8(x_273, sizeof(void*)*1, x_272); +x_274 = l_Repr_addAppParen(x_273, x_2); +return x_274; +} +} +case 17: +{ +lean_object* x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; +x_275 = lean_ctor_get(x_1, 0); +lean_inc(x_275); +lean_dec(x_1); +x_276 = lean_unsigned_to_nat(1024u); +x_277 = lean_nat_dec_le(x_276, x_2); +x_278 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_275, x_276); +x_279 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54; +x_280 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_278); +if (x_277 == 0) +{ +lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; +x_281 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_282 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_282, 0, x_281); +lean_ctor_set(x_282, 1, x_280); +x_283 = 0; +x_284 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_284, 0, x_282); +lean_ctor_set_uint8(x_284, sizeof(void*)*1, x_283); +x_285 = l_Repr_addAppParen(x_284, x_2); +return x_285; +} +else +{ +lean_object* x_286; lean_object* x_287; uint8_t x_288; lean_object* x_289; lean_object* x_290; +x_286 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_287 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_287, 0, x_286); +lean_ctor_set(x_287, 1, x_280); +x_288 = 0; +x_289 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_289, 0, x_287); +lean_ctor_set_uint8(x_289, sizeof(void*)*1, x_288); +x_290 = l_Repr_addAppParen(x_289, x_2); +return x_290; +} +} +case 18: +{ +lean_object* x_291; lean_object* x_292; uint8_t x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_291 = lean_ctor_get(x_1, 0); +lean_inc(x_291); +lean_dec(x_1); +x_292 = lean_unsigned_to_nat(1024u); +x_293 = lean_nat_dec_le(x_292, x_2); +x_294 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_291, x_292); +x_295 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57; +x_296 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_294); +if (x_293 == 0) +{ +lean_object* x_297; lean_object* x_298; uint8_t x_299; lean_object* x_300; lean_object* x_301; +x_297 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_298 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_298, 0, x_297); +lean_ctor_set(x_298, 1, x_296); +x_299 = 0; +x_300 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_300, 0, x_298); +lean_ctor_set_uint8(x_300, sizeof(void*)*1, x_299); +x_301 = l_Repr_addAppParen(x_300, x_2); +return x_301; +} +else +{ +lean_object* x_302; lean_object* x_303; uint8_t x_304; lean_object* x_305; lean_object* x_306; +x_302 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_303 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_303, 0, x_302); +lean_ctor_set(x_303, 1, x_296); +x_304 = 0; +x_305 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_305, 0, x_303); +lean_ctor_set_uint8(x_305, sizeof(void*)*1, x_304); +x_306 = l_Repr_addAppParen(x_305, x_2); +return x_306; +} +} +case 19: +{ +lean_object* x_307; lean_object* x_308; uint8_t x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_307 = lean_ctor_get(x_1, 0); +lean_inc(x_307); +lean_dec(x_1); +x_308 = lean_unsigned_to_nat(1024u); +x_309 = lean_nat_dec_le(x_308, x_2); +x_310 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313_(x_307, x_308); +x_311 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60; +x_312 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_312, 0, x_311); +lean_ctor_set(x_312, 1, x_310); +if (x_309 == 0) +{ +lean_object* x_313; lean_object* x_314; uint8_t x_315; lean_object* x_316; lean_object* x_317; +x_313 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_314 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_314, 0, x_313); +lean_ctor_set(x_314, 1, x_312); +x_315 = 0; +x_316 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_316, 0, x_314); +lean_ctor_set_uint8(x_316, sizeof(void*)*1, x_315); +x_317 = l_Repr_addAppParen(x_316, x_2); +return x_317; +} +else +{ +lean_object* x_318; lean_object* x_319; uint8_t x_320; lean_object* x_321; lean_object* x_322; +x_318 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_319 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_319, 0, x_318); +lean_ctor_set(x_319, 1, x_312); +x_320 = 0; +x_321 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_321, 0, x_319); +lean_ctor_set_uint8(x_321, sizeof(void*)*1, x_320); +x_322 = l_Repr_addAppParen(x_321, x_2); +return x_322; +} +} +case 20: +{ +lean_object* x_323; lean_object* x_324; uint8_t x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_323 = lean_ctor_get(x_1, 0); +lean_inc(x_323); +lean_dec(x_1); +x_324 = lean_unsigned_to_nat(1024u); +x_325 = lean_nat_dec_le(x_324, x_2); +x_326 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_323, x_324); +x_327 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63; +x_328 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_326); +if (x_325 == 0) +{ +lean_object* x_329; lean_object* x_330; uint8_t x_331; lean_object* x_332; lean_object* x_333; +x_329 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_330 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_328); +x_331 = 0; +x_332 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_332, 0, x_330); +lean_ctor_set_uint8(x_332, sizeof(void*)*1, x_331); +x_333 = l_Repr_addAppParen(x_332, x_2); +return x_333; +} +else +{ +lean_object* x_334; lean_object* x_335; uint8_t x_336; lean_object* x_337; lean_object* x_338; +x_334 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_335 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_335, 0, x_334); +lean_ctor_set(x_335, 1, x_328); +x_336 = 0; +x_337 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_337, 0, x_335); +lean_ctor_set_uint8(x_337, sizeof(void*)*1, x_336); +x_338 = l_Repr_addAppParen(x_337, x_2); +return x_338; +} +} +case 21: +{ +lean_object* x_339; lean_object* x_340; uint8_t x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; +x_339 = lean_ctor_get(x_1, 0); +lean_inc(x_339); +lean_dec(x_1); +x_340 = lean_unsigned_to_nat(1024u); +x_341 = lean_nat_dec_le(x_340, x_2); +x_342 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_339, x_340); +x_343 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66; +x_344 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_344, 0, x_343); +lean_ctor_set(x_344, 1, x_342); +if (x_341 == 0) +{ +lean_object* x_345; lean_object* x_346; uint8_t x_347; lean_object* x_348; lean_object* x_349; +x_345 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_346 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_346, 0, x_345); +lean_ctor_set(x_346, 1, x_344); +x_347 = 0; +x_348 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_348, 0, x_346); +lean_ctor_set_uint8(x_348, sizeof(void*)*1, x_347); +x_349 = l_Repr_addAppParen(x_348, x_2); +return x_349; +} +else +{ +lean_object* x_350; lean_object* x_351; uint8_t x_352; lean_object* x_353; lean_object* x_354; +x_350 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_351 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_351, 0, x_350); +lean_ctor_set(x_351, 1, x_344); +x_352 = 0; +x_353 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_353, 0, x_351); +lean_ctor_set_uint8(x_353, sizeof(void*)*1, x_352); +x_354 = l_Repr_addAppParen(x_353, x_2); +return x_354; +} +} +case 22: +{ +lean_object* x_355; lean_object* x_356; uint8_t x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; +x_355 = lean_ctor_get(x_1, 0); +lean_inc(x_355); +lean_dec(x_1); +x_356 = lean_unsigned_to_nat(1024u); +x_357 = lean_nat_dec_le(x_356, x_2); +x_358 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204_(x_355, x_356); +x_359 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69; +x_360 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_360, 0, x_359); +lean_ctor_set(x_360, 1, x_358); +if (x_357 == 0) +{ +lean_object* x_361; lean_object* x_362; uint8_t x_363; lean_object* x_364; lean_object* x_365; +x_361 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_362 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_362, 0, x_361); +lean_ctor_set(x_362, 1, x_360); +x_363 = 0; +x_364 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set_uint8(x_364, sizeof(void*)*1, x_363); +x_365 = l_Repr_addAppParen(x_364, x_2); +return x_365; +} +else +{ +lean_object* x_366; lean_object* x_367; uint8_t x_368; lean_object* x_369; lean_object* x_370; +x_366 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_367 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_367, 0, x_366); +lean_ctor_set(x_367, 1, x_360); +x_368 = 0; +x_369 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_369, 0, x_367); +lean_ctor_set_uint8(x_369, sizeof(void*)*1, x_368); +x_370 = l_Repr_addAppParen(x_369, x_2); +return x_370; +} +} +case 23: +{ +lean_object* x_371; uint8_t x_372; +x_371 = lean_unsigned_to_nat(1024u); +x_372 = lean_nat_dec_le(x_371, x_2); +if (x_372 == 0) +{ +lean_object* x_373; lean_object* x_374; +x_373 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73; +x_374 = l_Repr_addAppParen(x_373, x_2); +return x_374; +} +else +{ +lean_object* x_375; lean_object* x_376; +x_375 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75; +x_376 = l_Repr_addAppParen(x_375, x_2); +return x_376; +} +} +case 24: +{ +uint8_t x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; +x_377 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_378 = lean_unsigned_to_nat(1024u); +x_379 = lean_nat_dec_le(x_378, x_2); +x_380 = l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804_(x_377, x_378); +x_381 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78; +x_382 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_382, 0, x_381); +lean_ctor_set(x_382, 1, x_380); +if (x_379 == 0) +{ +lean_object* x_383; lean_object* x_384; uint8_t x_385; lean_object* x_386; lean_object* x_387; +x_383 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_384 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_382); +x_385 = 0; +x_386 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_386, 0, x_384); +lean_ctor_set_uint8(x_386, sizeof(void*)*1, x_385); +x_387 = l_Repr_addAppParen(x_386, x_2); +return x_387; +} +else +{ +lean_object* x_388; lean_object* x_389; uint8_t x_390; lean_object* x_391; lean_object* x_392; +x_388 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_389 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_389, 0, x_388); +lean_ctor_set(x_389, 1, x_382); +x_390 = 0; +x_391 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_391, 0, x_389); +lean_ctor_set_uint8(x_391, sizeof(void*)*1, x_390); +x_392 = l_Repr_addAppParen(x_391, x_2); +return x_392; +} +} +case 25: +{ +uint8_t x_393; lean_object* x_394; uint8_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; +x_393 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_394 = lean_unsigned_to_nat(1024u); +x_395 = lean_nat_dec_le(x_394, x_2); +x_396 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272_(x_393, x_394); +x_397 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81; +x_398 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_398, 0, x_397); +lean_ctor_set(x_398, 1, x_396); +if (x_395 == 0) +{ +lean_object* x_399; lean_object* x_400; uint8_t x_401; lean_object* x_402; lean_object* x_403; +x_399 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_400 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_400, 0, x_399); +lean_ctor_set(x_400, 1, x_398); +x_401 = 0; +x_402 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_402, 0, x_400); +lean_ctor_set_uint8(x_402, sizeof(void*)*1, x_401); +x_403 = l_Repr_addAppParen(x_402, x_2); +return x_403; +} +else +{ +lean_object* x_404; lean_object* x_405; uint8_t x_406; lean_object* x_407; lean_object* x_408; +x_404 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_405 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_405, 0, x_404); +lean_ctor_set(x_405, 1, x_398); +x_406 = 0; +x_407 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_407, 0, x_405); +lean_ctor_set_uint8(x_407, sizeof(void*)*1, x_406); +x_408 = l_Repr_addAppParen(x_407, x_2); +return x_408; +} +} +case 26: +{ +uint8_t x_409; lean_object* x_410; uint8_t x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; +x_409 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_410 = lean_unsigned_to_nat(1024u); +x_411 = lean_nat_dec_le(x_410, x_2); +x_412 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995_(x_409, x_410); +x_413 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84; +x_414 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_414, 0, x_413); +lean_ctor_set(x_414, 1, x_412); +if (x_411 == 0) +{ +lean_object* x_415; lean_object* x_416; uint8_t x_417; lean_object* x_418; lean_object* x_419; +x_415 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_416 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_416, 0, x_415); +lean_ctor_set(x_416, 1, x_414); +x_417 = 0; +x_418 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_418, 0, x_416); +lean_ctor_set_uint8(x_418, sizeof(void*)*1, x_417); +x_419 = l_Repr_addAppParen(x_418, x_2); +return x_419; +} +else +{ +lean_object* x_420; lean_object* x_421; uint8_t x_422; lean_object* x_423; lean_object* x_424; +x_420 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_421 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_421, 0, x_420); +lean_ctor_set(x_421, 1, x_414); +x_422 = 0; +x_423 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_423, 0, x_421); +lean_ctor_set_uint8(x_423, sizeof(void*)*1, x_422); +x_424 = l_Repr_addAppParen(x_423, x_2); +return x_424; +} +} +case 27: +{ +uint8_t x_425; lean_object* x_426; uint8_t x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +x_425 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_426 = lean_unsigned_to_nat(1024u); +x_427 = lean_nat_dec_le(x_426, x_2); +x_428 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995_(x_425, x_426); +x_429 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87; +x_430 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_430, 0, x_429); +lean_ctor_set(x_430, 1, x_428); +if (x_427 == 0) +{ +lean_object* x_431; lean_object* x_432; uint8_t x_433; lean_object* x_434; lean_object* x_435; +x_431 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_432 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_432, 0, x_431); +lean_ctor_set(x_432, 1, x_430); +x_433 = 0; +x_434 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_434, 0, x_432); +lean_ctor_set_uint8(x_434, sizeof(void*)*1, x_433); +x_435 = l_Repr_addAppParen(x_434, x_2); +return x_435; +} +else +{ +lean_object* x_436; lean_object* x_437; uint8_t x_438; lean_object* x_439; lean_object* x_440; +x_436 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_437 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_437, 0, x_436); +lean_ctor_set(x_437, 1, x_430); +x_438 = 0; +x_439 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_439, 0, x_437); +lean_ctor_set_uint8(x_439, sizeof(void*)*1, x_438); +x_440 = l_Repr_addAppParen(x_439, x_2); +return x_440; +} +} +default: +{ +uint8_t x_441; lean_object* x_442; uint8_t x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +x_441 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_442 = lean_unsigned_to_nat(1024u); +x_443 = lean_nat_dec_le(x_442, x_2); +x_444 = l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398_(x_441, x_442); +x_445 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90; +x_446 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_446, 0, x_445); +lean_ctor_set(x_446, 1, x_444); +if (x_443 == 0) +{ +lean_object* x_447; lean_object* x_448; uint8_t x_449; lean_object* x_450; lean_object* x_451; +x_447 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_448 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_448, 0, x_447); +lean_ctor_set(x_448, 1, x_446); +x_449 = 0; +x_450 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_450, 0, x_448); +lean_ctor_set_uint8(x_450, sizeof(void*)*1, x_449); +x_451 = l_Repr_addAppParen(x_450, x_2); +return x_451; +} +else +{ +lean_object* x_452; lean_object* x_453; uint8_t x_454; lean_object* x_455; lean_object* x_456; +x_452 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_453 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_453, 0, x_452); +lean_ctor_set(x_453, 1, x_446); +x_454 = 0; +x_455 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_455, 0, x_453); +lean_ctor_set_uint8(x_455, sizeof(void*)*1, x_454); +x_456 = l_Repr_addAppParen(x_455, x_2); +return x_456; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprModifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprModifier() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprModifier___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedModifier___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 0; +x_2 = lean_alloc_ctor(0, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedModifier() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedModifier___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid quantity of characters for '", 36, 36); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("'", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_string_length(x_3); +x_6 = lean_apply_1(x_2, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; uint32_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_1); +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_string_utf8_get(x_3, x_7); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_10 = lean_string_push(x_9, x_8); +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2; +x_12 = lean_string_append(x_11, x_10); +lean_dec(x_10); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_14 = lean_string_append(x_12, x_13); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +x_17 = lean_apply_1(x_1, x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_4); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Text_classify___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseText(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseText___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseText(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Fraction_classify), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_string_length(x_2); +x_5 = lean_apply_1(x_1, x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Year_classify), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseYear(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseYear(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_OffsetX_classify___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_OffsetZ_classify___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_OffsetO_classify___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_string_length(x_1); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_3); +if (x_5 == 0) +{ +lean_object* x_6; uint32_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_string_utf8_get(x_1, x_6); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_9 = lean_string_push(x_8, x_7); +x_10 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2; +x_11 = lean_string_append(x_10, x_9); +lean_dec(x_9); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_13 = lean_string_append(x_11, x_12); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(23); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_2); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_classifyNumberText), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg(x_1, x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint32_t x_6; lean_object* x_7; +x_4 = lean_string_length(x_2); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_string_utf8_get(x_2, x_5); +x_7 = l_Std_Time_ZoneName_classify(x_6, x_4); +lean_dec(x_4); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_1); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_9 = lean_string_push(x_8, x_6); +x_10 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2; +x_11 = lean_string_append(x_10, x_9); +lean_dec(x_9); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_13 = lean_string_append(x_11, x_12); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +lean_dec(x_7); +x_16 = lean_apply_1(x_1, x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_3); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +lean_inc(x_1); +lean_inc(x_3); +x_4 = lean_apply_1(x_1, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_unbox_uint32(x_6); +lean_dec(x_6); +x_8 = lean_string_push(x_2, x_7); +x_2 = x_8; +x_3 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_4, 0); +x_12 = lean_ctor_get(x_4, 1); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +lean_dec(x_3); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_2); +return x_4; +} +else +{ +lean_dec(x_12); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_4, 0); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_4); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); +lean_dec(x_3); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +x_20 = lean_nat_dec_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_2); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_17); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_17); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_2); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +lean_inc(x_1); +x_3 = lean_apply_1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_7 = lean_unbox_uint32(x_5); +lean_dec(x_5); +x_8 = lean_string_push(x_6, x_7); +x_9 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__2(x_1, x_8, x_4); +return x_9; +} +else +{ +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +return x_3; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_apply_1(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_apply_2(x_1, x_6, x_5); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_4); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(0, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__5(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__7(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(6, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__8(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(7, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__9(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(8, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(9, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__11(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(10, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__12(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(11, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(12, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__14(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(13, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__15(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(14, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__16(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(15, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__17(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(16, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__18(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(17, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__19(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__20(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(19, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__21(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(20, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__22(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(21, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__23(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(22, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(24, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(25, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(26, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(27, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(28, 0, 1); +lean_ctor_set_uint8(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 71; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__2), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 121; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__3), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 117; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__4), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 68; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__5), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 77; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 76; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__6), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 100; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__7), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 81; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 113; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__8), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 119; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__9), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 87; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 69; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__11), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 101; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 99; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__12), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 70; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseText___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 97; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__14), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 104; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__15), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 75; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__16), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 107; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__17), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 72; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__18), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 109; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__19), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 115; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__20), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 83; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__21), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 65; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__22), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 110; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__23), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumber___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 78; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 86; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneId___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseZoneName___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 122; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 79; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 88; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 120; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 90; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_String_pchar___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3; +x_3 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3, 0, x_2); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2; +lean_inc(x_1); +x_5 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4, x_3, x_1); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +lean_dec(x_1); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +return x_5; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_5); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_5, 0); +x_12 = lean_ctor_get(x_5, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_5; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_free_object(x_5); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6; +x_17 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_17, 0, x_16); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5; +x_19 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_18, x_17, x_11); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_free_object(x_19); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9; +x_30 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_30, 0, x_29); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8; +x_32 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_31, x_30, x_25); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +lean_dec(x_27); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +return x_32; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_32); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_32); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_32, 0); +x_39 = lean_ctor_get(x_32, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +lean_dec(x_40); +return x_32; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_free_object(x_32); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12; +x_43 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_43, 0, x_42); +x_44 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11; +x_45 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_44, x_43, x_38); +if (lean_obj_tag(x_45) == 0) +{ +uint8_t x_46; +lean_dec(x_40); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +return x_45; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_45); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_45); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_45, 0); +x_52 = lean_ctor_get(x_45, 1); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +x_54 = lean_nat_dec_eq(x_40, x_53); +lean_dec(x_40); +if (x_54 == 0) +{ +lean_dec(x_53); +return x_45; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_free_object(x_45); +lean_dec(x_52); +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +x_56 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_56, 0, x_55); +x_57 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +x_58 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_57, x_56, x_51); +if (lean_obj_tag(x_58) == 0) +{ +uint8_t x_59; +lean_dec(x_53); +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) +{ +return x_58; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_58, 0); +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_58); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_58, 0); +x_65 = lean_ctor_get(x_58, 1); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +x_67 = lean_nat_dec_eq(x_53, x_66); +lean_dec(x_53); +if (x_67 == 0) +{ +lean_dec(x_66); +return x_58; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_free_object(x_58); +lean_dec(x_65); +x_68 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_69 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_69, 0, x_68); +x_70 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_57, x_69, x_64); +if (lean_obj_tag(x_70) == 0) +{ +uint8_t x_71; +lean_dec(x_66); +x_71 = !lean_is_exclusive(x_70); +if (x_71 == 0) +{ +return x_70; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_70, 0); +x_73 = lean_ctor_get(x_70, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_70); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +else +{ +uint8_t x_75; +x_75 = !lean_is_exclusive(x_70); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = lean_ctor_get(x_70, 0); +x_77 = lean_ctor_get(x_70, 1); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +x_79 = lean_nat_dec_eq(x_66, x_78); +lean_dec(x_66); +if (x_79 == 0) +{ +lean_dec(x_78); +return x_70; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_free_object(x_70); +lean_dec(x_77); +x_80 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_81 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_81, 0, x_80); +x_82 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_83 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_82, x_81, x_76); +if (lean_obj_tag(x_83) == 0) +{ +uint8_t x_84; +lean_dec(x_78); +x_84 = !lean_is_exclusive(x_83); +if (x_84 == 0) +{ +return x_83; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_83, 0); +x_86 = lean_ctor_get(x_83, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_83); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_83); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_83, 0); +x_90 = lean_ctor_get(x_83, 1); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +x_92 = lean_nat_dec_eq(x_78, x_91); +lean_dec(x_78); +if (x_92 == 0) +{ +lean_dec(x_91); +return x_83; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +lean_free_object(x_83); +lean_dec(x_90); +x_93 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_94 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_94, 0, x_93); +x_95 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_96 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_95, x_94, x_89); +if (lean_obj_tag(x_96) == 0) +{ +uint8_t x_97; +lean_dec(x_91); +x_97 = !lean_is_exclusive(x_96); +if (x_97 == 0) +{ +return x_96; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_96, 0); +x_99 = lean_ctor_get(x_96, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_96); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; +} +} +else +{ +uint8_t x_101; +x_101 = !lean_is_exclusive(x_96); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; +x_102 = lean_ctor_get(x_96, 0); +x_103 = lean_ctor_get(x_96, 1); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +x_105 = lean_nat_dec_eq(x_91, x_104); +lean_dec(x_91); +if (x_105 == 0) +{ +lean_dec(x_104); +return x_96; +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +lean_free_object(x_96); +lean_dec(x_103); +x_106 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_107 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_107, 0, x_106); +x_108 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_95, x_107, x_102); +if (lean_obj_tag(x_108) == 0) +{ +uint8_t x_109; +lean_dec(x_104); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) +{ +return x_108; +} +else +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_108, 0); +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_108); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; +} +} +else +{ +uint8_t x_113; +x_113 = !lean_is_exclusive(x_108); +if (x_113 == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_114 = lean_ctor_get(x_108, 0); +x_115 = lean_ctor_get(x_108, 1); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +x_117 = lean_nat_dec_eq(x_104, x_116); +lean_dec(x_104); +if (x_117 == 0) +{ +lean_dec(x_116); +return x_108; +} +else +{ +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_free_object(x_108); +lean_dec(x_115); +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_119 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_119, 0, x_118); +x_120 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_121 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_120, x_119, x_114); +if (lean_obj_tag(x_121) == 0) +{ +uint8_t x_122; +lean_dec(x_116); +x_122 = !lean_is_exclusive(x_121); +if (x_122 == 0) +{ +return x_121; +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_123 = lean_ctor_get(x_121, 0); +x_124 = lean_ctor_get(x_121, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_121); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +return x_125; +} +} +else +{ +uint8_t x_126; +x_126 = !lean_is_exclusive(x_121); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; +x_127 = lean_ctor_get(x_121, 0); +x_128 = lean_ctor_get(x_121, 1); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +x_130 = lean_nat_dec_eq(x_116, x_129); +lean_dec(x_116); +if (x_130 == 0) +{ +lean_dec(x_129); +return x_121; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_free_object(x_121); +lean_dec(x_128); +x_131 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_132 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_132, 0, x_131); +x_133 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_134 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_133, x_132, x_127); +if (lean_obj_tag(x_134) == 0) +{ +uint8_t x_135; +lean_dec(x_129); +x_135 = !lean_is_exclusive(x_134); +if (x_135 == 0) +{ +return x_134; +} +else +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_134, 0); +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_134); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); +return x_138; +} +} +else +{ +uint8_t x_139; +x_139 = !lean_is_exclusive(x_134); +if (x_139 == 0) +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +x_140 = lean_ctor_get(x_134, 0); +x_141 = lean_ctor_get(x_134, 1); +x_142 = lean_ctor_get(x_140, 1); +lean_inc(x_142); +x_143 = lean_nat_dec_eq(x_129, x_142); +lean_dec(x_129); +if (x_143 == 0) +{ +lean_dec(x_142); +return x_134; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +lean_free_object(x_134); +lean_dec(x_141); +x_144 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_145 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_145, 0, x_144); +x_146 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_147 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_146, x_145, x_140); +if (lean_obj_tag(x_147) == 0) +{ +uint8_t x_148; +lean_dec(x_142); +x_148 = !lean_is_exclusive(x_147); +if (x_148 == 0) +{ +return x_147; +} +else +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_149 = lean_ctor_get(x_147, 0); +x_150 = lean_ctor_get(x_147, 1); +lean_inc(x_150); +lean_inc(x_149); +lean_dec(x_147); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_149); +lean_ctor_set(x_151, 1, x_150); +return x_151; +} +} +else +{ +uint8_t x_152; +x_152 = !lean_is_exclusive(x_147); +if (x_152 == 0) +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; +x_153 = lean_ctor_get(x_147, 0); +x_154 = lean_ctor_get(x_147, 1); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +x_156 = lean_nat_dec_eq(x_142, x_155); +lean_dec(x_142); +if (x_156 == 0) +{ +lean_dec(x_155); +return x_147; +} +else +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +lean_free_object(x_147); +lean_dec(x_154); +x_157 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_158 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_158, 0, x_157); +x_159 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_160 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_159, x_158, x_153); +if (lean_obj_tag(x_160) == 0) +{ +uint8_t x_161; +lean_dec(x_155); +x_161 = !lean_is_exclusive(x_160); +if (x_161 == 0) +{ +return x_160; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_160, 0); +x_163 = lean_ctor_get(x_160, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_160); +x_164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +return x_164; +} +} +else +{ +uint8_t x_165; +x_165 = !lean_is_exclusive(x_160); +if (x_165 == 0) +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; +x_166 = lean_ctor_get(x_160, 0); +x_167 = lean_ctor_get(x_160, 1); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +x_169 = lean_nat_dec_eq(x_155, x_168); +lean_dec(x_155); +if (x_169 == 0) +{ +lean_dec(x_168); +return x_160; +} +else +{ +lean_object* x_170; lean_object* x_171; lean_object* x_172; +lean_free_object(x_160); +lean_dec(x_167); +x_170 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_171 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_171, 0, x_170); +x_172 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_159, x_171, x_166); +if (lean_obj_tag(x_172) == 0) +{ +uint8_t x_173; +lean_dec(x_168); +x_173 = !lean_is_exclusive(x_172); +if (x_173 == 0) +{ +return x_172; +} +else +{ +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_ctor_get(x_172, 0); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +lean_inc(x_174); +lean_dec(x_172); +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +return x_176; +} +} +else +{ +uint8_t x_177; +x_177 = !lean_is_exclusive(x_172); +if (x_177 == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; +x_178 = lean_ctor_get(x_172, 0); +x_179 = lean_ctor_get(x_172, 1); +x_180 = lean_ctor_get(x_178, 1); +lean_inc(x_180); +x_181 = lean_nat_dec_eq(x_168, x_180); +lean_dec(x_168); +if (x_181 == 0) +{ +lean_dec(x_180); +return x_172; +} +else +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_free_object(x_172); +lean_dec(x_179); +x_182 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_183 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_183, 0, x_182); +x_184 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_185 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_184, x_183, x_178); +if (lean_obj_tag(x_185) == 0) +{ +uint8_t x_186; +lean_dec(x_180); +x_186 = !lean_is_exclusive(x_185); +if (x_186 == 0) +{ +return x_185; +} +else +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_185, 0); +x_188 = lean_ctor_get(x_185, 1); +lean_inc(x_188); +lean_inc(x_187); +lean_dec(x_185); +x_189 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +return x_189; +} +} +else +{ +uint8_t x_190; +x_190 = !lean_is_exclusive(x_185); +if (x_190 == 0) +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; +x_191 = lean_ctor_get(x_185, 0); +x_192 = lean_ctor_get(x_185, 1); +x_193 = lean_ctor_get(x_191, 1); +lean_inc(x_193); +x_194 = lean_nat_dec_eq(x_180, x_193); +lean_dec(x_180); +if (x_194 == 0) +{ +lean_dec(x_193); +return x_185; +} +else +{ +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_free_object(x_185); +lean_dec(x_192); +x_195 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_196 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_196, 0, x_195); +x_197 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_198 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_197, x_196, x_191); +if (lean_obj_tag(x_198) == 0) +{ +uint8_t x_199; +lean_dec(x_193); +x_199 = !lean_is_exclusive(x_198); +if (x_199 == 0) +{ +return x_198; +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_200 = lean_ctor_get(x_198, 0); +x_201 = lean_ctor_get(x_198, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_198); +x_202 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +return x_202; +} +} +else +{ +uint8_t x_203; +x_203 = !lean_is_exclusive(x_198); +if (x_203 == 0) +{ +lean_object* x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; +x_204 = lean_ctor_get(x_198, 0); +x_205 = lean_ctor_get(x_198, 1); +x_206 = lean_ctor_get(x_204, 1); +lean_inc(x_206); +x_207 = lean_nat_dec_eq(x_193, x_206); +lean_dec(x_193); +if (x_207 == 0) +{ +lean_dec(x_206); +return x_198; +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; +lean_free_object(x_198); +lean_dec(x_205); +x_208 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_209 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_209, 0, x_208); +x_210 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_211 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_210, x_209, x_204); +if (lean_obj_tag(x_211) == 0) +{ +uint8_t x_212; +lean_dec(x_206); +x_212 = !lean_is_exclusive(x_211); +if (x_212 == 0) +{ +return x_211; +} +else +{ +lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_213 = lean_ctor_get(x_211, 0); +x_214 = lean_ctor_get(x_211, 1); +lean_inc(x_214); +lean_inc(x_213); +lean_dec(x_211); +x_215 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_215, 0, x_213); +lean_ctor_set(x_215, 1, x_214); +return x_215; +} +} +else +{ +uint8_t x_216; +x_216 = !lean_is_exclusive(x_211); +if (x_216 == 0) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; +x_217 = lean_ctor_get(x_211, 0); +x_218 = lean_ctor_get(x_211, 1); +x_219 = lean_ctor_get(x_217, 1); +lean_inc(x_219); +x_220 = lean_nat_dec_eq(x_206, x_219); +lean_dec(x_206); +if (x_220 == 0) +{ +lean_dec(x_219); +return x_211; +} +else +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_free_object(x_211); +lean_dec(x_218); +x_221 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_222 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_222, 0, x_221); +x_223 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_224 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_223, x_222, x_217); +if (lean_obj_tag(x_224) == 0) +{ +uint8_t x_225; +lean_dec(x_219); +x_225 = !lean_is_exclusive(x_224); +if (x_225 == 0) +{ +return x_224; +} +else +{ +lean_object* x_226; lean_object* x_227; lean_object* x_228; +x_226 = lean_ctor_get(x_224, 0); +x_227 = lean_ctor_get(x_224, 1); +lean_inc(x_227); +lean_inc(x_226); +lean_dec(x_224); +x_228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_228, 0, x_226); +lean_ctor_set(x_228, 1, x_227); +return x_228; +} +} +else +{ +uint8_t x_229; +x_229 = !lean_is_exclusive(x_224); +if (x_229 == 0) +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; uint8_t x_233; +x_230 = lean_ctor_get(x_224, 0); +x_231 = lean_ctor_get(x_224, 1); +x_232 = lean_ctor_get(x_230, 1); +lean_inc(x_232); +x_233 = lean_nat_dec_eq(x_219, x_232); +lean_dec(x_219); +if (x_233 == 0) +{ +lean_dec(x_232); +return x_224; +} +else +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; +lean_free_object(x_224); +lean_dec(x_231); +x_234 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_235 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_235, 0, x_234); +x_236 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_237 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_236, x_235, x_230); +if (lean_obj_tag(x_237) == 0) +{ +uint8_t x_238; +lean_dec(x_232); +x_238 = !lean_is_exclusive(x_237); +if (x_238 == 0) +{ +return x_237; +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_239 = lean_ctor_get(x_237, 0); +x_240 = lean_ctor_get(x_237, 1); +lean_inc(x_240); +lean_inc(x_239); +lean_dec(x_237); +x_241 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_241, 0, x_239); +lean_ctor_set(x_241, 1, x_240); +return x_241; +} +} +else +{ +uint8_t x_242; +x_242 = !lean_is_exclusive(x_237); +if (x_242 == 0) +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; +x_243 = lean_ctor_get(x_237, 0); +x_244 = lean_ctor_get(x_237, 1); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +x_246 = lean_nat_dec_eq(x_232, x_245); +lean_dec(x_232); +if (x_246 == 0) +{ +lean_dec(x_245); +return x_237; +} +else +{ +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_free_object(x_237); +lean_dec(x_244); +x_247 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_248 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_248, 0, x_247); +x_249 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_250 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_249, x_248, x_243); +if (lean_obj_tag(x_250) == 0) +{ +uint8_t x_251; +lean_dec(x_245); +x_251 = !lean_is_exclusive(x_250); +if (x_251 == 0) +{ +return x_250; +} +else +{ +lean_object* x_252; lean_object* x_253; lean_object* x_254; +x_252 = lean_ctor_get(x_250, 0); +x_253 = lean_ctor_get(x_250, 1); +lean_inc(x_253); +lean_inc(x_252); +lean_dec(x_250); +x_254 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_254, 0, x_252); +lean_ctor_set(x_254, 1, x_253); +return x_254; +} +} +else +{ +uint8_t x_255; +x_255 = !lean_is_exclusive(x_250); +if (x_255 == 0) +{ +lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_256 = lean_ctor_get(x_250, 0); +x_257 = lean_ctor_get(x_250, 1); +x_258 = lean_ctor_get(x_256, 1); +lean_inc(x_258); +x_259 = lean_nat_dec_eq(x_245, x_258); +lean_dec(x_245); +if (x_259 == 0) +{ +lean_dec(x_258); +return x_250; +} +else +{ +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_free_object(x_250); +lean_dec(x_257); +x_260 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_261 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_261, 0, x_260); +x_262 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_263 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_262, x_261, x_256); +if (lean_obj_tag(x_263) == 0) +{ +uint8_t x_264; +lean_dec(x_258); +x_264 = !lean_is_exclusive(x_263); +if (x_264 == 0) +{ +return x_263; +} +else +{ +lean_object* x_265; lean_object* x_266; lean_object* x_267; +x_265 = lean_ctor_get(x_263, 0); +x_266 = lean_ctor_get(x_263, 1); +lean_inc(x_266); +lean_inc(x_265); +lean_dec(x_263); +x_267 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_267, 0, x_265); +lean_ctor_set(x_267, 1, x_266); +return x_267; +} +} +else +{ +uint8_t x_268; +x_268 = !lean_is_exclusive(x_263); +if (x_268 == 0) +{ +lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t x_272; +x_269 = lean_ctor_get(x_263, 0); +x_270 = lean_ctor_get(x_263, 1); +x_271 = lean_ctor_get(x_269, 1); +lean_inc(x_271); +x_272 = lean_nat_dec_eq(x_258, x_271); +lean_dec(x_258); +if (x_272 == 0) +{ +lean_dec(x_271); +return x_263; +} +else +{ +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +lean_free_object(x_263); +lean_dec(x_270); +x_273 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_274 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_274, 0, x_273); +x_275 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_276 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_275, x_274, x_269); +if (lean_obj_tag(x_276) == 0) +{ +uint8_t x_277; +lean_dec(x_271); +x_277 = !lean_is_exclusive(x_276); +if (x_277 == 0) +{ +return x_276; +} +else +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; +x_278 = lean_ctor_get(x_276, 0); +x_279 = lean_ctor_get(x_276, 1); +lean_inc(x_279); +lean_inc(x_278); +lean_dec(x_276); +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_278); +lean_ctor_set(x_280, 1, x_279); +return x_280; +} +} +else +{ +uint8_t x_281; +x_281 = !lean_is_exclusive(x_276); +if (x_281 == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; uint8_t x_285; +x_282 = lean_ctor_get(x_276, 0); +x_283 = lean_ctor_get(x_276, 1); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +x_285 = lean_nat_dec_eq(x_271, x_284); +lean_dec(x_271); +if (x_285 == 0) +{ +lean_dec(x_284); +return x_276; +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +lean_free_object(x_276); +lean_dec(x_283); +x_286 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_287 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_287, 0, x_286); +x_288 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_289 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_288, x_287, x_282); +if (lean_obj_tag(x_289) == 0) +{ +uint8_t x_290; +lean_dec(x_284); +x_290 = !lean_is_exclusive(x_289); +if (x_290 == 0) +{ +return x_289; +} +else +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_291 = lean_ctor_get(x_289, 0); +x_292 = lean_ctor_get(x_289, 1); +lean_inc(x_292); +lean_inc(x_291); +lean_dec(x_289); +x_293 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_293, 0, x_291); +lean_ctor_set(x_293, 1, x_292); +return x_293; +} +} +else +{ +uint8_t x_294; +x_294 = !lean_is_exclusive(x_289); +if (x_294 == 0) +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; uint8_t x_298; +x_295 = lean_ctor_get(x_289, 0); +x_296 = lean_ctor_get(x_289, 1); +x_297 = lean_ctor_get(x_295, 1); +lean_inc(x_297); +x_298 = lean_nat_dec_eq(x_284, x_297); +lean_dec(x_284); +if (x_298 == 0) +{ +lean_dec(x_297); +return x_289; +} +else +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; +lean_free_object(x_289); +lean_dec(x_296); +x_299 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_300 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_300, 0, x_299); +x_301 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_302 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_301, x_300, x_295); +if (lean_obj_tag(x_302) == 0) +{ +uint8_t x_303; +lean_dec(x_297); +x_303 = !lean_is_exclusive(x_302); +if (x_303 == 0) +{ +return x_302; +} +else +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; +x_304 = lean_ctor_get(x_302, 0); +x_305 = lean_ctor_get(x_302, 1); +lean_inc(x_305); +lean_inc(x_304); +lean_dec(x_302); +x_306 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_306, 0, x_304); +lean_ctor_set(x_306, 1, x_305); +return x_306; +} +} +else +{ +uint8_t x_307; +x_307 = !lean_is_exclusive(x_302); +if (x_307 == 0) +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; uint8_t x_311; +x_308 = lean_ctor_get(x_302, 0); +x_309 = lean_ctor_get(x_302, 1); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +x_311 = lean_nat_dec_eq(x_297, x_310); +lean_dec(x_297); +if (x_311 == 0) +{ +lean_dec(x_310); +return x_302; +} +else +{ +lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; +lean_free_object(x_302); +lean_dec(x_309); +x_312 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_313 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_313, 0, x_312); +x_314 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_315 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_314, x_313, x_308); +if (lean_obj_tag(x_315) == 0) +{ +uint8_t x_316; +lean_dec(x_310); +x_316 = !lean_is_exclusive(x_315); +if (x_316 == 0) +{ +return x_315; +} +else +{ +lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_317 = lean_ctor_get(x_315, 0); +x_318 = lean_ctor_get(x_315, 1); +lean_inc(x_318); +lean_inc(x_317); +lean_dec(x_315); +x_319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_319, 0, x_317); +lean_ctor_set(x_319, 1, x_318); +return x_319; +} +} +else +{ +uint8_t x_320; +x_320 = !lean_is_exclusive(x_315); +if (x_320 == 0) +{ +lean_object* x_321; lean_object* x_322; lean_object* x_323; uint8_t x_324; +x_321 = lean_ctor_get(x_315, 0); +x_322 = lean_ctor_get(x_315, 1); +x_323 = lean_ctor_get(x_321, 1); +lean_inc(x_323); +x_324 = lean_nat_dec_eq(x_310, x_323); +lean_dec(x_310); +if (x_324 == 0) +{ +lean_dec(x_323); +return x_315; +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_free_object(x_315); +lean_dec(x_322); +x_325 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_326 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_326, 0, x_325); +x_327 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_328 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_327, x_326, x_321); +if (lean_obj_tag(x_328) == 0) +{ +uint8_t x_329; +lean_dec(x_323); +x_329 = !lean_is_exclusive(x_328); +if (x_329 == 0) +{ +return x_328; +} +else +{ +lean_object* x_330; lean_object* x_331; lean_object* x_332; +x_330 = lean_ctor_get(x_328, 0); +x_331 = lean_ctor_get(x_328, 1); +lean_inc(x_331); +lean_inc(x_330); +lean_dec(x_328); +x_332 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_332, 0, x_330); +lean_ctor_set(x_332, 1, x_331); +return x_332; +} +} +else +{ +uint8_t x_333; +x_333 = !lean_is_exclusive(x_328); +if (x_333 == 0) +{ +lean_object* x_334; lean_object* x_335; lean_object* x_336; uint8_t x_337; +x_334 = lean_ctor_get(x_328, 0); +x_335 = lean_ctor_get(x_328, 1); +x_336 = lean_ctor_get(x_334, 1); +lean_inc(x_336); +x_337 = lean_nat_dec_eq(x_323, x_336); +lean_dec(x_323); +if (x_337 == 0) +{ +lean_dec(x_336); +return x_328; +} +else +{ +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_free_object(x_328); +lean_dec(x_335); +x_338 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_339 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_339, 0, x_338); +x_340 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_341 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_340, x_339, x_334); +if (lean_obj_tag(x_341) == 0) +{ +uint8_t x_342; +lean_dec(x_336); +x_342 = !lean_is_exclusive(x_341); +if (x_342 == 0) +{ +return x_341; +} +else +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; +x_343 = lean_ctor_get(x_341, 0); +x_344 = lean_ctor_get(x_341, 1); +lean_inc(x_344); +lean_inc(x_343); +lean_dec(x_341); +x_345 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_345, 0, x_343); +lean_ctor_set(x_345, 1, x_344); +return x_345; +} +} +else +{ +uint8_t x_346; +x_346 = !lean_is_exclusive(x_341); +if (x_346 == 0) +{ +lean_object* x_347; lean_object* x_348; lean_object* x_349; uint8_t x_350; +x_347 = lean_ctor_get(x_341, 0); +x_348 = lean_ctor_get(x_341, 1); +x_349 = lean_ctor_get(x_347, 1); +lean_inc(x_349); +x_350 = lean_nat_dec_eq(x_336, x_349); +lean_dec(x_336); +if (x_350 == 0) +{ +lean_dec(x_349); +return x_341; +} +else +{ +lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +lean_free_object(x_341); +lean_dec(x_348); +x_351 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_352 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_352, 0, x_351); +x_353 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_354 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_353, x_352, x_347); +if (lean_obj_tag(x_354) == 0) +{ +uint8_t x_355; +lean_dec(x_349); +x_355 = !lean_is_exclusive(x_354); +if (x_355 == 0) +{ +return x_354; +} +else +{ +lean_object* x_356; lean_object* x_357; lean_object* x_358; +x_356 = lean_ctor_get(x_354, 0); +x_357 = lean_ctor_get(x_354, 1); +lean_inc(x_357); +lean_inc(x_356); +lean_dec(x_354); +x_358 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_358, 0, x_356); +lean_ctor_set(x_358, 1, x_357); +return x_358; +} +} +else +{ +uint8_t x_359; +x_359 = !lean_is_exclusive(x_354); +if (x_359 == 0) +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; uint8_t x_363; +x_360 = lean_ctor_get(x_354, 0); +x_361 = lean_ctor_get(x_354, 1); +x_362 = lean_ctor_get(x_360, 1); +lean_inc(x_362); +x_363 = lean_nat_dec_eq(x_349, x_362); +lean_dec(x_349); +if (x_363 == 0) +{ +lean_dec(x_362); +return x_354; +} +else +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; +lean_free_object(x_354); +lean_dec(x_361); +x_364 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_365 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_365, 0, x_364); +x_366 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_367 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_366, x_365, x_360); +if (lean_obj_tag(x_367) == 0) +{ +uint8_t x_368; +lean_dec(x_362); +x_368 = !lean_is_exclusive(x_367); +if (x_368 == 0) +{ +return x_367; +} +else +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; +x_369 = lean_ctor_get(x_367, 0); +x_370 = lean_ctor_get(x_367, 1); +lean_inc(x_370); +lean_inc(x_369); +lean_dec(x_367); +x_371 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_371, 0, x_369); +lean_ctor_set(x_371, 1, x_370); +return x_371; +} +} +else +{ +uint8_t x_372; +x_372 = !lean_is_exclusive(x_367); +if (x_372 == 0) +{ +lean_object* x_373; lean_object* x_374; lean_object* x_375; uint8_t x_376; +x_373 = lean_ctor_get(x_367, 0); +x_374 = lean_ctor_get(x_367, 1); +x_375 = lean_ctor_get(x_373, 1); +lean_inc(x_375); +x_376 = lean_nat_dec_eq(x_362, x_375); +lean_dec(x_362); +if (x_376 == 0) +{ +lean_dec(x_375); +return x_367; +} +else +{ +lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; +lean_free_object(x_367); +lean_dec(x_374); +x_377 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_378 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_378, 0, x_377); +x_379 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_380 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_379, x_378, x_373); +if (lean_obj_tag(x_380) == 0) +{ +uint8_t x_381; +lean_dec(x_375); +x_381 = !lean_is_exclusive(x_380); +if (x_381 == 0) +{ +return x_380; +} +else +{ +lean_object* x_382; lean_object* x_383; lean_object* x_384; +x_382 = lean_ctor_get(x_380, 0); +x_383 = lean_ctor_get(x_380, 1); +lean_inc(x_383); +lean_inc(x_382); +lean_dec(x_380); +x_384 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_384, 0, x_382); +lean_ctor_set(x_384, 1, x_383); +return x_384; +} +} +else +{ +uint8_t x_385; +x_385 = !lean_is_exclusive(x_380); +if (x_385 == 0) +{ +lean_object* x_386; lean_object* x_387; lean_object* x_388; uint8_t x_389; +x_386 = lean_ctor_get(x_380, 0); +x_387 = lean_ctor_get(x_380, 1); +x_388 = lean_ctor_get(x_386, 1); +lean_inc(x_388); +x_389 = lean_nat_dec_eq(x_375, x_388); +lean_dec(x_375); +if (x_389 == 0) +{ +lean_dec(x_388); +return x_380; +} +else +{ +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; +lean_free_object(x_380); +lean_dec(x_387); +x_390 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_391 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_391, 0, x_390); +x_392 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_393 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_392, x_391, x_386); +if (lean_obj_tag(x_393) == 0) +{ +uint8_t x_394; +lean_dec(x_388); +x_394 = !lean_is_exclusive(x_393); +if (x_394 == 0) +{ +return x_393; +} +else +{ +lean_object* x_395; lean_object* x_396; lean_object* x_397; +x_395 = lean_ctor_get(x_393, 0); +x_396 = lean_ctor_get(x_393, 1); +lean_inc(x_396); +lean_inc(x_395); +lean_dec(x_393); +x_397 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_397, 0, x_395); +lean_ctor_set(x_397, 1, x_396); +return x_397; +} +} +else +{ +uint8_t x_398; +x_398 = !lean_is_exclusive(x_393); +if (x_398 == 0) +{ +lean_object* x_399; lean_object* x_400; lean_object* x_401; uint8_t x_402; +x_399 = lean_ctor_get(x_393, 0); +x_400 = lean_ctor_get(x_393, 1); +x_401 = lean_ctor_get(x_399, 1); +lean_inc(x_401); +x_402 = lean_nat_dec_eq(x_388, x_401); +lean_dec(x_401); +lean_dec(x_388); +if (x_402 == 0) +{ +return x_393; +} +else +{ +lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; +lean_free_object(x_393); +lean_dec(x_400); +x_403 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_404 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_404, 0, x_403); +x_405 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_406 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_405, x_404, x_399); +return x_406; +} +} +else +{ +lean_object* x_407; lean_object* x_408; lean_object* x_409; uint8_t x_410; +x_407 = lean_ctor_get(x_393, 0); +x_408 = lean_ctor_get(x_393, 1); +lean_inc(x_408); +lean_inc(x_407); +lean_dec(x_393); +x_409 = lean_ctor_get(x_407, 1); +lean_inc(x_409); +x_410 = lean_nat_dec_eq(x_388, x_409); +lean_dec(x_409); +lean_dec(x_388); +if (x_410 == 0) +{ +lean_object* x_411; +x_411 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_411, 0, x_407); +lean_ctor_set(x_411, 1, x_408); +return x_411; +} +else +{ +lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; +lean_dec(x_408); +x_412 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_413 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_413, 0, x_412); +x_414 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_415 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_414, x_413, x_407); +return x_415; +} +} +} +} +} +else +{ +lean_object* x_416; lean_object* x_417; lean_object* x_418; uint8_t x_419; +x_416 = lean_ctor_get(x_380, 0); +x_417 = lean_ctor_get(x_380, 1); +lean_inc(x_417); +lean_inc(x_416); +lean_dec(x_380); +x_418 = lean_ctor_get(x_416, 1); +lean_inc(x_418); +x_419 = lean_nat_dec_eq(x_375, x_418); +lean_dec(x_375); +if (x_419 == 0) +{ +lean_object* x_420; +lean_dec(x_418); +x_420 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_420, 0, x_416); +lean_ctor_set(x_420, 1, x_417); +return x_420; +} +else +{ +lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +lean_dec(x_417); +x_421 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_422 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_422, 0, x_421); +x_423 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_424 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_423, x_422, x_416); +if (lean_obj_tag(x_424) == 0) +{ +lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; +lean_dec(x_418); +x_425 = lean_ctor_get(x_424, 0); +lean_inc(x_425); +x_426 = lean_ctor_get(x_424, 1); +lean_inc(x_426); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + x_427 = x_424; +} else { + lean_dec_ref(x_424); + x_427 = lean_box(0); +} +if (lean_is_scalar(x_427)) { + x_428 = lean_alloc_ctor(0, 2, 0); +} else { + x_428 = x_427; +} +lean_ctor_set(x_428, 0, x_425); +lean_ctor_set(x_428, 1, x_426); +return x_428; +} +else +{ +lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; uint8_t x_433; +x_429 = lean_ctor_get(x_424, 0); +lean_inc(x_429); +x_430 = lean_ctor_get(x_424, 1); +lean_inc(x_430); +if (lean_is_exclusive(x_424)) { + lean_ctor_release(x_424, 0); + lean_ctor_release(x_424, 1); + x_431 = x_424; +} else { + lean_dec_ref(x_424); + x_431 = lean_box(0); +} +x_432 = lean_ctor_get(x_429, 1); +lean_inc(x_432); +x_433 = lean_nat_dec_eq(x_418, x_432); +lean_dec(x_432); +lean_dec(x_418); +if (x_433 == 0) +{ +lean_object* x_434; +if (lean_is_scalar(x_431)) { + x_434 = lean_alloc_ctor(1, 2, 0); +} else { + x_434 = x_431; +} +lean_ctor_set(x_434, 0, x_429); +lean_ctor_set(x_434, 1, x_430); +return x_434; +} +else +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; +lean_dec(x_431); +lean_dec(x_430); +x_435 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_436 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_436, 0, x_435); +x_437 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_438 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_437, x_436, x_429); +return x_438; +} +} +} +} +} +} +} +else +{ +lean_object* x_439; lean_object* x_440; lean_object* x_441; uint8_t x_442; +x_439 = lean_ctor_get(x_367, 0); +x_440 = lean_ctor_get(x_367, 1); +lean_inc(x_440); +lean_inc(x_439); +lean_dec(x_367); +x_441 = lean_ctor_get(x_439, 1); +lean_inc(x_441); +x_442 = lean_nat_dec_eq(x_362, x_441); +lean_dec(x_362); +if (x_442 == 0) +{ +lean_object* x_443; +lean_dec(x_441); +x_443 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_443, 0, x_439); +lean_ctor_set(x_443, 1, x_440); +return x_443; +} +else +{ +lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; +lean_dec(x_440); +x_444 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_445 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_445, 0, x_444); +x_446 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_447 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_446, x_445, x_439); +if (lean_obj_tag(x_447) == 0) +{ +lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; +lean_dec(x_441); +x_448 = lean_ctor_get(x_447, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_447, 1); +lean_inc(x_449); +if (lean_is_exclusive(x_447)) { + lean_ctor_release(x_447, 0); + lean_ctor_release(x_447, 1); + x_450 = x_447; +} else { + lean_dec_ref(x_447); + x_450 = lean_box(0); +} +if (lean_is_scalar(x_450)) { + x_451 = lean_alloc_ctor(0, 2, 0); +} else { + x_451 = x_450; +} +lean_ctor_set(x_451, 0, x_448); +lean_ctor_set(x_451, 1, x_449); +return x_451; +} +else +{ +lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; uint8_t x_456; +x_452 = lean_ctor_get(x_447, 0); +lean_inc(x_452); +x_453 = lean_ctor_get(x_447, 1); +lean_inc(x_453); +if (lean_is_exclusive(x_447)) { + lean_ctor_release(x_447, 0); + lean_ctor_release(x_447, 1); + x_454 = x_447; +} else { + lean_dec_ref(x_447); + x_454 = lean_box(0); +} +x_455 = lean_ctor_get(x_452, 1); +lean_inc(x_455); +x_456 = lean_nat_dec_eq(x_441, x_455); +lean_dec(x_441); +if (x_456 == 0) +{ +lean_object* x_457; +lean_dec(x_455); +if (lean_is_scalar(x_454)) { + x_457 = lean_alloc_ctor(1, 2, 0); +} else { + x_457 = x_454; +} +lean_ctor_set(x_457, 0, x_452); +lean_ctor_set(x_457, 1, x_453); +return x_457; +} +else +{ +lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; +lean_dec(x_454); +lean_dec(x_453); +x_458 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_459 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_459, 0, x_458); +x_460 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_461 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_460, x_459, x_452); +if (lean_obj_tag(x_461) == 0) +{ +lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; +lean_dec(x_455); +x_462 = lean_ctor_get(x_461, 0); +lean_inc(x_462); +x_463 = lean_ctor_get(x_461, 1); +lean_inc(x_463); +if (lean_is_exclusive(x_461)) { + lean_ctor_release(x_461, 0); + lean_ctor_release(x_461, 1); + x_464 = x_461; +} else { + lean_dec_ref(x_461); + x_464 = lean_box(0); +} +if (lean_is_scalar(x_464)) { + x_465 = lean_alloc_ctor(0, 2, 0); +} else { + x_465 = x_464; +} +lean_ctor_set(x_465, 0, x_462); +lean_ctor_set(x_465, 1, x_463); +return x_465; +} +else +{ +lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; uint8_t x_470; +x_466 = lean_ctor_get(x_461, 0); +lean_inc(x_466); +x_467 = lean_ctor_get(x_461, 1); +lean_inc(x_467); +if (lean_is_exclusive(x_461)) { + lean_ctor_release(x_461, 0); + lean_ctor_release(x_461, 1); + x_468 = x_461; +} else { + lean_dec_ref(x_461); + x_468 = lean_box(0); +} +x_469 = lean_ctor_get(x_466, 1); +lean_inc(x_469); +x_470 = lean_nat_dec_eq(x_455, x_469); +lean_dec(x_469); +lean_dec(x_455); +if (x_470 == 0) +{ +lean_object* x_471; +if (lean_is_scalar(x_468)) { + x_471 = lean_alloc_ctor(1, 2, 0); +} else { + x_471 = x_468; +} +lean_ctor_set(x_471, 0, x_466); +lean_ctor_set(x_471, 1, x_467); +return x_471; +} +else +{ +lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; +lean_dec(x_468); +lean_dec(x_467); +x_472 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_473 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_473, 0, x_472); +x_474 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_475 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_474, x_473, x_466); +return x_475; +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; +x_476 = lean_ctor_get(x_354, 0); +x_477 = lean_ctor_get(x_354, 1); +lean_inc(x_477); +lean_inc(x_476); +lean_dec(x_354); +x_478 = lean_ctor_get(x_476, 1); +lean_inc(x_478); +x_479 = lean_nat_dec_eq(x_349, x_478); +lean_dec(x_349); +if (x_479 == 0) +{ +lean_object* x_480; +lean_dec(x_478); +x_480 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_480, 0, x_476); +lean_ctor_set(x_480, 1, x_477); +return x_480; +} +else +{ +lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; +lean_dec(x_477); +x_481 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_482 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_482, 0, x_481); +x_483 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_484 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_483, x_482, x_476); +if (lean_obj_tag(x_484) == 0) +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; +lean_dec(x_478); +x_485 = lean_ctor_get(x_484, 0); +lean_inc(x_485); +x_486 = lean_ctor_get(x_484, 1); +lean_inc(x_486); +if (lean_is_exclusive(x_484)) { + lean_ctor_release(x_484, 0); + lean_ctor_release(x_484, 1); + x_487 = x_484; +} else { + lean_dec_ref(x_484); + x_487 = lean_box(0); +} +if (lean_is_scalar(x_487)) { + x_488 = lean_alloc_ctor(0, 2, 0); +} else { + x_488 = x_487; +} +lean_ctor_set(x_488, 0, x_485); +lean_ctor_set(x_488, 1, x_486); +return x_488; +} +else +{ +lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; uint8_t x_493; +x_489 = lean_ctor_get(x_484, 0); +lean_inc(x_489); +x_490 = lean_ctor_get(x_484, 1); +lean_inc(x_490); +if (lean_is_exclusive(x_484)) { + lean_ctor_release(x_484, 0); + lean_ctor_release(x_484, 1); + x_491 = x_484; +} else { + lean_dec_ref(x_484); + x_491 = lean_box(0); +} +x_492 = lean_ctor_get(x_489, 1); +lean_inc(x_492); +x_493 = lean_nat_dec_eq(x_478, x_492); +lean_dec(x_478); +if (x_493 == 0) +{ +lean_object* x_494; +lean_dec(x_492); +if (lean_is_scalar(x_491)) { + x_494 = lean_alloc_ctor(1, 2, 0); +} else { + x_494 = x_491; +} +lean_ctor_set(x_494, 0, x_489); +lean_ctor_set(x_494, 1, x_490); +return x_494; +} +else +{ +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; +lean_dec(x_491); +lean_dec(x_490); +x_495 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_496 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_496, 0, x_495); +x_497 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_498 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_497, x_496, x_489); +if (lean_obj_tag(x_498) == 0) +{ +lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; +lean_dec(x_492); +x_499 = lean_ctor_get(x_498, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_498, 1); +lean_inc(x_500); +if (lean_is_exclusive(x_498)) { + lean_ctor_release(x_498, 0); + lean_ctor_release(x_498, 1); + x_501 = x_498; +} else { + lean_dec_ref(x_498); + x_501 = lean_box(0); +} +if (lean_is_scalar(x_501)) { + x_502 = lean_alloc_ctor(0, 2, 0); +} else { + x_502 = x_501; +} +lean_ctor_set(x_502, 0, x_499); +lean_ctor_set(x_502, 1, x_500); +return x_502; +} +else +{ +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; uint8_t x_507; +x_503 = lean_ctor_get(x_498, 0); +lean_inc(x_503); +x_504 = lean_ctor_get(x_498, 1); +lean_inc(x_504); +if (lean_is_exclusive(x_498)) { + lean_ctor_release(x_498, 0); + lean_ctor_release(x_498, 1); + x_505 = x_498; +} else { + lean_dec_ref(x_498); + x_505 = lean_box(0); +} +x_506 = lean_ctor_get(x_503, 1); +lean_inc(x_506); +x_507 = lean_nat_dec_eq(x_492, x_506); +lean_dec(x_492); +if (x_507 == 0) +{ +lean_object* x_508; +lean_dec(x_506); +if (lean_is_scalar(x_505)) { + x_508 = lean_alloc_ctor(1, 2, 0); +} else { + x_508 = x_505; +} +lean_ctor_set(x_508, 0, x_503); +lean_ctor_set(x_508, 1, x_504); +return x_508; +} +else +{ +lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; +lean_dec(x_505); +lean_dec(x_504); +x_509 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_510 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_510, 0, x_509); +x_511 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_512 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_511, x_510, x_503); +if (lean_obj_tag(x_512) == 0) +{ +lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; +lean_dec(x_506); +x_513 = lean_ctor_get(x_512, 0); +lean_inc(x_513); +x_514 = lean_ctor_get(x_512, 1); +lean_inc(x_514); +if (lean_is_exclusive(x_512)) { + lean_ctor_release(x_512, 0); + lean_ctor_release(x_512, 1); + x_515 = x_512; +} else { + lean_dec_ref(x_512); + x_515 = lean_box(0); +} +if (lean_is_scalar(x_515)) { + x_516 = lean_alloc_ctor(0, 2, 0); +} else { + x_516 = x_515; +} +lean_ctor_set(x_516, 0, x_513); +lean_ctor_set(x_516, 1, x_514); +return x_516; +} +else +{ +lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; uint8_t x_521; +x_517 = lean_ctor_get(x_512, 0); +lean_inc(x_517); +x_518 = lean_ctor_get(x_512, 1); +lean_inc(x_518); +if (lean_is_exclusive(x_512)) { + lean_ctor_release(x_512, 0); + lean_ctor_release(x_512, 1); + x_519 = x_512; +} else { + lean_dec_ref(x_512); + x_519 = lean_box(0); +} +x_520 = lean_ctor_get(x_517, 1); +lean_inc(x_520); +x_521 = lean_nat_dec_eq(x_506, x_520); +lean_dec(x_520); +lean_dec(x_506); +if (x_521 == 0) +{ +lean_object* x_522; +if (lean_is_scalar(x_519)) { + x_522 = lean_alloc_ctor(1, 2, 0); +} else { + x_522 = x_519; +} +lean_ctor_set(x_522, 0, x_517); +lean_ctor_set(x_522, 1, x_518); +return x_522; +} +else +{ +lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; +lean_dec(x_519); +lean_dec(x_518); +x_523 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_524 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_524, 0, x_523); +x_525 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_526 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_525, x_524, x_517); +return x_526; +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_527; lean_object* x_528; lean_object* x_529; uint8_t x_530; +x_527 = lean_ctor_get(x_341, 0); +x_528 = lean_ctor_get(x_341, 1); +lean_inc(x_528); +lean_inc(x_527); +lean_dec(x_341); +x_529 = lean_ctor_get(x_527, 1); +lean_inc(x_529); +x_530 = lean_nat_dec_eq(x_336, x_529); +lean_dec(x_336); +if (x_530 == 0) +{ +lean_object* x_531; +lean_dec(x_529); +x_531 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_531, 0, x_527); +lean_ctor_set(x_531, 1, x_528); +return x_531; +} +else +{ +lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; +lean_dec(x_528); +x_532 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_533 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_533, 0, x_532); +x_534 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_535 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_534, x_533, x_527); +if (lean_obj_tag(x_535) == 0) +{ +lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; +lean_dec(x_529); +x_536 = lean_ctor_get(x_535, 0); +lean_inc(x_536); +x_537 = lean_ctor_get(x_535, 1); +lean_inc(x_537); +if (lean_is_exclusive(x_535)) { + lean_ctor_release(x_535, 0); + lean_ctor_release(x_535, 1); + x_538 = x_535; +} else { + lean_dec_ref(x_535); + x_538 = lean_box(0); +} +if (lean_is_scalar(x_538)) { + x_539 = lean_alloc_ctor(0, 2, 0); +} else { + x_539 = x_538; +} +lean_ctor_set(x_539, 0, x_536); +lean_ctor_set(x_539, 1, x_537); +return x_539; +} +else +{ +lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; uint8_t x_544; +x_540 = lean_ctor_get(x_535, 0); +lean_inc(x_540); +x_541 = lean_ctor_get(x_535, 1); +lean_inc(x_541); +if (lean_is_exclusive(x_535)) { + lean_ctor_release(x_535, 0); + lean_ctor_release(x_535, 1); + x_542 = x_535; +} else { + lean_dec_ref(x_535); + x_542 = lean_box(0); +} +x_543 = lean_ctor_get(x_540, 1); +lean_inc(x_543); +x_544 = lean_nat_dec_eq(x_529, x_543); +lean_dec(x_529); +if (x_544 == 0) +{ +lean_object* x_545; +lean_dec(x_543); +if (lean_is_scalar(x_542)) { + x_545 = lean_alloc_ctor(1, 2, 0); +} else { + x_545 = x_542; +} +lean_ctor_set(x_545, 0, x_540); +lean_ctor_set(x_545, 1, x_541); +return x_545; +} +else +{ +lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; +lean_dec(x_542); +lean_dec(x_541); +x_546 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_547 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_547, 0, x_546); +x_548 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_549 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_548, x_547, x_540); +if (lean_obj_tag(x_549) == 0) +{ +lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; +lean_dec(x_543); +x_550 = lean_ctor_get(x_549, 0); +lean_inc(x_550); +x_551 = lean_ctor_get(x_549, 1); +lean_inc(x_551); +if (lean_is_exclusive(x_549)) { + lean_ctor_release(x_549, 0); + lean_ctor_release(x_549, 1); + x_552 = x_549; +} else { + lean_dec_ref(x_549); + x_552 = lean_box(0); +} +if (lean_is_scalar(x_552)) { + x_553 = lean_alloc_ctor(0, 2, 0); +} else { + x_553 = x_552; +} +lean_ctor_set(x_553, 0, x_550); +lean_ctor_set(x_553, 1, x_551); +return x_553; +} +else +{ +lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; uint8_t x_558; +x_554 = lean_ctor_get(x_549, 0); +lean_inc(x_554); +x_555 = lean_ctor_get(x_549, 1); +lean_inc(x_555); +if (lean_is_exclusive(x_549)) { + lean_ctor_release(x_549, 0); + lean_ctor_release(x_549, 1); + x_556 = x_549; +} else { + lean_dec_ref(x_549); + x_556 = lean_box(0); +} +x_557 = lean_ctor_get(x_554, 1); +lean_inc(x_557); +x_558 = lean_nat_dec_eq(x_543, x_557); +lean_dec(x_543); +if (x_558 == 0) +{ +lean_object* x_559; +lean_dec(x_557); +if (lean_is_scalar(x_556)) { + x_559 = lean_alloc_ctor(1, 2, 0); +} else { + x_559 = x_556; +} +lean_ctor_set(x_559, 0, x_554); +lean_ctor_set(x_559, 1, x_555); +return x_559; +} +else +{ +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; +lean_dec(x_556); +lean_dec(x_555); +x_560 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_561 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_561, 0, x_560); +x_562 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_563 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_562, x_561, x_554); +if (lean_obj_tag(x_563) == 0) +{ +lean_object* x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; +lean_dec(x_557); +x_564 = lean_ctor_get(x_563, 0); +lean_inc(x_564); +x_565 = lean_ctor_get(x_563, 1); +lean_inc(x_565); +if (lean_is_exclusive(x_563)) { + lean_ctor_release(x_563, 0); + lean_ctor_release(x_563, 1); + x_566 = x_563; +} else { + lean_dec_ref(x_563); + x_566 = lean_box(0); +} +if (lean_is_scalar(x_566)) { + x_567 = lean_alloc_ctor(0, 2, 0); +} else { + x_567 = x_566; +} +lean_ctor_set(x_567, 0, x_564); +lean_ctor_set(x_567, 1, x_565); +return x_567; +} +else +{ +lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; uint8_t x_572; +x_568 = lean_ctor_get(x_563, 0); +lean_inc(x_568); +x_569 = lean_ctor_get(x_563, 1); +lean_inc(x_569); +if (lean_is_exclusive(x_563)) { + lean_ctor_release(x_563, 0); + lean_ctor_release(x_563, 1); + x_570 = x_563; +} else { + lean_dec_ref(x_563); + x_570 = lean_box(0); +} +x_571 = lean_ctor_get(x_568, 1); +lean_inc(x_571); +x_572 = lean_nat_dec_eq(x_557, x_571); +lean_dec(x_557); +if (x_572 == 0) +{ +lean_object* x_573; +lean_dec(x_571); +if (lean_is_scalar(x_570)) { + x_573 = lean_alloc_ctor(1, 2, 0); +} else { + x_573 = x_570; +} +lean_ctor_set(x_573, 0, x_568); +lean_ctor_set(x_573, 1, x_569); +return x_573; +} +else +{ +lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; +lean_dec(x_570); +lean_dec(x_569); +x_574 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_575 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_575, 0, x_574); +x_576 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_577 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_576, x_575, x_568); +if (lean_obj_tag(x_577) == 0) +{ +lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; +lean_dec(x_571); +x_578 = lean_ctor_get(x_577, 0); +lean_inc(x_578); +x_579 = lean_ctor_get(x_577, 1); +lean_inc(x_579); +if (lean_is_exclusive(x_577)) { + lean_ctor_release(x_577, 0); + lean_ctor_release(x_577, 1); + x_580 = x_577; +} else { + lean_dec_ref(x_577); + x_580 = lean_box(0); +} +if (lean_is_scalar(x_580)) { + x_581 = lean_alloc_ctor(0, 2, 0); +} else { + x_581 = x_580; +} +lean_ctor_set(x_581, 0, x_578); +lean_ctor_set(x_581, 1, x_579); +return x_581; +} +else +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; uint8_t x_586; +x_582 = lean_ctor_get(x_577, 0); +lean_inc(x_582); +x_583 = lean_ctor_get(x_577, 1); +lean_inc(x_583); +if (lean_is_exclusive(x_577)) { + lean_ctor_release(x_577, 0); + lean_ctor_release(x_577, 1); + x_584 = x_577; +} else { + lean_dec_ref(x_577); + x_584 = lean_box(0); +} +x_585 = lean_ctor_get(x_582, 1); +lean_inc(x_585); +x_586 = lean_nat_dec_eq(x_571, x_585); +lean_dec(x_585); +lean_dec(x_571); +if (x_586 == 0) +{ +lean_object* x_587; +if (lean_is_scalar(x_584)) { + x_587 = lean_alloc_ctor(1, 2, 0); +} else { + x_587 = x_584; +} +lean_ctor_set(x_587, 0, x_582); +lean_ctor_set(x_587, 1, x_583); +return x_587; +} +else +{ +lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; +lean_dec(x_584); +lean_dec(x_583); +x_588 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_589 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_589, 0, x_588); +x_590 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_591 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_590, x_589, x_582); +return x_591; +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_592; lean_object* x_593; lean_object* x_594; uint8_t x_595; +x_592 = lean_ctor_get(x_328, 0); +x_593 = lean_ctor_get(x_328, 1); +lean_inc(x_593); +lean_inc(x_592); +lean_dec(x_328); +x_594 = lean_ctor_get(x_592, 1); +lean_inc(x_594); +x_595 = lean_nat_dec_eq(x_323, x_594); +lean_dec(x_323); +if (x_595 == 0) +{ +lean_object* x_596; +lean_dec(x_594); +x_596 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_596, 0, x_592); +lean_ctor_set(x_596, 1, x_593); +return x_596; +} +else +{ +lean_object* x_597; lean_object* x_598; lean_object* x_599; lean_object* x_600; +lean_dec(x_593); +x_597 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_598 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_598, 0, x_597); +x_599 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_600 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_599, x_598, x_592); +if (lean_obj_tag(x_600) == 0) +{ +lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; +lean_dec(x_594); +x_601 = lean_ctor_get(x_600, 0); +lean_inc(x_601); +x_602 = lean_ctor_get(x_600, 1); +lean_inc(x_602); +if (lean_is_exclusive(x_600)) { + lean_ctor_release(x_600, 0); + lean_ctor_release(x_600, 1); + x_603 = x_600; +} else { + lean_dec_ref(x_600); + x_603 = lean_box(0); +} +if (lean_is_scalar(x_603)) { + x_604 = lean_alloc_ctor(0, 2, 0); +} else { + x_604 = x_603; +} +lean_ctor_set(x_604, 0, x_601); +lean_ctor_set(x_604, 1, x_602); +return x_604; +} +else +{ +lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; uint8_t x_609; +x_605 = lean_ctor_get(x_600, 0); +lean_inc(x_605); +x_606 = lean_ctor_get(x_600, 1); +lean_inc(x_606); +if (lean_is_exclusive(x_600)) { + lean_ctor_release(x_600, 0); + lean_ctor_release(x_600, 1); + x_607 = x_600; +} else { + lean_dec_ref(x_600); + x_607 = lean_box(0); +} +x_608 = lean_ctor_get(x_605, 1); +lean_inc(x_608); +x_609 = lean_nat_dec_eq(x_594, x_608); +lean_dec(x_594); +if (x_609 == 0) +{ +lean_object* x_610; +lean_dec(x_608); +if (lean_is_scalar(x_607)) { + x_610 = lean_alloc_ctor(1, 2, 0); +} else { + x_610 = x_607; +} +lean_ctor_set(x_610, 0, x_605); +lean_ctor_set(x_610, 1, x_606); +return x_610; +} +else +{ +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; +lean_dec(x_607); +lean_dec(x_606); +x_611 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_612 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_612, 0, x_611); +x_613 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_614 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_613, x_612, x_605); +if (lean_obj_tag(x_614) == 0) +{ +lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; +lean_dec(x_608); +x_615 = lean_ctor_get(x_614, 0); +lean_inc(x_615); +x_616 = lean_ctor_get(x_614, 1); +lean_inc(x_616); +if (lean_is_exclusive(x_614)) { + lean_ctor_release(x_614, 0); + lean_ctor_release(x_614, 1); + x_617 = x_614; +} else { + lean_dec_ref(x_614); + x_617 = lean_box(0); +} +if (lean_is_scalar(x_617)) { + x_618 = lean_alloc_ctor(0, 2, 0); +} else { + x_618 = x_617; +} +lean_ctor_set(x_618, 0, x_615); +lean_ctor_set(x_618, 1, x_616); +return x_618; +} +else +{ +lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; uint8_t x_623; +x_619 = lean_ctor_get(x_614, 0); +lean_inc(x_619); +x_620 = lean_ctor_get(x_614, 1); +lean_inc(x_620); +if (lean_is_exclusive(x_614)) { + lean_ctor_release(x_614, 0); + lean_ctor_release(x_614, 1); + x_621 = x_614; +} else { + lean_dec_ref(x_614); + x_621 = lean_box(0); +} +x_622 = lean_ctor_get(x_619, 1); +lean_inc(x_622); +x_623 = lean_nat_dec_eq(x_608, x_622); +lean_dec(x_608); +if (x_623 == 0) +{ +lean_object* x_624; +lean_dec(x_622); +if (lean_is_scalar(x_621)) { + x_624 = lean_alloc_ctor(1, 2, 0); +} else { + x_624 = x_621; +} +lean_ctor_set(x_624, 0, x_619); +lean_ctor_set(x_624, 1, x_620); +return x_624; +} +else +{ +lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; +lean_dec(x_621); +lean_dec(x_620); +x_625 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_626 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_626, 0, x_625); +x_627 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_628 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_627, x_626, x_619); +if (lean_obj_tag(x_628) == 0) +{ +lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; +lean_dec(x_622); +x_629 = lean_ctor_get(x_628, 0); +lean_inc(x_629); +x_630 = lean_ctor_get(x_628, 1); +lean_inc(x_630); +if (lean_is_exclusive(x_628)) { + lean_ctor_release(x_628, 0); + lean_ctor_release(x_628, 1); + x_631 = x_628; +} else { + lean_dec_ref(x_628); + x_631 = lean_box(0); +} +if (lean_is_scalar(x_631)) { + x_632 = lean_alloc_ctor(0, 2, 0); +} else { + x_632 = x_631; +} +lean_ctor_set(x_632, 0, x_629); +lean_ctor_set(x_632, 1, x_630); +return x_632; +} +else +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; uint8_t x_637; +x_633 = lean_ctor_get(x_628, 0); +lean_inc(x_633); +x_634 = lean_ctor_get(x_628, 1); +lean_inc(x_634); +if (lean_is_exclusive(x_628)) { + lean_ctor_release(x_628, 0); + lean_ctor_release(x_628, 1); + x_635 = x_628; +} else { + lean_dec_ref(x_628); + x_635 = lean_box(0); +} +x_636 = lean_ctor_get(x_633, 1); +lean_inc(x_636); +x_637 = lean_nat_dec_eq(x_622, x_636); +lean_dec(x_622); +if (x_637 == 0) +{ +lean_object* x_638; +lean_dec(x_636); +if (lean_is_scalar(x_635)) { + x_638 = lean_alloc_ctor(1, 2, 0); +} else { + x_638 = x_635; +} +lean_ctor_set(x_638, 0, x_633); +lean_ctor_set(x_638, 1, x_634); +return x_638; +} +else +{ +lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; +lean_dec(x_635); +lean_dec(x_634); +x_639 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_640 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_640, 0, x_639); +x_641 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_642 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_641, x_640, x_633); +if (lean_obj_tag(x_642) == 0) +{ +lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; +lean_dec(x_636); +x_643 = lean_ctor_get(x_642, 0); +lean_inc(x_643); +x_644 = lean_ctor_get(x_642, 1); +lean_inc(x_644); +if (lean_is_exclusive(x_642)) { + lean_ctor_release(x_642, 0); + lean_ctor_release(x_642, 1); + x_645 = x_642; +} else { + lean_dec_ref(x_642); + x_645 = lean_box(0); +} +if (lean_is_scalar(x_645)) { + x_646 = lean_alloc_ctor(0, 2, 0); +} else { + x_646 = x_645; +} +lean_ctor_set(x_646, 0, x_643); +lean_ctor_set(x_646, 1, x_644); +return x_646; +} +else +{ +lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; uint8_t x_651; +x_647 = lean_ctor_get(x_642, 0); +lean_inc(x_647); +x_648 = lean_ctor_get(x_642, 1); +lean_inc(x_648); +if (lean_is_exclusive(x_642)) { + lean_ctor_release(x_642, 0); + lean_ctor_release(x_642, 1); + x_649 = x_642; +} else { + lean_dec_ref(x_642); + x_649 = lean_box(0); +} +x_650 = lean_ctor_get(x_647, 1); +lean_inc(x_650); +x_651 = lean_nat_dec_eq(x_636, x_650); +lean_dec(x_636); +if (x_651 == 0) +{ +lean_object* x_652; +lean_dec(x_650); +if (lean_is_scalar(x_649)) { + x_652 = lean_alloc_ctor(1, 2, 0); +} else { + x_652 = x_649; +} +lean_ctor_set(x_652, 0, x_647); +lean_ctor_set(x_652, 1, x_648); +return x_652; +} +else +{ +lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; +lean_dec(x_649); +lean_dec(x_648); +x_653 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_654 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_654, 0, x_653); +x_655 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_656 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_655, x_654, x_647); +if (lean_obj_tag(x_656) == 0) +{ +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; +lean_dec(x_650); +x_657 = lean_ctor_get(x_656, 0); +lean_inc(x_657); +x_658 = lean_ctor_get(x_656, 1); +lean_inc(x_658); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_659 = x_656; +} else { + lean_dec_ref(x_656); + x_659 = lean_box(0); +} +if (lean_is_scalar(x_659)) { + x_660 = lean_alloc_ctor(0, 2, 0); +} else { + x_660 = x_659; +} +lean_ctor_set(x_660, 0, x_657); +lean_ctor_set(x_660, 1, x_658); +return x_660; +} +else +{ +lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; uint8_t x_665; +x_661 = lean_ctor_get(x_656, 0); +lean_inc(x_661); +x_662 = lean_ctor_get(x_656, 1); +lean_inc(x_662); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_663 = x_656; +} else { + lean_dec_ref(x_656); + x_663 = lean_box(0); +} +x_664 = lean_ctor_get(x_661, 1); +lean_inc(x_664); +x_665 = lean_nat_dec_eq(x_650, x_664); +lean_dec(x_664); +lean_dec(x_650); +if (x_665 == 0) +{ +lean_object* x_666; +if (lean_is_scalar(x_663)) { + x_666 = lean_alloc_ctor(1, 2, 0); +} else { + x_666 = x_663; +} +lean_ctor_set(x_666, 0, x_661); +lean_ctor_set(x_666, 1, x_662); +return x_666; +} +else +{ +lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; +lean_dec(x_663); +lean_dec(x_662); +x_667 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_668 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_668, 0, x_667); +x_669 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_670 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_669, x_668, x_661); +return x_670; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_671; lean_object* x_672; lean_object* x_673; uint8_t x_674; +x_671 = lean_ctor_get(x_315, 0); +x_672 = lean_ctor_get(x_315, 1); +lean_inc(x_672); +lean_inc(x_671); +lean_dec(x_315); +x_673 = lean_ctor_get(x_671, 1); +lean_inc(x_673); +x_674 = lean_nat_dec_eq(x_310, x_673); +lean_dec(x_310); +if (x_674 == 0) +{ +lean_object* x_675; +lean_dec(x_673); +x_675 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_675, 0, x_671); +lean_ctor_set(x_675, 1, x_672); +return x_675; +} +else +{ +lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; +lean_dec(x_672); +x_676 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_677 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_677, 0, x_676); +x_678 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_679 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_678, x_677, x_671); +if (lean_obj_tag(x_679) == 0) +{ +lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; +lean_dec(x_673); +x_680 = lean_ctor_get(x_679, 0); +lean_inc(x_680); +x_681 = lean_ctor_get(x_679, 1); +lean_inc(x_681); +if (lean_is_exclusive(x_679)) { + lean_ctor_release(x_679, 0); + lean_ctor_release(x_679, 1); + x_682 = x_679; +} else { + lean_dec_ref(x_679); + x_682 = lean_box(0); +} +if (lean_is_scalar(x_682)) { + x_683 = lean_alloc_ctor(0, 2, 0); +} else { + x_683 = x_682; +} +lean_ctor_set(x_683, 0, x_680); +lean_ctor_set(x_683, 1, x_681); +return x_683; +} +else +{ +lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; uint8_t x_688; +x_684 = lean_ctor_get(x_679, 0); +lean_inc(x_684); +x_685 = lean_ctor_get(x_679, 1); +lean_inc(x_685); +if (lean_is_exclusive(x_679)) { + lean_ctor_release(x_679, 0); + lean_ctor_release(x_679, 1); + x_686 = x_679; +} else { + lean_dec_ref(x_679); + x_686 = lean_box(0); +} +x_687 = lean_ctor_get(x_684, 1); +lean_inc(x_687); +x_688 = lean_nat_dec_eq(x_673, x_687); +lean_dec(x_673); +if (x_688 == 0) +{ +lean_object* x_689; +lean_dec(x_687); +if (lean_is_scalar(x_686)) { + x_689 = lean_alloc_ctor(1, 2, 0); +} else { + x_689 = x_686; +} +lean_ctor_set(x_689, 0, x_684); +lean_ctor_set(x_689, 1, x_685); +return x_689; +} +else +{ +lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; +lean_dec(x_686); +lean_dec(x_685); +x_690 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_691 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_691, 0, x_690); +x_692 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_693 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_692, x_691, x_684); +if (lean_obj_tag(x_693) == 0) +{ +lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; +lean_dec(x_687); +x_694 = lean_ctor_get(x_693, 0); +lean_inc(x_694); +x_695 = lean_ctor_get(x_693, 1); +lean_inc(x_695); +if (lean_is_exclusive(x_693)) { + lean_ctor_release(x_693, 0); + lean_ctor_release(x_693, 1); + x_696 = x_693; +} else { + lean_dec_ref(x_693); + x_696 = lean_box(0); +} +if (lean_is_scalar(x_696)) { + x_697 = lean_alloc_ctor(0, 2, 0); +} else { + x_697 = x_696; +} +lean_ctor_set(x_697, 0, x_694); +lean_ctor_set(x_697, 1, x_695); +return x_697; +} +else +{ +lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; uint8_t x_702; +x_698 = lean_ctor_get(x_693, 0); +lean_inc(x_698); +x_699 = lean_ctor_get(x_693, 1); +lean_inc(x_699); +if (lean_is_exclusive(x_693)) { + lean_ctor_release(x_693, 0); + lean_ctor_release(x_693, 1); + x_700 = x_693; +} else { + lean_dec_ref(x_693); + x_700 = lean_box(0); +} +x_701 = lean_ctor_get(x_698, 1); +lean_inc(x_701); +x_702 = lean_nat_dec_eq(x_687, x_701); +lean_dec(x_687); +if (x_702 == 0) +{ +lean_object* x_703; +lean_dec(x_701); +if (lean_is_scalar(x_700)) { + x_703 = lean_alloc_ctor(1, 2, 0); +} else { + x_703 = x_700; +} +lean_ctor_set(x_703, 0, x_698); +lean_ctor_set(x_703, 1, x_699); +return x_703; +} +else +{ +lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; +lean_dec(x_700); +lean_dec(x_699); +x_704 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_705 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_705, 0, x_704); +x_706 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_707 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_706, x_705, x_698); +if (lean_obj_tag(x_707) == 0) +{ +lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; +lean_dec(x_701); +x_708 = lean_ctor_get(x_707, 0); +lean_inc(x_708); +x_709 = lean_ctor_get(x_707, 1); +lean_inc(x_709); +if (lean_is_exclusive(x_707)) { + lean_ctor_release(x_707, 0); + lean_ctor_release(x_707, 1); + x_710 = x_707; +} else { + lean_dec_ref(x_707); + x_710 = lean_box(0); +} +if (lean_is_scalar(x_710)) { + x_711 = lean_alloc_ctor(0, 2, 0); +} else { + x_711 = x_710; +} +lean_ctor_set(x_711, 0, x_708); +lean_ctor_set(x_711, 1, x_709); +return x_711; +} +else +{ +lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; uint8_t x_716; +x_712 = lean_ctor_get(x_707, 0); +lean_inc(x_712); +x_713 = lean_ctor_get(x_707, 1); +lean_inc(x_713); +if (lean_is_exclusive(x_707)) { + lean_ctor_release(x_707, 0); + lean_ctor_release(x_707, 1); + x_714 = x_707; +} else { + lean_dec_ref(x_707); + x_714 = lean_box(0); +} +x_715 = lean_ctor_get(x_712, 1); +lean_inc(x_715); +x_716 = lean_nat_dec_eq(x_701, x_715); +lean_dec(x_701); +if (x_716 == 0) +{ +lean_object* x_717; +lean_dec(x_715); +if (lean_is_scalar(x_714)) { + x_717 = lean_alloc_ctor(1, 2, 0); +} else { + x_717 = x_714; +} +lean_ctor_set(x_717, 0, x_712); +lean_ctor_set(x_717, 1, x_713); +return x_717; +} +else +{ +lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; +lean_dec(x_714); +lean_dec(x_713); +x_718 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_719 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_719, 0, x_718); +x_720 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_721 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_720, x_719, x_712); +if (lean_obj_tag(x_721) == 0) +{ +lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; +lean_dec(x_715); +x_722 = lean_ctor_get(x_721, 0); +lean_inc(x_722); +x_723 = lean_ctor_get(x_721, 1); +lean_inc(x_723); +if (lean_is_exclusive(x_721)) { + lean_ctor_release(x_721, 0); + lean_ctor_release(x_721, 1); + x_724 = x_721; +} else { + lean_dec_ref(x_721); + x_724 = lean_box(0); +} +if (lean_is_scalar(x_724)) { + x_725 = lean_alloc_ctor(0, 2, 0); +} else { + x_725 = x_724; +} +lean_ctor_set(x_725, 0, x_722); +lean_ctor_set(x_725, 1, x_723); +return x_725; +} +else +{ +lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; uint8_t x_730; +x_726 = lean_ctor_get(x_721, 0); +lean_inc(x_726); +x_727 = lean_ctor_get(x_721, 1); +lean_inc(x_727); +if (lean_is_exclusive(x_721)) { + lean_ctor_release(x_721, 0); + lean_ctor_release(x_721, 1); + x_728 = x_721; +} else { + lean_dec_ref(x_721); + x_728 = lean_box(0); +} +x_729 = lean_ctor_get(x_726, 1); +lean_inc(x_729); +x_730 = lean_nat_dec_eq(x_715, x_729); +lean_dec(x_715); +if (x_730 == 0) +{ +lean_object* x_731; +lean_dec(x_729); +if (lean_is_scalar(x_728)) { + x_731 = lean_alloc_ctor(1, 2, 0); +} else { + x_731 = x_728; +} +lean_ctor_set(x_731, 0, x_726); +lean_ctor_set(x_731, 1, x_727); +return x_731; +} +else +{ +lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +lean_dec(x_728); +lean_dec(x_727); +x_732 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_733 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_733, 0, x_732); +x_734 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_735 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_734, x_733, x_726); +if (lean_obj_tag(x_735) == 0) +{ +lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; +lean_dec(x_729); +x_736 = lean_ctor_get(x_735, 0); +lean_inc(x_736); +x_737 = lean_ctor_get(x_735, 1); +lean_inc(x_737); +if (lean_is_exclusive(x_735)) { + lean_ctor_release(x_735, 0); + lean_ctor_release(x_735, 1); + x_738 = x_735; +} else { + lean_dec_ref(x_735); + x_738 = lean_box(0); +} +if (lean_is_scalar(x_738)) { + x_739 = lean_alloc_ctor(0, 2, 0); +} else { + x_739 = x_738; +} +lean_ctor_set(x_739, 0, x_736); +lean_ctor_set(x_739, 1, x_737); +return x_739; +} +else +{ +lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; uint8_t x_744; +x_740 = lean_ctor_get(x_735, 0); +lean_inc(x_740); +x_741 = lean_ctor_get(x_735, 1); +lean_inc(x_741); +if (lean_is_exclusive(x_735)) { + lean_ctor_release(x_735, 0); + lean_ctor_release(x_735, 1); + x_742 = x_735; +} else { + lean_dec_ref(x_735); + x_742 = lean_box(0); +} +x_743 = lean_ctor_get(x_740, 1); +lean_inc(x_743); +x_744 = lean_nat_dec_eq(x_729, x_743); +lean_dec(x_729); +if (x_744 == 0) +{ +lean_object* x_745; +lean_dec(x_743); +if (lean_is_scalar(x_742)) { + x_745 = lean_alloc_ctor(1, 2, 0); +} else { + x_745 = x_742; +} +lean_ctor_set(x_745, 0, x_740); +lean_ctor_set(x_745, 1, x_741); +return x_745; +} +else +{ +lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; +lean_dec(x_742); +lean_dec(x_741); +x_746 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_747 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_747, 0, x_746); +x_748 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_749 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_748, x_747, x_740); +if (lean_obj_tag(x_749) == 0) +{ +lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; +lean_dec(x_743); +x_750 = lean_ctor_get(x_749, 0); +lean_inc(x_750); +x_751 = lean_ctor_get(x_749, 1); +lean_inc(x_751); +if (lean_is_exclusive(x_749)) { + lean_ctor_release(x_749, 0); + lean_ctor_release(x_749, 1); + x_752 = x_749; +} else { + lean_dec_ref(x_749); + x_752 = lean_box(0); +} +if (lean_is_scalar(x_752)) { + x_753 = lean_alloc_ctor(0, 2, 0); +} else { + x_753 = x_752; +} +lean_ctor_set(x_753, 0, x_750); +lean_ctor_set(x_753, 1, x_751); +return x_753; +} +else +{ +lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; uint8_t x_758; +x_754 = lean_ctor_get(x_749, 0); +lean_inc(x_754); +x_755 = lean_ctor_get(x_749, 1); +lean_inc(x_755); +if (lean_is_exclusive(x_749)) { + lean_ctor_release(x_749, 0); + lean_ctor_release(x_749, 1); + x_756 = x_749; +} else { + lean_dec_ref(x_749); + x_756 = lean_box(0); +} +x_757 = lean_ctor_get(x_754, 1); +lean_inc(x_757); +x_758 = lean_nat_dec_eq(x_743, x_757); +lean_dec(x_757); +lean_dec(x_743); +if (x_758 == 0) +{ +lean_object* x_759; +if (lean_is_scalar(x_756)) { + x_759 = lean_alloc_ctor(1, 2, 0); +} else { + x_759 = x_756; +} +lean_ctor_set(x_759, 0, x_754); +lean_ctor_set(x_759, 1, x_755); +return x_759; +} +else +{ +lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; +lean_dec(x_756); +lean_dec(x_755); +x_760 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_761 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_761, 0, x_760); +x_762 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_763 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_762, x_761, x_754); +return x_763; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_764; lean_object* x_765; lean_object* x_766; uint8_t x_767; +x_764 = lean_ctor_get(x_302, 0); +x_765 = lean_ctor_get(x_302, 1); +lean_inc(x_765); +lean_inc(x_764); +lean_dec(x_302); +x_766 = lean_ctor_get(x_764, 1); +lean_inc(x_766); +x_767 = lean_nat_dec_eq(x_297, x_766); +lean_dec(x_297); +if (x_767 == 0) +{ +lean_object* x_768; +lean_dec(x_766); +x_768 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_768, 0, x_764); +lean_ctor_set(x_768, 1, x_765); +return x_768; +} +else +{ +lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; +lean_dec(x_765); +x_769 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_770 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_770, 0, x_769); +x_771 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_772 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_771, x_770, x_764); +if (lean_obj_tag(x_772) == 0) +{ +lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; +lean_dec(x_766); +x_773 = lean_ctor_get(x_772, 0); +lean_inc(x_773); +x_774 = lean_ctor_get(x_772, 1); +lean_inc(x_774); +if (lean_is_exclusive(x_772)) { + lean_ctor_release(x_772, 0); + lean_ctor_release(x_772, 1); + x_775 = x_772; +} else { + lean_dec_ref(x_772); + x_775 = lean_box(0); +} +if (lean_is_scalar(x_775)) { + x_776 = lean_alloc_ctor(0, 2, 0); +} else { + x_776 = x_775; +} +lean_ctor_set(x_776, 0, x_773); +lean_ctor_set(x_776, 1, x_774); +return x_776; +} +else +{ +lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; uint8_t x_781; +x_777 = lean_ctor_get(x_772, 0); +lean_inc(x_777); +x_778 = lean_ctor_get(x_772, 1); +lean_inc(x_778); +if (lean_is_exclusive(x_772)) { + lean_ctor_release(x_772, 0); + lean_ctor_release(x_772, 1); + x_779 = x_772; +} else { + lean_dec_ref(x_772); + x_779 = lean_box(0); +} +x_780 = lean_ctor_get(x_777, 1); +lean_inc(x_780); +x_781 = lean_nat_dec_eq(x_766, x_780); +lean_dec(x_766); +if (x_781 == 0) +{ +lean_object* x_782; +lean_dec(x_780); +if (lean_is_scalar(x_779)) { + x_782 = lean_alloc_ctor(1, 2, 0); +} else { + x_782 = x_779; +} +lean_ctor_set(x_782, 0, x_777); +lean_ctor_set(x_782, 1, x_778); +return x_782; +} +else +{ +lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; +lean_dec(x_779); +lean_dec(x_778); +x_783 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_784 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_784, 0, x_783); +x_785 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_786 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_785, x_784, x_777); +if (lean_obj_tag(x_786) == 0) +{ +lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; +lean_dec(x_780); +x_787 = lean_ctor_get(x_786, 0); +lean_inc(x_787); +x_788 = lean_ctor_get(x_786, 1); +lean_inc(x_788); +if (lean_is_exclusive(x_786)) { + lean_ctor_release(x_786, 0); + lean_ctor_release(x_786, 1); + x_789 = x_786; +} else { + lean_dec_ref(x_786); + x_789 = lean_box(0); +} +if (lean_is_scalar(x_789)) { + x_790 = lean_alloc_ctor(0, 2, 0); +} else { + x_790 = x_789; +} +lean_ctor_set(x_790, 0, x_787); +lean_ctor_set(x_790, 1, x_788); +return x_790; +} +else +{ +lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; uint8_t x_795; +x_791 = lean_ctor_get(x_786, 0); +lean_inc(x_791); +x_792 = lean_ctor_get(x_786, 1); +lean_inc(x_792); +if (lean_is_exclusive(x_786)) { + lean_ctor_release(x_786, 0); + lean_ctor_release(x_786, 1); + x_793 = x_786; +} else { + lean_dec_ref(x_786); + x_793 = lean_box(0); +} +x_794 = lean_ctor_get(x_791, 1); +lean_inc(x_794); +x_795 = lean_nat_dec_eq(x_780, x_794); +lean_dec(x_780); +if (x_795 == 0) +{ +lean_object* x_796; +lean_dec(x_794); +if (lean_is_scalar(x_793)) { + x_796 = lean_alloc_ctor(1, 2, 0); +} else { + x_796 = x_793; +} +lean_ctor_set(x_796, 0, x_791); +lean_ctor_set(x_796, 1, x_792); +return x_796; +} +else +{ +lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; +lean_dec(x_793); +lean_dec(x_792); +x_797 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_798 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_798, 0, x_797); +x_799 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_800 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_799, x_798, x_791); +if (lean_obj_tag(x_800) == 0) +{ +lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; +lean_dec(x_794); +x_801 = lean_ctor_get(x_800, 0); +lean_inc(x_801); +x_802 = lean_ctor_get(x_800, 1); +lean_inc(x_802); +if (lean_is_exclusive(x_800)) { + lean_ctor_release(x_800, 0); + lean_ctor_release(x_800, 1); + x_803 = x_800; +} else { + lean_dec_ref(x_800); + x_803 = lean_box(0); +} +if (lean_is_scalar(x_803)) { + x_804 = lean_alloc_ctor(0, 2, 0); +} else { + x_804 = x_803; +} +lean_ctor_set(x_804, 0, x_801); +lean_ctor_set(x_804, 1, x_802); +return x_804; +} +else +{ +lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; uint8_t x_809; +x_805 = lean_ctor_get(x_800, 0); +lean_inc(x_805); +x_806 = lean_ctor_get(x_800, 1); +lean_inc(x_806); +if (lean_is_exclusive(x_800)) { + lean_ctor_release(x_800, 0); + lean_ctor_release(x_800, 1); + x_807 = x_800; +} else { + lean_dec_ref(x_800); + x_807 = lean_box(0); +} +x_808 = lean_ctor_get(x_805, 1); +lean_inc(x_808); +x_809 = lean_nat_dec_eq(x_794, x_808); +lean_dec(x_794); +if (x_809 == 0) +{ +lean_object* x_810; +lean_dec(x_808); +if (lean_is_scalar(x_807)) { + x_810 = lean_alloc_ctor(1, 2, 0); +} else { + x_810 = x_807; +} +lean_ctor_set(x_810, 0, x_805); +lean_ctor_set(x_810, 1, x_806); +return x_810; +} +else +{ +lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; +lean_dec(x_807); +lean_dec(x_806); +x_811 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_812 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_812, 0, x_811); +x_813 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_814 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_813, x_812, x_805); +if (lean_obj_tag(x_814) == 0) +{ +lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; +lean_dec(x_808); +x_815 = lean_ctor_get(x_814, 0); +lean_inc(x_815); +x_816 = lean_ctor_get(x_814, 1); +lean_inc(x_816); +if (lean_is_exclusive(x_814)) { + lean_ctor_release(x_814, 0); + lean_ctor_release(x_814, 1); + x_817 = x_814; +} else { + lean_dec_ref(x_814); + x_817 = lean_box(0); +} +if (lean_is_scalar(x_817)) { + x_818 = lean_alloc_ctor(0, 2, 0); +} else { + x_818 = x_817; +} +lean_ctor_set(x_818, 0, x_815); +lean_ctor_set(x_818, 1, x_816); +return x_818; +} +else +{ +lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; uint8_t x_823; +x_819 = lean_ctor_get(x_814, 0); +lean_inc(x_819); +x_820 = lean_ctor_get(x_814, 1); +lean_inc(x_820); +if (lean_is_exclusive(x_814)) { + lean_ctor_release(x_814, 0); + lean_ctor_release(x_814, 1); + x_821 = x_814; +} else { + lean_dec_ref(x_814); + x_821 = lean_box(0); +} +x_822 = lean_ctor_get(x_819, 1); +lean_inc(x_822); +x_823 = lean_nat_dec_eq(x_808, x_822); +lean_dec(x_808); +if (x_823 == 0) +{ +lean_object* x_824; +lean_dec(x_822); +if (lean_is_scalar(x_821)) { + x_824 = lean_alloc_ctor(1, 2, 0); +} else { + x_824 = x_821; +} +lean_ctor_set(x_824, 0, x_819); +lean_ctor_set(x_824, 1, x_820); +return x_824; +} +else +{ +lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; +lean_dec(x_821); +lean_dec(x_820); +x_825 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_826 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_826, 0, x_825); +x_827 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_828 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_827, x_826, x_819); +if (lean_obj_tag(x_828) == 0) +{ +lean_object* x_829; lean_object* x_830; lean_object* x_831; lean_object* x_832; +lean_dec(x_822); +x_829 = lean_ctor_get(x_828, 0); +lean_inc(x_829); +x_830 = lean_ctor_get(x_828, 1); +lean_inc(x_830); +if (lean_is_exclusive(x_828)) { + lean_ctor_release(x_828, 0); + lean_ctor_release(x_828, 1); + x_831 = x_828; +} else { + lean_dec_ref(x_828); + x_831 = lean_box(0); +} +if (lean_is_scalar(x_831)) { + x_832 = lean_alloc_ctor(0, 2, 0); +} else { + x_832 = x_831; +} +lean_ctor_set(x_832, 0, x_829); +lean_ctor_set(x_832, 1, x_830); +return x_832; +} +else +{ +lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; uint8_t x_837; +x_833 = lean_ctor_get(x_828, 0); +lean_inc(x_833); +x_834 = lean_ctor_get(x_828, 1); +lean_inc(x_834); +if (lean_is_exclusive(x_828)) { + lean_ctor_release(x_828, 0); + lean_ctor_release(x_828, 1); + x_835 = x_828; +} else { + lean_dec_ref(x_828); + x_835 = lean_box(0); +} +x_836 = lean_ctor_get(x_833, 1); +lean_inc(x_836); +x_837 = lean_nat_dec_eq(x_822, x_836); +lean_dec(x_822); +if (x_837 == 0) +{ +lean_object* x_838; +lean_dec(x_836); +if (lean_is_scalar(x_835)) { + x_838 = lean_alloc_ctor(1, 2, 0); +} else { + x_838 = x_835; +} +lean_ctor_set(x_838, 0, x_833); +lean_ctor_set(x_838, 1, x_834); +return x_838; +} +else +{ +lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; +lean_dec(x_835); +lean_dec(x_834); +x_839 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_840 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_840, 0, x_839); +x_841 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_842 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_841, x_840, x_833); +if (lean_obj_tag(x_842) == 0) +{ +lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; +lean_dec(x_836); +x_843 = lean_ctor_get(x_842, 0); +lean_inc(x_843); +x_844 = lean_ctor_get(x_842, 1); +lean_inc(x_844); +if (lean_is_exclusive(x_842)) { + lean_ctor_release(x_842, 0); + lean_ctor_release(x_842, 1); + x_845 = x_842; +} else { + lean_dec_ref(x_842); + x_845 = lean_box(0); +} +if (lean_is_scalar(x_845)) { + x_846 = lean_alloc_ctor(0, 2, 0); +} else { + x_846 = x_845; +} +lean_ctor_set(x_846, 0, x_843); +lean_ctor_set(x_846, 1, x_844); +return x_846; +} +else +{ +lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; uint8_t x_851; +x_847 = lean_ctor_get(x_842, 0); +lean_inc(x_847); +x_848 = lean_ctor_get(x_842, 1); +lean_inc(x_848); +if (lean_is_exclusive(x_842)) { + lean_ctor_release(x_842, 0); + lean_ctor_release(x_842, 1); + x_849 = x_842; +} else { + lean_dec_ref(x_842); + x_849 = lean_box(0); +} +x_850 = lean_ctor_get(x_847, 1); +lean_inc(x_850); +x_851 = lean_nat_dec_eq(x_836, x_850); +lean_dec(x_836); +if (x_851 == 0) +{ +lean_object* x_852; +lean_dec(x_850); +if (lean_is_scalar(x_849)) { + x_852 = lean_alloc_ctor(1, 2, 0); +} else { + x_852 = x_849; +} +lean_ctor_set(x_852, 0, x_847); +lean_ctor_set(x_852, 1, x_848); +return x_852; +} +else +{ +lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; +lean_dec(x_849); +lean_dec(x_848); +x_853 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_854 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_854, 0, x_853); +x_855 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_856 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_855, x_854, x_847); +if (lean_obj_tag(x_856) == 0) +{ +lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; +lean_dec(x_850); +x_857 = lean_ctor_get(x_856, 0); +lean_inc(x_857); +x_858 = lean_ctor_get(x_856, 1); +lean_inc(x_858); +if (lean_is_exclusive(x_856)) { + lean_ctor_release(x_856, 0); + lean_ctor_release(x_856, 1); + x_859 = x_856; +} else { + lean_dec_ref(x_856); + x_859 = lean_box(0); +} +if (lean_is_scalar(x_859)) { + x_860 = lean_alloc_ctor(0, 2, 0); +} else { + x_860 = x_859; +} +lean_ctor_set(x_860, 0, x_857); +lean_ctor_set(x_860, 1, x_858); +return x_860; +} +else +{ +lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; uint8_t x_865; +x_861 = lean_ctor_get(x_856, 0); +lean_inc(x_861); +x_862 = lean_ctor_get(x_856, 1); +lean_inc(x_862); +if (lean_is_exclusive(x_856)) { + lean_ctor_release(x_856, 0); + lean_ctor_release(x_856, 1); + x_863 = x_856; +} else { + lean_dec_ref(x_856); + x_863 = lean_box(0); +} +x_864 = lean_ctor_get(x_861, 1); +lean_inc(x_864); +x_865 = lean_nat_dec_eq(x_850, x_864); +lean_dec(x_864); +lean_dec(x_850); +if (x_865 == 0) +{ +lean_object* x_866; +if (lean_is_scalar(x_863)) { + x_866 = lean_alloc_ctor(1, 2, 0); +} else { + x_866 = x_863; +} +lean_ctor_set(x_866, 0, x_861); +lean_ctor_set(x_866, 1, x_862); +return x_866; +} +else +{ +lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; +lean_dec(x_863); +lean_dec(x_862); +x_867 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_868 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_868, 0, x_867); +x_869 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_870 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_869, x_868, x_861); +return x_870; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_871; lean_object* x_872; lean_object* x_873; uint8_t x_874; +x_871 = lean_ctor_get(x_289, 0); +x_872 = lean_ctor_get(x_289, 1); +lean_inc(x_872); +lean_inc(x_871); +lean_dec(x_289); +x_873 = lean_ctor_get(x_871, 1); +lean_inc(x_873); +x_874 = lean_nat_dec_eq(x_284, x_873); +lean_dec(x_284); +if (x_874 == 0) +{ +lean_object* x_875; +lean_dec(x_873); +x_875 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_875, 0, x_871); +lean_ctor_set(x_875, 1, x_872); +return x_875; +} +else +{ +lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; +lean_dec(x_872); +x_876 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_877 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_877, 0, x_876); +x_878 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_879 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_878, x_877, x_871); +if (lean_obj_tag(x_879) == 0) +{ +lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_dec(x_873); +x_880 = lean_ctor_get(x_879, 0); +lean_inc(x_880); +x_881 = lean_ctor_get(x_879, 1); +lean_inc(x_881); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_882 = x_879; +} else { + lean_dec_ref(x_879); + x_882 = lean_box(0); +} +if (lean_is_scalar(x_882)) { + x_883 = lean_alloc_ctor(0, 2, 0); +} else { + x_883 = x_882; +} +lean_ctor_set(x_883, 0, x_880); +lean_ctor_set(x_883, 1, x_881); +return x_883; +} +else +{ +lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; uint8_t x_888; +x_884 = lean_ctor_get(x_879, 0); +lean_inc(x_884); +x_885 = lean_ctor_get(x_879, 1); +lean_inc(x_885); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_886 = x_879; +} else { + lean_dec_ref(x_879); + x_886 = lean_box(0); +} +x_887 = lean_ctor_get(x_884, 1); +lean_inc(x_887); +x_888 = lean_nat_dec_eq(x_873, x_887); +lean_dec(x_873); +if (x_888 == 0) +{ +lean_object* x_889; +lean_dec(x_887); +if (lean_is_scalar(x_886)) { + x_889 = lean_alloc_ctor(1, 2, 0); +} else { + x_889 = x_886; +} +lean_ctor_set(x_889, 0, x_884); +lean_ctor_set(x_889, 1, x_885); +return x_889; +} +else +{ +lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; +lean_dec(x_886); +lean_dec(x_885); +x_890 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_891 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_891, 0, x_890); +x_892 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_893 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_892, x_891, x_884); +if (lean_obj_tag(x_893) == 0) +{ +lean_object* x_894; lean_object* x_895; lean_object* x_896; lean_object* x_897; +lean_dec(x_887); +x_894 = lean_ctor_get(x_893, 0); +lean_inc(x_894); +x_895 = lean_ctor_get(x_893, 1); +lean_inc(x_895); +if (lean_is_exclusive(x_893)) { + lean_ctor_release(x_893, 0); + lean_ctor_release(x_893, 1); + x_896 = x_893; +} else { + lean_dec_ref(x_893); + x_896 = lean_box(0); +} +if (lean_is_scalar(x_896)) { + x_897 = lean_alloc_ctor(0, 2, 0); +} else { + x_897 = x_896; +} +lean_ctor_set(x_897, 0, x_894); +lean_ctor_set(x_897, 1, x_895); +return x_897; +} +else +{ +lean_object* x_898; lean_object* x_899; lean_object* x_900; lean_object* x_901; uint8_t x_902; +x_898 = lean_ctor_get(x_893, 0); +lean_inc(x_898); +x_899 = lean_ctor_get(x_893, 1); +lean_inc(x_899); +if (lean_is_exclusive(x_893)) { + lean_ctor_release(x_893, 0); + lean_ctor_release(x_893, 1); + x_900 = x_893; +} else { + lean_dec_ref(x_893); + x_900 = lean_box(0); +} +x_901 = lean_ctor_get(x_898, 1); +lean_inc(x_901); +x_902 = lean_nat_dec_eq(x_887, x_901); +lean_dec(x_887); +if (x_902 == 0) +{ +lean_object* x_903; +lean_dec(x_901); +if (lean_is_scalar(x_900)) { + x_903 = lean_alloc_ctor(1, 2, 0); +} else { + x_903 = x_900; +} +lean_ctor_set(x_903, 0, x_898); +lean_ctor_set(x_903, 1, x_899); +return x_903; +} +else +{ +lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; +lean_dec(x_900); +lean_dec(x_899); +x_904 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_905 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_905, 0, x_904); +x_906 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_907 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_906, x_905, x_898); +if (lean_obj_tag(x_907) == 0) +{ +lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; +lean_dec(x_901); +x_908 = lean_ctor_get(x_907, 0); +lean_inc(x_908); +x_909 = lean_ctor_get(x_907, 1); +lean_inc(x_909); +if (lean_is_exclusive(x_907)) { + lean_ctor_release(x_907, 0); + lean_ctor_release(x_907, 1); + x_910 = x_907; +} else { + lean_dec_ref(x_907); + x_910 = lean_box(0); +} +if (lean_is_scalar(x_910)) { + x_911 = lean_alloc_ctor(0, 2, 0); +} else { + x_911 = x_910; +} +lean_ctor_set(x_911, 0, x_908); +lean_ctor_set(x_911, 1, x_909); +return x_911; +} +else +{ +lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; uint8_t x_916; +x_912 = lean_ctor_get(x_907, 0); +lean_inc(x_912); +x_913 = lean_ctor_get(x_907, 1); +lean_inc(x_913); +if (lean_is_exclusive(x_907)) { + lean_ctor_release(x_907, 0); + lean_ctor_release(x_907, 1); + x_914 = x_907; +} else { + lean_dec_ref(x_907); + x_914 = lean_box(0); +} +x_915 = lean_ctor_get(x_912, 1); +lean_inc(x_915); +x_916 = lean_nat_dec_eq(x_901, x_915); +lean_dec(x_901); +if (x_916 == 0) +{ +lean_object* x_917; +lean_dec(x_915); +if (lean_is_scalar(x_914)) { + x_917 = lean_alloc_ctor(1, 2, 0); +} else { + x_917 = x_914; +} +lean_ctor_set(x_917, 0, x_912); +lean_ctor_set(x_917, 1, x_913); +return x_917; +} +else +{ +lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; +lean_dec(x_914); +lean_dec(x_913); +x_918 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_919 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_919, 0, x_918); +x_920 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_921 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_920, x_919, x_912); +if (lean_obj_tag(x_921) == 0) +{ +lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; +lean_dec(x_915); +x_922 = lean_ctor_get(x_921, 0); +lean_inc(x_922); +x_923 = lean_ctor_get(x_921, 1); +lean_inc(x_923); +if (lean_is_exclusive(x_921)) { + lean_ctor_release(x_921, 0); + lean_ctor_release(x_921, 1); + x_924 = x_921; +} else { + lean_dec_ref(x_921); + x_924 = lean_box(0); +} +if (lean_is_scalar(x_924)) { + x_925 = lean_alloc_ctor(0, 2, 0); +} else { + x_925 = x_924; +} +lean_ctor_set(x_925, 0, x_922); +lean_ctor_set(x_925, 1, x_923); +return x_925; +} +else +{ +lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; uint8_t x_930; +x_926 = lean_ctor_get(x_921, 0); +lean_inc(x_926); +x_927 = lean_ctor_get(x_921, 1); +lean_inc(x_927); +if (lean_is_exclusive(x_921)) { + lean_ctor_release(x_921, 0); + lean_ctor_release(x_921, 1); + x_928 = x_921; +} else { + lean_dec_ref(x_921); + x_928 = lean_box(0); +} +x_929 = lean_ctor_get(x_926, 1); +lean_inc(x_929); +x_930 = lean_nat_dec_eq(x_915, x_929); +lean_dec(x_915); +if (x_930 == 0) +{ +lean_object* x_931; +lean_dec(x_929); +if (lean_is_scalar(x_928)) { + x_931 = lean_alloc_ctor(1, 2, 0); +} else { + x_931 = x_928; +} +lean_ctor_set(x_931, 0, x_926); +lean_ctor_set(x_931, 1, x_927); +return x_931; +} +else +{ +lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; +lean_dec(x_928); +lean_dec(x_927); +x_932 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_933 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_933, 0, x_932); +x_934 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_935 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_934, x_933, x_926); +if (lean_obj_tag(x_935) == 0) +{ +lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; +lean_dec(x_929); +x_936 = lean_ctor_get(x_935, 0); +lean_inc(x_936); +x_937 = lean_ctor_get(x_935, 1); +lean_inc(x_937); +if (lean_is_exclusive(x_935)) { + lean_ctor_release(x_935, 0); + lean_ctor_release(x_935, 1); + x_938 = x_935; +} else { + lean_dec_ref(x_935); + x_938 = lean_box(0); +} +if (lean_is_scalar(x_938)) { + x_939 = lean_alloc_ctor(0, 2, 0); +} else { + x_939 = x_938; +} +lean_ctor_set(x_939, 0, x_936); +lean_ctor_set(x_939, 1, x_937); +return x_939; +} +else +{ +lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; uint8_t x_944; +x_940 = lean_ctor_get(x_935, 0); +lean_inc(x_940); +x_941 = lean_ctor_get(x_935, 1); +lean_inc(x_941); +if (lean_is_exclusive(x_935)) { + lean_ctor_release(x_935, 0); + lean_ctor_release(x_935, 1); + x_942 = x_935; +} else { + lean_dec_ref(x_935); + x_942 = lean_box(0); +} +x_943 = lean_ctor_get(x_940, 1); +lean_inc(x_943); +x_944 = lean_nat_dec_eq(x_929, x_943); +lean_dec(x_929); +if (x_944 == 0) +{ +lean_object* x_945; +lean_dec(x_943); +if (lean_is_scalar(x_942)) { + x_945 = lean_alloc_ctor(1, 2, 0); +} else { + x_945 = x_942; +} +lean_ctor_set(x_945, 0, x_940); +lean_ctor_set(x_945, 1, x_941); +return x_945; +} +else +{ +lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; +lean_dec(x_942); +lean_dec(x_941); +x_946 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_947 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_947, 0, x_946); +x_948 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_949 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_948, x_947, x_940); +if (lean_obj_tag(x_949) == 0) +{ +lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; +lean_dec(x_943); +x_950 = lean_ctor_get(x_949, 0); +lean_inc(x_950); +x_951 = lean_ctor_get(x_949, 1); +lean_inc(x_951); +if (lean_is_exclusive(x_949)) { + lean_ctor_release(x_949, 0); + lean_ctor_release(x_949, 1); + x_952 = x_949; +} else { + lean_dec_ref(x_949); + x_952 = lean_box(0); +} +if (lean_is_scalar(x_952)) { + x_953 = lean_alloc_ctor(0, 2, 0); +} else { + x_953 = x_952; +} +lean_ctor_set(x_953, 0, x_950); +lean_ctor_set(x_953, 1, x_951); +return x_953; +} +else +{ +lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; uint8_t x_958; +x_954 = lean_ctor_get(x_949, 0); +lean_inc(x_954); +x_955 = lean_ctor_get(x_949, 1); +lean_inc(x_955); +if (lean_is_exclusive(x_949)) { + lean_ctor_release(x_949, 0); + lean_ctor_release(x_949, 1); + x_956 = x_949; +} else { + lean_dec_ref(x_949); + x_956 = lean_box(0); +} +x_957 = lean_ctor_get(x_954, 1); +lean_inc(x_957); +x_958 = lean_nat_dec_eq(x_943, x_957); +lean_dec(x_943); +if (x_958 == 0) +{ +lean_object* x_959; +lean_dec(x_957); +if (lean_is_scalar(x_956)) { + x_959 = lean_alloc_ctor(1, 2, 0); +} else { + x_959 = x_956; +} +lean_ctor_set(x_959, 0, x_954); +lean_ctor_set(x_959, 1, x_955); +return x_959; +} +else +{ +lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; +lean_dec(x_956); +lean_dec(x_955); +x_960 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_961 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_961, 0, x_960); +x_962 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_963 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_962, x_961, x_954); +if (lean_obj_tag(x_963) == 0) +{ +lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; +lean_dec(x_957); +x_964 = lean_ctor_get(x_963, 0); +lean_inc(x_964); +x_965 = lean_ctor_get(x_963, 1); +lean_inc(x_965); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_966 = x_963; +} else { + lean_dec_ref(x_963); + x_966 = lean_box(0); +} +if (lean_is_scalar(x_966)) { + x_967 = lean_alloc_ctor(0, 2, 0); +} else { + x_967 = x_966; +} +lean_ctor_set(x_967, 0, x_964); +lean_ctor_set(x_967, 1, x_965); +return x_967; +} +else +{ +lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; uint8_t x_972; +x_968 = lean_ctor_get(x_963, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_963, 1); +lean_inc(x_969); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_970 = x_963; +} else { + lean_dec_ref(x_963); + x_970 = lean_box(0); +} +x_971 = lean_ctor_get(x_968, 1); +lean_inc(x_971); +x_972 = lean_nat_dec_eq(x_957, x_971); +lean_dec(x_957); +if (x_972 == 0) +{ +lean_object* x_973; +lean_dec(x_971); +if (lean_is_scalar(x_970)) { + x_973 = lean_alloc_ctor(1, 2, 0); +} else { + x_973 = x_970; +} +lean_ctor_set(x_973, 0, x_968); +lean_ctor_set(x_973, 1, x_969); +return x_973; +} +else +{ +lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; +lean_dec(x_970); +lean_dec(x_969); +x_974 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_975 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_975, 0, x_974); +x_976 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_977 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_976, x_975, x_968); +if (lean_obj_tag(x_977) == 0) +{ +lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; +lean_dec(x_971); +x_978 = lean_ctor_get(x_977, 0); +lean_inc(x_978); +x_979 = lean_ctor_get(x_977, 1); +lean_inc(x_979); +if (lean_is_exclusive(x_977)) { + lean_ctor_release(x_977, 0); + lean_ctor_release(x_977, 1); + x_980 = x_977; +} else { + lean_dec_ref(x_977); + x_980 = lean_box(0); +} +if (lean_is_scalar(x_980)) { + x_981 = lean_alloc_ctor(0, 2, 0); +} else { + x_981 = x_980; +} +lean_ctor_set(x_981, 0, x_978); +lean_ctor_set(x_981, 1, x_979); +return x_981; +} +else +{ +lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; uint8_t x_986; +x_982 = lean_ctor_get(x_977, 0); +lean_inc(x_982); +x_983 = lean_ctor_get(x_977, 1); +lean_inc(x_983); +if (lean_is_exclusive(x_977)) { + lean_ctor_release(x_977, 0); + lean_ctor_release(x_977, 1); + x_984 = x_977; +} else { + lean_dec_ref(x_977); + x_984 = lean_box(0); +} +x_985 = lean_ctor_get(x_982, 1); +lean_inc(x_985); +x_986 = lean_nat_dec_eq(x_971, x_985); +lean_dec(x_985); +lean_dec(x_971); +if (x_986 == 0) +{ +lean_object* x_987; +if (lean_is_scalar(x_984)) { + x_987 = lean_alloc_ctor(1, 2, 0); +} else { + x_987 = x_984; +} +lean_ctor_set(x_987, 0, x_982); +lean_ctor_set(x_987, 1, x_983); +return x_987; +} +else +{ +lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; +lean_dec(x_984); +lean_dec(x_983); +x_988 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_989 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_989, 0, x_988); +x_990 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_991 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_990, x_989, x_982); +return x_991; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_992; lean_object* x_993; lean_object* x_994; uint8_t x_995; +x_992 = lean_ctor_get(x_276, 0); +x_993 = lean_ctor_get(x_276, 1); +lean_inc(x_993); +lean_inc(x_992); +lean_dec(x_276); +x_994 = lean_ctor_get(x_992, 1); +lean_inc(x_994); +x_995 = lean_nat_dec_eq(x_271, x_994); +lean_dec(x_271); +if (x_995 == 0) +{ +lean_object* x_996; +lean_dec(x_994); +x_996 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_996, 0, x_992); +lean_ctor_set(x_996, 1, x_993); +return x_996; +} +else +{ +lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; +lean_dec(x_993); +x_997 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_998 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_998, 0, x_997); +x_999 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1000 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_999, x_998, x_992); +if (lean_obj_tag(x_1000) == 0) +{ +lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; +lean_dec(x_994); +x_1001 = lean_ctor_get(x_1000, 0); +lean_inc(x_1001); +x_1002 = lean_ctor_get(x_1000, 1); +lean_inc(x_1002); +if (lean_is_exclusive(x_1000)) { + lean_ctor_release(x_1000, 0); + lean_ctor_release(x_1000, 1); + x_1003 = x_1000; +} else { + lean_dec_ref(x_1000); + x_1003 = lean_box(0); +} +if (lean_is_scalar(x_1003)) { + x_1004 = lean_alloc_ctor(0, 2, 0); +} else { + x_1004 = x_1003; +} +lean_ctor_set(x_1004, 0, x_1001); +lean_ctor_set(x_1004, 1, x_1002); +return x_1004; +} +else +{ +lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; uint8_t x_1009; +x_1005 = lean_ctor_get(x_1000, 0); +lean_inc(x_1005); +x_1006 = lean_ctor_get(x_1000, 1); +lean_inc(x_1006); +if (lean_is_exclusive(x_1000)) { + lean_ctor_release(x_1000, 0); + lean_ctor_release(x_1000, 1); + x_1007 = x_1000; +} else { + lean_dec_ref(x_1000); + x_1007 = lean_box(0); +} +x_1008 = lean_ctor_get(x_1005, 1); +lean_inc(x_1008); +x_1009 = lean_nat_dec_eq(x_994, x_1008); +lean_dec(x_994); +if (x_1009 == 0) +{ +lean_object* x_1010; +lean_dec(x_1008); +if (lean_is_scalar(x_1007)) { + x_1010 = lean_alloc_ctor(1, 2, 0); +} else { + x_1010 = x_1007; +} +lean_ctor_set(x_1010, 0, x_1005); +lean_ctor_set(x_1010, 1, x_1006); +return x_1010; +} +else +{ +lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; +lean_dec(x_1007); +lean_dec(x_1006); +x_1011 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1012 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1012, 0, x_1011); +x_1013 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1014 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1013, x_1012, x_1005); +if (lean_obj_tag(x_1014) == 0) +{ +lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; +lean_dec(x_1008); +x_1015 = lean_ctor_get(x_1014, 0); +lean_inc(x_1015); +x_1016 = lean_ctor_get(x_1014, 1); +lean_inc(x_1016); +if (lean_is_exclusive(x_1014)) { + lean_ctor_release(x_1014, 0); + lean_ctor_release(x_1014, 1); + x_1017 = x_1014; +} else { + lean_dec_ref(x_1014); + x_1017 = lean_box(0); +} +if (lean_is_scalar(x_1017)) { + x_1018 = lean_alloc_ctor(0, 2, 0); +} else { + x_1018 = x_1017; +} +lean_ctor_set(x_1018, 0, x_1015); +lean_ctor_set(x_1018, 1, x_1016); +return x_1018; +} +else +{ +lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; uint8_t x_1023; +x_1019 = lean_ctor_get(x_1014, 0); +lean_inc(x_1019); +x_1020 = lean_ctor_get(x_1014, 1); +lean_inc(x_1020); +if (lean_is_exclusive(x_1014)) { + lean_ctor_release(x_1014, 0); + lean_ctor_release(x_1014, 1); + x_1021 = x_1014; +} else { + lean_dec_ref(x_1014); + x_1021 = lean_box(0); +} +x_1022 = lean_ctor_get(x_1019, 1); +lean_inc(x_1022); +x_1023 = lean_nat_dec_eq(x_1008, x_1022); +lean_dec(x_1008); +if (x_1023 == 0) +{ +lean_object* x_1024; +lean_dec(x_1022); +if (lean_is_scalar(x_1021)) { + x_1024 = lean_alloc_ctor(1, 2, 0); +} else { + x_1024 = x_1021; +} +lean_ctor_set(x_1024, 0, x_1019); +lean_ctor_set(x_1024, 1, x_1020); +return x_1024; +} +else +{ +lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; +lean_dec(x_1021); +lean_dec(x_1020); +x_1025 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1026 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1026, 0, x_1025); +x_1027 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1028 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1027, x_1026, x_1019); +if (lean_obj_tag(x_1028) == 0) +{ +lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; +lean_dec(x_1022); +x_1029 = lean_ctor_get(x_1028, 0); +lean_inc(x_1029); +x_1030 = lean_ctor_get(x_1028, 1); +lean_inc(x_1030); +if (lean_is_exclusive(x_1028)) { + lean_ctor_release(x_1028, 0); + lean_ctor_release(x_1028, 1); + x_1031 = x_1028; +} else { + lean_dec_ref(x_1028); + x_1031 = lean_box(0); +} +if (lean_is_scalar(x_1031)) { + x_1032 = lean_alloc_ctor(0, 2, 0); +} else { + x_1032 = x_1031; +} +lean_ctor_set(x_1032, 0, x_1029); +lean_ctor_set(x_1032, 1, x_1030); +return x_1032; +} +else +{ +lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; uint8_t x_1037; +x_1033 = lean_ctor_get(x_1028, 0); +lean_inc(x_1033); +x_1034 = lean_ctor_get(x_1028, 1); +lean_inc(x_1034); +if (lean_is_exclusive(x_1028)) { + lean_ctor_release(x_1028, 0); + lean_ctor_release(x_1028, 1); + x_1035 = x_1028; +} else { + lean_dec_ref(x_1028); + x_1035 = lean_box(0); +} +x_1036 = lean_ctor_get(x_1033, 1); +lean_inc(x_1036); +x_1037 = lean_nat_dec_eq(x_1022, x_1036); +lean_dec(x_1022); +if (x_1037 == 0) +{ +lean_object* x_1038; +lean_dec(x_1036); +if (lean_is_scalar(x_1035)) { + x_1038 = lean_alloc_ctor(1, 2, 0); +} else { + x_1038 = x_1035; +} +lean_ctor_set(x_1038, 0, x_1033); +lean_ctor_set(x_1038, 1, x_1034); +return x_1038; +} +else +{ +lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; +lean_dec(x_1035); +lean_dec(x_1034); +x_1039 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1040 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1040, 0, x_1039); +x_1041 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1042 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1041, x_1040, x_1033); +if (lean_obj_tag(x_1042) == 0) +{ +lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; +lean_dec(x_1036); +x_1043 = lean_ctor_get(x_1042, 0); +lean_inc(x_1043); +x_1044 = lean_ctor_get(x_1042, 1); +lean_inc(x_1044); +if (lean_is_exclusive(x_1042)) { + lean_ctor_release(x_1042, 0); + lean_ctor_release(x_1042, 1); + x_1045 = x_1042; +} else { + lean_dec_ref(x_1042); + x_1045 = lean_box(0); +} +if (lean_is_scalar(x_1045)) { + x_1046 = lean_alloc_ctor(0, 2, 0); +} else { + x_1046 = x_1045; +} +lean_ctor_set(x_1046, 0, x_1043); +lean_ctor_set(x_1046, 1, x_1044); +return x_1046; +} +else +{ +lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; uint8_t x_1051; +x_1047 = lean_ctor_get(x_1042, 0); +lean_inc(x_1047); +x_1048 = lean_ctor_get(x_1042, 1); +lean_inc(x_1048); +if (lean_is_exclusive(x_1042)) { + lean_ctor_release(x_1042, 0); + lean_ctor_release(x_1042, 1); + x_1049 = x_1042; +} else { + lean_dec_ref(x_1042); + x_1049 = lean_box(0); +} +x_1050 = lean_ctor_get(x_1047, 1); +lean_inc(x_1050); +x_1051 = lean_nat_dec_eq(x_1036, x_1050); +lean_dec(x_1036); +if (x_1051 == 0) +{ +lean_object* x_1052; +lean_dec(x_1050); +if (lean_is_scalar(x_1049)) { + x_1052 = lean_alloc_ctor(1, 2, 0); +} else { + x_1052 = x_1049; +} +lean_ctor_set(x_1052, 0, x_1047); +lean_ctor_set(x_1052, 1, x_1048); +return x_1052; +} +else +{ +lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; +lean_dec(x_1049); +lean_dec(x_1048); +x_1053 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1054 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1054, 0, x_1053); +x_1055 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1056 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1055, x_1054, x_1047); +if (lean_obj_tag(x_1056) == 0) +{ +lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; +lean_dec(x_1050); +x_1057 = lean_ctor_get(x_1056, 0); +lean_inc(x_1057); +x_1058 = lean_ctor_get(x_1056, 1); +lean_inc(x_1058); +if (lean_is_exclusive(x_1056)) { + lean_ctor_release(x_1056, 0); + lean_ctor_release(x_1056, 1); + x_1059 = x_1056; +} else { + lean_dec_ref(x_1056); + x_1059 = lean_box(0); +} +if (lean_is_scalar(x_1059)) { + x_1060 = lean_alloc_ctor(0, 2, 0); +} else { + x_1060 = x_1059; +} +lean_ctor_set(x_1060, 0, x_1057); +lean_ctor_set(x_1060, 1, x_1058); +return x_1060; +} +else +{ +lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; uint8_t x_1065; +x_1061 = lean_ctor_get(x_1056, 0); +lean_inc(x_1061); +x_1062 = lean_ctor_get(x_1056, 1); +lean_inc(x_1062); +if (lean_is_exclusive(x_1056)) { + lean_ctor_release(x_1056, 0); + lean_ctor_release(x_1056, 1); + x_1063 = x_1056; +} else { + lean_dec_ref(x_1056); + x_1063 = lean_box(0); +} +x_1064 = lean_ctor_get(x_1061, 1); +lean_inc(x_1064); +x_1065 = lean_nat_dec_eq(x_1050, x_1064); +lean_dec(x_1050); +if (x_1065 == 0) +{ +lean_object* x_1066; +lean_dec(x_1064); +if (lean_is_scalar(x_1063)) { + x_1066 = lean_alloc_ctor(1, 2, 0); +} else { + x_1066 = x_1063; +} +lean_ctor_set(x_1066, 0, x_1061); +lean_ctor_set(x_1066, 1, x_1062); +return x_1066; +} +else +{ +lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; +lean_dec(x_1063); +lean_dec(x_1062); +x_1067 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1068 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1068, 0, x_1067); +x_1069 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1070 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1069, x_1068, x_1061); +if (lean_obj_tag(x_1070) == 0) +{ +lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; +lean_dec(x_1064); +x_1071 = lean_ctor_get(x_1070, 0); +lean_inc(x_1071); +x_1072 = lean_ctor_get(x_1070, 1); +lean_inc(x_1072); +if (lean_is_exclusive(x_1070)) { + lean_ctor_release(x_1070, 0); + lean_ctor_release(x_1070, 1); + x_1073 = x_1070; +} else { + lean_dec_ref(x_1070); + x_1073 = lean_box(0); +} +if (lean_is_scalar(x_1073)) { + x_1074 = lean_alloc_ctor(0, 2, 0); +} else { + x_1074 = x_1073; +} +lean_ctor_set(x_1074, 0, x_1071); +lean_ctor_set(x_1074, 1, x_1072); +return x_1074; +} +else +{ +lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; uint8_t x_1079; +x_1075 = lean_ctor_get(x_1070, 0); +lean_inc(x_1075); +x_1076 = lean_ctor_get(x_1070, 1); +lean_inc(x_1076); +if (lean_is_exclusive(x_1070)) { + lean_ctor_release(x_1070, 0); + lean_ctor_release(x_1070, 1); + x_1077 = x_1070; +} else { + lean_dec_ref(x_1070); + x_1077 = lean_box(0); +} +x_1078 = lean_ctor_get(x_1075, 1); +lean_inc(x_1078); +x_1079 = lean_nat_dec_eq(x_1064, x_1078); +lean_dec(x_1064); +if (x_1079 == 0) +{ +lean_object* x_1080; +lean_dec(x_1078); +if (lean_is_scalar(x_1077)) { + x_1080 = lean_alloc_ctor(1, 2, 0); +} else { + x_1080 = x_1077; +} +lean_ctor_set(x_1080, 0, x_1075); +lean_ctor_set(x_1080, 1, x_1076); +return x_1080; +} +else +{ +lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; +lean_dec(x_1077); +lean_dec(x_1076); +x_1081 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1082 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1082, 0, x_1081); +x_1083 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1084 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1083, x_1082, x_1075); +if (lean_obj_tag(x_1084) == 0) +{ +lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; +lean_dec(x_1078); +x_1085 = lean_ctor_get(x_1084, 0); +lean_inc(x_1085); +x_1086 = lean_ctor_get(x_1084, 1); +lean_inc(x_1086); +if (lean_is_exclusive(x_1084)) { + lean_ctor_release(x_1084, 0); + lean_ctor_release(x_1084, 1); + x_1087 = x_1084; +} else { + lean_dec_ref(x_1084); + x_1087 = lean_box(0); +} +if (lean_is_scalar(x_1087)) { + x_1088 = lean_alloc_ctor(0, 2, 0); +} else { + x_1088 = x_1087; +} +lean_ctor_set(x_1088, 0, x_1085); +lean_ctor_set(x_1088, 1, x_1086); +return x_1088; +} +else +{ +lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; uint8_t x_1093; +x_1089 = lean_ctor_get(x_1084, 0); +lean_inc(x_1089); +x_1090 = lean_ctor_get(x_1084, 1); +lean_inc(x_1090); +if (lean_is_exclusive(x_1084)) { + lean_ctor_release(x_1084, 0); + lean_ctor_release(x_1084, 1); + x_1091 = x_1084; +} else { + lean_dec_ref(x_1084); + x_1091 = lean_box(0); +} +x_1092 = lean_ctor_get(x_1089, 1); +lean_inc(x_1092); +x_1093 = lean_nat_dec_eq(x_1078, x_1092); +lean_dec(x_1078); +if (x_1093 == 0) +{ +lean_object* x_1094; +lean_dec(x_1092); +if (lean_is_scalar(x_1091)) { + x_1094 = lean_alloc_ctor(1, 2, 0); +} else { + x_1094 = x_1091; +} +lean_ctor_set(x_1094, 0, x_1089); +lean_ctor_set(x_1094, 1, x_1090); +return x_1094; +} +else +{ +lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; +lean_dec(x_1091); +lean_dec(x_1090); +x_1095 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1096 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1096, 0, x_1095); +x_1097 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1098 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1097, x_1096, x_1089); +if (lean_obj_tag(x_1098) == 0) +{ +lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; +lean_dec(x_1092); +x_1099 = lean_ctor_get(x_1098, 0); +lean_inc(x_1099); +x_1100 = lean_ctor_get(x_1098, 1); +lean_inc(x_1100); +if (lean_is_exclusive(x_1098)) { + lean_ctor_release(x_1098, 0); + lean_ctor_release(x_1098, 1); + x_1101 = x_1098; +} else { + lean_dec_ref(x_1098); + x_1101 = lean_box(0); +} +if (lean_is_scalar(x_1101)) { + x_1102 = lean_alloc_ctor(0, 2, 0); +} else { + x_1102 = x_1101; +} +lean_ctor_set(x_1102, 0, x_1099); +lean_ctor_set(x_1102, 1, x_1100); +return x_1102; +} +else +{ +lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; uint8_t x_1107; +x_1103 = lean_ctor_get(x_1098, 0); +lean_inc(x_1103); +x_1104 = lean_ctor_get(x_1098, 1); +lean_inc(x_1104); +if (lean_is_exclusive(x_1098)) { + lean_ctor_release(x_1098, 0); + lean_ctor_release(x_1098, 1); + x_1105 = x_1098; +} else { + lean_dec_ref(x_1098); + x_1105 = lean_box(0); +} +x_1106 = lean_ctor_get(x_1103, 1); +lean_inc(x_1106); +x_1107 = lean_nat_dec_eq(x_1092, x_1106); +lean_dec(x_1092); +if (x_1107 == 0) +{ +lean_object* x_1108; +lean_dec(x_1106); +if (lean_is_scalar(x_1105)) { + x_1108 = lean_alloc_ctor(1, 2, 0); +} else { + x_1108 = x_1105; +} +lean_ctor_set(x_1108, 0, x_1103); +lean_ctor_set(x_1108, 1, x_1104); +return x_1108; +} +else +{ +lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; +lean_dec(x_1105); +lean_dec(x_1104); +x_1109 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1110 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1110, 0, x_1109); +x_1111 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1112 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1111, x_1110, x_1103); +if (lean_obj_tag(x_1112) == 0) +{ +lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; +lean_dec(x_1106); +x_1113 = lean_ctor_get(x_1112, 0); +lean_inc(x_1113); +x_1114 = lean_ctor_get(x_1112, 1); +lean_inc(x_1114); +if (lean_is_exclusive(x_1112)) { + lean_ctor_release(x_1112, 0); + lean_ctor_release(x_1112, 1); + x_1115 = x_1112; +} else { + lean_dec_ref(x_1112); + x_1115 = lean_box(0); +} +if (lean_is_scalar(x_1115)) { + x_1116 = lean_alloc_ctor(0, 2, 0); +} else { + x_1116 = x_1115; +} +lean_ctor_set(x_1116, 0, x_1113); +lean_ctor_set(x_1116, 1, x_1114); +return x_1116; +} +else +{ +lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; uint8_t x_1121; +x_1117 = lean_ctor_get(x_1112, 0); +lean_inc(x_1117); +x_1118 = lean_ctor_get(x_1112, 1); +lean_inc(x_1118); +if (lean_is_exclusive(x_1112)) { + lean_ctor_release(x_1112, 0); + lean_ctor_release(x_1112, 1); + x_1119 = x_1112; +} else { + lean_dec_ref(x_1112); + x_1119 = lean_box(0); +} +x_1120 = lean_ctor_get(x_1117, 1); +lean_inc(x_1120); +x_1121 = lean_nat_dec_eq(x_1106, x_1120); +lean_dec(x_1120); +lean_dec(x_1106); +if (x_1121 == 0) +{ +lean_object* x_1122; +if (lean_is_scalar(x_1119)) { + x_1122 = lean_alloc_ctor(1, 2, 0); +} else { + x_1122 = x_1119; +} +lean_ctor_set(x_1122, 0, x_1117); +lean_ctor_set(x_1122, 1, x_1118); +return x_1122; +} +else +{ +lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; +lean_dec(x_1119); +lean_dec(x_1118); +x_1123 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_1124 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1124, 0, x_1123); +x_1125 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_1126 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1125, x_1124, x_1117); +return x_1126; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; uint8_t x_1130; +x_1127 = lean_ctor_get(x_263, 0); +x_1128 = lean_ctor_get(x_263, 1); +lean_inc(x_1128); +lean_inc(x_1127); +lean_dec(x_263); +x_1129 = lean_ctor_get(x_1127, 1); +lean_inc(x_1129); +x_1130 = lean_nat_dec_eq(x_258, x_1129); +lean_dec(x_258); +if (x_1130 == 0) +{ +lean_object* x_1131; +lean_dec(x_1129); +x_1131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1131, 0, x_1127); +lean_ctor_set(x_1131, 1, x_1128); +return x_1131; +} +else +{ +lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; +lean_dec(x_1128); +x_1132 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_1133 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1133, 0, x_1132); +x_1134 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_1135 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1134, x_1133, x_1127); +if (lean_obj_tag(x_1135) == 0) +{ +lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; +lean_dec(x_1129); +x_1136 = lean_ctor_get(x_1135, 0); +lean_inc(x_1136); +x_1137 = lean_ctor_get(x_1135, 1); +lean_inc(x_1137); +if (lean_is_exclusive(x_1135)) { + lean_ctor_release(x_1135, 0); + lean_ctor_release(x_1135, 1); + x_1138 = x_1135; +} else { + lean_dec_ref(x_1135); + x_1138 = lean_box(0); +} +if (lean_is_scalar(x_1138)) { + x_1139 = lean_alloc_ctor(0, 2, 0); +} else { + x_1139 = x_1138; +} +lean_ctor_set(x_1139, 0, x_1136); +lean_ctor_set(x_1139, 1, x_1137); +return x_1139; +} +else +{ +lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; uint8_t x_1144; +x_1140 = lean_ctor_get(x_1135, 0); +lean_inc(x_1140); +x_1141 = lean_ctor_get(x_1135, 1); +lean_inc(x_1141); +if (lean_is_exclusive(x_1135)) { + lean_ctor_release(x_1135, 0); + lean_ctor_release(x_1135, 1); + x_1142 = x_1135; +} else { + lean_dec_ref(x_1135); + x_1142 = lean_box(0); +} +x_1143 = lean_ctor_get(x_1140, 1); +lean_inc(x_1143); +x_1144 = lean_nat_dec_eq(x_1129, x_1143); +lean_dec(x_1129); +if (x_1144 == 0) +{ +lean_object* x_1145; +lean_dec(x_1143); +if (lean_is_scalar(x_1142)) { + x_1145 = lean_alloc_ctor(1, 2, 0); +} else { + x_1145 = x_1142; +} +lean_ctor_set(x_1145, 0, x_1140); +lean_ctor_set(x_1145, 1, x_1141); +return x_1145; +} +else +{ +lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; +lean_dec(x_1142); +lean_dec(x_1141); +x_1146 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_1147 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1147, 0, x_1146); +x_1148 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1149 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1148, x_1147, x_1140); +if (lean_obj_tag(x_1149) == 0) +{ +lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; +lean_dec(x_1143); +x_1150 = lean_ctor_get(x_1149, 0); +lean_inc(x_1150); +x_1151 = lean_ctor_get(x_1149, 1); +lean_inc(x_1151); +if (lean_is_exclusive(x_1149)) { + lean_ctor_release(x_1149, 0); + lean_ctor_release(x_1149, 1); + x_1152 = x_1149; +} else { + lean_dec_ref(x_1149); + x_1152 = lean_box(0); +} +if (lean_is_scalar(x_1152)) { + x_1153 = lean_alloc_ctor(0, 2, 0); +} else { + x_1153 = x_1152; +} +lean_ctor_set(x_1153, 0, x_1150); +lean_ctor_set(x_1153, 1, x_1151); +return x_1153; +} +else +{ +lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; uint8_t x_1158; +x_1154 = lean_ctor_get(x_1149, 0); +lean_inc(x_1154); +x_1155 = lean_ctor_get(x_1149, 1); +lean_inc(x_1155); +if (lean_is_exclusive(x_1149)) { + lean_ctor_release(x_1149, 0); + lean_ctor_release(x_1149, 1); + x_1156 = x_1149; +} else { + lean_dec_ref(x_1149); + x_1156 = lean_box(0); +} +x_1157 = lean_ctor_get(x_1154, 1); +lean_inc(x_1157); +x_1158 = lean_nat_dec_eq(x_1143, x_1157); +lean_dec(x_1143); +if (x_1158 == 0) +{ +lean_object* x_1159; +lean_dec(x_1157); +if (lean_is_scalar(x_1156)) { + x_1159 = lean_alloc_ctor(1, 2, 0); +} else { + x_1159 = x_1156; +} +lean_ctor_set(x_1159, 0, x_1154); +lean_ctor_set(x_1159, 1, x_1155); +return x_1159; +} +else +{ +lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; +lean_dec(x_1156); +lean_dec(x_1155); +x_1160 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1161 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1161, 0, x_1160); +x_1162 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1163 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1162, x_1161, x_1154); +if (lean_obj_tag(x_1163) == 0) +{ +lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; +lean_dec(x_1157); +x_1164 = lean_ctor_get(x_1163, 0); +lean_inc(x_1164); +x_1165 = lean_ctor_get(x_1163, 1); +lean_inc(x_1165); +if (lean_is_exclusive(x_1163)) { + lean_ctor_release(x_1163, 0); + lean_ctor_release(x_1163, 1); + x_1166 = x_1163; +} else { + lean_dec_ref(x_1163); + x_1166 = lean_box(0); +} +if (lean_is_scalar(x_1166)) { + x_1167 = lean_alloc_ctor(0, 2, 0); +} else { + x_1167 = x_1166; +} +lean_ctor_set(x_1167, 0, x_1164); +lean_ctor_set(x_1167, 1, x_1165); +return x_1167; +} +else +{ +lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; uint8_t x_1172; +x_1168 = lean_ctor_get(x_1163, 0); +lean_inc(x_1168); +x_1169 = lean_ctor_get(x_1163, 1); +lean_inc(x_1169); +if (lean_is_exclusive(x_1163)) { + lean_ctor_release(x_1163, 0); + lean_ctor_release(x_1163, 1); + x_1170 = x_1163; +} else { + lean_dec_ref(x_1163); + x_1170 = lean_box(0); +} +x_1171 = lean_ctor_get(x_1168, 1); +lean_inc(x_1171); +x_1172 = lean_nat_dec_eq(x_1157, x_1171); +lean_dec(x_1157); +if (x_1172 == 0) +{ +lean_object* x_1173; +lean_dec(x_1171); +if (lean_is_scalar(x_1170)) { + x_1173 = lean_alloc_ctor(1, 2, 0); +} else { + x_1173 = x_1170; +} +lean_ctor_set(x_1173, 0, x_1168); +lean_ctor_set(x_1173, 1, x_1169); +return x_1173; +} +else +{ +lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; +lean_dec(x_1170); +lean_dec(x_1169); +x_1174 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1175 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1175, 0, x_1174); +x_1176 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1177 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1176, x_1175, x_1168); +if (lean_obj_tag(x_1177) == 0) +{ +lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; +lean_dec(x_1171); +x_1178 = lean_ctor_get(x_1177, 0); +lean_inc(x_1178); +x_1179 = lean_ctor_get(x_1177, 1); +lean_inc(x_1179); +if (lean_is_exclusive(x_1177)) { + lean_ctor_release(x_1177, 0); + lean_ctor_release(x_1177, 1); + x_1180 = x_1177; +} else { + lean_dec_ref(x_1177); + x_1180 = lean_box(0); +} +if (lean_is_scalar(x_1180)) { + x_1181 = lean_alloc_ctor(0, 2, 0); +} else { + x_1181 = x_1180; +} +lean_ctor_set(x_1181, 0, x_1178); +lean_ctor_set(x_1181, 1, x_1179); +return x_1181; +} +else +{ +lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; uint8_t x_1186; +x_1182 = lean_ctor_get(x_1177, 0); +lean_inc(x_1182); +x_1183 = lean_ctor_get(x_1177, 1); +lean_inc(x_1183); +if (lean_is_exclusive(x_1177)) { + lean_ctor_release(x_1177, 0); + lean_ctor_release(x_1177, 1); + x_1184 = x_1177; +} else { + lean_dec_ref(x_1177); + x_1184 = lean_box(0); +} +x_1185 = lean_ctor_get(x_1182, 1); +lean_inc(x_1185); +x_1186 = lean_nat_dec_eq(x_1171, x_1185); +lean_dec(x_1171); +if (x_1186 == 0) +{ +lean_object* x_1187; +lean_dec(x_1185); +if (lean_is_scalar(x_1184)) { + x_1187 = lean_alloc_ctor(1, 2, 0); +} else { + x_1187 = x_1184; +} +lean_ctor_set(x_1187, 0, x_1182); +lean_ctor_set(x_1187, 1, x_1183); +return x_1187; +} +else +{ +lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; +lean_dec(x_1184); +lean_dec(x_1183); +x_1188 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1189 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1189, 0, x_1188); +x_1190 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1191 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1190, x_1189, x_1182); +if (lean_obj_tag(x_1191) == 0) +{ +lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; +lean_dec(x_1185); +x_1192 = lean_ctor_get(x_1191, 0); +lean_inc(x_1192); +x_1193 = lean_ctor_get(x_1191, 1); +lean_inc(x_1193); +if (lean_is_exclusive(x_1191)) { + lean_ctor_release(x_1191, 0); + lean_ctor_release(x_1191, 1); + x_1194 = x_1191; +} else { + lean_dec_ref(x_1191); + x_1194 = lean_box(0); +} +if (lean_is_scalar(x_1194)) { + x_1195 = lean_alloc_ctor(0, 2, 0); +} else { + x_1195 = x_1194; +} +lean_ctor_set(x_1195, 0, x_1192); +lean_ctor_set(x_1195, 1, x_1193); +return x_1195; +} +else +{ +lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; uint8_t x_1200; +x_1196 = lean_ctor_get(x_1191, 0); +lean_inc(x_1196); +x_1197 = lean_ctor_get(x_1191, 1); +lean_inc(x_1197); +if (lean_is_exclusive(x_1191)) { + lean_ctor_release(x_1191, 0); + lean_ctor_release(x_1191, 1); + x_1198 = x_1191; +} else { + lean_dec_ref(x_1191); + x_1198 = lean_box(0); +} +x_1199 = lean_ctor_get(x_1196, 1); +lean_inc(x_1199); +x_1200 = lean_nat_dec_eq(x_1185, x_1199); +lean_dec(x_1185); +if (x_1200 == 0) +{ +lean_object* x_1201; +lean_dec(x_1199); +if (lean_is_scalar(x_1198)) { + x_1201 = lean_alloc_ctor(1, 2, 0); +} else { + x_1201 = x_1198; +} +lean_ctor_set(x_1201, 0, x_1196); +lean_ctor_set(x_1201, 1, x_1197); +return x_1201; +} +else +{ +lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; +lean_dec(x_1198); +lean_dec(x_1197); +x_1202 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1203 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1203, 0, x_1202); +x_1204 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1205 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1204, x_1203, x_1196); +if (lean_obj_tag(x_1205) == 0) +{ +lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; +lean_dec(x_1199); +x_1206 = lean_ctor_get(x_1205, 0); +lean_inc(x_1206); +x_1207 = lean_ctor_get(x_1205, 1); +lean_inc(x_1207); +if (lean_is_exclusive(x_1205)) { + lean_ctor_release(x_1205, 0); + lean_ctor_release(x_1205, 1); + x_1208 = x_1205; +} else { + lean_dec_ref(x_1205); + x_1208 = lean_box(0); +} +if (lean_is_scalar(x_1208)) { + x_1209 = lean_alloc_ctor(0, 2, 0); +} else { + x_1209 = x_1208; +} +lean_ctor_set(x_1209, 0, x_1206); +lean_ctor_set(x_1209, 1, x_1207); +return x_1209; +} +else +{ +lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; uint8_t x_1214; +x_1210 = lean_ctor_get(x_1205, 0); +lean_inc(x_1210); +x_1211 = lean_ctor_get(x_1205, 1); +lean_inc(x_1211); +if (lean_is_exclusive(x_1205)) { + lean_ctor_release(x_1205, 0); + lean_ctor_release(x_1205, 1); + x_1212 = x_1205; +} else { + lean_dec_ref(x_1205); + x_1212 = lean_box(0); +} +x_1213 = lean_ctor_get(x_1210, 1); +lean_inc(x_1213); +x_1214 = lean_nat_dec_eq(x_1199, x_1213); +lean_dec(x_1199); +if (x_1214 == 0) +{ +lean_object* x_1215; +lean_dec(x_1213); +if (lean_is_scalar(x_1212)) { + x_1215 = lean_alloc_ctor(1, 2, 0); +} else { + x_1215 = x_1212; +} +lean_ctor_set(x_1215, 0, x_1210); +lean_ctor_set(x_1215, 1, x_1211); +return x_1215; +} +else +{ +lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; +lean_dec(x_1212); +lean_dec(x_1211); +x_1216 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1217 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1217, 0, x_1216); +x_1218 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1219 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1218, x_1217, x_1210); +if (lean_obj_tag(x_1219) == 0) +{ +lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; +lean_dec(x_1213); +x_1220 = lean_ctor_get(x_1219, 0); +lean_inc(x_1220); +x_1221 = lean_ctor_get(x_1219, 1); +lean_inc(x_1221); +if (lean_is_exclusive(x_1219)) { + lean_ctor_release(x_1219, 0); + lean_ctor_release(x_1219, 1); + x_1222 = x_1219; +} else { + lean_dec_ref(x_1219); + x_1222 = lean_box(0); +} +if (lean_is_scalar(x_1222)) { + x_1223 = lean_alloc_ctor(0, 2, 0); +} else { + x_1223 = x_1222; +} +lean_ctor_set(x_1223, 0, x_1220); +lean_ctor_set(x_1223, 1, x_1221); +return x_1223; +} +else +{ +lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; lean_object* x_1227; uint8_t x_1228; +x_1224 = lean_ctor_get(x_1219, 0); +lean_inc(x_1224); +x_1225 = lean_ctor_get(x_1219, 1); +lean_inc(x_1225); +if (lean_is_exclusive(x_1219)) { + lean_ctor_release(x_1219, 0); + lean_ctor_release(x_1219, 1); + x_1226 = x_1219; +} else { + lean_dec_ref(x_1219); + x_1226 = lean_box(0); +} +x_1227 = lean_ctor_get(x_1224, 1); +lean_inc(x_1227); +x_1228 = lean_nat_dec_eq(x_1213, x_1227); +lean_dec(x_1213); +if (x_1228 == 0) +{ +lean_object* x_1229; +lean_dec(x_1227); +if (lean_is_scalar(x_1226)) { + x_1229 = lean_alloc_ctor(1, 2, 0); +} else { + x_1229 = x_1226; +} +lean_ctor_set(x_1229, 0, x_1224); +lean_ctor_set(x_1229, 1, x_1225); +return x_1229; +} +else +{ +lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; +lean_dec(x_1226); +lean_dec(x_1225); +x_1230 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1231 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1231, 0, x_1230); +x_1232 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1233 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1232, x_1231, x_1224); +if (lean_obj_tag(x_1233) == 0) +{ +lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; +lean_dec(x_1227); +x_1234 = lean_ctor_get(x_1233, 0); +lean_inc(x_1234); +x_1235 = lean_ctor_get(x_1233, 1); +lean_inc(x_1235); +if (lean_is_exclusive(x_1233)) { + lean_ctor_release(x_1233, 0); + lean_ctor_release(x_1233, 1); + x_1236 = x_1233; +} else { + lean_dec_ref(x_1233); + x_1236 = lean_box(0); +} +if (lean_is_scalar(x_1236)) { + x_1237 = lean_alloc_ctor(0, 2, 0); +} else { + x_1237 = x_1236; +} +lean_ctor_set(x_1237, 0, x_1234); +lean_ctor_set(x_1237, 1, x_1235); +return x_1237; +} +else +{ +lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; uint8_t x_1242; +x_1238 = lean_ctor_get(x_1233, 0); +lean_inc(x_1238); +x_1239 = lean_ctor_get(x_1233, 1); +lean_inc(x_1239); +if (lean_is_exclusive(x_1233)) { + lean_ctor_release(x_1233, 0); + lean_ctor_release(x_1233, 1); + x_1240 = x_1233; +} else { + lean_dec_ref(x_1233); + x_1240 = lean_box(0); +} +x_1241 = lean_ctor_get(x_1238, 1); +lean_inc(x_1241); +x_1242 = lean_nat_dec_eq(x_1227, x_1241); +lean_dec(x_1227); +if (x_1242 == 0) +{ +lean_object* x_1243; +lean_dec(x_1241); +if (lean_is_scalar(x_1240)) { + x_1243 = lean_alloc_ctor(1, 2, 0); +} else { + x_1243 = x_1240; +} +lean_ctor_set(x_1243, 0, x_1238); +lean_ctor_set(x_1243, 1, x_1239); +return x_1243; +} +else +{ +lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; +lean_dec(x_1240); +lean_dec(x_1239); +x_1244 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1245 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1245, 0, x_1244); +x_1246 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1247 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1246, x_1245, x_1238); +if (lean_obj_tag(x_1247) == 0) +{ +lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; +lean_dec(x_1241); +x_1248 = lean_ctor_get(x_1247, 0); +lean_inc(x_1248); +x_1249 = lean_ctor_get(x_1247, 1); +lean_inc(x_1249); +if (lean_is_exclusive(x_1247)) { + lean_ctor_release(x_1247, 0); + lean_ctor_release(x_1247, 1); + x_1250 = x_1247; +} else { + lean_dec_ref(x_1247); + x_1250 = lean_box(0); +} +if (lean_is_scalar(x_1250)) { + x_1251 = lean_alloc_ctor(0, 2, 0); +} else { + x_1251 = x_1250; +} +lean_ctor_set(x_1251, 0, x_1248); +lean_ctor_set(x_1251, 1, x_1249); +return x_1251; +} +else +{ +lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; uint8_t x_1256; +x_1252 = lean_ctor_get(x_1247, 0); +lean_inc(x_1252); +x_1253 = lean_ctor_get(x_1247, 1); +lean_inc(x_1253); +if (lean_is_exclusive(x_1247)) { + lean_ctor_release(x_1247, 0); + lean_ctor_release(x_1247, 1); + x_1254 = x_1247; +} else { + lean_dec_ref(x_1247); + x_1254 = lean_box(0); +} +x_1255 = lean_ctor_get(x_1252, 1); +lean_inc(x_1255); +x_1256 = lean_nat_dec_eq(x_1241, x_1255); +lean_dec(x_1241); +if (x_1256 == 0) +{ +lean_object* x_1257; +lean_dec(x_1255); +if (lean_is_scalar(x_1254)) { + x_1257 = lean_alloc_ctor(1, 2, 0); +} else { + x_1257 = x_1254; +} +lean_ctor_set(x_1257, 0, x_1252); +lean_ctor_set(x_1257, 1, x_1253); +return x_1257; +} +else +{ +lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; lean_object* x_1261; +lean_dec(x_1254); +lean_dec(x_1253); +x_1258 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1259 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1259, 0, x_1258); +x_1260 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1261 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1260, x_1259, x_1252); +if (lean_obj_tag(x_1261) == 0) +{ +lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; +lean_dec(x_1255); +x_1262 = lean_ctor_get(x_1261, 0); +lean_inc(x_1262); +x_1263 = lean_ctor_get(x_1261, 1); +lean_inc(x_1263); +if (lean_is_exclusive(x_1261)) { + lean_ctor_release(x_1261, 0); + lean_ctor_release(x_1261, 1); + x_1264 = x_1261; +} else { + lean_dec_ref(x_1261); + x_1264 = lean_box(0); +} +if (lean_is_scalar(x_1264)) { + x_1265 = lean_alloc_ctor(0, 2, 0); +} else { + x_1265 = x_1264; +} +lean_ctor_set(x_1265, 0, x_1262); +lean_ctor_set(x_1265, 1, x_1263); +return x_1265; +} +else +{ +lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; uint8_t x_1270; +x_1266 = lean_ctor_get(x_1261, 0); +lean_inc(x_1266); +x_1267 = lean_ctor_get(x_1261, 1); +lean_inc(x_1267); +if (lean_is_exclusive(x_1261)) { + lean_ctor_release(x_1261, 0); + lean_ctor_release(x_1261, 1); + x_1268 = x_1261; +} else { + lean_dec_ref(x_1261); + x_1268 = lean_box(0); +} +x_1269 = lean_ctor_get(x_1266, 1); +lean_inc(x_1269); +x_1270 = lean_nat_dec_eq(x_1255, x_1269); +lean_dec(x_1269); +lean_dec(x_1255); +if (x_1270 == 0) +{ +lean_object* x_1271; +if (lean_is_scalar(x_1268)) { + x_1271 = lean_alloc_ctor(1, 2, 0); +} else { + x_1271 = x_1268; +} +lean_ctor_set(x_1271, 0, x_1266); +lean_ctor_set(x_1271, 1, x_1267); +return x_1271; +} +else +{ +lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; +lean_dec(x_1268); +lean_dec(x_1267); +x_1272 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_1273 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1273, 0, x_1272); +x_1274 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_1275 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1274, x_1273, x_1266); +return x_1275; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; uint8_t x_1279; +x_1276 = lean_ctor_get(x_250, 0); +x_1277 = lean_ctor_get(x_250, 1); +lean_inc(x_1277); +lean_inc(x_1276); +lean_dec(x_250); +x_1278 = lean_ctor_get(x_1276, 1); +lean_inc(x_1278); +x_1279 = lean_nat_dec_eq(x_245, x_1278); +lean_dec(x_245); +if (x_1279 == 0) +{ +lean_object* x_1280; +lean_dec(x_1278); +x_1280 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1280, 0, x_1276); +lean_ctor_set(x_1280, 1, x_1277); +return x_1280; +} +else +{ +lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; +lean_dec(x_1277); +x_1281 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_1282 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1282, 0, x_1281); +x_1283 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_1284 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1283, x_1282, x_1276); +if (lean_obj_tag(x_1284) == 0) +{ +lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; +lean_dec(x_1278); +x_1285 = lean_ctor_get(x_1284, 0); +lean_inc(x_1285); +x_1286 = lean_ctor_get(x_1284, 1); +lean_inc(x_1286); +if (lean_is_exclusive(x_1284)) { + lean_ctor_release(x_1284, 0); + lean_ctor_release(x_1284, 1); + x_1287 = x_1284; +} else { + lean_dec_ref(x_1284); + x_1287 = lean_box(0); +} +if (lean_is_scalar(x_1287)) { + x_1288 = lean_alloc_ctor(0, 2, 0); +} else { + x_1288 = x_1287; +} +lean_ctor_set(x_1288, 0, x_1285); +lean_ctor_set(x_1288, 1, x_1286); +return x_1288; +} +else +{ +lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; uint8_t x_1293; +x_1289 = lean_ctor_get(x_1284, 0); +lean_inc(x_1289); +x_1290 = lean_ctor_get(x_1284, 1); +lean_inc(x_1290); +if (lean_is_exclusive(x_1284)) { + lean_ctor_release(x_1284, 0); + lean_ctor_release(x_1284, 1); + x_1291 = x_1284; +} else { + lean_dec_ref(x_1284); + x_1291 = lean_box(0); +} +x_1292 = lean_ctor_get(x_1289, 1); +lean_inc(x_1292); +x_1293 = lean_nat_dec_eq(x_1278, x_1292); +lean_dec(x_1278); +if (x_1293 == 0) +{ +lean_object* x_1294; +lean_dec(x_1292); +if (lean_is_scalar(x_1291)) { + x_1294 = lean_alloc_ctor(1, 2, 0); +} else { + x_1294 = x_1291; +} +lean_ctor_set(x_1294, 0, x_1289); +lean_ctor_set(x_1294, 1, x_1290); +return x_1294; +} +else +{ +lean_object* x_1295; lean_object* x_1296; lean_object* x_1297; lean_object* x_1298; +lean_dec(x_1291); +lean_dec(x_1290); +x_1295 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_1296 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1296, 0, x_1295); +x_1297 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_1298 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1297, x_1296, x_1289); +if (lean_obj_tag(x_1298) == 0) +{ +lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; +lean_dec(x_1292); +x_1299 = lean_ctor_get(x_1298, 0); +lean_inc(x_1299); +x_1300 = lean_ctor_get(x_1298, 1); +lean_inc(x_1300); +if (lean_is_exclusive(x_1298)) { + lean_ctor_release(x_1298, 0); + lean_ctor_release(x_1298, 1); + x_1301 = x_1298; +} else { + lean_dec_ref(x_1298); + x_1301 = lean_box(0); +} +if (lean_is_scalar(x_1301)) { + x_1302 = lean_alloc_ctor(0, 2, 0); +} else { + x_1302 = x_1301; +} +lean_ctor_set(x_1302, 0, x_1299); +lean_ctor_set(x_1302, 1, x_1300); +return x_1302; +} +else +{ +lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; uint8_t x_1307; +x_1303 = lean_ctor_get(x_1298, 0); +lean_inc(x_1303); +x_1304 = lean_ctor_get(x_1298, 1); +lean_inc(x_1304); +if (lean_is_exclusive(x_1298)) { + lean_ctor_release(x_1298, 0); + lean_ctor_release(x_1298, 1); + x_1305 = x_1298; +} else { + lean_dec_ref(x_1298); + x_1305 = lean_box(0); +} +x_1306 = lean_ctor_get(x_1303, 1); +lean_inc(x_1306); +x_1307 = lean_nat_dec_eq(x_1292, x_1306); +lean_dec(x_1292); +if (x_1307 == 0) +{ +lean_object* x_1308; +lean_dec(x_1306); +if (lean_is_scalar(x_1305)) { + x_1308 = lean_alloc_ctor(1, 2, 0); +} else { + x_1308 = x_1305; +} +lean_ctor_set(x_1308, 0, x_1303); +lean_ctor_set(x_1308, 1, x_1304); +return x_1308; +} +else +{ +lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; +lean_dec(x_1305); +lean_dec(x_1304); +x_1309 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_1310 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1310, 0, x_1309); +x_1311 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1312 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1311, x_1310, x_1303); +if (lean_obj_tag(x_1312) == 0) +{ +lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; +lean_dec(x_1306); +x_1313 = lean_ctor_get(x_1312, 0); +lean_inc(x_1313); +x_1314 = lean_ctor_get(x_1312, 1); +lean_inc(x_1314); +if (lean_is_exclusive(x_1312)) { + lean_ctor_release(x_1312, 0); + lean_ctor_release(x_1312, 1); + x_1315 = x_1312; +} else { + lean_dec_ref(x_1312); + x_1315 = lean_box(0); +} +if (lean_is_scalar(x_1315)) { + x_1316 = lean_alloc_ctor(0, 2, 0); +} else { + x_1316 = x_1315; +} +lean_ctor_set(x_1316, 0, x_1313); +lean_ctor_set(x_1316, 1, x_1314); +return x_1316; +} +else +{ +lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; uint8_t x_1321; +x_1317 = lean_ctor_get(x_1312, 0); +lean_inc(x_1317); +x_1318 = lean_ctor_get(x_1312, 1); +lean_inc(x_1318); +if (lean_is_exclusive(x_1312)) { + lean_ctor_release(x_1312, 0); + lean_ctor_release(x_1312, 1); + x_1319 = x_1312; +} else { + lean_dec_ref(x_1312); + x_1319 = lean_box(0); +} +x_1320 = lean_ctor_get(x_1317, 1); +lean_inc(x_1320); +x_1321 = lean_nat_dec_eq(x_1306, x_1320); +lean_dec(x_1306); +if (x_1321 == 0) +{ +lean_object* x_1322; +lean_dec(x_1320); +if (lean_is_scalar(x_1319)) { + x_1322 = lean_alloc_ctor(1, 2, 0); +} else { + x_1322 = x_1319; +} +lean_ctor_set(x_1322, 0, x_1317); +lean_ctor_set(x_1322, 1, x_1318); +return x_1322; +} +else +{ +lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; +lean_dec(x_1319); +lean_dec(x_1318); +x_1323 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1324 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1324, 0, x_1323); +x_1325 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1326 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1325, x_1324, x_1317); +if (lean_obj_tag(x_1326) == 0) +{ +lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; +lean_dec(x_1320); +x_1327 = lean_ctor_get(x_1326, 0); +lean_inc(x_1327); +x_1328 = lean_ctor_get(x_1326, 1); +lean_inc(x_1328); +if (lean_is_exclusive(x_1326)) { + lean_ctor_release(x_1326, 0); + lean_ctor_release(x_1326, 1); + x_1329 = x_1326; +} else { + lean_dec_ref(x_1326); + x_1329 = lean_box(0); +} +if (lean_is_scalar(x_1329)) { + x_1330 = lean_alloc_ctor(0, 2, 0); +} else { + x_1330 = x_1329; +} +lean_ctor_set(x_1330, 0, x_1327); +lean_ctor_set(x_1330, 1, x_1328); +return x_1330; +} +else +{ +lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; uint8_t x_1335; +x_1331 = lean_ctor_get(x_1326, 0); +lean_inc(x_1331); +x_1332 = lean_ctor_get(x_1326, 1); +lean_inc(x_1332); +if (lean_is_exclusive(x_1326)) { + lean_ctor_release(x_1326, 0); + lean_ctor_release(x_1326, 1); + x_1333 = x_1326; +} else { + lean_dec_ref(x_1326); + x_1333 = lean_box(0); +} +x_1334 = lean_ctor_get(x_1331, 1); +lean_inc(x_1334); +x_1335 = lean_nat_dec_eq(x_1320, x_1334); +lean_dec(x_1320); +if (x_1335 == 0) +{ +lean_object* x_1336; +lean_dec(x_1334); +if (lean_is_scalar(x_1333)) { + x_1336 = lean_alloc_ctor(1, 2, 0); +} else { + x_1336 = x_1333; +} +lean_ctor_set(x_1336, 0, x_1331); +lean_ctor_set(x_1336, 1, x_1332); +return x_1336; +} +else +{ +lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; +lean_dec(x_1333); +lean_dec(x_1332); +x_1337 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1338 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1338, 0, x_1337); +x_1339 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1340 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1339, x_1338, x_1331); +if (lean_obj_tag(x_1340) == 0) +{ +lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; +lean_dec(x_1334); +x_1341 = lean_ctor_get(x_1340, 0); +lean_inc(x_1341); +x_1342 = lean_ctor_get(x_1340, 1); +lean_inc(x_1342); +if (lean_is_exclusive(x_1340)) { + lean_ctor_release(x_1340, 0); + lean_ctor_release(x_1340, 1); + x_1343 = x_1340; +} else { + lean_dec_ref(x_1340); + x_1343 = lean_box(0); +} +if (lean_is_scalar(x_1343)) { + x_1344 = lean_alloc_ctor(0, 2, 0); +} else { + x_1344 = x_1343; +} +lean_ctor_set(x_1344, 0, x_1341); +lean_ctor_set(x_1344, 1, x_1342); +return x_1344; +} +else +{ +lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; uint8_t x_1349; +x_1345 = lean_ctor_get(x_1340, 0); +lean_inc(x_1345); +x_1346 = lean_ctor_get(x_1340, 1); +lean_inc(x_1346); +if (lean_is_exclusive(x_1340)) { + lean_ctor_release(x_1340, 0); + lean_ctor_release(x_1340, 1); + x_1347 = x_1340; +} else { + lean_dec_ref(x_1340); + x_1347 = lean_box(0); +} +x_1348 = lean_ctor_get(x_1345, 1); +lean_inc(x_1348); +x_1349 = lean_nat_dec_eq(x_1334, x_1348); +lean_dec(x_1334); +if (x_1349 == 0) +{ +lean_object* x_1350; +lean_dec(x_1348); +if (lean_is_scalar(x_1347)) { + x_1350 = lean_alloc_ctor(1, 2, 0); +} else { + x_1350 = x_1347; +} +lean_ctor_set(x_1350, 0, x_1345); +lean_ctor_set(x_1350, 1, x_1346); +return x_1350; +} +else +{ +lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; +lean_dec(x_1347); +lean_dec(x_1346); +x_1351 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1352 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1352, 0, x_1351); +x_1353 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1354 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1353, x_1352, x_1345); +if (lean_obj_tag(x_1354) == 0) +{ +lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; +lean_dec(x_1348); +x_1355 = lean_ctor_get(x_1354, 0); +lean_inc(x_1355); +x_1356 = lean_ctor_get(x_1354, 1); +lean_inc(x_1356); +if (lean_is_exclusive(x_1354)) { + lean_ctor_release(x_1354, 0); + lean_ctor_release(x_1354, 1); + x_1357 = x_1354; +} else { + lean_dec_ref(x_1354); + x_1357 = lean_box(0); +} +if (lean_is_scalar(x_1357)) { + x_1358 = lean_alloc_ctor(0, 2, 0); +} else { + x_1358 = x_1357; +} +lean_ctor_set(x_1358, 0, x_1355); +lean_ctor_set(x_1358, 1, x_1356); +return x_1358; +} +else +{ +lean_object* x_1359; lean_object* x_1360; lean_object* x_1361; lean_object* x_1362; uint8_t x_1363; +x_1359 = lean_ctor_get(x_1354, 0); +lean_inc(x_1359); +x_1360 = lean_ctor_get(x_1354, 1); +lean_inc(x_1360); +if (lean_is_exclusive(x_1354)) { + lean_ctor_release(x_1354, 0); + lean_ctor_release(x_1354, 1); + x_1361 = x_1354; +} else { + lean_dec_ref(x_1354); + x_1361 = lean_box(0); +} +x_1362 = lean_ctor_get(x_1359, 1); +lean_inc(x_1362); +x_1363 = lean_nat_dec_eq(x_1348, x_1362); +lean_dec(x_1348); +if (x_1363 == 0) +{ +lean_object* x_1364; +lean_dec(x_1362); +if (lean_is_scalar(x_1361)) { + x_1364 = lean_alloc_ctor(1, 2, 0); +} else { + x_1364 = x_1361; +} +lean_ctor_set(x_1364, 0, x_1359); +lean_ctor_set(x_1364, 1, x_1360); +return x_1364; +} +else +{ +lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; +lean_dec(x_1361); +lean_dec(x_1360); +x_1365 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1366 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1366, 0, x_1365); +x_1367 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1368 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1367, x_1366, x_1359); +if (lean_obj_tag(x_1368) == 0) +{ +lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; +lean_dec(x_1362); +x_1369 = lean_ctor_get(x_1368, 0); +lean_inc(x_1369); +x_1370 = lean_ctor_get(x_1368, 1); +lean_inc(x_1370); +if (lean_is_exclusive(x_1368)) { + lean_ctor_release(x_1368, 0); + lean_ctor_release(x_1368, 1); + x_1371 = x_1368; +} else { + lean_dec_ref(x_1368); + x_1371 = lean_box(0); +} +if (lean_is_scalar(x_1371)) { + x_1372 = lean_alloc_ctor(0, 2, 0); +} else { + x_1372 = x_1371; +} +lean_ctor_set(x_1372, 0, x_1369); +lean_ctor_set(x_1372, 1, x_1370); +return x_1372; +} +else +{ +lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; uint8_t x_1377; +x_1373 = lean_ctor_get(x_1368, 0); +lean_inc(x_1373); +x_1374 = lean_ctor_get(x_1368, 1); +lean_inc(x_1374); +if (lean_is_exclusive(x_1368)) { + lean_ctor_release(x_1368, 0); + lean_ctor_release(x_1368, 1); + x_1375 = x_1368; +} else { + lean_dec_ref(x_1368); + x_1375 = lean_box(0); +} +x_1376 = lean_ctor_get(x_1373, 1); +lean_inc(x_1376); +x_1377 = lean_nat_dec_eq(x_1362, x_1376); +lean_dec(x_1362); +if (x_1377 == 0) +{ +lean_object* x_1378; +lean_dec(x_1376); +if (lean_is_scalar(x_1375)) { + x_1378 = lean_alloc_ctor(1, 2, 0); +} else { + x_1378 = x_1375; +} +lean_ctor_set(x_1378, 0, x_1373); +lean_ctor_set(x_1378, 1, x_1374); +return x_1378; +} +else +{ +lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; +lean_dec(x_1375); +lean_dec(x_1374); +x_1379 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1380 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1380, 0, x_1379); +x_1381 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1382 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1381, x_1380, x_1373); +if (lean_obj_tag(x_1382) == 0) +{ +lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; +lean_dec(x_1376); +x_1383 = lean_ctor_get(x_1382, 0); +lean_inc(x_1383); +x_1384 = lean_ctor_get(x_1382, 1); +lean_inc(x_1384); +if (lean_is_exclusive(x_1382)) { + lean_ctor_release(x_1382, 0); + lean_ctor_release(x_1382, 1); + x_1385 = x_1382; +} else { + lean_dec_ref(x_1382); + x_1385 = lean_box(0); +} +if (lean_is_scalar(x_1385)) { + x_1386 = lean_alloc_ctor(0, 2, 0); +} else { + x_1386 = x_1385; +} +lean_ctor_set(x_1386, 0, x_1383); +lean_ctor_set(x_1386, 1, x_1384); +return x_1386; +} +else +{ +lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; uint8_t x_1391; +x_1387 = lean_ctor_get(x_1382, 0); +lean_inc(x_1387); +x_1388 = lean_ctor_get(x_1382, 1); +lean_inc(x_1388); +if (lean_is_exclusive(x_1382)) { + lean_ctor_release(x_1382, 0); + lean_ctor_release(x_1382, 1); + x_1389 = x_1382; +} else { + lean_dec_ref(x_1382); + x_1389 = lean_box(0); +} +x_1390 = lean_ctor_get(x_1387, 1); +lean_inc(x_1390); +x_1391 = lean_nat_dec_eq(x_1376, x_1390); +lean_dec(x_1376); +if (x_1391 == 0) +{ +lean_object* x_1392; +lean_dec(x_1390); +if (lean_is_scalar(x_1389)) { + x_1392 = lean_alloc_ctor(1, 2, 0); +} else { + x_1392 = x_1389; +} +lean_ctor_set(x_1392, 0, x_1387); +lean_ctor_set(x_1392, 1, x_1388); +return x_1392; +} +else +{ +lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; +lean_dec(x_1389); +lean_dec(x_1388); +x_1393 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1394 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1394, 0, x_1393); +x_1395 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1396 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1395, x_1394, x_1387); +if (lean_obj_tag(x_1396) == 0) +{ +lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; +lean_dec(x_1390); +x_1397 = lean_ctor_get(x_1396, 0); +lean_inc(x_1397); +x_1398 = lean_ctor_get(x_1396, 1); +lean_inc(x_1398); +if (lean_is_exclusive(x_1396)) { + lean_ctor_release(x_1396, 0); + lean_ctor_release(x_1396, 1); + x_1399 = x_1396; +} else { + lean_dec_ref(x_1396); + x_1399 = lean_box(0); +} +if (lean_is_scalar(x_1399)) { + x_1400 = lean_alloc_ctor(0, 2, 0); +} else { + x_1400 = x_1399; +} +lean_ctor_set(x_1400, 0, x_1397); +lean_ctor_set(x_1400, 1, x_1398); +return x_1400; +} +else +{ +lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; uint8_t x_1405; +x_1401 = lean_ctor_get(x_1396, 0); +lean_inc(x_1401); +x_1402 = lean_ctor_get(x_1396, 1); +lean_inc(x_1402); +if (lean_is_exclusive(x_1396)) { + lean_ctor_release(x_1396, 0); + lean_ctor_release(x_1396, 1); + x_1403 = x_1396; +} else { + lean_dec_ref(x_1396); + x_1403 = lean_box(0); +} +x_1404 = lean_ctor_get(x_1401, 1); +lean_inc(x_1404); +x_1405 = lean_nat_dec_eq(x_1390, x_1404); +lean_dec(x_1390); +if (x_1405 == 0) +{ +lean_object* x_1406; +lean_dec(x_1404); +if (lean_is_scalar(x_1403)) { + x_1406 = lean_alloc_ctor(1, 2, 0); +} else { + x_1406 = x_1403; +} +lean_ctor_set(x_1406, 0, x_1401); +lean_ctor_set(x_1406, 1, x_1402); +return x_1406; +} +else +{ +lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; +lean_dec(x_1403); +lean_dec(x_1402); +x_1407 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1408 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1408, 0, x_1407); +x_1409 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1410 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1409, x_1408, x_1401); +if (lean_obj_tag(x_1410) == 0) +{ +lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; +lean_dec(x_1404); +x_1411 = lean_ctor_get(x_1410, 0); +lean_inc(x_1411); +x_1412 = lean_ctor_get(x_1410, 1); +lean_inc(x_1412); +if (lean_is_exclusive(x_1410)) { + lean_ctor_release(x_1410, 0); + lean_ctor_release(x_1410, 1); + x_1413 = x_1410; +} else { + lean_dec_ref(x_1410); + x_1413 = lean_box(0); +} +if (lean_is_scalar(x_1413)) { + x_1414 = lean_alloc_ctor(0, 2, 0); +} else { + x_1414 = x_1413; +} +lean_ctor_set(x_1414, 0, x_1411); +lean_ctor_set(x_1414, 1, x_1412); +return x_1414; +} +else +{ +lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; uint8_t x_1419; +x_1415 = lean_ctor_get(x_1410, 0); +lean_inc(x_1415); +x_1416 = lean_ctor_get(x_1410, 1); +lean_inc(x_1416); +if (lean_is_exclusive(x_1410)) { + lean_ctor_release(x_1410, 0); + lean_ctor_release(x_1410, 1); + x_1417 = x_1410; +} else { + lean_dec_ref(x_1410); + x_1417 = lean_box(0); +} +x_1418 = lean_ctor_get(x_1415, 1); +lean_inc(x_1418); +x_1419 = lean_nat_dec_eq(x_1404, x_1418); +lean_dec(x_1404); +if (x_1419 == 0) +{ +lean_object* x_1420; +lean_dec(x_1418); +if (lean_is_scalar(x_1417)) { + x_1420 = lean_alloc_ctor(1, 2, 0); +} else { + x_1420 = x_1417; +} +lean_ctor_set(x_1420, 0, x_1415); +lean_ctor_set(x_1420, 1, x_1416); +return x_1420; +} +else +{ +lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; +lean_dec(x_1417); +lean_dec(x_1416); +x_1421 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1422 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1422, 0, x_1421); +x_1423 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1424 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1423, x_1422, x_1415); +if (lean_obj_tag(x_1424) == 0) +{ +lean_object* x_1425; lean_object* x_1426; lean_object* x_1427; lean_object* x_1428; +lean_dec(x_1418); +x_1425 = lean_ctor_get(x_1424, 0); +lean_inc(x_1425); +x_1426 = lean_ctor_get(x_1424, 1); +lean_inc(x_1426); +if (lean_is_exclusive(x_1424)) { + lean_ctor_release(x_1424, 0); + lean_ctor_release(x_1424, 1); + x_1427 = x_1424; +} else { + lean_dec_ref(x_1424); + x_1427 = lean_box(0); +} +if (lean_is_scalar(x_1427)) { + x_1428 = lean_alloc_ctor(0, 2, 0); +} else { + x_1428 = x_1427; +} +lean_ctor_set(x_1428, 0, x_1425); +lean_ctor_set(x_1428, 1, x_1426); +return x_1428; +} +else +{ +lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; uint8_t x_1433; +x_1429 = lean_ctor_get(x_1424, 0); +lean_inc(x_1429); +x_1430 = lean_ctor_get(x_1424, 1); +lean_inc(x_1430); +if (lean_is_exclusive(x_1424)) { + lean_ctor_release(x_1424, 0); + lean_ctor_release(x_1424, 1); + x_1431 = x_1424; +} else { + lean_dec_ref(x_1424); + x_1431 = lean_box(0); +} +x_1432 = lean_ctor_get(x_1429, 1); +lean_inc(x_1432); +x_1433 = lean_nat_dec_eq(x_1418, x_1432); +lean_dec(x_1432); +lean_dec(x_1418); +if (x_1433 == 0) +{ +lean_object* x_1434; +if (lean_is_scalar(x_1431)) { + x_1434 = lean_alloc_ctor(1, 2, 0); +} else { + x_1434 = x_1431; +} +lean_ctor_set(x_1434, 0, x_1429); +lean_ctor_set(x_1434, 1, x_1430); +return x_1434; +} +else +{ +lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; +lean_dec(x_1431); +lean_dec(x_1430); +x_1435 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_1436 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1436, 0, x_1435); +x_1437 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_1438 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1437, x_1436, x_1429); +return x_1438; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; uint8_t x_1442; +x_1439 = lean_ctor_get(x_237, 0); +x_1440 = lean_ctor_get(x_237, 1); +lean_inc(x_1440); +lean_inc(x_1439); +lean_dec(x_237); +x_1441 = lean_ctor_get(x_1439, 1); +lean_inc(x_1441); +x_1442 = lean_nat_dec_eq(x_232, x_1441); +lean_dec(x_232); +if (x_1442 == 0) +{ +lean_object* x_1443; +lean_dec(x_1441); +x_1443 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1443, 0, x_1439); +lean_ctor_set(x_1443, 1, x_1440); +return x_1443; +} +else +{ +lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; lean_object* x_1447; +lean_dec(x_1440); +x_1444 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_1445 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1445, 0, x_1444); +x_1446 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_1447 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1446, x_1445, x_1439); +if (lean_obj_tag(x_1447) == 0) +{ +lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; lean_object* x_1451; +lean_dec(x_1441); +x_1448 = lean_ctor_get(x_1447, 0); +lean_inc(x_1448); +x_1449 = lean_ctor_get(x_1447, 1); +lean_inc(x_1449); +if (lean_is_exclusive(x_1447)) { + lean_ctor_release(x_1447, 0); + lean_ctor_release(x_1447, 1); + x_1450 = x_1447; +} else { + lean_dec_ref(x_1447); + x_1450 = lean_box(0); +} +if (lean_is_scalar(x_1450)) { + x_1451 = lean_alloc_ctor(0, 2, 0); +} else { + x_1451 = x_1450; +} +lean_ctor_set(x_1451, 0, x_1448); +lean_ctor_set(x_1451, 1, x_1449); +return x_1451; +} +else +{ +lean_object* x_1452; lean_object* x_1453; lean_object* x_1454; lean_object* x_1455; uint8_t x_1456; +x_1452 = lean_ctor_get(x_1447, 0); +lean_inc(x_1452); +x_1453 = lean_ctor_get(x_1447, 1); +lean_inc(x_1453); +if (lean_is_exclusive(x_1447)) { + lean_ctor_release(x_1447, 0); + lean_ctor_release(x_1447, 1); + x_1454 = x_1447; +} else { + lean_dec_ref(x_1447); + x_1454 = lean_box(0); +} +x_1455 = lean_ctor_get(x_1452, 1); +lean_inc(x_1455); +x_1456 = lean_nat_dec_eq(x_1441, x_1455); +lean_dec(x_1441); +if (x_1456 == 0) +{ +lean_object* x_1457; +lean_dec(x_1455); +if (lean_is_scalar(x_1454)) { + x_1457 = lean_alloc_ctor(1, 2, 0); +} else { + x_1457 = x_1454; +} +lean_ctor_set(x_1457, 0, x_1452); +lean_ctor_set(x_1457, 1, x_1453); +return x_1457; +} +else +{ +lean_object* x_1458; lean_object* x_1459; lean_object* x_1460; lean_object* x_1461; +lean_dec(x_1454); +lean_dec(x_1453); +x_1458 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_1459 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1459, 0, x_1458); +x_1460 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_1461 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1460, x_1459, x_1452); +if (lean_obj_tag(x_1461) == 0) +{ +lean_object* x_1462; lean_object* x_1463; lean_object* x_1464; lean_object* x_1465; +lean_dec(x_1455); +x_1462 = lean_ctor_get(x_1461, 0); +lean_inc(x_1462); +x_1463 = lean_ctor_get(x_1461, 1); +lean_inc(x_1463); +if (lean_is_exclusive(x_1461)) { + lean_ctor_release(x_1461, 0); + lean_ctor_release(x_1461, 1); + x_1464 = x_1461; +} else { + lean_dec_ref(x_1461); + x_1464 = lean_box(0); +} +if (lean_is_scalar(x_1464)) { + x_1465 = lean_alloc_ctor(0, 2, 0); +} else { + x_1465 = x_1464; +} +lean_ctor_set(x_1465, 0, x_1462); +lean_ctor_set(x_1465, 1, x_1463); +return x_1465; +} +else +{ +lean_object* x_1466; lean_object* x_1467; lean_object* x_1468; lean_object* x_1469; uint8_t x_1470; +x_1466 = lean_ctor_get(x_1461, 0); +lean_inc(x_1466); +x_1467 = lean_ctor_get(x_1461, 1); +lean_inc(x_1467); +if (lean_is_exclusive(x_1461)) { + lean_ctor_release(x_1461, 0); + lean_ctor_release(x_1461, 1); + x_1468 = x_1461; +} else { + lean_dec_ref(x_1461); + x_1468 = lean_box(0); +} +x_1469 = lean_ctor_get(x_1466, 1); +lean_inc(x_1469); +x_1470 = lean_nat_dec_eq(x_1455, x_1469); +lean_dec(x_1455); +if (x_1470 == 0) +{ +lean_object* x_1471; +lean_dec(x_1469); +if (lean_is_scalar(x_1468)) { + x_1471 = lean_alloc_ctor(1, 2, 0); +} else { + x_1471 = x_1468; +} +lean_ctor_set(x_1471, 0, x_1466); +lean_ctor_set(x_1471, 1, x_1467); +return x_1471; +} +else +{ +lean_object* x_1472; lean_object* x_1473; lean_object* x_1474; lean_object* x_1475; +lean_dec(x_1468); +lean_dec(x_1467); +x_1472 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_1473 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1473, 0, x_1472); +x_1474 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_1475 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1474, x_1473, x_1466); +if (lean_obj_tag(x_1475) == 0) +{ +lean_object* x_1476; lean_object* x_1477; lean_object* x_1478; lean_object* x_1479; +lean_dec(x_1469); +x_1476 = lean_ctor_get(x_1475, 0); +lean_inc(x_1476); +x_1477 = lean_ctor_get(x_1475, 1); +lean_inc(x_1477); +if (lean_is_exclusive(x_1475)) { + lean_ctor_release(x_1475, 0); + lean_ctor_release(x_1475, 1); + x_1478 = x_1475; +} else { + lean_dec_ref(x_1475); + x_1478 = lean_box(0); +} +if (lean_is_scalar(x_1478)) { + x_1479 = lean_alloc_ctor(0, 2, 0); +} else { + x_1479 = x_1478; +} +lean_ctor_set(x_1479, 0, x_1476); +lean_ctor_set(x_1479, 1, x_1477); +return x_1479; +} +else +{ +lean_object* x_1480; lean_object* x_1481; lean_object* x_1482; lean_object* x_1483; uint8_t x_1484; +x_1480 = lean_ctor_get(x_1475, 0); +lean_inc(x_1480); +x_1481 = lean_ctor_get(x_1475, 1); +lean_inc(x_1481); +if (lean_is_exclusive(x_1475)) { + lean_ctor_release(x_1475, 0); + lean_ctor_release(x_1475, 1); + x_1482 = x_1475; +} else { + lean_dec_ref(x_1475); + x_1482 = lean_box(0); +} +x_1483 = lean_ctor_get(x_1480, 1); +lean_inc(x_1483); +x_1484 = lean_nat_dec_eq(x_1469, x_1483); +lean_dec(x_1469); +if (x_1484 == 0) +{ +lean_object* x_1485; +lean_dec(x_1483); +if (lean_is_scalar(x_1482)) { + x_1485 = lean_alloc_ctor(1, 2, 0); +} else { + x_1485 = x_1482; +} +lean_ctor_set(x_1485, 0, x_1480); +lean_ctor_set(x_1485, 1, x_1481); +return x_1485; +} +else +{ +lean_object* x_1486; lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; +lean_dec(x_1482); +lean_dec(x_1481); +x_1486 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_1487 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1487, 0, x_1486); +x_1488 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1489 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1488, x_1487, x_1480); +if (lean_obj_tag(x_1489) == 0) +{ +lean_object* x_1490; lean_object* x_1491; lean_object* x_1492; lean_object* x_1493; +lean_dec(x_1483); +x_1490 = lean_ctor_get(x_1489, 0); +lean_inc(x_1490); +x_1491 = lean_ctor_get(x_1489, 1); +lean_inc(x_1491); +if (lean_is_exclusive(x_1489)) { + lean_ctor_release(x_1489, 0); + lean_ctor_release(x_1489, 1); + x_1492 = x_1489; +} else { + lean_dec_ref(x_1489); + x_1492 = lean_box(0); +} +if (lean_is_scalar(x_1492)) { + x_1493 = lean_alloc_ctor(0, 2, 0); +} else { + x_1493 = x_1492; +} +lean_ctor_set(x_1493, 0, x_1490); +lean_ctor_set(x_1493, 1, x_1491); +return x_1493; +} +else +{ +lean_object* x_1494; lean_object* x_1495; lean_object* x_1496; lean_object* x_1497; uint8_t x_1498; +x_1494 = lean_ctor_get(x_1489, 0); +lean_inc(x_1494); +x_1495 = lean_ctor_get(x_1489, 1); +lean_inc(x_1495); +if (lean_is_exclusive(x_1489)) { + lean_ctor_release(x_1489, 0); + lean_ctor_release(x_1489, 1); + x_1496 = x_1489; +} else { + lean_dec_ref(x_1489); + x_1496 = lean_box(0); +} +x_1497 = lean_ctor_get(x_1494, 1); +lean_inc(x_1497); +x_1498 = lean_nat_dec_eq(x_1483, x_1497); +lean_dec(x_1483); +if (x_1498 == 0) +{ +lean_object* x_1499; +lean_dec(x_1497); +if (lean_is_scalar(x_1496)) { + x_1499 = lean_alloc_ctor(1, 2, 0); +} else { + x_1499 = x_1496; +} +lean_ctor_set(x_1499, 0, x_1494); +lean_ctor_set(x_1499, 1, x_1495); +return x_1499; +} +else +{ +lean_object* x_1500; lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; +lean_dec(x_1496); +lean_dec(x_1495); +x_1500 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1501 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1501, 0, x_1500); +x_1502 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1503 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1502, x_1501, x_1494); +if (lean_obj_tag(x_1503) == 0) +{ +lean_object* x_1504; lean_object* x_1505; lean_object* x_1506; lean_object* x_1507; +lean_dec(x_1497); +x_1504 = lean_ctor_get(x_1503, 0); +lean_inc(x_1504); +x_1505 = lean_ctor_get(x_1503, 1); +lean_inc(x_1505); +if (lean_is_exclusive(x_1503)) { + lean_ctor_release(x_1503, 0); + lean_ctor_release(x_1503, 1); + x_1506 = x_1503; +} else { + lean_dec_ref(x_1503); + x_1506 = lean_box(0); +} +if (lean_is_scalar(x_1506)) { + x_1507 = lean_alloc_ctor(0, 2, 0); +} else { + x_1507 = x_1506; +} +lean_ctor_set(x_1507, 0, x_1504); +lean_ctor_set(x_1507, 1, x_1505); +return x_1507; +} +else +{ +lean_object* x_1508; lean_object* x_1509; lean_object* x_1510; lean_object* x_1511; uint8_t x_1512; +x_1508 = lean_ctor_get(x_1503, 0); +lean_inc(x_1508); +x_1509 = lean_ctor_get(x_1503, 1); +lean_inc(x_1509); +if (lean_is_exclusive(x_1503)) { + lean_ctor_release(x_1503, 0); + lean_ctor_release(x_1503, 1); + x_1510 = x_1503; +} else { + lean_dec_ref(x_1503); + x_1510 = lean_box(0); +} +x_1511 = lean_ctor_get(x_1508, 1); +lean_inc(x_1511); +x_1512 = lean_nat_dec_eq(x_1497, x_1511); +lean_dec(x_1497); +if (x_1512 == 0) +{ +lean_object* x_1513; +lean_dec(x_1511); +if (lean_is_scalar(x_1510)) { + x_1513 = lean_alloc_ctor(1, 2, 0); +} else { + x_1513 = x_1510; +} +lean_ctor_set(x_1513, 0, x_1508); +lean_ctor_set(x_1513, 1, x_1509); +return x_1513; +} +else +{ +lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; +lean_dec(x_1510); +lean_dec(x_1509); +x_1514 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1515 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1515, 0, x_1514); +x_1516 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1517 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1516, x_1515, x_1508); +if (lean_obj_tag(x_1517) == 0) +{ +lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; lean_object* x_1521; +lean_dec(x_1511); +x_1518 = lean_ctor_get(x_1517, 0); +lean_inc(x_1518); +x_1519 = lean_ctor_get(x_1517, 1); +lean_inc(x_1519); +if (lean_is_exclusive(x_1517)) { + lean_ctor_release(x_1517, 0); + lean_ctor_release(x_1517, 1); + x_1520 = x_1517; +} else { + lean_dec_ref(x_1517); + x_1520 = lean_box(0); +} +if (lean_is_scalar(x_1520)) { + x_1521 = lean_alloc_ctor(0, 2, 0); +} else { + x_1521 = x_1520; +} +lean_ctor_set(x_1521, 0, x_1518); +lean_ctor_set(x_1521, 1, x_1519); +return x_1521; +} +else +{ +lean_object* x_1522; lean_object* x_1523; lean_object* x_1524; lean_object* x_1525; uint8_t x_1526; +x_1522 = lean_ctor_get(x_1517, 0); +lean_inc(x_1522); +x_1523 = lean_ctor_get(x_1517, 1); +lean_inc(x_1523); +if (lean_is_exclusive(x_1517)) { + lean_ctor_release(x_1517, 0); + lean_ctor_release(x_1517, 1); + x_1524 = x_1517; +} else { + lean_dec_ref(x_1517); + x_1524 = lean_box(0); +} +x_1525 = lean_ctor_get(x_1522, 1); +lean_inc(x_1525); +x_1526 = lean_nat_dec_eq(x_1511, x_1525); +lean_dec(x_1511); +if (x_1526 == 0) +{ +lean_object* x_1527; +lean_dec(x_1525); +if (lean_is_scalar(x_1524)) { + x_1527 = lean_alloc_ctor(1, 2, 0); +} else { + x_1527 = x_1524; +} +lean_ctor_set(x_1527, 0, x_1522); +lean_ctor_set(x_1527, 1, x_1523); +return x_1527; +} +else +{ +lean_object* x_1528; lean_object* x_1529; lean_object* x_1530; lean_object* x_1531; +lean_dec(x_1524); +lean_dec(x_1523); +x_1528 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1529 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1529, 0, x_1528); +x_1530 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1531 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1530, x_1529, x_1522); +if (lean_obj_tag(x_1531) == 0) +{ +lean_object* x_1532; lean_object* x_1533; lean_object* x_1534; lean_object* x_1535; +lean_dec(x_1525); +x_1532 = lean_ctor_get(x_1531, 0); +lean_inc(x_1532); +x_1533 = lean_ctor_get(x_1531, 1); +lean_inc(x_1533); +if (lean_is_exclusive(x_1531)) { + lean_ctor_release(x_1531, 0); + lean_ctor_release(x_1531, 1); + x_1534 = x_1531; +} else { + lean_dec_ref(x_1531); + x_1534 = lean_box(0); +} +if (lean_is_scalar(x_1534)) { + x_1535 = lean_alloc_ctor(0, 2, 0); +} else { + x_1535 = x_1534; +} +lean_ctor_set(x_1535, 0, x_1532); +lean_ctor_set(x_1535, 1, x_1533); +return x_1535; +} +else +{ +lean_object* x_1536; lean_object* x_1537; lean_object* x_1538; lean_object* x_1539; uint8_t x_1540; +x_1536 = lean_ctor_get(x_1531, 0); +lean_inc(x_1536); +x_1537 = lean_ctor_get(x_1531, 1); +lean_inc(x_1537); +if (lean_is_exclusive(x_1531)) { + lean_ctor_release(x_1531, 0); + lean_ctor_release(x_1531, 1); + x_1538 = x_1531; +} else { + lean_dec_ref(x_1531); + x_1538 = lean_box(0); +} +x_1539 = lean_ctor_get(x_1536, 1); +lean_inc(x_1539); +x_1540 = lean_nat_dec_eq(x_1525, x_1539); +lean_dec(x_1525); +if (x_1540 == 0) +{ +lean_object* x_1541; +lean_dec(x_1539); +if (lean_is_scalar(x_1538)) { + x_1541 = lean_alloc_ctor(1, 2, 0); +} else { + x_1541 = x_1538; +} +lean_ctor_set(x_1541, 0, x_1536); +lean_ctor_set(x_1541, 1, x_1537); +return x_1541; +} +else +{ +lean_object* x_1542; lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; +lean_dec(x_1538); +lean_dec(x_1537); +x_1542 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1543 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1543, 0, x_1542); +x_1544 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1545 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1544, x_1543, x_1536); +if (lean_obj_tag(x_1545) == 0) +{ +lean_object* x_1546; lean_object* x_1547; lean_object* x_1548; lean_object* x_1549; +lean_dec(x_1539); +x_1546 = lean_ctor_get(x_1545, 0); +lean_inc(x_1546); +x_1547 = lean_ctor_get(x_1545, 1); +lean_inc(x_1547); +if (lean_is_exclusive(x_1545)) { + lean_ctor_release(x_1545, 0); + lean_ctor_release(x_1545, 1); + x_1548 = x_1545; +} else { + lean_dec_ref(x_1545); + x_1548 = lean_box(0); +} +if (lean_is_scalar(x_1548)) { + x_1549 = lean_alloc_ctor(0, 2, 0); +} else { + x_1549 = x_1548; +} +lean_ctor_set(x_1549, 0, x_1546); +lean_ctor_set(x_1549, 1, x_1547); +return x_1549; +} +else +{ +lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_object* x_1553; uint8_t x_1554; +x_1550 = lean_ctor_get(x_1545, 0); +lean_inc(x_1550); +x_1551 = lean_ctor_get(x_1545, 1); +lean_inc(x_1551); +if (lean_is_exclusive(x_1545)) { + lean_ctor_release(x_1545, 0); + lean_ctor_release(x_1545, 1); + x_1552 = x_1545; +} else { + lean_dec_ref(x_1545); + x_1552 = lean_box(0); +} +x_1553 = lean_ctor_get(x_1550, 1); +lean_inc(x_1553); +x_1554 = lean_nat_dec_eq(x_1539, x_1553); +lean_dec(x_1539); +if (x_1554 == 0) +{ +lean_object* x_1555; +lean_dec(x_1553); +if (lean_is_scalar(x_1552)) { + x_1555 = lean_alloc_ctor(1, 2, 0); +} else { + x_1555 = x_1552; +} +lean_ctor_set(x_1555, 0, x_1550); +lean_ctor_set(x_1555, 1, x_1551); +return x_1555; +} +else +{ +lean_object* x_1556; lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; +lean_dec(x_1552); +lean_dec(x_1551); +x_1556 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1557 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1557, 0, x_1556); +x_1558 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1559 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1558, x_1557, x_1550); +if (lean_obj_tag(x_1559) == 0) +{ +lean_object* x_1560; lean_object* x_1561; lean_object* x_1562; lean_object* x_1563; +lean_dec(x_1553); +x_1560 = lean_ctor_get(x_1559, 0); +lean_inc(x_1560); +x_1561 = lean_ctor_get(x_1559, 1); +lean_inc(x_1561); +if (lean_is_exclusive(x_1559)) { + lean_ctor_release(x_1559, 0); + lean_ctor_release(x_1559, 1); + x_1562 = x_1559; +} else { + lean_dec_ref(x_1559); + x_1562 = lean_box(0); +} +if (lean_is_scalar(x_1562)) { + x_1563 = lean_alloc_ctor(0, 2, 0); +} else { + x_1563 = x_1562; +} +lean_ctor_set(x_1563, 0, x_1560); +lean_ctor_set(x_1563, 1, x_1561); +return x_1563; +} +else +{ +lean_object* x_1564; lean_object* x_1565; lean_object* x_1566; lean_object* x_1567; uint8_t x_1568; +x_1564 = lean_ctor_get(x_1559, 0); +lean_inc(x_1564); +x_1565 = lean_ctor_get(x_1559, 1); +lean_inc(x_1565); +if (lean_is_exclusive(x_1559)) { + lean_ctor_release(x_1559, 0); + lean_ctor_release(x_1559, 1); + x_1566 = x_1559; +} else { + lean_dec_ref(x_1559); + x_1566 = lean_box(0); +} +x_1567 = lean_ctor_get(x_1564, 1); +lean_inc(x_1567); +x_1568 = lean_nat_dec_eq(x_1553, x_1567); +lean_dec(x_1553); +if (x_1568 == 0) +{ +lean_object* x_1569; +lean_dec(x_1567); +if (lean_is_scalar(x_1566)) { + x_1569 = lean_alloc_ctor(1, 2, 0); +} else { + x_1569 = x_1566; +} +lean_ctor_set(x_1569, 0, x_1564); +lean_ctor_set(x_1569, 1, x_1565); +return x_1569; +} +else +{ +lean_object* x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; +lean_dec(x_1566); +lean_dec(x_1565); +x_1570 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1571 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1571, 0, x_1570); +x_1572 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1573 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1572, x_1571, x_1564); +if (lean_obj_tag(x_1573) == 0) +{ +lean_object* x_1574; lean_object* x_1575; lean_object* x_1576; lean_object* x_1577; +lean_dec(x_1567); +x_1574 = lean_ctor_get(x_1573, 0); +lean_inc(x_1574); +x_1575 = lean_ctor_get(x_1573, 1); +lean_inc(x_1575); +if (lean_is_exclusive(x_1573)) { + lean_ctor_release(x_1573, 0); + lean_ctor_release(x_1573, 1); + x_1576 = x_1573; +} else { + lean_dec_ref(x_1573); + x_1576 = lean_box(0); +} +if (lean_is_scalar(x_1576)) { + x_1577 = lean_alloc_ctor(0, 2, 0); +} else { + x_1577 = x_1576; +} +lean_ctor_set(x_1577, 0, x_1574); +lean_ctor_set(x_1577, 1, x_1575); +return x_1577; +} +else +{ +lean_object* x_1578; lean_object* x_1579; lean_object* x_1580; lean_object* x_1581; uint8_t x_1582; +x_1578 = lean_ctor_get(x_1573, 0); +lean_inc(x_1578); +x_1579 = lean_ctor_get(x_1573, 1); +lean_inc(x_1579); +if (lean_is_exclusive(x_1573)) { + lean_ctor_release(x_1573, 0); + lean_ctor_release(x_1573, 1); + x_1580 = x_1573; +} else { + lean_dec_ref(x_1573); + x_1580 = lean_box(0); +} +x_1581 = lean_ctor_get(x_1578, 1); +lean_inc(x_1581); +x_1582 = lean_nat_dec_eq(x_1567, x_1581); +lean_dec(x_1567); +if (x_1582 == 0) +{ +lean_object* x_1583; +lean_dec(x_1581); +if (lean_is_scalar(x_1580)) { + x_1583 = lean_alloc_ctor(1, 2, 0); +} else { + x_1583 = x_1580; +} +lean_ctor_set(x_1583, 0, x_1578); +lean_ctor_set(x_1583, 1, x_1579); +return x_1583; +} +else +{ +lean_object* x_1584; lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; +lean_dec(x_1580); +lean_dec(x_1579); +x_1584 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1585 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1585, 0, x_1584); +x_1586 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1587 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1586, x_1585, x_1578); +if (lean_obj_tag(x_1587) == 0) +{ +lean_object* x_1588; lean_object* x_1589; lean_object* x_1590; lean_object* x_1591; +lean_dec(x_1581); +x_1588 = lean_ctor_get(x_1587, 0); +lean_inc(x_1588); +x_1589 = lean_ctor_get(x_1587, 1); +lean_inc(x_1589); +if (lean_is_exclusive(x_1587)) { + lean_ctor_release(x_1587, 0); + lean_ctor_release(x_1587, 1); + x_1590 = x_1587; +} else { + lean_dec_ref(x_1587); + x_1590 = lean_box(0); +} +if (lean_is_scalar(x_1590)) { + x_1591 = lean_alloc_ctor(0, 2, 0); +} else { + x_1591 = x_1590; +} +lean_ctor_set(x_1591, 0, x_1588); +lean_ctor_set(x_1591, 1, x_1589); +return x_1591; +} +else +{ +lean_object* x_1592; lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; uint8_t x_1596; +x_1592 = lean_ctor_get(x_1587, 0); +lean_inc(x_1592); +x_1593 = lean_ctor_get(x_1587, 1); +lean_inc(x_1593); +if (lean_is_exclusive(x_1587)) { + lean_ctor_release(x_1587, 0); + lean_ctor_release(x_1587, 1); + x_1594 = x_1587; +} else { + lean_dec_ref(x_1587); + x_1594 = lean_box(0); +} +x_1595 = lean_ctor_get(x_1592, 1); +lean_inc(x_1595); +x_1596 = lean_nat_dec_eq(x_1581, x_1595); +lean_dec(x_1581); +if (x_1596 == 0) +{ +lean_object* x_1597; +lean_dec(x_1595); +if (lean_is_scalar(x_1594)) { + x_1597 = lean_alloc_ctor(1, 2, 0); +} else { + x_1597 = x_1594; +} +lean_ctor_set(x_1597, 0, x_1592); +lean_ctor_set(x_1597, 1, x_1593); +return x_1597; +} +else +{ +lean_object* x_1598; lean_object* x_1599; lean_object* x_1600; lean_object* x_1601; +lean_dec(x_1594); +lean_dec(x_1593); +x_1598 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1599 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1599, 0, x_1598); +x_1600 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1601 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1600, x_1599, x_1592); +if (lean_obj_tag(x_1601) == 0) +{ +lean_object* x_1602; lean_object* x_1603; lean_object* x_1604; lean_object* x_1605; +lean_dec(x_1595); +x_1602 = lean_ctor_get(x_1601, 0); +lean_inc(x_1602); +x_1603 = lean_ctor_get(x_1601, 1); +lean_inc(x_1603); +if (lean_is_exclusive(x_1601)) { + lean_ctor_release(x_1601, 0); + lean_ctor_release(x_1601, 1); + x_1604 = x_1601; +} else { + lean_dec_ref(x_1601); + x_1604 = lean_box(0); +} +if (lean_is_scalar(x_1604)) { + x_1605 = lean_alloc_ctor(0, 2, 0); +} else { + x_1605 = x_1604; +} +lean_ctor_set(x_1605, 0, x_1602); +lean_ctor_set(x_1605, 1, x_1603); +return x_1605; +} +else +{ +lean_object* x_1606; lean_object* x_1607; lean_object* x_1608; lean_object* x_1609; uint8_t x_1610; +x_1606 = lean_ctor_get(x_1601, 0); +lean_inc(x_1606); +x_1607 = lean_ctor_get(x_1601, 1); +lean_inc(x_1607); +if (lean_is_exclusive(x_1601)) { + lean_ctor_release(x_1601, 0); + lean_ctor_release(x_1601, 1); + x_1608 = x_1601; +} else { + lean_dec_ref(x_1601); + x_1608 = lean_box(0); +} +x_1609 = lean_ctor_get(x_1606, 1); +lean_inc(x_1609); +x_1610 = lean_nat_dec_eq(x_1595, x_1609); +lean_dec(x_1609); +lean_dec(x_1595); +if (x_1610 == 0) +{ +lean_object* x_1611; +if (lean_is_scalar(x_1608)) { + x_1611 = lean_alloc_ctor(1, 2, 0); +} else { + x_1611 = x_1608; +} +lean_ctor_set(x_1611, 0, x_1606); +lean_ctor_set(x_1611, 1, x_1607); +return x_1611; +} +else +{ +lean_object* x_1612; lean_object* x_1613; lean_object* x_1614; lean_object* x_1615; +lean_dec(x_1608); +lean_dec(x_1607); +x_1612 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_1613 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1613, 0, x_1612); +x_1614 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_1615 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1614, x_1613, x_1606); +return x_1615; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_1616; lean_object* x_1617; lean_object* x_1618; uint8_t x_1619; +x_1616 = lean_ctor_get(x_224, 0); +x_1617 = lean_ctor_get(x_224, 1); +lean_inc(x_1617); +lean_inc(x_1616); +lean_dec(x_224); +x_1618 = lean_ctor_get(x_1616, 1); +lean_inc(x_1618); +x_1619 = lean_nat_dec_eq(x_219, x_1618); +lean_dec(x_219); +if (x_1619 == 0) +{ +lean_object* x_1620; +lean_dec(x_1618); +x_1620 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1620, 0, x_1616); +lean_ctor_set(x_1620, 1, x_1617); +return x_1620; +} +else +{ +lean_object* x_1621; lean_object* x_1622; lean_object* x_1623; lean_object* x_1624; +lean_dec(x_1617); +x_1621 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_1622 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1622, 0, x_1621); +x_1623 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_1624 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1623, x_1622, x_1616); +if (lean_obj_tag(x_1624) == 0) +{ +lean_object* x_1625; lean_object* x_1626; lean_object* x_1627; lean_object* x_1628; +lean_dec(x_1618); +x_1625 = lean_ctor_get(x_1624, 0); +lean_inc(x_1625); +x_1626 = lean_ctor_get(x_1624, 1); +lean_inc(x_1626); +if (lean_is_exclusive(x_1624)) { + lean_ctor_release(x_1624, 0); + lean_ctor_release(x_1624, 1); + x_1627 = x_1624; +} else { + lean_dec_ref(x_1624); + x_1627 = lean_box(0); +} +if (lean_is_scalar(x_1627)) { + x_1628 = lean_alloc_ctor(0, 2, 0); +} else { + x_1628 = x_1627; +} +lean_ctor_set(x_1628, 0, x_1625); +lean_ctor_set(x_1628, 1, x_1626); +return x_1628; +} +else +{ +lean_object* x_1629; lean_object* x_1630; lean_object* x_1631; lean_object* x_1632; uint8_t x_1633; +x_1629 = lean_ctor_get(x_1624, 0); +lean_inc(x_1629); +x_1630 = lean_ctor_get(x_1624, 1); +lean_inc(x_1630); +if (lean_is_exclusive(x_1624)) { + lean_ctor_release(x_1624, 0); + lean_ctor_release(x_1624, 1); + x_1631 = x_1624; +} else { + lean_dec_ref(x_1624); + x_1631 = lean_box(0); +} +x_1632 = lean_ctor_get(x_1629, 1); +lean_inc(x_1632); +x_1633 = lean_nat_dec_eq(x_1618, x_1632); +lean_dec(x_1618); +if (x_1633 == 0) +{ +lean_object* x_1634; +lean_dec(x_1632); +if (lean_is_scalar(x_1631)) { + x_1634 = lean_alloc_ctor(1, 2, 0); +} else { + x_1634 = x_1631; +} +lean_ctor_set(x_1634, 0, x_1629); +lean_ctor_set(x_1634, 1, x_1630); +return x_1634; +} +else +{ +lean_object* x_1635; lean_object* x_1636; lean_object* x_1637; lean_object* x_1638; +lean_dec(x_1631); +lean_dec(x_1630); +x_1635 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_1636 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1636, 0, x_1635); +x_1637 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_1638 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1637, x_1636, x_1629); +if (lean_obj_tag(x_1638) == 0) +{ +lean_object* x_1639; lean_object* x_1640; lean_object* x_1641; lean_object* x_1642; +lean_dec(x_1632); +x_1639 = lean_ctor_get(x_1638, 0); +lean_inc(x_1639); +x_1640 = lean_ctor_get(x_1638, 1); +lean_inc(x_1640); +if (lean_is_exclusive(x_1638)) { + lean_ctor_release(x_1638, 0); + lean_ctor_release(x_1638, 1); + x_1641 = x_1638; +} else { + lean_dec_ref(x_1638); + x_1641 = lean_box(0); +} +if (lean_is_scalar(x_1641)) { + x_1642 = lean_alloc_ctor(0, 2, 0); +} else { + x_1642 = x_1641; +} +lean_ctor_set(x_1642, 0, x_1639); +lean_ctor_set(x_1642, 1, x_1640); +return x_1642; +} +else +{ +lean_object* x_1643; lean_object* x_1644; lean_object* x_1645; lean_object* x_1646; uint8_t x_1647; +x_1643 = lean_ctor_get(x_1638, 0); +lean_inc(x_1643); +x_1644 = lean_ctor_get(x_1638, 1); +lean_inc(x_1644); +if (lean_is_exclusive(x_1638)) { + lean_ctor_release(x_1638, 0); + lean_ctor_release(x_1638, 1); + x_1645 = x_1638; +} else { + lean_dec_ref(x_1638); + x_1645 = lean_box(0); +} +x_1646 = lean_ctor_get(x_1643, 1); +lean_inc(x_1646); +x_1647 = lean_nat_dec_eq(x_1632, x_1646); +lean_dec(x_1632); +if (x_1647 == 0) +{ +lean_object* x_1648; +lean_dec(x_1646); +if (lean_is_scalar(x_1645)) { + x_1648 = lean_alloc_ctor(1, 2, 0); +} else { + x_1648 = x_1645; +} +lean_ctor_set(x_1648, 0, x_1643); +lean_ctor_set(x_1648, 1, x_1644); +return x_1648; +} +else +{ +lean_object* x_1649; lean_object* x_1650; lean_object* x_1651; lean_object* x_1652; +lean_dec(x_1645); +lean_dec(x_1644); +x_1649 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_1650 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1650, 0, x_1649); +x_1651 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_1652 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1651, x_1650, x_1643); +if (lean_obj_tag(x_1652) == 0) +{ +lean_object* x_1653; lean_object* x_1654; lean_object* x_1655; lean_object* x_1656; +lean_dec(x_1646); +x_1653 = lean_ctor_get(x_1652, 0); +lean_inc(x_1653); +x_1654 = lean_ctor_get(x_1652, 1); +lean_inc(x_1654); +if (lean_is_exclusive(x_1652)) { + lean_ctor_release(x_1652, 0); + lean_ctor_release(x_1652, 1); + x_1655 = x_1652; +} else { + lean_dec_ref(x_1652); + x_1655 = lean_box(0); +} +if (lean_is_scalar(x_1655)) { + x_1656 = lean_alloc_ctor(0, 2, 0); +} else { + x_1656 = x_1655; +} +lean_ctor_set(x_1656, 0, x_1653); +lean_ctor_set(x_1656, 1, x_1654); +return x_1656; +} +else +{ +lean_object* x_1657; lean_object* x_1658; lean_object* x_1659; lean_object* x_1660; uint8_t x_1661; +x_1657 = lean_ctor_get(x_1652, 0); +lean_inc(x_1657); +x_1658 = lean_ctor_get(x_1652, 1); +lean_inc(x_1658); +if (lean_is_exclusive(x_1652)) { + lean_ctor_release(x_1652, 0); + lean_ctor_release(x_1652, 1); + x_1659 = x_1652; +} else { + lean_dec_ref(x_1652); + x_1659 = lean_box(0); +} +x_1660 = lean_ctor_get(x_1657, 1); +lean_inc(x_1660); +x_1661 = lean_nat_dec_eq(x_1646, x_1660); +lean_dec(x_1646); +if (x_1661 == 0) +{ +lean_object* x_1662; +lean_dec(x_1660); +if (lean_is_scalar(x_1659)) { + x_1662 = lean_alloc_ctor(1, 2, 0); +} else { + x_1662 = x_1659; +} +lean_ctor_set(x_1662, 0, x_1657); +lean_ctor_set(x_1662, 1, x_1658); +return x_1662; +} +else +{ +lean_object* x_1663; lean_object* x_1664; lean_object* x_1665; lean_object* x_1666; +lean_dec(x_1659); +lean_dec(x_1658); +x_1663 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_1664 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1664, 0, x_1663); +x_1665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_1666 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1665, x_1664, x_1657); +if (lean_obj_tag(x_1666) == 0) +{ +lean_object* x_1667; lean_object* x_1668; lean_object* x_1669; lean_object* x_1670; +lean_dec(x_1660); +x_1667 = lean_ctor_get(x_1666, 0); +lean_inc(x_1667); +x_1668 = lean_ctor_get(x_1666, 1); +lean_inc(x_1668); +if (lean_is_exclusive(x_1666)) { + lean_ctor_release(x_1666, 0); + lean_ctor_release(x_1666, 1); + x_1669 = x_1666; +} else { + lean_dec_ref(x_1666); + x_1669 = lean_box(0); +} +if (lean_is_scalar(x_1669)) { + x_1670 = lean_alloc_ctor(0, 2, 0); +} else { + x_1670 = x_1669; +} +lean_ctor_set(x_1670, 0, x_1667); +lean_ctor_set(x_1670, 1, x_1668); +return x_1670; +} +else +{ +lean_object* x_1671; lean_object* x_1672; lean_object* x_1673; lean_object* x_1674; uint8_t x_1675; +x_1671 = lean_ctor_get(x_1666, 0); +lean_inc(x_1671); +x_1672 = lean_ctor_get(x_1666, 1); +lean_inc(x_1672); +if (lean_is_exclusive(x_1666)) { + lean_ctor_release(x_1666, 0); + lean_ctor_release(x_1666, 1); + x_1673 = x_1666; +} else { + lean_dec_ref(x_1666); + x_1673 = lean_box(0); +} +x_1674 = lean_ctor_get(x_1671, 1); +lean_inc(x_1674); +x_1675 = lean_nat_dec_eq(x_1660, x_1674); +lean_dec(x_1660); +if (x_1675 == 0) +{ +lean_object* x_1676; +lean_dec(x_1674); +if (lean_is_scalar(x_1673)) { + x_1676 = lean_alloc_ctor(1, 2, 0); +} else { + x_1676 = x_1673; +} +lean_ctor_set(x_1676, 0, x_1671); +lean_ctor_set(x_1676, 1, x_1672); +return x_1676; +} +else +{ +lean_object* x_1677; lean_object* x_1678; lean_object* x_1679; lean_object* x_1680; +lean_dec(x_1673); +lean_dec(x_1672); +x_1677 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_1678 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1678, 0, x_1677); +x_1679 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1680 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1679, x_1678, x_1671); +if (lean_obj_tag(x_1680) == 0) +{ +lean_object* x_1681; lean_object* x_1682; lean_object* x_1683; lean_object* x_1684; +lean_dec(x_1674); +x_1681 = lean_ctor_get(x_1680, 0); +lean_inc(x_1681); +x_1682 = lean_ctor_get(x_1680, 1); +lean_inc(x_1682); +if (lean_is_exclusive(x_1680)) { + lean_ctor_release(x_1680, 0); + lean_ctor_release(x_1680, 1); + x_1683 = x_1680; +} else { + lean_dec_ref(x_1680); + x_1683 = lean_box(0); +} +if (lean_is_scalar(x_1683)) { + x_1684 = lean_alloc_ctor(0, 2, 0); +} else { + x_1684 = x_1683; +} +lean_ctor_set(x_1684, 0, x_1681); +lean_ctor_set(x_1684, 1, x_1682); +return x_1684; +} +else +{ +lean_object* x_1685; lean_object* x_1686; lean_object* x_1687; lean_object* x_1688; uint8_t x_1689; +x_1685 = lean_ctor_get(x_1680, 0); +lean_inc(x_1685); +x_1686 = lean_ctor_get(x_1680, 1); +lean_inc(x_1686); +if (lean_is_exclusive(x_1680)) { + lean_ctor_release(x_1680, 0); + lean_ctor_release(x_1680, 1); + x_1687 = x_1680; +} else { + lean_dec_ref(x_1680); + x_1687 = lean_box(0); +} +x_1688 = lean_ctor_get(x_1685, 1); +lean_inc(x_1688); +x_1689 = lean_nat_dec_eq(x_1674, x_1688); +lean_dec(x_1674); +if (x_1689 == 0) +{ +lean_object* x_1690; +lean_dec(x_1688); +if (lean_is_scalar(x_1687)) { + x_1690 = lean_alloc_ctor(1, 2, 0); +} else { + x_1690 = x_1687; +} +lean_ctor_set(x_1690, 0, x_1685); +lean_ctor_set(x_1690, 1, x_1686); +return x_1690; +} +else +{ +lean_object* x_1691; lean_object* x_1692; lean_object* x_1693; lean_object* x_1694; +lean_dec(x_1687); +lean_dec(x_1686); +x_1691 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1692 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1692, 0, x_1691); +x_1693 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1694 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1693, x_1692, x_1685); +if (lean_obj_tag(x_1694) == 0) +{ +lean_object* x_1695; lean_object* x_1696; lean_object* x_1697; lean_object* x_1698; +lean_dec(x_1688); +x_1695 = lean_ctor_get(x_1694, 0); +lean_inc(x_1695); +x_1696 = lean_ctor_get(x_1694, 1); +lean_inc(x_1696); +if (lean_is_exclusive(x_1694)) { + lean_ctor_release(x_1694, 0); + lean_ctor_release(x_1694, 1); + x_1697 = x_1694; +} else { + lean_dec_ref(x_1694); + x_1697 = lean_box(0); +} +if (lean_is_scalar(x_1697)) { + x_1698 = lean_alloc_ctor(0, 2, 0); +} else { + x_1698 = x_1697; +} +lean_ctor_set(x_1698, 0, x_1695); +lean_ctor_set(x_1698, 1, x_1696); +return x_1698; +} +else +{ +lean_object* x_1699; lean_object* x_1700; lean_object* x_1701; lean_object* x_1702; uint8_t x_1703; +x_1699 = lean_ctor_get(x_1694, 0); +lean_inc(x_1699); +x_1700 = lean_ctor_get(x_1694, 1); +lean_inc(x_1700); +if (lean_is_exclusive(x_1694)) { + lean_ctor_release(x_1694, 0); + lean_ctor_release(x_1694, 1); + x_1701 = x_1694; +} else { + lean_dec_ref(x_1694); + x_1701 = lean_box(0); +} +x_1702 = lean_ctor_get(x_1699, 1); +lean_inc(x_1702); +x_1703 = lean_nat_dec_eq(x_1688, x_1702); +lean_dec(x_1688); +if (x_1703 == 0) +{ +lean_object* x_1704; +lean_dec(x_1702); +if (lean_is_scalar(x_1701)) { + x_1704 = lean_alloc_ctor(1, 2, 0); +} else { + x_1704 = x_1701; +} +lean_ctor_set(x_1704, 0, x_1699); +lean_ctor_set(x_1704, 1, x_1700); +return x_1704; +} +else +{ +lean_object* x_1705; lean_object* x_1706; lean_object* x_1707; lean_object* x_1708; +lean_dec(x_1701); +lean_dec(x_1700); +x_1705 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1706 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1706, 0, x_1705); +x_1707 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1708 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1707, x_1706, x_1699); +if (lean_obj_tag(x_1708) == 0) +{ +lean_object* x_1709; lean_object* x_1710; lean_object* x_1711; lean_object* x_1712; +lean_dec(x_1702); +x_1709 = lean_ctor_get(x_1708, 0); +lean_inc(x_1709); +x_1710 = lean_ctor_get(x_1708, 1); +lean_inc(x_1710); +if (lean_is_exclusive(x_1708)) { + lean_ctor_release(x_1708, 0); + lean_ctor_release(x_1708, 1); + x_1711 = x_1708; +} else { + lean_dec_ref(x_1708); + x_1711 = lean_box(0); +} +if (lean_is_scalar(x_1711)) { + x_1712 = lean_alloc_ctor(0, 2, 0); +} else { + x_1712 = x_1711; +} +lean_ctor_set(x_1712, 0, x_1709); +lean_ctor_set(x_1712, 1, x_1710); +return x_1712; +} +else +{ +lean_object* x_1713; lean_object* x_1714; lean_object* x_1715; lean_object* x_1716; uint8_t x_1717; +x_1713 = lean_ctor_get(x_1708, 0); +lean_inc(x_1713); +x_1714 = lean_ctor_get(x_1708, 1); +lean_inc(x_1714); +if (lean_is_exclusive(x_1708)) { + lean_ctor_release(x_1708, 0); + lean_ctor_release(x_1708, 1); + x_1715 = x_1708; +} else { + lean_dec_ref(x_1708); + x_1715 = lean_box(0); +} +x_1716 = lean_ctor_get(x_1713, 1); +lean_inc(x_1716); +x_1717 = lean_nat_dec_eq(x_1702, x_1716); +lean_dec(x_1702); +if (x_1717 == 0) +{ +lean_object* x_1718; +lean_dec(x_1716); +if (lean_is_scalar(x_1715)) { + x_1718 = lean_alloc_ctor(1, 2, 0); +} else { + x_1718 = x_1715; +} +lean_ctor_set(x_1718, 0, x_1713); +lean_ctor_set(x_1718, 1, x_1714); +return x_1718; +} +else +{ +lean_object* x_1719; lean_object* x_1720; lean_object* x_1721; lean_object* x_1722; +lean_dec(x_1715); +lean_dec(x_1714); +x_1719 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1720 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1720, 0, x_1719); +x_1721 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1722 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1721, x_1720, x_1713); +if (lean_obj_tag(x_1722) == 0) +{ +lean_object* x_1723; lean_object* x_1724; lean_object* x_1725; lean_object* x_1726; +lean_dec(x_1716); +x_1723 = lean_ctor_get(x_1722, 0); +lean_inc(x_1723); +x_1724 = lean_ctor_get(x_1722, 1); +lean_inc(x_1724); +if (lean_is_exclusive(x_1722)) { + lean_ctor_release(x_1722, 0); + lean_ctor_release(x_1722, 1); + x_1725 = x_1722; +} else { + lean_dec_ref(x_1722); + x_1725 = lean_box(0); +} +if (lean_is_scalar(x_1725)) { + x_1726 = lean_alloc_ctor(0, 2, 0); +} else { + x_1726 = x_1725; +} +lean_ctor_set(x_1726, 0, x_1723); +lean_ctor_set(x_1726, 1, x_1724); +return x_1726; +} +else +{ +lean_object* x_1727; lean_object* x_1728; lean_object* x_1729; lean_object* x_1730; uint8_t x_1731; +x_1727 = lean_ctor_get(x_1722, 0); +lean_inc(x_1727); +x_1728 = lean_ctor_get(x_1722, 1); +lean_inc(x_1728); +if (lean_is_exclusive(x_1722)) { + lean_ctor_release(x_1722, 0); + lean_ctor_release(x_1722, 1); + x_1729 = x_1722; +} else { + lean_dec_ref(x_1722); + x_1729 = lean_box(0); +} +x_1730 = lean_ctor_get(x_1727, 1); +lean_inc(x_1730); +x_1731 = lean_nat_dec_eq(x_1716, x_1730); +lean_dec(x_1716); +if (x_1731 == 0) +{ +lean_object* x_1732; +lean_dec(x_1730); +if (lean_is_scalar(x_1729)) { + x_1732 = lean_alloc_ctor(1, 2, 0); +} else { + x_1732 = x_1729; +} +lean_ctor_set(x_1732, 0, x_1727); +lean_ctor_set(x_1732, 1, x_1728); +return x_1732; +} +else +{ +lean_object* x_1733; lean_object* x_1734; lean_object* x_1735; lean_object* x_1736; +lean_dec(x_1729); +lean_dec(x_1728); +x_1733 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1734 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1734, 0, x_1733); +x_1735 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1736 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1735, x_1734, x_1727); +if (lean_obj_tag(x_1736) == 0) +{ +lean_object* x_1737; lean_object* x_1738; lean_object* x_1739; lean_object* x_1740; +lean_dec(x_1730); +x_1737 = lean_ctor_get(x_1736, 0); +lean_inc(x_1737); +x_1738 = lean_ctor_get(x_1736, 1); +lean_inc(x_1738); +if (lean_is_exclusive(x_1736)) { + lean_ctor_release(x_1736, 0); + lean_ctor_release(x_1736, 1); + x_1739 = x_1736; +} else { + lean_dec_ref(x_1736); + x_1739 = lean_box(0); +} +if (lean_is_scalar(x_1739)) { + x_1740 = lean_alloc_ctor(0, 2, 0); +} else { + x_1740 = x_1739; +} +lean_ctor_set(x_1740, 0, x_1737); +lean_ctor_set(x_1740, 1, x_1738); +return x_1740; +} +else +{ +lean_object* x_1741; lean_object* x_1742; lean_object* x_1743; lean_object* x_1744; uint8_t x_1745; +x_1741 = lean_ctor_get(x_1736, 0); +lean_inc(x_1741); +x_1742 = lean_ctor_get(x_1736, 1); +lean_inc(x_1742); +if (lean_is_exclusive(x_1736)) { + lean_ctor_release(x_1736, 0); + lean_ctor_release(x_1736, 1); + x_1743 = x_1736; +} else { + lean_dec_ref(x_1736); + x_1743 = lean_box(0); +} +x_1744 = lean_ctor_get(x_1741, 1); +lean_inc(x_1744); +x_1745 = lean_nat_dec_eq(x_1730, x_1744); +lean_dec(x_1730); +if (x_1745 == 0) +{ +lean_object* x_1746; +lean_dec(x_1744); +if (lean_is_scalar(x_1743)) { + x_1746 = lean_alloc_ctor(1, 2, 0); +} else { + x_1746 = x_1743; +} +lean_ctor_set(x_1746, 0, x_1741); +lean_ctor_set(x_1746, 1, x_1742); +return x_1746; +} +else +{ +lean_object* x_1747; lean_object* x_1748; lean_object* x_1749; lean_object* x_1750; +lean_dec(x_1743); +lean_dec(x_1742); +x_1747 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1748 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1748, 0, x_1747); +x_1749 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1750 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1749, x_1748, x_1741); +if (lean_obj_tag(x_1750) == 0) +{ +lean_object* x_1751; lean_object* x_1752; lean_object* x_1753; lean_object* x_1754; +lean_dec(x_1744); +x_1751 = lean_ctor_get(x_1750, 0); +lean_inc(x_1751); +x_1752 = lean_ctor_get(x_1750, 1); +lean_inc(x_1752); +if (lean_is_exclusive(x_1750)) { + lean_ctor_release(x_1750, 0); + lean_ctor_release(x_1750, 1); + x_1753 = x_1750; +} else { + lean_dec_ref(x_1750); + x_1753 = lean_box(0); +} +if (lean_is_scalar(x_1753)) { + x_1754 = lean_alloc_ctor(0, 2, 0); +} else { + x_1754 = x_1753; +} +lean_ctor_set(x_1754, 0, x_1751); +lean_ctor_set(x_1754, 1, x_1752); +return x_1754; +} +else +{ +lean_object* x_1755; lean_object* x_1756; lean_object* x_1757; lean_object* x_1758; uint8_t x_1759; +x_1755 = lean_ctor_get(x_1750, 0); +lean_inc(x_1755); +x_1756 = lean_ctor_get(x_1750, 1); +lean_inc(x_1756); +if (lean_is_exclusive(x_1750)) { + lean_ctor_release(x_1750, 0); + lean_ctor_release(x_1750, 1); + x_1757 = x_1750; +} else { + lean_dec_ref(x_1750); + x_1757 = lean_box(0); +} +x_1758 = lean_ctor_get(x_1755, 1); +lean_inc(x_1758); +x_1759 = lean_nat_dec_eq(x_1744, x_1758); +lean_dec(x_1744); +if (x_1759 == 0) +{ +lean_object* x_1760; +lean_dec(x_1758); +if (lean_is_scalar(x_1757)) { + x_1760 = lean_alloc_ctor(1, 2, 0); +} else { + x_1760 = x_1757; +} +lean_ctor_set(x_1760, 0, x_1755); +lean_ctor_set(x_1760, 1, x_1756); +return x_1760; +} +else +{ +lean_object* x_1761; lean_object* x_1762; lean_object* x_1763; lean_object* x_1764; +lean_dec(x_1757); +lean_dec(x_1756); +x_1761 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1762 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1762, 0, x_1761); +x_1763 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1764 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1763, x_1762, x_1755); +if (lean_obj_tag(x_1764) == 0) +{ +lean_object* x_1765; lean_object* x_1766; lean_object* x_1767; lean_object* x_1768; +lean_dec(x_1758); +x_1765 = lean_ctor_get(x_1764, 0); +lean_inc(x_1765); +x_1766 = lean_ctor_get(x_1764, 1); +lean_inc(x_1766); +if (lean_is_exclusive(x_1764)) { + lean_ctor_release(x_1764, 0); + lean_ctor_release(x_1764, 1); + x_1767 = x_1764; +} else { + lean_dec_ref(x_1764); + x_1767 = lean_box(0); +} +if (lean_is_scalar(x_1767)) { + x_1768 = lean_alloc_ctor(0, 2, 0); +} else { + x_1768 = x_1767; +} +lean_ctor_set(x_1768, 0, x_1765); +lean_ctor_set(x_1768, 1, x_1766); +return x_1768; +} +else +{ +lean_object* x_1769; lean_object* x_1770; lean_object* x_1771; lean_object* x_1772; uint8_t x_1773; +x_1769 = lean_ctor_get(x_1764, 0); +lean_inc(x_1769); +x_1770 = lean_ctor_get(x_1764, 1); +lean_inc(x_1770); +if (lean_is_exclusive(x_1764)) { + lean_ctor_release(x_1764, 0); + lean_ctor_release(x_1764, 1); + x_1771 = x_1764; +} else { + lean_dec_ref(x_1764); + x_1771 = lean_box(0); +} +x_1772 = lean_ctor_get(x_1769, 1); +lean_inc(x_1772); +x_1773 = lean_nat_dec_eq(x_1758, x_1772); +lean_dec(x_1758); +if (x_1773 == 0) +{ +lean_object* x_1774; +lean_dec(x_1772); +if (lean_is_scalar(x_1771)) { + x_1774 = lean_alloc_ctor(1, 2, 0); +} else { + x_1774 = x_1771; +} +lean_ctor_set(x_1774, 0, x_1769); +lean_ctor_set(x_1774, 1, x_1770); +return x_1774; +} +else +{ +lean_object* x_1775; lean_object* x_1776; lean_object* x_1777; lean_object* x_1778; +lean_dec(x_1771); +lean_dec(x_1770); +x_1775 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1776 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1776, 0, x_1775); +x_1777 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1778 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1777, x_1776, x_1769); +if (lean_obj_tag(x_1778) == 0) +{ +lean_object* x_1779; lean_object* x_1780; lean_object* x_1781; lean_object* x_1782; +lean_dec(x_1772); +x_1779 = lean_ctor_get(x_1778, 0); +lean_inc(x_1779); +x_1780 = lean_ctor_get(x_1778, 1); +lean_inc(x_1780); +if (lean_is_exclusive(x_1778)) { + lean_ctor_release(x_1778, 0); + lean_ctor_release(x_1778, 1); + x_1781 = x_1778; +} else { + lean_dec_ref(x_1778); + x_1781 = lean_box(0); +} +if (lean_is_scalar(x_1781)) { + x_1782 = lean_alloc_ctor(0, 2, 0); +} else { + x_1782 = x_1781; +} +lean_ctor_set(x_1782, 0, x_1779); +lean_ctor_set(x_1782, 1, x_1780); +return x_1782; +} +else +{ +lean_object* x_1783; lean_object* x_1784; lean_object* x_1785; lean_object* x_1786; uint8_t x_1787; +x_1783 = lean_ctor_get(x_1778, 0); +lean_inc(x_1783); +x_1784 = lean_ctor_get(x_1778, 1); +lean_inc(x_1784); +if (lean_is_exclusive(x_1778)) { + lean_ctor_release(x_1778, 0); + lean_ctor_release(x_1778, 1); + x_1785 = x_1778; +} else { + lean_dec_ref(x_1778); + x_1785 = lean_box(0); +} +x_1786 = lean_ctor_get(x_1783, 1); +lean_inc(x_1786); +x_1787 = lean_nat_dec_eq(x_1772, x_1786); +lean_dec(x_1772); +if (x_1787 == 0) +{ +lean_object* x_1788; +lean_dec(x_1786); +if (lean_is_scalar(x_1785)) { + x_1788 = lean_alloc_ctor(1, 2, 0); +} else { + x_1788 = x_1785; +} +lean_ctor_set(x_1788, 0, x_1783); +lean_ctor_set(x_1788, 1, x_1784); +return x_1788; +} +else +{ +lean_object* x_1789; lean_object* x_1790; lean_object* x_1791; lean_object* x_1792; +lean_dec(x_1785); +lean_dec(x_1784); +x_1789 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1790 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1790, 0, x_1789); +x_1791 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1792 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1791, x_1790, x_1783); +if (lean_obj_tag(x_1792) == 0) +{ +lean_object* x_1793; lean_object* x_1794; lean_object* x_1795; lean_object* x_1796; +lean_dec(x_1786); +x_1793 = lean_ctor_get(x_1792, 0); +lean_inc(x_1793); +x_1794 = lean_ctor_get(x_1792, 1); +lean_inc(x_1794); +if (lean_is_exclusive(x_1792)) { + lean_ctor_release(x_1792, 0); + lean_ctor_release(x_1792, 1); + x_1795 = x_1792; +} else { + lean_dec_ref(x_1792); + x_1795 = lean_box(0); +} +if (lean_is_scalar(x_1795)) { + x_1796 = lean_alloc_ctor(0, 2, 0); +} else { + x_1796 = x_1795; +} +lean_ctor_set(x_1796, 0, x_1793); +lean_ctor_set(x_1796, 1, x_1794); +return x_1796; +} +else +{ +lean_object* x_1797; lean_object* x_1798; lean_object* x_1799; lean_object* x_1800; uint8_t x_1801; +x_1797 = lean_ctor_get(x_1792, 0); +lean_inc(x_1797); +x_1798 = lean_ctor_get(x_1792, 1); +lean_inc(x_1798); +if (lean_is_exclusive(x_1792)) { + lean_ctor_release(x_1792, 0); + lean_ctor_release(x_1792, 1); + x_1799 = x_1792; +} else { + lean_dec_ref(x_1792); + x_1799 = lean_box(0); +} +x_1800 = lean_ctor_get(x_1797, 1); +lean_inc(x_1800); +x_1801 = lean_nat_dec_eq(x_1786, x_1800); +lean_dec(x_1800); +lean_dec(x_1786); +if (x_1801 == 0) +{ +lean_object* x_1802; +if (lean_is_scalar(x_1799)) { + x_1802 = lean_alloc_ctor(1, 2, 0); +} else { + x_1802 = x_1799; +} +lean_ctor_set(x_1802, 0, x_1797); +lean_ctor_set(x_1802, 1, x_1798); +return x_1802; +} +else +{ +lean_object* x_1803; lean_object* x_1804; lean_object* x_1805; lean_object* x_1806; +lean_dec(x_1799); +lean_dec(x_1798); +x_1803 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_1804 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1804, 0, x_1803); +x_1805 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_1806 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1805, x_1804, x_1797); +return x_1806; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_1807; lean_object* x_1808; lean_object* x_1809; uint8_t x_1810; +x_1807 = lean_ctor_get(x_211, 0); +x_1808 = lean_ctor_get(x_211, 1); +lean_inc(x_1808); +lean_inc(x_1807); +lean_dec(x_211); +x_1809 = lean_ctor_get(x_1807, 1); +lean_inc(x_1809); +x_1810 = lean_nat_dec_eq(x_206, x_1809); +lean_dec(x_206); +if (x_1810 == 0) +{ +lean_object* x_1811; +lean_dec(x_1809); +x_1811 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_1811, 0, x_1807); +lean_ctor_set(x_1811, 1, x_1808); +return x_1811; +} +else +{ +lean_object* x_1812; lean_object* x_1813; lean_object* x_1814; lean_object* x_1815; +lean_dec(x_1808); +x_1812 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_1813 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1813, 0, x_1812); +x_1814 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_1815 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1814, x_1813, x_1807); +if (lean_obj_tag(x_1815) == 0) +{ +lean_object* x_1816; lean_object* x_1817; lean_object* x_1818; lean_object* x_1819; +lean_dec(x_1809); +x_1816 = lean_ctor_get(x_1815, 0); +lean_inc(x_1816); +x_1817 = lean_ctor_get(x_1815, 1); +lean_inc(x_1817); +if (lean_is_exclusive(x_1815)) { + lean_ctor_release(x_1815, 0); + lean_ctor_release(x_1815, 1); + x_1818 = x_1815; +} else { + lean_dec_ref(x_1815); + x_1818 = lean_box(0); +} +if (lean_is_scalar(x_1818)) { + x_1819 = lean_alloc_ctor(0, 2, 0); +} else { + x_1819 = x_1818; +} +lean_ctor_set(x_1819, 0, x_1816); +lean_ctor_set(x_1819, 1, x_1817); +return x_1819; +} +else +{ +lean_object* x_1820; lean_object* x_1821; lean_object* x_1822; lean_object* x_1823; uint8_t x_1824; +x_1820 = lean_ctor_get(x_1815, 0); +lean_inc(x_1820); +x_1821 = lean_ctor_get(x_1815, 1); +lean_inc(x_1821); +if (lean_is_exclusive(x_1815)) { + lean_ctor_release(x_1815, 0); + lean_ctor_release(x_1815, 1); + x_1822 = x_1815; +} else { + lean_dec_ref(x_1815); + x_1822 = lean_box(0); +} +x_1823 = lean_ctor_get(x_1820, 1); +lean_inc(x_1823); +x_1824 = lean_nat_dec_eq(x_1809, x_1823); +lean_dec(x_1809); +if (x_1824 == 0) +{ +lean_object* x_1825; +lean_dec(x_1823); +if (lean_is_scalar(x_1822)) { + x_1825 = lean_alloc_ctor(1, 2, 0); +} else { + x_1825 = x_1822; +} +lean_ctor_set(x_1825, 0, x_1820); +lean_ctor_set(x_1825, 1, x_1821); +return x_1825; +} +else +{ +lean_object* x_1826; lean_object* x_1827; lean_object* x_1828; lean_object* x_1829; +lean_dec(x_1822); +lean_dec(x_1821); +x_1826 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_1827 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1827, 0, x_1826); +x_1828 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_1829 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1828, x_1827, x_1820); +if (lean_obj_tag(x_1829) == 0) +{ +lean_object* x_1830; lean_object* x_1831; lean_object* x_1832; lean_object* x_1833; +lean_dec(x_1823); +x_1830 = lean_ctor_get(x_1829, 0); +lean_inc(x_1830); +x_1831 = lean_ctor_get(x_1829, 1); +lean_inc(x_1831); +if (lean_is_exclusive(x_1829)) { + lean_ctor_release(x_1829, 0); + lean_ctor_release(x_1829, 1); + x_1832 = x_1829; +} else { + lean_dec_ref(x_1829); + x_1832 = lean_box(0); +} +if (lean_is_scalar(x_1832)) { + x_1833 = lean_alloc_ctor(0, 2, 0); +} else { + x_1833 = x_1832; +} +lean_ctor_set(x_1833, 0, x_1830); +lean_ctor_set(x_1833, 1, x_1831); +return x_1833; +} +else +{ +lean_object* x_1834; lean_object* x_1835; lean_object* x_1836; lean_object* x_1837; uint8_t x_1838; +x_1834 = lean_ctor_get(x_1829, 0); +lean_inc(x_1834); +x_1835 = lean_ctor_get(x_1829, 1); +lean_inc(x_1835); +if (lean_is_exclusive(x_1829)) { + lean_ctor_release(x_1829, 0); + lean_ctor_release(x_1829, 1); + x_1836 = x_1829; +} else { + lean_dec_ref(x_1829); + x_1836 = lean_box(0); +} +x_1837 = lean_ctor_get(x_1834, 1); +lean_inc(x_1837); +x_1838 = lean_nat_dec_eq(x_1823, x_1837); +lean_dec(x_1823); +if (x_1838 == 0) +{ +lean_object* x_1839; +lean_dec(x_1837); +if (lean_is_scalar(x_1836)) { + x_1839 = lean_alloc_ctor(1, 2, 0); +} else { + x_1839 = x_1836; +} +lean_ctor_set(x_1839, 0, x_1834); +lean_ctor_set(x_1839, 1, x_1835); +return x_1839; +} +else +{ +lean_object* x_1840; lean_object* x_1841; lean_object* x_1842; lean_object* x_1843; +lean_dec(x_1836); +lean_dec(x_1835); +x_1840 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_1841 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1841, 0, x_1840); +x_1842 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_1843 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1842, x_1841, x_1834); +if (lean_obj_tag(x_1843) == 0) +{ +lean_object* x_1844; lean_object* x_1845; lean_object* x_1846; lean_object* x_1847; +lean_dec(x_1837); +x_1844 = lean_ctor_get(x_1843, 0); +lean_inc(x_1844); +x_1845 = lean_ctor_get(x_1843, 1); +lean_inc(x_1845); +if (lean_is_exclusive(x_1843)) { + lean_ctor_release(x_1843, 0); + lean_ctor_release(x_1843, 1); + x_1846 = x_1843; +} else { + lean_dec_ref(x_1843); + x_1846 = lean_box(0); +} +if (lean_is_scalar(x_1846)) { + x_1847 = lean_alloc_ctor(0, 2, 0); +} else { + x_1847 = x_1846; +} +lean_ctor_set(x_1847, 0, x_1844); +lean_ctor_set(x_1847, 1, x_1845); +return x_1847; +} +else +{ +lean_object* x_1848; lean_object* x_1849; lean_object* x_1850; lean_object* x_1851; uint8_t x_1852; +x_1848 = lean_ctor_get(x_1843, 0); +lean_inc(x_1848); +x_1849 = lean_ctor_get(x_1843, 1); +lean_inc(x_1849); +if (lean_is_exclusive(x_1843)) { + lean_ctor_release(x_1843, 0); + lean_ctor_release(x_1843, 1); + x_1850 = x_1843; +} else { + lean_dec_ref(x_1843); + x_1850 = lean_box(0); +} +x_1851 = lean_ctor_get(x_1848, 1); +lean_inc(x_1851); +x_1852 = lean_nat_dec_eq(x_1837, x_1851); +lean_dec(x_1837); +if (x_1852 == 0) +{ +lean_object* x_1853; +lean_dec(x_1851); +if (lean_is_scalar(x_1850)) { + x_1853 = lean_alloc_ctor(1, 2, 0); +} else { + x_1853 = x_1850; +} +lean_ctor_set(x_1853, 0, x_1848); +lean_ctor_set(x_1853, 1, x_1849); +return x_1853; +} +else +{ +lean_object* x_1854; lean_object* x_1855; lean_object* x_1856; lean_object* x_1857; +lean_dec(x_1850); +lean_dec(x_1849); +x_1854 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_1855 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1855, 0, x_1854); +x_1856 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_1857 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1856, x_1855, x_1848); +if (lean_obj_tag(x_1857) == 0) +{ +lean_object* x_1858; lean_object* x_1859; lean_object* x_1860; lean_object* x_1861; +lean_dec(x_1851); +x_1858 = lean_ctor_get(x_1857, 0); +lean_inc(x_1858); +x_1859 = lean_ctor_get(x_1857, 1); +lean_inc(x_1859); +if (lean_is_exclusive(x_1857)) { + lean_ctor_release(x_1857, 0); + lean_ctor_release(x_1857, 1); + x_1860 = x_1857; +} else { + lean_dec_ref(x_1857); + x_1860 = lean_box(0); +} +if (lean_is_scalar(x_1860)) { + x_1861 = lean_alloc_ctor(0, 2, 0); +} else { + x_1861 = x_1860; +} +lean_ctor_set(x_1861, 0, x_1858); +lean_ctor_set(x_1861, 1, x_1859); +return x_1861; +} +else +{ +lean_object* x_1862; lean_object* x_1863; lean_object* x_1864; lean_object* x_1865; uint8_t x_1866; +x_1862 = lean_ctor_get(x_1857, 0); +lean_inc(x_1862); +x_1863 = lean_ctor_get(x_1857, 1); +lean_inc(x_1863); +if (lean_is_exclusive(x_1857)) { + lean_ctor_release(x_1857, 0); + lean_ctor_release(x_1857, 1); + x_1864 = x_1857; +} else { + lean_dec_ref(x_1857); + x_1864 = lean_box(0); +} +x_1865 = lean_ctor_get(x_1862, 1); +lean_inc(x_1865); +x_1866 = lean_nat_dec_eq(x_1851, x_1865); +lean_dec(x_1851); +if (x_1866 == 0) +{ +lean_object* x_1867; +lean_dec(x_1865); +if (lean_is_scalar(x_1864)) { + x_1867 = lean_alloc_ctor(1, 2, 0); +} else { + x_1867 = x_1864; +} +lean_ctor_set(x_1867, 0, x_1862); +lean_ctor_set(x_1867, 1, x_1863); +return x_1867; +} +else +{ +lean_object* x_1868; lean_object* x_1869; lean_object* x_1870; lean_object* x_1871; +lean_dec(x_1864); +lean_dec(x_1863); +x_1868 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_1869 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1869, 0, x_1868); +x_1870 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_1871 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1870, x_1869, x_1862); +if (lean_obj_tag(x_1871) == 0) +{ +lean_object* x_1872; lean_object* x_1873; lean_object* x_1874; lean_object* x_1875; +lean_dec(x_1865); +x_1872 = lean_ctor_get(x_1871, 0); +lean_inc(x_1872); +x_1873 = lean_ctor_get(x_1871, 1); +lean_inc(x_1873); +if (lean_is_exclusive(x_1871)) { + lean_ctor_release(x_1871, 0); + lean_ctor_release(x_1871, 1); + x_1874 = x_1871; +} else { + lean_dec_ref(x_1871); + x_1874 = lean_box(0); +} +if (lean_is_scalar(x_1874)) { + x_1875 = lean_alloc_ctor(0, 2, 0); +} else { + x_1875 = x_1874; +} +lean_ctor_set(x_1875, 0, x_1872); +lean_ctor_set(x_1875, 1, x_1873); +return x_1875; +} +else +{ +lean_object* x_1876; lean_object* x_1877; lean_object* x_1878; lean_object* x_1879; uint8_t x_1880; +x_1876 = lean_ctor_get(x_1871, 0); +lean_inc(x_1876); +x_1877 = lean_ctor_get(x_1871, 1); +lean_inc(x_1877); +if (lean_is_exclusive(x_1871)) { + lean_ctor_release(x_1871, 0); + lean_ctor_release(x_1871, 1); + x_1878 = x_1871; +} else { + lean_dec_ref(x_1871); + x_1878 = lean_box(0); +} +x_1879 = lean_ctor_get(x_1876, 1); +lean_inc(x_1879); +x_1880 = lean_nat_dec_eq(x_1865, x_1879); +lean_dec(x_1865); +if (x_1880 == 0) +{ +lean_object* x_1881; +lean_dec(x_1879); +if (lean_is_scalar(x_1878)) { + x_1881 = lean_alloc_ctor(1, 2, 0); +} else { + x_1881 = x_1878; +} +lean_ctor_set(x_1881, 0, x_1876); +lean_ctor_set(x_1881, 1, x_1877); +return x_1881; +} +else +{ +lean_object* x_1882; lean_object* x_1883; lean_object* x_1884; lean_object* x_1885; +lean_dec(x_1878); +lean_dec(x_1877); +x_1882 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_1883 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1883, 0, x_1882); +x_1884 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_1885 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1884, x_1883, x_1876); +if (lean_obj_tag(x_1885) == 0) +{ +lean_object* x_1886; lean_object* x_1887; lean_object* x_1888; lean_object* x_1889; +lean_dec(x_1879); +x_1886 = lean_ctor_get(x_1885, 0); +lean_inc(x_1886); +x_1887 = lean_ctor_get(x_1885, 1); +lean_inc(x_1887); +if (lean_is_exclusive(x_1885)) { + lean_ctor_release(x_1885, 0); + lean_ctor_release(x_1885, 1); + x_1888 = x_1885; +} else { + lean_dec_ref(x_1885); + x_1888 = lean_box(0); +} +if (lean_is_scalar(x_1888)) { + x_1889 = lean_alloc_ctor(0, 2, 0); +} else { + x_1889 = x_1888; +} +lean_ctor_set(x_1889, 0, x_1886); +lean_ctor_set(x_1889, 1, x_1887); +return x_1889; +} +else +{ +lean_object* x_1890; lean_object* x_1891; lean_object* x_1892; lean_object* x_1893; uint8_t x_1894; +x_1890 = lean_ctor_get(x_1885, 0); +lean_inc(x_1890); +x_1891 = lean_ctor_get(x_1885, 1); +lean_inc(x_1891); +if (lean_is_exclusive(x_1885)) { + lean_ctor_release(x_1885, 0); + lean_ctor_release(x_1885, 1); + x_1892 = x_1885; +} else { + lean_dec_ref(x_1885); + x_1892 = lean_box(0); +} +x_1893 = lean_ctor_get(x_1890, 1); +lean_inc(x_1893); +x_1894 = lean_nat_dec_eq(x_1879, x_1893); +lean_dec(x_1879); +if (x_1894 == 0) +{ +lean_object* x_1895; +lean_dec(x_1893); +if (lean_is_scalar(x_1892)) { + x_1895 = lean_alloc_ctor(1, 2, 0); +} else { + x_1895 = x_1892; +} +lean_ctor_set(x_1895, 0, x_1890); +lean_ctor_set(x_1895, 1, x_1891); +return x_1895; +} +else +{ +lean_object* x_1896; lean_object* x_1897; lean_object* x_1898; lean_object* x_1899; +lean_dec(x_1892); +lean_dec(x_1891); +x_1896 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_1897 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1897, 0, x_1896); +x_1898 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_1899 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1898, x_1897, x_1890); +if (lean_obj_tag(x_1899) == 0) +{ +lean_object* x_1900; lean_object* x_1901; lean_object* x_1902; lean_object* x_1903; +lean_dec(x_1893); +x_1900 = lean_ctor_get(x_1899, 0); +lean_inc(x_1900); +x_1901 = lean_ctor_get(x_1899, 1); +lean_inc(x_1901); +if (lean_is_exclusive(x_1899)) { + lean_ctor_release(x_1899, 0); + lean_ctor_release(x_1899, 1); + x_1902 = x_1899; +} else { + lean_dec_ref(x_1899); + x_1902 = lean_box(0); +} +if (lean_is_scalar(x_1902)) { + x_1903 = lean_alloc_ctor(0, 2, 0); +} else { + x_1903 = x_1902; +} +lean_ctor_set(x_1903, 0, x_1900); +lean_ctor_set(x_1903, 1, x_1901); +return x_1903; +} +else +{ +lean_object* x_1904; lean_object* x_1905; lean_object* x_1906; lean_object* x_1907; uint8_t x_1908; +x_1904 = lean_ctor_get(x_1899, 0); +lean_inc(x_1904); +x_1905 = lean_ctor_get(x_1899, 1); +lean_inc(x_1905); +if (lean_is_exclusive(x_1899)) { + lean_ctor_release(x_1899, 0); + lean_ctor_release(x_1899, 1); + x_1906 = x_1899; +} else { + lean_dec_ref(x_1899); + x_1906 = lean_box(0); +} +x_1907 = lean_ctor_get(x_1904, 1); +lean_inc(x_1907); +x_1908 = lean_nat_dec_eq(x_1893, x_1907); +lean_dec(x_1893); +if (x_1908 == 0) +{ +lean_object* x_1909; +lean_dec(x_1907); +if (lean_is_scalar(x_1906)) { + x_1909 = lean_alloc_ctor(1, 2, 0); +} else { + x_1909 = x_1906; +} +lean_ctor_set(x_1909, 0, x_1904); +lean_ctor_set(x_1909, 1, x_1905); +return x_1909; +} +else +{ +lean_object* x_1910; lean_object* x_1911; lean_object* x_1912; lean_object* x_1913; +lean_dec(x_1906); +lean_dec(x_1905); +x_1910 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_1911 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1911, 0, x_1910); +x_1912 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_1913 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1912, x_1911, x_1904); +if (lean_obj_tag(x_1913) == 0) +{ +lean_object* x_1914; lean_object* x_1915; lean_object* x_1916; lean_object* x_1917; +lean_dec(x_1907); +x_1914 = lean_ctor_get(x_1913, 0); +lean_inc(x_1914); +x_1915 = lean_ctor_get(x_1913, 1); +lean_inc(x_1915); +if (lean_is_exclusive(x_1913)) { + lean_ctor_release(x_1913, 0); + lean_ctor_release(x_1913, 1); + x_1916 = x_1913; +} else { + lean_dec_ref(x_1913); + x_1916 = lean_box(0); +} +if (lean_is_scalar(x_1916)) { + x_1917 = lean_alloc_ctor(0, 2, 0); +} else { + x_1917 = x_1916; +} +lean_ctor_set(x_1917, 0, x_1914); +lean_ctor_set(x_1917, 1, x_1915); +return x_1917; +} +else +{ +lean_object* x_1918; lean_object* x_1919; lean_object* x_1920; lean_object* x_1921; uint8_t x_1922; +x_1918 = lean_ctor_get(x_1913, 0); +lean_inc(x_1918); +x_1919 = lean_ctor_get(x_1913, 1); +lean_inc(x_1919); +if (lean_is_exclusive(x_1913)) { + lean_ctor_release(x_1913, 0); + lean_ctor_release(x_1913, 1); + x_1920 = x_1913; +} else { + lean_dec_ref(x_1913); + x_1920 = lean_box(0); +} +x_1921 = lean_ctor_get(x_1918, 1); +lean_inc(x_1921); +x_1922 = lean_nat_dec_eq(x_1907, x_1921); +lean_dec(x_1907); +if (x_1922 == 0) +{ +lean_object* x_1923; +lean_dec(x_1921); +if (lean_is_scalar(x_1920)) { + x_1923 = lean_alloc_ctor(1, 2, 0); +} else { + x_1923 = x_1920; +} +lean_ctor_set(x_1923, 0, x_1918); +lean_ctor_set(x_1923, 1, x_1919); +return x_1923; +} +else +{ +lean_object* x_1924; lean_object* x_1925; lean_object* x_1926; lean_object* x_1927; +lean_dec(x_1920); +lean_dec(x_1919); +x_1924 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_1925 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1925, 0, x_1924); +x_1926 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_1927 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1926, x_1925, x_1918); +if (lean_obj_tag(x_1927) == 0) +{ +lean_object* x_1928; lean_object* x_1929; lean_object* x_1930; lean_object* x_1931; +lean_dec(x_1921); +x_1928 = lean_ctor_get(x_1927, 0); +lean_inc(x_1928); +x_1929 = lean_ctor_get(x_1927, 1); +lean_inc(x_1929); +if (lean_is_exclusive(x_1927)) { + lean_ctor_release(x_1927, 0); + lean_ctor_release(x_1927, 1); + x_1930 = x_1927; +} else { + lean_dec_ref(x_1927); + x_1930 = lean_box(0); +} +if (lean_is_scalar(x_1930)) { + x_1931 = lean_alloc_ctor(0, 2, 0); +} else { + x_1931 = x_1930; +} +lean_ctor_set(x_1931, 0, x_1928); +lean_ctor_set(x_1931, 1, x_1929); +return x_1931; +} +else +{ +lean_object* x_1932; lean_object* x_1933; lean_object* x_1934; lean_object* x_1935; uint8_t x_1936; +x_1932 = lean_ctor_get(x_1927, 0); +lean_inc(x_1932); +x_1933 = lean_ctor_get(x_1927, 1); +lean_inc(x_1933); +if (lean_is_exclusive(x_1927)) { + lean_ctor_release(x_1927, 0); + lean_ctor_release(x_1927, 1); + x_1934 = x_1927; +} else { + lean_dec_ref(x_1927); + x_1934 = lean_box(0); +} +x_1935 = lean_ctor_get(x_1932, 1); +lean_inc(x_1935); +x_1936 = lean_nat_dec_eq(x_1921, x_1935); +lean_dec(x_1921); +if (x_1936 == 0) +{ +lean_object* x_1937; +lean_dec(x_1935); +if (lean_is_scalar(x_1934)) { + x_1937 = lean_alloc_ctor(1, 2, 0); +} else { + x_1937 = x_1934; +} +lean_ctor_set(x_1937, 0, x_1932); +lean_ctor_set(x_1937, 1, x_1933); +return x_1937; +} +else +{ +lean_object* x_1938; lean_object* x_1939; lean_object* x_1940; lean_object* x_1941; +lean_dec(x_1934); +lean_dec(x_1933); +x_1938 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_1939 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1939, 0, x_1938); +x_1940 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_1941 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1940, x_1939, x_1932); +if (lean_obj_tag(x_1941) == 0) +{ +lean_object* x_1942; lean_object* x_1943; lean_object* x_1944; lean_object* x_1945; +lean_dec(x_1935); +x_1942 = lean_ctor_get(x_1941, 0); +lean_inc(x_1942); +x_1943 = lean_ctor_get(x_1941, 1); +lean_inc(x_1943); +if (lean_is_exclusive(x_1941)) { + lean_ctor_release(x_1941, 0); + lean_ctor_release(x_1941, 1); + x_1944 = x_1941; +} else { + lean_dec_ref(x_1941); + x_1944 = lean_box(0); +} +if (lean_is_scalar(x_1944)) { + x_1945 = lean_alloc_ctor(0, 2, 0); +} else { + x_1945 = x_1944; +} +lean_ctor_set(x_1945, 0, x_1942); +lean_ctor_set(x_1945, 1, x_1943); +return x_1945; +} +else +{ +lean_object* x_1946; lean_object* x_1947; lean_object* x_1948; lean_object* x_1949; uint8_t x_1950; +x_1946 = lean_ctor_get(x_1941, 0); +lean_inc(x_1946); +x_1947 = lean_ctor_get(x_1941, 1); +lean_inc(x_1947); +if (lean_is_exclusive(x_1941)) { + lean_ctor_release(x_1941, 0); + lean_ctor_release(x_1941, 1); + x_1948 = x_1941; +} else { + lean_dec_ref(x_1941); + x_1948 = lean_box(0); +} +x_1949 = lean_ctor_get(x_1946, 1); +lean_inc(x_1949); +x_1950 = lean_nat_dec_eq(x_1935, x_1949); +lean_dec(x_1935); +if (x_1950 == 0) +{ +lean_object* x_1951; +lean_dec(x_1949); +if (lean_is_scalar(x_1948)) { + x_1951 = lean_alloc_ctor(1, 2, 0); +} else { + x_1951 = x_1948; +} +lean_ctor_set(x_1951, 0, x_1946); +lean_ctor_set(x_1951, 1, x_1947); +return x_1951; +} +else +{ +lean_object* x_1952; lean_object* x_1953; lean_object* x_1954; lean_object* x_1955; +lean_dec(x_1948); +lean_dec(x_1947); +x_1952 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_1953 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1953, 0, x_1952); +x_1954 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_1955 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1954, x_1953, x_1946); +if (lean_obj_tag(x_1955) == 0) +{ +lean_object* x_1956; lean_object* x_1957; lean_object* x_1958; lean_object* x_1959; +lean_dec(x_1949); +x_1956 = lean_ctor_get(x_1955, 0); +lean_inc(x_1956); +x_1957 = lean_ctor_get(x_1955, 1); +lean_inc(x_1957); +if (lean_is_exclusive(x_1955)) { + lean_ctor_release(x_1955, 0); + lean_ctor_release(x_1955, 1); + x_1958 = x_1955; +} else { + lean_dec_ref(x_1955); + x_1958 = lean_box(0); +} +if (lean_is_scalar(x_1958)) { + x_1959 = lean_alloc_ctor(0, 2, 0); +} else { + x_1959 = x_1958; +} +lean_ctor_set(x_1959, 0, x_1956); +lean_ctor_set(x_1959, 1, x_1957); +return x_1959; +} +else +{ +lean_object* x_1960; lean_object* x_1961; lean_object* x_1962; lean_object* x_1963; uint8_t x_1964; +x_1960 = lean_ctor_get(x_1955, 0); +lean_inc(x_1960); +x_1961 = lean_ctor_get(x_1955, 1); +lean_inc(x_1961); +if (lean_is_exclusive(x_1955)) { + lean_ctor_release(x_1955, 0); + lean_ctor_release(x_1955, 1); + x_1962 = x_1955; +} else { + lean_dec_ref(x_1955); + x_1962 = lean_box(0); +} +x_1963 = lean_ctor_get(x_1960, 1); +lean_inc(x_1963); +x_1964 = lean_nat_dec_eq(x_1949, x_1963); +lean_dec(x_1949); +if (x_1964 == 0) +{ +lean_object* x_1965; +lean_dec(x_1963); +if (lean_is_scalar(x_1962)) { + x_1965 = lean_alloc_ctor(1, 2, 0); +} else { + x_1965 = x_1962; +} +lean_ctor_set(x_1965, 0, x_1960); +lean_ctor_set(x_1965, 1, x_1961); +return x_1965; +} +else +{ +lean_object* x_1966; lean_object* x_1967; lean_object* x_1968; lean_object* x_1969; +lean_dec(x_1962); +lean_dec(x_1961); +x_1966 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_1967 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1967, 0, x_1966); +x_1968 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_1969 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1968, x_1967, x_1960); +if (lean_obj_tag(x_1969) == 0) +{ +lean_object* x_1970; lean_object* x_1971; lean_object* x_1972; lean_object* x_1973; +lean_dec(x_1963); +x_1970 = lean_ctor_get(x_1969, 0); +lean_inc(x_1970); +x_1971 = lean_ctor_get(x_1969, 1); +lean_inc(x_1971); +if (lean_is_exclusive(x_1969)) { + lean_ctor_release(x_1969, 0); + lean_ctor_release(x_1969, 1); + x_1972 = x_1969; +} else { + lean_dec_ref(x_1969); + x_1972 = lean_box(0); +} +if (lean_is_scalar(x_1972)) { + x_1973 = lean_alloc_ctor(0, 2, 0); +} else { + x_1973 = x_1972; +} +lean_ctor_set(x_1973, 0, x_1970); +lean_ctor_set(x_1973, 1, x_1971); +return x_1973; +} +else +{ +lean_object* x_1974; lean_object* x_1975; lean_object* x_1976; lean_object* x_1977; uint8_t x_1978; +x_1974 = lean_ctor_get(x_1969, 0); +lean_inc(x_1974); +x_1975 = lean_ctor_get(x_1969, 1); +lean_inc(x_1975); +if (lean_is_exclusive(x_1969)) { + lean_ctor_release(x_1969, 0); + lean_ctor_release(x_1969, 1); + x_1976 = x_1969; +} else { + lean_dec_ref(x_1969); + x_1976 = lean_box(0); +} +x_1977 = lean_ctor_get(x_1974, 1); +lean_inc(x_1977); +x_1978 = lean_nat_dec_eq(x_1963, x_1977); +lean_dec(x_1963); +if (x_1978 == 0) +{ +lean_object* x_1979; +lean_dec(x_1977); +if (lean_is_scalar(x_1976)) { + x_1979 = lean_alloc_ctor(1, 2, 0); +} else { + x_1979 = x_1976; +} +lean_ctor_set(x_1979, 0, x_1974); +lean_ctor_set(x_1979, 1, x_1975); +return x_1979; +} +else +{ +lean_object* x_1980; lean_object* x_1981; lean_object* x_1982; lean_object* x_1983; +lean_dec(x_1976); +lean_dec(x_1975); +x_1980 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_1981 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1981, 0, x_1980); +x_1982 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_1983 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1982, x_1981, x_1974); +if (lean_obj_tag(x_1983) == 0) +{ +lean_object* x_1984; lean_object* x_1985; lean_object* x_1986; lean_object* x_1987; +lean_dec(x_1977); +x_1984 = lean_ctor_get(x_1983, 0); +lean_inc(x_1984); +x_1985 = lean_ctor_get(x_1983, 1); +lean_inc(x_1985); +if (lean_is_exclusive(x_1983)) { + lean_ctor_release(x_1983, 0); + lean_ctor_release(x_1983, 1); + x_1986 = x_1983; +} else { + lean_dec_ref(x_1983); + x_1986 = lean_box(0); +} +if (lean_is_scalar(x_1986)) { + x_1987 = lean_alloc_ctor(0, 2, 0); +} else { + x_1987 = x_1986; +} +lean_ctor_set(x_1987, 0, x_1984); +lean_ctor_set(x_1987, 1, x_1985); +return x_1987; +} +else +{ +lean_object* x_1988; lean_object* x_1989; lean_object* x_1990; lean_object* x_1991; uint8_t x_1992; +x_1988 = lean_ctor_get(x_1983, 0); +lean_inc(x_1988); +x_1989 = lean_ctor_get(x_1983, 1); +lean_inc(x_1989); +if (lean_is_exclusive(x_1983)) { + lean_ctor_release(x_1983, 0); + lean_ctor_release(x_1983, 1); + x_1990 = x_1983; +} else { + lean_dec_ref(x_1983); + x_1990 = lean_box(0); +} +x_1991 = lean_ctor_get(x_1988, 1); +lean_inc(x_1991); +x_1992 = lean_nat_dec_eq(x_1977, x_1991); +lean_dec(x_1977); +if (x_1992 == 0) +{ +lean_object* x_1993; +lean_dec(x_1991); +if (lean_is_scalar(x_1990)) { + x_1993 = lean_alloc_ctor(1, 2, 0); +} else { + x_1993 = x_1990; +} +lean_ctor_set(x_1993, 0, x_1988); +lean_ctor_set(x_1993, 1, x_1989); +return x_1993; +} +else +{ +lean_object* x_1994; lean_object* x_1995; lean_object* x_1996; lean_object* x_1997; +lean_dec(x_1990); +lean_dec(x_1989); +x_1994 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_1995 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_1995, 0, x_1994); +x_1996 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_1997 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_1996, x_1995, x_1988); +if (lean_obj_tag(x_1997) == 0) +{ +lean_object* x_1998; lean_object* x_1999; lean_object* x_2000; lean_object* x_2001; +lean_dec(x_1991); +x_1998 = lean_ctor_get(x_1997, 0); +lean_inc(x_1998); +x_1999 = lean_ctor_get(x_1997, 1); +lean_inc(x_1999); +if (lean_is_exclusive(x_1997)) { + lean_ctor_release(x_1997, 0); + lean_ctor_release(x_1997, 1); + x_2000 = x_1997; +} else { + lean_dec_ref(x_1997); + x_2000 = lean_box(0); +} +if (lean_is_scalar(x_2000)) { + x_2001 = lean_alloc_ctor(0, 2, 0); +} else { + x_2001 = x_2000; +} +lean_ctor_set(x_2001, 0, x_1998); +lean_ctor_set(x_2001, 1, x_1999); +return x_2001; +} +else +{ +lean_object* x_2002; lean_object* x_2003; lean_object* x_2004; lean_object* x_2005; uint8_t x_2006; +x_2002 = lean_ctor_get(x_1997, 0); +lean_inc(x_2002); +x_2003 = lean_ctor_get(x_1997, 1); +lean_inc(x_2003); +if (lean_is_exclusive(x_1997)) { + lean_ctor_release(x_1997, 0); + lean_ctor_release(x_1997, 1); + x_2004 = x_1997; +} else { + lean_dec_ref(x_1997); + x_2004 = lean_box(0); +} +x_2005 = lean_ctor_get(x_2002, 1); +lean_inc(x_2005); +x_2006 = lean_nat_dec_eq(x_1991, x_2005); +lean_dec(x_2005); +lean_dec(x_1991); +if (x_2006 == 0) +{ +lean_object* x_2007; +if (lean_is_scalar(x_2004)) { + x_2007 = lean_alloc_ctor(1, 2, 0); +} else { + x_2007 = x_2004; +} +lean_ctor_set(x_2007, 0, x_2002); +lean_ctor_set(x_2007, 1, x_2003); +return x_2007; +} +else +{ +lean_object* x_2008; lean_object* x_2009; lean_object* x_2010; lean_object* x_2011; +lean_dec(x_2004); +lean_dec(x_2003); +x_2008 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_2009 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2009, 0, x_2008); +x_2010 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_2011 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2010, x_2009, x_2002); +return x_2011; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_2012; lean_object* x_2013; lean_object* x_2014; uint8_t x_2015; +x_2012 = lean_ctor_get(x_198, 0); +x_2013 = lean_ctor_get(x_198, 1); +lean_inc(x_2013); +lean_inc(x_2012); +lean_dec(x_198); +x_2014 = lean_ctor_get(x_2012, 1); +lean_inc(x_2014); +x_2015 = lean_nat_dec_eq(x_193, x_2014); +lean_dec(x_193); +if (x_2015 == 0) +{ +lean_object* x_2016; +lean_dec(x_2014); +x_2016 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2016, 0, x_2012); +lean_ctor_set(x_2016, 1, x_2013); +return x_2016; +} +else +{ +lean_object* x_2017; lean_object* x_2018; lean_object* x_2019; lean_object* x_2020; +lean_dec(x_2013); +x_2017 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_2018 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2018, 0, x_2017); +x_2019 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_2020 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2019, x_2018, x_2012); +if (lean_obj_tag(x_2020) == 0) +{ +lean_object* x_2021; lean_object* x_2022; lean_object* x_2023; lean_object* x_2024; +lean_dec(x_2014); +x_2021 = lean_ctor_get(x_2020, 0); +lean_inc(x_2021); +x_2022 = lean_ctor_get(x_2020, 1); +lean_inc(x_2022); +if (lean_is_exclusive(x_2020)) { + lean_ctor_release(x_2020, 0); + lean_ctor_release(x_2020, 1); + x_2023 = x_2020; +} else { + lean_dec_ref(x_2020); + x_2023 = lean_box(0); +} +if (lean_is_scalar(x_2023)) { + x_2024 = lean_alloc_ctor(0, 2, 0); +} else { + x_2024 = x_2023; +} +lean_ctor_set(x_2024, 0, x_2021); +lean_ctor_set(x_2024, 1, x_2022); +return x_2024; +} +else +{ +lean_object* x_2025; lean_object* x_2026; lean_object* x_2027; lean_object* x_2028; uint8_t x_2029; +x_2025 = lean_ctor_get(x_2020, 0); +lean_inc(x_2025); +x_2026 = lean_ctor_get(x_2020, 1); +lean_inc(x_2026); +if (lean_is_exclusive(x_2020)) { + lean_ctor_release(x_2020, 0); + lean_ctor_release(x_2020, 1); + x_2027 = x_2020; +} else { + lean_dec_ref(x_2020); + x_2027 = lean_box(0); +} +x_2028 = lean_ctor_get(x_2025, 1); +lean_inc(x_2028); +x_2029 = lean_nat_dec_eq(x_2014, x_2028); +lean_dec(x_2014); +if (x_2029 == 0) +{ +lean_object* x_2030; +lean_dec(x_2028); +if (lean_is_scalar(x_2027)) { + x_2030 = lean_alloc_ctor(1, 2, 0); +} else { + x_2030 = x_2027; +} +lean_ctor_set(x_2030, 0, x_2025); +lean_ctor_set(x_2030, 1, x_2026); +return x_2030; +} +else +{ +lean_object* x_2031; lean_object* x_2032; lean_object* x_2033; lean_object* x_2034; +lean_dec(x_2027); +lean_dec(x_2026); +x_2031 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_2032 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2032, 0, x_2031); +x_2033 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_2034 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2033, x_2032, x_2025); +if (lean_obj_tag(x_2034) == 0) +{ +lean_object* x_2035; lean_object* x_2036; lean_object* x_2037; lean_object* x_2038; +lean_dec(x_2028); +x_2035 = lean_ctor_get(x_2034, 0); +lean_inc(x_2035); +x_2036 = lean_ctor_get(x_2034, 1); +lean_inc(x_2036); +if (lean_is_exclusive(x_2034)) { + lean_ctor_release(x_2034, 0); + lean_ctor_release(x_2034, 1); + x_2037 = x_2034; +} else { + lean_dec_ref(x_2034); + x_2037 = lean_box(0); +} +if (lean_is_scalar(x_2037)) { + x_2038 = lean_alloc_ctor(0, 2, 0); +} else { + x_2038 = x_2037; +} +lean_ctor_set(x_2038, 0, x_2035); +lean_ctor_set(x_2038, 1, x_2036); +return x_2038; +} +else +{ +lean_object* x_2039; lean_object* x_2040; lean_object* x_2041; lean_object* x_2042; uint8_t x_2043; +x_2039 = lean_ctor_get(x_2034, 0); +lean_inc(x_2039); +x_2040 = lean_ctor_get(x_2034, 1); +lean_inc(x_2040); +if (lean_is_exclusive(x_2034)) { + lean_ctor_release(x_2034, 0); + lean_ctor_release(x_2034, 1); + x_2041 = x_2034; +} else { + lean_dec_ref(x_2034); + x_2041 = lean_box(0); +} +x_2042 = lean_ctor_get(x_2039, 1); +lean_inc(x_2042); +x_2043 = lean_nat_dec_eq(x_2028, x_2042); +lean_dec(x_2028); +if (x_2043 == 0) +{ +lean_object* x_2044; +lean_dec(x_2042); +if (lean_is_scalar(x_2041)) { + x_2044 = lean_alloc_ctor(1, 2, 0); +} else { + x_2044 = x_2041; +} +lean_ctor_set(x_2044, 0, x_2039); +lean_ctor_set(x_2044, 1, x_2040); +return x_2044; +} +else +{ +lean_object* x_2045; lean_object* x_2046; lean_object* x_2047; lean_object* x_2048; +lean_dec(x_2041); +lean_dec(x_2040); +x_2045 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_2046 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2046, 0, x_2045); +x_2047 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_2048 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2047, x_2046, x_2039); +if (lean_obj_tag(x_2048) == 0) +{ +lean_object* x_2049; lean_object* x_2050; lean_object* x_2051; lean_object* x_2052; +lean_dec(x_2042); +x_2049 = lean_ctor_get(x_2048, 0); +lean_inc(x_2049); +x_2050 = lean_ctor_get(x_2048, 1); +lean_inc(x_2050); +if (lean_is_exclusive(x_2048)) { + lean_ctor_release(x_2048, 0); + lean_ctor_release(x_2048, 1); + x_2051 = x_2048; +} else { + lean_dec_ref(x_2048); + x_2051 = lean_box(0); +} +if (lean_is_scalar(x_2051)) { + x_2052 = lean_alloc_ctor(0, 2, 0); +} else { + x_2052 = x_2051; +} +lean_ctor_set(x_2052, 0, x_2049); +lean_ctor_set(x_2052, 1, x_2050); +return x_2052; +} +else +{ +lean_object* x_2053; lean_object* x_2054; lean_object* x_2055; lean_object* x_2056; uint8_t x_2057; +x_2053 = lean_ctor_get(x_2048, 0); +lean_inc(x_2053); +x_2054 = lean_ctor_get(x_2048, 1); +lean_inc(x_2054); +if (lean_is_exclusive(x_2048)) { + lean_ctor_release(x_2048, 0); + lean_ctor_release(x_2048, 1); + x_2055 = x_2048; +} else { + lean_dec_ref(x_2048); + x_2055 = lean_box(0); +} +x_2056 = lean_ctor_get(x_2053, 1); +lean_inc(x_2056); +x_2057 = lean_nat_dec_eq(x_2042, x_2056); +lean_dec(x_2042); +if (x_2057 == 0) +{ +lean_object* x_2058; +lean_dec(x_2056); +if (lean_is_scalar(x_2055)) { + x_2058 = lean_alloc_ctor(1, 2, 0); +} else { + x_2058 = x_2055; +} +lean_ctor_set(x_2058, 0, x_2053); +lean_ctor_set(x_2058, 1, x_2054); +return x_2058; +} +else +{ +lean_object* x_2059; lean_object* x_2060; lean_object* x_2061; lean_object* x_2062; +lean_dec(x_2055); +lean_dec(x_2054); +x_2059 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_2060 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2060, 0, x_2059); +x_2061 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_2062 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2061, x_2060, x_2053); +if (lean_obj_tag(x_2062) == 0) +{ +lean_object* x_2063; lean_object* x_2064; lean_object* x_2065; lean_object* x_2066; +lean_dec(x_2056); +x_2063 = lean_ctor_get(x_2062, 0); +lean_inc(x_2063); +x_2064 = lean_ctor_get(x_2062, 1); +lean_inc(x_2064); +if (lean_is_exclusive(x_2062)) { + lean_ctor_release(x_2062, 0); + lean_ctor_release(x_2062, 1); + x_2065 = x_2062; +} else { + lean_dec_ref(x_2062); + x_2065 = lean_box(0); +} +if (lean_is_scalar(x_2065)) { + x_2066 = lean_alloc_ctor(0, 2, 0); +} else { + x_2066 = x_2065; +} +lean_ctor_set(x_2066, 0, x_2063); +lean_ctor_set(x_2066, 1, x_2064); +return x_2066; +} +else +{ +lean_object* x_2067; lean_object* x_2068; lean_object* x_2069; lean_object* x_2070; uint8_t x_2071; +x_2067 = lean_ctor_get(x_2062, 0); +lean_inc(x_2067); +x_2068 = lean_ctor_get(x_2062, 1); +lean_inc(x_2068); +if (lean_is_exclusive(x_2062)) { + lean_ctor_release(x_2062, 0); + lean_ctor_release(x_2062, 1); + x_2069 = x_2062; +} else { + lean_dec_ref(x_2062); + x_2069 = lean_box(0); +} +x_2070 = lean_ctor_get(x_2067, 1); +lean_inc(x_2070); +x_2071 = lean_nat_dec_eq(x_2056, x_2070); +lean_dec(x_2056); +if (x_2071 == 0) +{ +lean_object* x_2072; +lean_dec(x_2070); +if (lean_is_scalar(x_2069)) { + x_2072 = lean_alloc_ctor(1, 2, 0); +} else { + x_2072 = x_2069; +} +lean_ctor_set(x_2072, 0, x_2067); +lean_ctor_set(x_2072, 1, x_2068); +return x_2072; +} +else +{ +lean_object* x_2073; lean_object* x_2074; lean_object* x_2075; lean_object* x_2076; +lean_dec(x_2069); +lean_dec(x_2068); +x_2073 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_2074 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2074, 0, x_2073); +x_2075 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_2076 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2075, x_2074, x_2067); +if (lean_obj_tag(x_2076) == 0) +{ +lean_object* x_2077; lean_object* x_2078; lean_object* x_2079; lean_object* x_2080; +lean_dec(x_2070); +x_2077 = lean_ctor_get(x_2076, 0); +lean_inc(x_2077); +x_2078 = lean_ctor_get(x_2076, 1); +lean_inc(x_2078); +if (lean_is_exclusive(x_2076)) { + lean_ctor_release(x_2076, 0); + lean_ctor_release(x_2076, 1); + x_2079 = x_2076; +} else { + lean_dec_ref(x_2076); + x_2079 = lean_box(0); +} +if (lean_is_scalar(x_2079)) { + x_2080 = lean_alloc_ctor(0, 2, 0); +} else { + x_2080 = x_2079; +} +lean_ctor_set(x_2080, 0, x_2077); +lean_ctor_set(x_2080, 1, x_2078); +return x_2080; +} +else +{ +lean_object* x_2081; lean_object* x_2082; lean_object* x_2083; lean_object* x_2084; uint8_t x_2085; +x_2081 = lean_ctor_get(x_2076, 0); +lean_inc(x_2081); +x_2082 = lean_ctor_get(x_2076, 1); +lean_inc(x_2082); +if (lean_is_exclusive(x_2076)) { + lean_ctor_release(x_2076, 0); + lean_ctor_release(x_2076, 1); + x_2083 = x_2076; +} else { + lean_dec_ref(x_2076); + x_2083 = lean_box(0); +} +x_2084 = lean_ctor_get(x_2081, 1); +lean_inc(x_2084); +x_2085 = lean_nat_dec_eq(x_2070, x_2084); +lean_dec(x_2070); +if (x_2085 == 0) +{ +lean_object* x_2086; +lean_dec(x_2084); +if (lean_is_scalar(x_2083)) { + x_2086 = lean_alloc_ctor(1, 2, 0); +} else { + x_2086 = x_2083; +} +lean_ctor_set(x_2086, 0, x_2081); +lean_ctor_set(x_2086, 1, x_2082); +return x_2086; +} +else +{ +lean_object* x_2087; lean_object* x_2088; lean_object* x_2089; lean_object* x_2090; +lean_dec(x_2083); +lean_dec(x_2082); +x_2087 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_2088 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2088, 0, x_2087); +x_2089 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_2090 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2089, x_2088, x_2081); +if (lean_obj_tag(x_2090) == 0) +{ +lean_object* x_2091; lean_object* x_2092; lean_object* x_2093; lean_object* x_2094; +lean_dec(x_2084); +x_2091 = lean_ctor_get(x_2090, 0); +lean_inc(x_2091); +x_2092 = lean_ctor_get(x_2090, 1); +lean_inc(x_2092); +if (lean_is_exclusive(x_2090)) { + lean_ctor_release(x_2090, 0); + lean_ctor_release(x_2090, 1); + x_2093 = x_2090; +} else { + lean_dec_ref(x_2090); + x_2093 = lean_box(0); +} +if (lean_is_scalar(x_2093)) { + x_2094 = lean_alloc_ctor(0, 2, 0); +} else { + x_2094 = x_2093; +} +lean_ctor_set(x_2094, 0, x_2091); +lean_ctor_set(x_2094, 1, x_2092); +return x_2094; +} +else +{ +lean_object* x_2095; lean_object* x_2096; lean_object* x_2097; lean_object* x_2098; uint8_t x_2099; +x_2095 = lean_ctor_get(x_2090, 0); +lean_inc(x_2095); +x_2096 = lean_ctor_get(x_2090, 1); +lean_inc(x_2096); +if (lean_is_exclusive(x_2090)) { + lean_ctor_release(x_2090, 0); + lean_ctor_release(x_2090, 1); + x_2097 = x_2090; +} else { + lean_dec_ref(x_2090); + x_2097 = lean_box(0); +} +x_2098 = lean_ctor_get(x_2095, 1); +lean_inc(x_2098); +x_2099 = lean_nat_dec_eq(x_2084, x_2098); +lean_dec(x_2084); +if (x_2099 == 0) +{ +lean_object* x_2100; +lean_dec(x_2098); +if (lean_is_scalar(x_2097)) { + x_2100 = lean_alloc_ctor(1, 2, 0); +} else { + x_2100 = x_2097; +} +lean_ctor_set(x_2100, 0, x_2095); +lean_ctor_set(x_2100, 1, x_2096); +return x_2100; +} +else +{ +lean_object* x_2101; lean_object* x_2102; lean_object* x_2103; lean_object* x_2104; +lean_dec(x_2097); +lean_dec(x_2096); +x_2101 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_2102 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2102, 0, x_2101); +x_2103 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_2104 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2103, x_2102, x_2095); +if (lean_obj_tag(x_2104) == 0) +{ +lean_object* x_2105; lean_object* x_2106; lean_object* x_2107; lean_object* x_2108; +lean_dec(x_2098); +x_2105 = lean_ctor_get(x_2104, 0); +lean_inc(x_2105); +x_2106 = lean_ctor_get(x_2104, 1); +lean_inc(x_2106); +if (lean_is_exclusive(x_2104)) { + lean_ctor_release(x_2104, 0); + lean_ctor_release(x_2104, 1); + x_2107 = x_2104; +} else { + lean_dec_ref(x_2104); + x_2107 = lean_box(0); +} +if (lean_is_scalar(x_2107)) { + x_2108 = lean_alloc_ctor(0, 2, 0); +} else { + x_2108 = x_2107; +} +lean_ctor_set(x_2108, 0, x_2105); +lean_ctor_set(x_2108, 1, x_2106); +return x_2108; +} +else +{ +lean_object* x_2109; lean_object* x_2110; lean_object* x_2111; lean_object* x_2112; uint8_t x_2113; +x_2109 = lean_ctor_get(x_2104, 0); +lean_inc(x_2109); +x_2110 = lean_ctor_get(x_2104, 1); +lean_inc(x_2110); +if (lean_is_exclusive(x_2104)) { + lean_ctor_release(x_2104, 0); + lean_ctor_release(x_2104, 1); + x_2111 = x_2104; +} else { + lean_dec_ref(x_2104); + x_2111 = lean_box(0); +} +x_2112 = lean_ctor_get(x_2109, 1); +lean_inc(x_2112); +x_2113 = lean_nat_dec_eq(x_2098, x_2112); +lean_dec(x_2098); +if (x_2113 == 0) +{ +lean_object* x_2114; +lean_dec(x_2112); +if (lean_is_scalar(x_2111)) { + x_2114 = lean_alloc_ctor(1, 2, 0); +} else { + x_2114 = x_2111; +} +lean_ctor_set(x_2114, 0, x_2109); +lean_ctor_set(x_2114, 1, x_2110); +return x_2114; +} +else +{ +lean_object* x_2115; lean_object* x_2116; lean_object* x_2117; lean_object* x_2118; +lean_dec(x_2111); +lean_dec(x_2110); +x_2115 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_2116 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2116, 0, x_2115); +x_2117 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_2118 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2117, x_2116, x_2109); +if (lean_obj_tag(x_2118) == 0) +{ +lean_object* x_2119; lean_object* x_2120; lean_object* x_2121; lean_object* x_2122; +lean_dec(x_2112); +x_2119 = lean_ctor_get(x_2118, 0); +lean_inc(x_2119); +x_2120 = lean_ctor_get(x_2118, 1); +lean_inc(x_2120); +if (lean_is_exclusive(x_2118)) { + lean_ctor_release(x_2118, 0); + lean_ctor_release(x_2118, 1); + x_2121 = x_2118; +} else { + lean_dec_ref(x_2118); + x_2121 = lean_box(0); +} +if (lean_is_scalar(x_2121)) { + x_2122 = lean_alloc_ctor(0, 2, 0); +} else { + x_2122 = x_2121; +} +lean_ctor_set(x_2122, 0, x_2119); +lean_ctor_set(x_2122, 1, x_2120); +return x_2122; +} +else +{ +lean_object* x_2123; lean_object* x_2124; lean_object* x_2125; lean_object* x_2126; uint8_t x_2127; +x_2123 = lean_ctor_get(x_2118, 0); +lean_inc(x_2123); +x_2124 = lean_ctor_get(x_2118, 1); +lean_inc(x_2124); +if (lean_is_exclusive(x_2118)) { + lean_ctor_release(x_2118, 0); + lean_ctor_release(x_2118, 1); + x_2125 = x_2118; +} else { + lean_dec_ref(x_2118); + x_2125 = lean_box(0); +} +x_2126 = lean_ctor_get(x_2123, 1); +lean_inc(x_2126); +x_2127 = lean_nat_dec_eq(x_2112, x_2126); +lean_dec(x_2112); +if (x_2127 == 0) +{ +lean_object* x_2128; +lean_dec(x_2126); +if (lean_is_scalar(x_2125)) { + x_2128 = lean_alloc_ctor(1, 2, 0); +} else { + x_2128 = x_2125; +} +lean_ctor_set(x_2128, 0, x_2123); +lean_ctor_set(x_2128, 1, x_2124); +return x_2128; +} +else +{ +lean_object* x_2129; lean_object* x_2130; lean_object* x_2131; lean_object* x_2132; +lean_dec(x_2125); +lean_dec(x_2124); +x_2129 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_2130 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2130, 0, x_2129); +x_2131 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_2132 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2131, x_2130, x_2123); +if (lean_obj_tag(x_2132) == 0) +{ +lean_object* x_2133; lean_object* x_2134; lean_object* x_2135; lean_object* x_2136; +lean_dec(x_2126); +x_2133 = lean_ctor_get(x_2132, 0); +lean_inc(x_2133); +x_2134 = lean_ctor_get(x_2132, 1); +lean_inc(x_2134); +if (lean_is_exclusive(x_2132)) { + lean_ctor_release(x_2132, 0); + lean_ctor_release(x_2132, 1); + x_2135 = x_2132; +} else { + lean_dec_ref(x_2132); + x_2135 = lean_box(0); +} +if (lean_is_scalar(x_2135)) { + x_2136 = lean_alloc_ctor(0, 2, 0); +} else { + x_2136 = x_2135; +} +lean_ctor_set(x_2136, 0, x_2133); +lean_ctor_set(x_2136, 1, x_2134); +return x_2136; +} +else +{ +lean_object* x_2137; lean_object* x_2138; lean_object* x_2139; lean_object* x_2140; uint8_t x_2141; +x_2137 = lean_ctor_get(x_2132, 0); +lean_inc(x_2137); +x_2138 = lean_ctor_get(x_2132, 1); +lean_inc(x_2138); +if (lean_is_exclusive(x_2132)) { + lean_ctor_release(x_2132, 0); + lean_ctor_release(x_2132, 1); + x_2139 = x_2132; +} else { + lean_dec_ref(x_2132); + x_2139 = lean_box(0); +} +x_2140 = lean_ctor_get(x_2137, 1); +lean_inc(x_2140); +x_2141 = lean_nat_dec_eq(x_2126, x_2140); +lean_dec(x_2126); +if (x_2141 == 0) +{ +lean_object* x_2142; +lean_dec(x_2140); +if (lean_is_scalar(x_2139)) { + x_2142 = lean_alloc_ctor(1, 2, 0); +} else { + x_2142 = x_2139; +} +lean_ctor_set(x_2142, 0, x_2137); +lean_ctor_set(x_2142, 1, x_2138); +return x_2142; +} +else +{ +lean_object* x_2143; lean_object* x_2144; lean_object* x_2145; lean_object* x_2146; +lean_dec(x_2139); +lean_dec(x_2138); +x_2143 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_2144 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2144, 0, x_2143); +x_2145 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_2146 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2145, x_2144, x_2137); +if (lean_obj_tag(x_2146) == 0) +{ +lean_object* x_2147; lean_object* x_2148; lean_object* x_2149; lean_object* x_2150; +lean_dec(x_2140); +x_2147 = lean_ctor_get(x_2146, 0); +lean_inc(x_2147); +x_2148 = lean_ctor_get(x_2146, 1); +lean_inc(x_2148); +if (lean_is_exclusive(x_2146)) { + lean_ctor_release(x_2146, 0); + lean_ctor_release(x_2146, 1); + x_2149 = x_2146; +} else { + lean_dec_ref(x_2146); + x_2149 = lean_box(0); +} +if (lean_is_scalar(x_2149)) { + x_2150 = lean_alloc_ctor(0, 2, 0); +} else { + x_2150 = x_2149; +} +lean_ctor_set(x_2150, 0, x_2147); +lean_ctor_set(x_2150, 1, x_2148); +return x_2150; +} +else +{ +lean_object* x_2151; lean_object* x_2152; lean_object* x_2153; lean_object* x_2154; uint8_t x_2155; +x_2151 = lean_ctor_get(x_2146, 0); +lean_inc(x_2151); +x_2152 = lean_ctor_get(x_2146, 1); +lean_inc(x_2152); +if (lean_is_exclusive(x_2146)) { + lean_ctor_release(x_2146, 0); + lean_ctor_release(x_2146, 1); + x_2153 = x_2146; +} else { + lean_dec_ref(x_2146); + x_2153 = lean_box(0); +} +x_2154 = lean_ctor_get(x_2151, 1); +lean_inc(x_2154); +x_2155 = lean_nat_dec_eq(x_2140, x_2154); +lean_dec(x_2140); +if (x_2155 == 0) +{ +lean_object* x_2156; +lean_dec(x_2154); +if (lean_is_scalar(x_2153)) { + x_2156 = lean_alloc_ctor(1, 2, 0); +} else { + x_2156 = x_2153; +} +lean_ctor_set(x_2156, 0, x_2151); +lean_ctor_set(x_2156, 1, x_2152); +return x_2156; +} +else +{ +lean_object* x_2157; lean_object* x_2158; lean_object* x_2159; lean_object* x_2160; +lean_dec(x_2153); +lean_dec(x_2152); +x_2157 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_2158 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2158, 0, x_2157); +x_2159 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_2160 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2159, x_2158, x_2151); +if (lean_obj_tag(x_2160) == 0) +{ +lean_object* x_2161; lean_object* x_2162; lean_object* x_2163; lean_object* x_2164; +lean_dec(x_2154); +x_2161 = lean_ctor_get(x_2160, 0); +lean_inc(x_2161); +x_2162 = lean_ctor_get(x_2160, 1); +lean_inc(x_2162); +if (lean_is_exclusive(x_2160)) { + lean_ctor_release(x_2160, 0); + lean_ctor_release(x_2160, 1); + x_2163 = x_2160; +} else { + lean_dec_ref(x_2160); + x_2163 = lean_box(0); +} +if (lean_is_scalar(x_2163)) { + x_2164 = lean_alloc_ctor(0, 2, 0); +} else { + x_2164 = x_2163; +} +lean_ctor_set(x_2164, 0, x_2161); +lean_ctor_set(x_2164, 1, x_2162); +return x_2164; +} +else +{ +lean_object* x_2165; lean_object* x_2166; lean_object* x_2167; lean_object* x_2168; uint8_t x_2169; +x_2165 = lean_ctor_get(x_2160, 0); +lean_inc(x_2165); +x_2166 = lean_ctor_get(x_2160, 1); +lean_inc(x_2166); +if (lean_is_exclusive(x_2160)) { + lean_ctor_release(x_2160, 0); + lean_ctor_release(x_2160, 1); + x_2167 = x_2160; +} else { + lean_dec_ref(x_2160); + x_2167 = lean_box(0); +} +x_2168 = lean_ctor_get(x_2165, 1); +lean_inc(x_2168); +x_2169 = lean_nat_dec_eq(x_2154, x_2168); +lean_dec(x_2154); +if (x_2169 == 0) +{ +lean_object* x_2170; +lean_dec(x_2168); +if (lean_is_scalar(x_2167)) { + x_2170 = lean_alloc_ctor(1, 2, 0); +} else { + x_2170 = x_2167; +} +lean_ctor_set(x_2170, 0, x_2165); +lean_ctor_set(x_2170, 1, x_2166); +return x_2170; +} +else +{ +lean_object* x_2171; lean_object* x_2172; lean_object* x_2173; lean_object* x_2174; +lean_dec(x_2167); +lean_dec(x_2166); +x_2171 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_2172 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2172, 0, x_2171); +x_2173 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_2174 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2173, x_2172, x_2165); +if (lean_obj_tag(x_2174) == 0) +{ +lean_object* x_2175; lean_object* x_2176; lean_object* x_2177; lean_object* x_2178; +lean_dec(x_2168); +x_2175 = lean_ctor_get(x_2174, 0); +lean_inc(x_2175); +x_2176 = lean_ctor_get(x_2174, 1); +lean_inc(x_2176); +if (lean_is_exclusive(x_2174)) { + lean_ctor_release(x_2174, 0); + lean_ctor_release(x_2174, 1); + x_2177 = x_2174; +} else { + lean_dec_ref(x_2174); + x_2177 = lean_box(0); +} +if (lean_is_scalar(x_2177)) { + x_2178 = lean_alloc_ctor(0, 2, 0); +} else { + x_2178 = x_2177; +} +lean_ctor_set(x_2178, 0, x_2175); +lean_ctor_set(x_2178, 1, x_2176); +return x_2178; +} +else +{ +lean_object* x_2179; lean_object* x_2180; lean_object* x_2181; lean_object* x_2182; uint8_t x_2183; +x_2179 = lean_ctor_get(x_2174, 0); +lean_inc(x_2179); +x_2180 = lean_ctor_get(x_2174, 1); +lean_inc(x_2180); +if (lean_is_exclusive(x_2174)) { + lean_ctor_release(x_2174, 0); + lean_ctor_release(x_2174, 1); + x_2181 = x_2174; +} else { + lean_dec_ref(x_2174); + x_2181 = lean_box(0); +} +x_2182 = lean_ctor_get(x_2179, 1); +lean_inc(x_2182); +x_2183 = lean_nat_dec_eq(x_2168, x_2182); +lean_dec(x_2168); +if (x_2183 == 0) +{ +lean_object* x_2184; +lean_dec(x_2182); +if (lean_is_scalar(x_2181)) { + x_2184 = lean_alloc_ctor(1, 2, 0); +} else { + x_2184 = x_2181; +} +lean_ctor_set(x_2184, 0, x_2179); +lean_ctor_set(x_2184, 1, x_2180); +return x_2184; +} +else +{ +lean_object* x_2185; lean_object* x_2186; lean_object* x_2187; lean_object* x_2188; +lean_dec(x_2181); +lean_dec(x_2180); +x_2185 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_2186 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2186, 0, x_2185); +x_2187 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_2188 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2187, x_2186, x_2179); +if (lean_obj_tag(x_2188) == 0) +{ +lean_object* x_2189; lean_object* x_2190; lean_object* x_2191; lean_object* x_2192; +lean_dec(x_2182); +x_2189 = lean_ctor_get(x_2188, 0); +lean_inc(x_2189); +x_2190 = lean_ctor_get(x_2188, 1); +lean_inc(x_2190); +if (lean_is_exclusive(x_2188)) { + lean_ctor_release(x_2188, 0); + lean_ctor_release(x_2188, 1); + x_2191 = x_2188; +} else { + lean_dec_ref(x_2188); + x_2191 = lean_box(0); +} +if (lean_is_scalar(x_2191)) { + x_2192 = lean_alloc_ctor(0, 2, 0); +} else { + x_2192 = x_2191; +} +lean_ctor_set(x_2192, 0, x_2189); +lean_ctor_set(x_2192, 1, x_2190); +return x_2192; +} +else +{ +lean_object* x_2193; lean_object* x_2194; lean_object* x_2195; lean_object* x_2196; uint8_t x_2197; +x_2193 = lean_ctor_get(x_2188, 0); +lean_inc(x_2193); +x_2194 = lean_ctor_get(x_2188, 1); +lean_inc(x_2194); +if (lean_is_exclusive(x_2188)) { + lean_ctor_release(x_2188, 0); + lean_ctor_release(x_2188, 1); + x_2195 = x_2188; +} else { + lean_dec_ref(x_2188); + x_2195 = lean_box(0); +} +x_2196 = lean_ctor_get(x_2193, 1); +lean_inc(x_2196); +x_2197 = lean_nat_dec_eq(x_2182, x_2196); +lean_dec(x_2182); +if (x_2197 == 0) +{ +lean_object* x_2198; +lean_dec(x_2196); +if (lean_is_scalar(x_2195)) { + x_2198 = lean_alloc_ctor(1, 2, 0); +} else { + x_2198 = x_2195; +} +lean_ctor_set(x_2198, 0, x_2193); +lean_ctor_set(x_2198, 1, x_2194); +return x_2198; +} +else +{ +lean_object* x_2199; lean_object* x_2200; lean_object* x_2201; lean_object* x_2202; +lean_dec(x_2195); +lean_dec(x_2194); +x_2199 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_2200 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2200, 0, x_2199); +x_2201 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_2202 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2201, x_2200, x_2193); +if (lean_obj_tag(x_2202) == 0) +{ +lean_object* x_2203; lean_object* x_2204; lean_object* x_2205; lean_object* x_2206; +lean_dec(x_2196); +x_2203 = lean_ctor_get(x_2202, 0); +lean_inc(x_2203); +x_2204 = lean_ctor_get(x_2202, 1); +lean_inc(x_2204); +if (lean_is_exclusive(x_2202)) { + lean_ctor_release(x_2202, 0); + lean_ctor_release(x_2202, 1); + x_2205 = x_2202; +} else { + lean_dec_ref(x_2202); + x_2205 = lean_box(0); +} +if (lean_is_scalar(x_2205)) { + x_2206 = lean_alloc_ctor(0, 2, 0); +} else { + x_2206 = x_2205; +} +lean_ctor_set(x_2206, 0, x_2203); +lean_ctor_set(x_2206, 1, x_2204); +return x_2206; +} +else +{ +lean_object* x_2207; lean_object* x_2208; lean_object* x_2209; lean_object* x_2210; uint8_t x_2211; +x_2207 = lean_ctor_get(x_2202, 0); +lean_inc(x_2207); +x_2208 = lean_ctor_get(x_2202, 1); +lean_inc(x_2208); +if (lean_is_exclusive(x_2202)) { + lean_ctor_release(x_2202, 0); + lean_ctor_release(x_2202, 1); + x_2209 = x_2202; +} else { + lean_dec_ref(x_2202); + x_2209 = lean_box(0); +} +x_2210 = lean_ctor_get(x_2207, 1); +lean_inc(x_2210); +x_2211 = lean_nat_dec_eq(x_2196, x_2210); +lean_dec(x_2196); +if (x_2211 == 0) +{ +lean_object* x_2212; +lean_dec(x_2210); +if (lean_is_scalar(x_2209)) { + x_2212 = lean_alloc_ctor(1, 2, 0); +} else { + x_2212 = x_2209; +} +lean_ctor_set(x_2212, 0, x_2207); +lean_ctor_set(x_2212, 1, x_2208); +return x_2212; +} +else +{ +lean_object* x_2213; lean_object* x_2214; lean_object* x_2215; lean_object* x_2216; +lean_dec(x_2209); +lean_dec(x_2208); +x_2213 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_2214 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2214, 0, x_2213); +x_2215 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_2216 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2215, x_2214, x_2207); +if (lean_obj_tag(x_2216) == 0) +{ +lean_object* x_2217; lean_object* x_2218; lean_object* x_2219; lean_object* x_2220; +lean_dec(x_2210); +x_2217 = lean_ctor_get(x_2216, 0); +lean_inc(x_2217); +x_2218 = lean_ctor_get(x_2216, 1); +lean_inc(x_2218); +if (lean_is_exclusive(x_2216)) { + lean_ctor_release(x_2216, 0); + lean_ctor_release(x_2216, 1); + x_2219 = x_2216; +} else { + lean_dec_ref(x_2216); + x_2219 = lean_box(0); +} +if (lean_is_scalar(x_2219)) { + x_2220 = lean_alloc_ctor(0, 2, 0); +} else { + x_2220 = x_2219; +} +lean_ctor_set(x_2220, 0, x_2217); +lean_ctor_set(x_2220, 1, x_2218); +return x_2220; +} +else +{ +lean_object* x_2221; lean_object* x_2222; lean_object* x_2223; lean_object* x_2224; uint8_t x_2225; +x_2221 = lean_ctor_get(x_2216, 0); +lean_inc(x_2221); +x_2222 = lean_ctor_get(x_2216, 1); +lean_inc(x_2222); +if (lean_is_exclusive(x_2216)) { + lean_ctor_release(x_2216, 0); + lean_ctor_release(x_2216, 1); + x_2223 = x_2216; +} else { + lean_dec_ref(x_2216); + x_2223 = lean_box(0); +} +x_2224 = lean_ctor_get(x_2221, 1); +lean_inc(x_2224); +x_2225 = lean_nat_dec_eq(x_2210, x_2224); +lean_dec(x_2224); +lean_dec(x_2210); +if (x_2225 == 0) +{ +lean_object* x_2226; +if (lean_is_scalar(x_2223)) { + x_2226 = lean_alloc_ctor(1, 2, 0); +} else { + x_2226 = x_2223; +} +lean_ctor_set(x_2226, 0, x_2221); +lean_ctor_set(x_2226, 1, x_2222); +return x_2226; +} +else +{ +lean_object* x_2227; lean_object* x_2228; lean_object* x_2229; lean_object* x_2230; +lean_dec(x_2223); +lean_dec(x_2222); +x_2227 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_2228 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2228, 0, x_2227); +x_2229 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_2230 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2229, x_2228, x_2221); +return x_2230; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_2231; lean_object* x_2232; lean_object* x_2233; uint8_t x_2234; +x_2231 = lean_ctor_get(x_185, 0); +x_2232 = lean_ctor_get(x_185, 1); +lean_inc(x_2232); +lean_inc(x_2231); +lean_dec(x_185); +x_2233 = lean_ctor_get(x_2231, 1); +lean_inc(x_2233); +x_2234 = lean_nat_dec_eq(x_180, x_2233); +lean_dec(x_180); +if (x_2234 == 0) +{ +lean_object* x_2235; +lean_dec(x_2233); +x_2235 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2235, 0, x_2231); +lean_ctor_set(x_2235, 1, x_2232); +return x_2235; +} +else +{ +lean_object* x_2236; lean_object* x_2237; lean_object* x_2238; lean_object* x_2239; +lean_dec(x_2232); +x_2236 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_2237 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2237, 0, x_2236); +x_2238 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_2239 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2238, x_2237, x_2231); +if (lean_obj_tag(x_2239) == 0) +{ +lean_object* x_2240; lean_object* x_2241; lean_object* x_2242; lean_object* x_2243; +lean_dec(x_2233); +x_2240 = lean_ctor_get(x_2239, 0); +lean_inc(x_2240); +x_2241 = lean_ctor_get(x_2239, 1); +lean_inc(x_2241); +if (lean_is_exclusive(x_2239)) { + lean_ctor_release(x_2239, 0); + lean_ctor_release(x_2239, 1); + x_2242 = x_2239; +} else { + lean_dec_ref(x_2239); + x_2242 = lean_box(0); +} +if (lean_is_scalar(x_2242)) { + x_2243 = lean_alloc_ctor(0, 2, 0); +} else { + x_2243 = x_2242; +} +lean_ctor_set(x_2243, 0, x_2240); +lean_ctor_set(x_2243, 1, x_2241); +return x_2243; +} +else +{ +lean_object* x_2244; lean_object* x_2245; lean_object* x_2246; lean_object* x_2247; uint8_t x_2248; +x_2244 = lean_ctor_get(x_2239, 0); +lean_inc(x_2244); +x_2245 = lean_ctor_get(x_2239, 1); +lean_inc(x_2245); +if (lean_is_exclusive(x_2239)) { + lean_ctor_release(x_2239, 0); + lean_ctor_release(x_2239, 1); + x_2246 = x_2239; +} else { + lean_dec_ref(x_2239); + x_2246 = lean_box(0); +} +x_2247 = lean_ctor_get(x_2244, 1); +lean_inc(x_2247); +x_2248 = lean_nat_dec_eq(x_2233, x_2247); +lean_dec(x_2233); +if (x_2248 == 0) +{ +lean_object* x_2249; +lean_dec(x_2247); +if (lean_is_scalar(x_2246)) { + x_2249 = lean_alloc_ctor(1, 2, 0); +} else { + x_2249 = x_2246; +} +lean_ctor_set(x_2249, 0, x_2244); +lean_ctor_set(x_2249, 1, x_2245); +return x_2249; +} +else +{ +lean_object* x_2250; lean_object* x_2251; lean_object* x_2252; lean_object* x_2253; +lean_dec(x_2246); +lean_dec(x_2245); +x_2250 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_2251 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2251, 0, x_2250); +x_2252 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_2253 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2252, x_2251, x_2244); +if (lean_obj_tag(x_2253) == 0) +{ +lean_object* x_2254; lean_object* x_2255; lean_object* x_2256; lean_object* x_2257; +lean_dec(x_2247); +x_2254 = lean_ctor_get(x_2253, 0); +lean_inc(x_2254); +x_2255 = lean_ctor_get(x_2253, 1); +lean_inc(x_2255); +if (lean_is_exclusive(x_2253)) { + lean_ctor_release(x_2253, 0); + lean_ctor_release(x_2253, 1); + x_2256 = x_2253; +} else { + lean_dec_ref(x_2253); + x_2256 = lean_box(0); +} +if (lean_is_scalar(x_2256)) { + x_2257 = lean_alloc_ctor(0, 2, 0); +} else { + x_2257 = x_2256; +} +lean_ctor_set(x_2257, 0, x_2254); +lean_ctor_set(x_2257, 1, x_2255); +return x_2257; +} +else +{ +lean_object* x_2258; lean_object* x_2259; lean_object* x_2260; lean_object* x_2261; uint8_t x_2262; +x_2258 = lean_ctor_get(x_2253, 0); +lean_inc(x_2258); +x_2259 = lean_ctor_get(x_2253, 1); +lean_inc(x_2259); +if (lean_is_exclusive(x_2253)) { + lean_ctor_release(x_2253, 0); + lean_ctor_release(x_2253, 1); + x_2260 = x_2253; +} else { + lean_dec_ref(x_2253); + x_2260 = lean_box(0); +} +x_2261 = lean_ctor_get(x_2258, 1); +lean_inc(x_2261); +x_2262 = lean_nat_dec_eq(x_2247, x_2261); +lean_dec(x_2247); +if (x_2262 == 0) +{ +lean_object* x_2263; +lean_dec(x_2261); +if (lean_is_scalar(x_2260)) { + x_2263 = lean_alloc_ctor(1, 2, 0); +} else { + x_2263 = x_2260; +} +lean_ctor_set(x_2263, 0, x_2258); +lean_ctor_set(x_2263, 1, x_2259); +return x_2263; +} +else +{ +lean_object* x_2264; lean_object* x_2265; lean_object* x_2266; lean_object* x_2267; +lean_dec(x_2260); +lean_dec(x_2259); +x_2264 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_2265 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2265, 0, x_2264); +x_2266 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_2267 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2266, x_2265, x_2258); +if (lean_obj_tag(x_2267) == 0) +{ +lean_object* x_2268; lean_object* x_2269; lean_object* x_2270; lean_object* x_2271; +lean_dec(x_2261); +x_2268 = lean_ctor_get(x_2267, 0); +lean_inc(x_2268); +x_2269 = lean_ctor_get(x_2267, 1); +lean_inc(x_2269); +if (lean_is_exclusive(x_2267)) { + lean_ctor_release(x_2267, 0); + lean_ctor_release(x_2267, 1); + x_2270 = x_2267; +} else { + lean_dec_ref(x_2267); + x_2270 = lean_box(0); +} +if (lean_is_scalar(x_2270)) { + x_2271 = lean_alloc_ctor(0, 2, 0); +} else { + x_2271 = x_2270; +} +lean_ctor_set(x_2271, 0, x_2268); +lean_ctor_set(x_2271, 1, x_2269); +return x_2271; +} +else +{ +lean_object* x_2272; lean_object* x_2273; lean_object* x_2274; lean_object* x_2275; uint8_t x_2276; +x_2272 = lean_ctor_get(x_2267, 0); +lean_inc(x_2272); +x_2273 = lean_ctor_get(x_2267, 1); +lean_inc(x_2273); +if (lean_is_exclusive(x_2267)) { + lean_ctor_release(x_2267, 0); + lean_ctor_release(x_2267, 1); + x_2274 = x_2267; +} else { + lean_dec_ref(x_2267); + x_2274 = lean_box(0); +} +x_2275 = lean_ctor_get(x_2272, 1); +lean_inc(x_2275); +x_2276 = lean_nat_dec_eq(x_2261, x_2275); +lean_dec(x_2261); +if (x_2276 == 0) +{ +lean_object* x_2277; +lean_dec(x_2275); +if (lean_is_scalar(x_2274)) { + x_2277 = lean_alloc_ctor(1, 2, 0); +} else { + x_2277 = x_2274; +} +lean_ctor_set(x_2277, 0, x_2272); +lean_ctor_set(x_2277, 1, x_2273); +return x_2277; +} +else +{ +lean_object* x_2278; lean_object* x_2279; lean_object* x_2280; lean_object* x_2281; +lean_dec(x_2274); +lean_dec(x_2273); +x_2278 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_2279 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2279, 0, x_2278); +x_2280 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_2281 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2280, x_2279, x_2272); +if (lean_obj_tag(x_2281) == 0) +{ +lean_object* x_2282; lean_object* x_2283; lean_object* x_2284; lean_object* x_2285; +lean_dec(x_2275); +x_2282 = lean_ctor_get(x_2281, 0); +lean_inc(x_2282); +x_2283 = lean_ctor_get(x_2281, 1); +lean_inc(x_2283); +if (lean_is_exclusive(x_2281)) { + lean_ctor_release(x_2281, 0); + lean_ctor_release(x_2281, 1); + x_2284 = x_2281; +} else { + lean_dec_ref(x_2281); + x_2284 = lean_box(0); +} +if (lean_is_scalar(x_2284)) { + x_2285 = lean_alloc_ctor(0, 2, 0); +} else { + x_2285 = x_2284; +} +lean_ctor_set(x_2285, 0, x_2282); +lean_ctor_set(x_2285, 1, x_2283); +return x_2285; +} +else +{ +lean_object* x_2286; lean_object* x_2287; lean_object* x_2288; lean_object* x_2289; uint8_t x_2290; +x_2286 = lean_ctor_get(x_2281, 0); +lean_inc(x_2286); +x_2287 = lean_ctor_get(x_2281, 1); +lean_inc(x_2287); +if (lean_is_exclusive(x_2281)) { + lean_ctor_release(x_2281, 0); + lean_ctor_release(x_2281, 1); + x_2288 = x_2281; +} else { + lean_dec_ref(x_2281); + x_2288 = lean_box(0); +} +x_2289 = lean_ctor_get(x_2286, 1); +lean_inc(x_2289); +x_2290 = lean_nat_dec_eq(x_2275, x_2289); +lean_dec(x_2275); +if (x_2290 == 0) +{ +lean_object* x_2291; +lean_dec(x_2289); +if (lean_is_scalar(x_2288)) { + x_2291 = lean_alloc_ctor(1, 2, 0); +} else { + x_2291 = x_2288; +} +lean_ctor_set(x_2291, 0, x_2286); +lean_ctor_set(x_2291, 1, x_2287); +return x_2291; +} +else +{ +lean_object* x_2292; lean_object* x_2293; lean_object* x_2294; lean_object* x_2295; +lean_dec(x_2288); +lean_dec(x_2287); +x_2292 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_2293 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2293, 0, x_2292); +x_2294 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_2295 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2294, x_2293, x_2286); +if (lean_obj_tag(x_2295) == 0) +{ +lean_object* x_2296; lean_object* x_2297; lean_object* x_2298; lean_object* x_2299; +lean_dec(x_2289); +x_2296 = lean_ctor_get(x_2295, 0); +lean_inc(x_2296); +x_2297 = lean_ctor_get(x_2295, 1); +lean_inc(x_2297); +if (lean_is_exclusive(x_2295)) { + lean_ctor_release(x_2295, 0); + lean_ctor_release(x_2295, 1); + x_2298 = x_2295; +} else { + lean_dec_ref(x_2295); + x_2298 = lean_box(0); +} +if (lean_is_scalar(x_2298)) { + x_2299 = lean_alloc_ctor(0, 2, 0); +} else { + x_2299 = x_2298; +} +lean_ctor_set(x_2299, 0, x_2296); +lean_ctor_set(x_2299, 1, x_2297); +return x_2299; +} +else +{ +lean_object* x_2300; lean_object* x_2301; lean_object* x_2302; lean_object* x_2303; uint8_t x_2304; +x_2300 = lean_ctor_get(x_2295, 0); +lean_inc(x_2300); +x_2301 = lean_ctor_get(x_2295, 1); +lean_inc(x_2301); +if (lean_is_exclusive(x_2295)) { + lean_ctor_release(x_2295, 0); + lean_ctor_release(x_2295, 1); + x_2302 = x_2295; +} else { + lean_dec_ref(x_2295); + x_2302 = lean_box(0); +} +x_2303 = lean_ctor_get(x_2300, 1); +lean_inc(x_2303); +x_2304 = lean_nat_dec_eq(x_2289, x_2303); +lean_dec(x_2289); +if (x_2304 == 0) +{ +lean_object* x_2305; +lean_dec(x_2303); +if (lean_is_scalar(x_2302)) { + x_2305 = lean_alloc_ctor(1, 2, 0); +} else { + x_2305 = x_2302; +} +lean_ctor_set(x_2305, 0, x_2300); +lean_ctor_set(x_2305, 1, x_2301); +return x_2305; +} +else +{ +lean_object* x_2306; lean_object* x_2307; lean_object* x_2308; lean_object* x_2309; +lean_dec(x_2302); +lean_dec(x_2301); +x_2306 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_2307 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2307, 0, x_2306); +x_2308 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_2309 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2308, x_2307, x_2300); +if (lean_obj_tag(x_2309) == 0) +{ +lean_object* x_2310; lean_object* x_2311; lean_object* x_2312; lean_object* x_2313; +lean_dec(x_2303); +x_2310 = lean_ctor_get(x_2309, 0); +lean_inc(x_2310); +x_2311 = lean_ctor_get(x_2309, 1); +lean_inc(x_2311); +if (lean_is_exclusive(x_2309)) { + lean_ctor_release(x_2309, 0); + lean_ctor_release(x_2309, 1); + x_2312 = x_2309; +} else { + lean_dec_ref(x_2309); + x_2312 = lean_box(0); +} +if (lean_is_scalar(x_2312)) { + x_2313 = lean_alloc_ctor(0, 2, 0); +} else { + x_2313 = x_2312; +} +lean_ctor_set(x_2313, 0, x_2310); +lean_ctor_set(x_2313, 1, x_2311); +return x_2313; +} +else +{ +lean_object* x_2314; lean_object* x_2315; lean_object* x_2316; lean_object* x_2317; uint8_t x_2318; +x_2314 = lean_ctor_get(x_2309, 0); +lean_inc(x_2314); +x_2315 = lean_ctor_get(x_2309, 1); +lean_inc(x_2315); +if (lean_is_exclusive(x_2309)) { + lean_ctor_release(x_2309, 0); + lean_ctor_release(x_2309, 1); + x_2316 = x_2309; +} else { + lean_dec_ref(x_2309); + x_2316 = lean_box(0); +} +x_2317 = lean_ctor_get(x_2314, 1); +lean_inc(x_2317); +x_2318 = lean_nat_dec_eq(x_2303, x_2317); +lean_dec(x_2303); +if (x_2318 == 0) +{ +lean_object* x_2319; +lean_dec(x_2317); +if (lean_is_scalar(x_2316)) { + x_2319 = lean_alloc_ctor(1, 2, 0); +} else { + x_2319 = x_2316; +} +lean_ctor_set(x_2319, 0, x_2314); +lean_ctor_set(x_2319, 1, x_2315); +return x_2319; +} +else +{ +lean_object* x_2320; lean_object* x_2321; lean_object* x_2322; lean_object* x_2323; +lean_dec(x_2316); +lean_dec(x_2315); +x_2320 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_2321 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2321, 0, x_2320); +x_2322 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_2323 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2322, x_2321, x_2314); +if (lean_obj_tag(x_2323) == 0) +{ +lean_object* x_2324; lean_object* x_2325; lean_object* x_2326; lean_object* x_2327; +lean_dec(x_2317); +x_2324 = lean_ctor_get(x_2323, 0); +lean_inc(x_2324); +x_2325 = lean_ctor_get(x_2323, 1); +lean_inc(x_2325); +if (lean_is_exclusive(x_2323)) { + lean_ctor_release(x_2323, 0); + lean_ctor_release(x_2323, 1); + x_2326 = x_2323; +} else { + lean_dec_ref(x_2323); + x_2326 = lean_box(0); +} +if (lean_is_scalar(x_2326)) { + x_2327 = lean_alloc_ctor(0, 2, 0); +} else { + x_2327 = x_2326; +} +lean_ctor_set(x_2327, 0, x_2324); +lean_ctor_set(x_2327, 1, x_2325); +return x_2327; +} +else +{ +lean_object* x_2328; lean_object* x_2329; lean_object* x_2330; lean_object* x_2331; uint8_t x_2332; +x_2328 = lean_ctor_get(x_2323, 0); +lean_inc(x_2328); +x_2329 = lean_ctor_get(x_2323, 1); +lean_inc(x_2329); +if (lean_is_exclusive(x_2323)) { + lean_ctor_release(x_2323, 0); + lean_ctor_release(x_2323, 1); + x_2330 = x_2323; +} else { + lean_dec_ref(x_2323); + x_2330 = lean_box(0); +} +x_2331 = lean_ctor_get(x_2328, 1); +lean_inc(x_2331); +x_2332 = lean_nat_dec_eq(x_2317, x_2331); +lean_dec(x_2317); +if (x_2332 == 0) +{ +lean_object* x_2333; +lean_dec(x_2331); +if (lean_is_scalar(x_2330)) { + x_2333 = lean_alloc_ctor(1, 2, 0); +} else { + x_2333 = x_2330; +} +lean_ctor_set(x_2333, 0, x_2328); +lean_ctor_set(x_2333, 1, x_2329); +return x_2333; +} +else +{ +lean_object* x_2334; lean_object* x_2335; lean_object* x_2336; lean_object* x_2337; +lean_dec(x_2330); +lean_dec(x_2329); +x_2334 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_2335 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2335, 0, x_2334); +x_2336 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_2337 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2336, x_2335, x_2328); +if (lean_obj_tag(x_2337) == 0) +{ +lean_object* x_2338; lean_object* x_2339; lean_object* x_2340; lean_object* x_2341; +lean_dec(x_2331); +x_2338 = lean_ctor_get(x_2337, 0); +lean_inc(x_2338); +x_2339 = lean_ctor_get(x_2337, 1); +lean_inc(x_2339); +if (lean_is_exclusive(x_2337)) { + lean_ctor_release(x_2337, 0); + lean_ctor_release(x_2337, 1); + x_2340 = x_2337; +} else { + lean_dec_ref(x_2337); + x_2340 = lean_box(0); +} +if (lean_is_scalar(x_2340)) { + x_2341 = lean_alloc_ctor(0, 2, 0); +} else { + x_2341 = x_2340; +} +lean_ctor_set(x_2341, 0, x_2338); +lean_ctor_set(x_2341, 1, x_2339); +return x_2341; +} +else +{ +lean_object* x_2342; lean_object* x_2343; lean_object* x_2344; lean_object* x_2345; uint8_t x_2346; +x_2342 = lean_ctor_get(x_2337, 0); +lean_inc(x_2342); +x_2343 = lean_ctor_get(x_2337, 1); +lean_inc(x_2343); +if (lean_is_exclusive(x_2337)) { + lean_ctor_release(x_2337, 0); + lean_ctor_release(x_2337, 1); + x_2344 = x_2337; +} else { + lean_dec_ref(x_2337); + x_2344 = lean_box(0); +} +x_2345 = lean_ctor_get(x_2342, 1); +lean_inc(x_2345); +x_2346 = lean_nat_dec_eq(x_2331, x_2345); +lean_dec(x_2331); +if (x_2346 == 0) +{ +lean_object* x_2347; +lean_dec(x_2345); +if (lean_is_scalar(x_2344)) { + x_2347 = lean_alloc_ctor(1, 2, 0); +} else { + x_2347 = x_2344; +} +lean_ctor_set(x_2347, 0, x_2342); +lean_ctor_set(x_2347, 1, x_2343); +return x_2347; +} +else +{ +lean_object* x_2348; lean_object* x_2349; lean_object* x_2350; lean_object* x_2351; +lean_dec(x_2344); +lean_dec(x_2343); +x_2348 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_2349 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2349, 0, x_2348); +x_2350 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_2351 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2350, x_2349, x_2342); +if (lean_obj_tag(x_2351) == 0) +{ +lean_object* x_2352; lean_object* x_2353; lean_object* x_2354; lean_object* x_2355; +lean_dec(x_2345); +x_2352 = lean_ctor_get(x_2351, 0); +lean_inc(x_2352); +x_2353 = lean_ctor_get(x_2351, 1); +lean_inc(x_2353); +if (lean_is_exclusive(x_2351)) { + lean_ctor_release(x_2351, 0); + lean_ctor_release(x_2351, 1); + x_2354 = x_2351; +} else { + lean_dec_ref(x_2351); + x_2354 = lean_box(0); +} +if (lean_is_scalar(x_2354)) { + x_2355 = lean_alloc_ctor(0, 2, 0); +} else { + x_2355 = x_2354; +} +lean_ctor_set(x_2355, 0, x_2352); +lean_ctor_set(x_2355, 1, x_2353); +return x_2355; +} +else +{ +lean_object* x_2356; lean_object* x_2357; lean_object* x_2358; lean_object* x_2359; uint8_t x_2360; +x_2356 = lean_ctor_get(x_2351, 0); +lean_inc(x_2356); +x_2357 = lean_ctor_get(x_2351, 1); +lean_inc(x_2357); +if (lean_is_exclusive(x_2351)) { + lean_ctor_release(x_2351, 0); + lean_ctor_release(x_2351, 1); + x_2358 = x_2351; +} else { + lean_dec_ref(x_2351); + x_2358 = lean_box(0); +} +x_2359 = lean_ctor_get(x_2356, 1); +lean_inc(x_2359); +x_2360 = lean_nat_dec_eq(x_2345, x_2359); +lean_dec(x_2345); +if (x_2360 == 0) +{ +lean_object* x_2361; +lean_dec(x_2359); +if (lean_is_scalar(x_2358)) { + x_2361 = lean_alloc_ctor(1, 2, 0); +} else { + x_2361 = x_2358; +} +lean_ctor_set(x_2361, 0, x_2356); +lean_ctor_set(x_2361, 1, x_2357); +return x_2361; +} +else +{ +lean_object* x_2362; lean_object* x_2363; lean_object* x_2364; lean_object* x_2365; +lean_dec(x_2358); +lean_dec(x_2357); +x_2362 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_2363 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2363, 0, x_2362); +x_2364 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_2365 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2364, x_2363, x_2356); +if (lean_obj_tag(x_2365) == 0) +{ +lean_object* x_2366; lean_object* x_2367; lean_object* x_2368; lean_object* x_2369; +lean_dec(x_2359); +x_2366 = lean_ctor_get(x_2365, 0); +lean_inc(x_2366); +x_2367 = lean_ctor_get(x_2365, 1); +lean_inc(x_2367); +if (lean_is_exclusive(x_2365)) { + lean_ctor_release(x_2365, 0); + lean_ctor_release(x_2365, 1); + x_2368 = x_2365; +} else { + lean_dec_ref(x_2365); + x_2368 = lean_box(0); +} +if (lean_is_scalar(x_2368)) { + x_2369 = lean_alloc_ctor(0, 2, 0); +} else { + x_2369 = x_2368; +} +lean_ctor_set(x_2369, 0, x_2366); +lean_ctor_set(x_2369, 1, x_2367); +return x_2369; +} +else +{ +lean_object* x_2370; lean_object* x_2371; lean_object* x_2372; lean_object* x_2373; uint8_t x_2374; +x_2370 = lean_ctor_get(x_2365, 0); +lean_inc(x_2370); +x_2371 = lean_ctor_get(x_2365, 1); +lean_inc(x_2371); +if (lean_is_exclusive(x_2365)) { + lean_ctor_release(x_2365, 0); + lean_ctor_release(x_2365, 1); + x_2372 = x_2365; +} else { + lean_dec_ref(x_2365); + x_2372 = lean_box(0); +} +x_2373 = lean_ctor_get(x_2370, 1); +lean_inc(x_2373); +x_2374 = lean_nat_dec_eq(x_2359, x_2373); +lean_dec(x_2359); +if (x_2374 == 0) +{ +lean_object* x_2375; +lean_dec(x_2373); +if (lean_is_scalar(x_2372)) { + x_2375 = lean_alloc_ctor(1, 2, 0); +} else { + x_2375 = x_2372; +} +lean_ctor_set(x_2375, 0, x_2370); +lean_ctor_set(x_2375, 1, x_2371); +return x_2375; +} +else +{ +lean_object* x_2376; lean_object* x_2377; lean_object* x_2378; lean_object* x_2379; +lean_dec(x_2372); +lean_dec(x_2371); +x_2376 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_2377 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2377, 0, x_2376); +x_2378 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_2379 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2378, x_2377, x_2370); +if (lean_obj_tag(x_2379) == 0) +{ +lean_object* x_2380; lean_object* x_2381; lean_object* x_2382; lean_object* x_2383; +lean_dec(x_2373); +x_2380 = lean_ctor_get(x_2379, 0); +lean_inc(x_2380); +x_2381 = lean_ctor_get(x_2379, 1); +lean_inc(x_2381); +if (lean_is_exclusive(x_2379)) { + lean_ctor_release(x_2379, 0); + lean_ctor_release(x_2379, 1); + x_2382 = x_2379; +} else { + lean_dec_ref(x_2379); + x_2382 = lean_box(0); +} +if (lean_is_scalar(x_2382)) { + x_2383 = lean_alloc_ctor(0, 2, 0); +} else { + x_2383 = x_2382; +} +lean_ctor_set(x_2383, 0, x_2380); +lean_ctor_set(x_2383, 1, x_2381); +return x_2383; +} +else +{ +lean_object* x_2384; lean_object* x_2385; lean_object* x_2386; lean_object* x_2387; uint8_t x_2388; +x_2384 = lean_ctor_get(x_2379, 0); +lean_inc(x_2384); +x_2385 = lean_ctor_get(x_2379, 1); +lean_inc(x_2385); +if (lean_is_exclusive(x_2379)) { + lean_ctor_release(x_2379, 0); + lean_ctor_release(x_2379, 1); + x_2386 = x_2379; +} else { + lean_dec_ref(x_2379); + x_2386 = lean_box(0); +} +x_2387 = lean_ctor_get(x_2384, 1); +lean_inc(x_2387); +x_2388 = lean_nat_dec_eq(x_2373, x_2387); +lean_dec(x_2373); +if (x_2388 == 0) +{ +lean_object* x_2389; +lean_dec(x_2387); +if (lean_is_scalar(x_2386)) { + x_2389 = lean_alloc_ctor(1, 2, 0); +} else { + x_2389 = x_2386; +} +lean_ctor_set(x_2389, 0, x_2384); +lean_ctor_set(x_2389, 1, x_2385); +return x_2389; +} +else +{ +lean_object* x_2390; lean_object* x_2391; lean_object* x_2392; lean_object* x_2393; +lean_dec(x_2386); +lean_dec(x_2385); +x_2390 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_2391 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2391, 0, x_2390); +x_2392 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_2393 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2392, x_2391, x_2384); +if (lean_obj_tag(x_2393) == 0) +{ +lean_object* x_2394; lean_object* x_2395; lean_object* x_2396; lean_object* x_2397; +lean_dec(x_2387); +x_2394 = lean_ctor_get(x_2393, 0); +lean_inc(x_2394); +x_2395 = lean_ctor_get(x_2393, 1); +lean_inc(x_2395); +if (lean_is_exclusive(x_2393)) { + lean_ctor_release(x_2393, 0); + lean_ctor_release(x_2393, 1); + x_2396 = x_2393; +} else { + lean_dec_ref(x_2393); + x_2396 = lean_box(0); +} +if (lean_is_scalar(x_2396)) { + x_2397 = lean_alloc_ctor(0, 2, 0); +} else { + x_2397 = x_2396; +} +lean_ctor_set(x_2397, 0, x_2394); +lean_ctor_set(x_2397, 1, x_2395); +return x_2397; +} +else +{ +lean_object* x_2398; lean_object* x_2399; lean_object* x_2400; lean_object* x_2401; uint8_t x_2402; +x_2398 = lean_ctor_get(x_2393, 0); +lean_inc(x_2398); +x_2399 = lean_ctor_get(x_2393, 1); +lean_inc(x_2399); +if (lean_is_exclusive(x_2393)) { + lean_ctor_release(x_2393, 0); + lean_ctor_release(x_2393, 1); + x_2400 = x_2393; +} else { + lean_dec_ref(x_2393); + x_2400 = lean_box(0); +} +x_2401 = lean_ctor_get(x_2398, 1); +lean_inc(x_2401); +x_2402 = lean_nat_dec_eq(x_2387, x_2401); +lean_dec(x_2387); +if (x_2402 == 0) +{ +lean_object* x_2403; +lean_dec(x_2401); +if (lean_is_scalar(x_2400)) { + x_2403 = lean_alloc_ctor(1, 2, 0); +} else { + x_2403 = x_2400; +} +lean_ctor_set(x_2403, 0, x_2398); +lean_ctor_set(x_2403, 1, x_2399); +return x_2403; +} +else +{ +lean_object* x_2404; lean_object* x_2405; lean_object* x_2406; lean_object* x_2407; +lean_dec(x_2400); +lean_dec(x_2399); +x_2404 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_2405 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2405, 0, x_2404); +x_2406 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_2407 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2406, x_2405, x_2398); +if (lean_obj_tag(x_2407) == 0) +{ +lean_object* x_2408; lean_object* x_2409; lean_object* x_2410; lean_object* x_2411; +lean_dec(x_2401); +x_2408 = lean_ctor_get(x_2407, 0); +lean_inc(x_2408); +x_2409 = lean_ctor_get(x_2407, 1); +lean_inc(x_2409); +if (lean_is_exclusive(x_2407)) { + lean_ctor_release(x_2407, 0); + lean_ctor_release(x_2407, 1); + x_2410 = x_2407; +} else { + lean_dec_ref(x_2407); + x_2410 = lean_box(0); +} +if (lean_is_scalar(x_2410)) { + x_2411 = lean_alloc_ctor(0, 2, 0); +} else { + x_2411 = x_2410; +} +lean_ctor_set(x_2411, 0, x_2408); +lean_ctor_set(x_2411, 1, x_2409); +return x_2411; +} +else +{ +lean_object* x_2412; lean_object* x_2413; lean_object* x_2414; lean_object* x_2415; uint8_t x_2416; +x_2412 = lean_ctor_get(x_2407, 0); +lean_inc(x_2412); +x_2413 = lean_ctor_get(x_2407, 1); +lean_inc(x_2413); +if (lean_is_exclusive(x_2407)) { + lean_ctor_release(x_2407, 0); + lean_ctor_release(x_2407, 1); + x_2414 = x_2407; +} else { + lean_dec_ref(x_2407); + x_2414 = lean_box(0); +} +x_2415 = lean_ctor_get(x_2412, 1); +lean_inc(x_2415); +x_2416 = lean_nat_dec_eq(x_2401, x_2415); +lean_dec(x_2401); +if (x_2416 == 0) +{ +lean_object* x_2417; +lean_dec(x_2415); +if (lean_is_scalar(x_2414)) { + x_2417 = lean_alloc_ctor(1, 2, 0); +} else { + x_2417 = x_2414; +} +lean_ctor_set(x_2417, 0, x_2412); +lean_ctor_set(x_2417, 1, x_2413); +return x_2417; +} +else +{ +lean_object* x_2418; lean_object* x_2419; lean_object* x_2420; lean_object* x_2421; +lean_dec(x_2414); +lean_dec(x_2413); +x_2418 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_2419 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2419, 0, x_2418); +x_2420 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_2421 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2420, x_2419, x_2412); +if (lean_obj_tag(x_2421) == 0) +{ +lean_object* x_2422; lean_object* x_2423; lean_object* x_2424; lean_object* x_2425; +lean_dec(x_2415); +x_2422 = lean_ctor_get(x_2421, 0); +lean_inc(x_2422); +x_2423 = lean_ctor_get(x_2421, 1); +lean_inc(x_2423); +if (lean_is_exclusive(x_2421)) { + lean_ctor_release(x_2421, 0); + lean_ctor_release(x_2421, 1); + x_2424 = x_2421; +} else { + lean_dec_ref(x_2421); + x_2424 = lean_box(0); +} +if (lean_is_scalar(x_2424)) { + x_2425 = lean_alloc_ctor(0, 2, 0); +} else { + x_2425 = x_2424; +} +lean_ctor_set(x_2425, 0, x_2422); +lean_ctor_set(x_2425, 1, x_2423); +return x_2425; +} +else +{ +lean_object* x_2426; lean_object* x_2427; lean_object* x_2428; lean_object* x_2429; uint8_t x_2430; +x_2426 = lean_ctor_get(x_2421, 0); +lean_inc(x_2426); +x_2427 = lean_ctor_get(x_2421, 1); +lean_inc(x_2427); +if (lean_is_exclusive(x_2421)) { + lean_ctor_release(x_2421, 0); + lean_ctor_release(x_2421, 1); + x_2428 = x_2421; +} else { + lean_dec_ref(x_2421); + x_2428 = lean_box(0); +} +x_2429 = lean_ctor_get(x_2426, 1); +lean_inc(x_2429); +x_2430 = lean_nat_dec_eq(x_2415, x_2429); +lean_dec(x_2415); +if (x_2430 == 0) +{ +lean_object* x_2431; +lean_dec(x_2429); +if (lean_is_scalar(x_2428)) { + x_2431 = lean_alloc_ctor(1, 2, 0); +} else { + x_2431 = x_2428; +} +lean_ctor_set(x_2431, 0, x_2426); +lean_ctor_set(x_2431, 1, x_2427); +return x_2431; +} +else +{ +lean_object* x_2432; lean_object* x_2433; lean_object* x_2434; lean_object* x_2435; +lean_dec(x_2428); +lean_dec(x_2427); +x_2432 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_2433 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2433, 0, x_2432); +x_2434 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_2435 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2434, x_2433, x_2426); +if (lean_obj_tag(x_2435) == 0) +{ +lean_object* x_2436; lean_object* x_2437; lean_object* x_2438; lean_object* x_2439; +lean_dec(x_2429); +x_2436 = lean_ctor_get(x_2435, 0); +lean_inc(x_2436); +x_2437 = lean_ctor_get(x_2435, 1); +lean_inc(x_2437); +if (lean_is_exclusive(x_2435)) { + lean_ctor_release(x_2435, 0); + lean_ctor_release(x_2435, 1); + x_2438 = x_2435; +} else { + lean_dec_ref(x_2435); + x_2438 = lean_box(0); +} +if (lean_is_scalar(x_2438)) { + x_2439 = lean_alloc_ctor(0, 2, 0); +} else { + x_2439 = x_2438; +} +lean_ctor_set(x_2439, 0, x_2436); +lean_ctor_set(x_2439, 1, x_2437); +return x_2439; +} +else +{ +lean_object* x_2440; lean_object* x_2441; lean_object* x_2442; lean_object* x_2443; uint8_t x_2444; +x_2440 = lean_ctor_get(x_2435, 0); +lean_inc(x_2440); +x_2441 = lean_ctor_get(x_2435, 1); +lean_inc(x_2441); +if (lean_is_exclusive(x_2435)) { + lean_ctor_release(x_2435, 0); + lean_ctor_release(x_2435, 1); + x_2442 = x_2435; +} else { + lean_dec_ref(x_2435); + x_2442 = lean_box(0); +} +x_2443 = lean_ctor_get(x_2440, 1); +lean_inc(x_2443); +x_2444 = lean_nat_dec_eq(x_2429, x_2443); +lean_dec(x_2429); +if (x_2444 == 0) +{ +lean_object* x_2445; +lean_dec(x_2443); +if (lean_is_scalar(x_2442)) { + x_2445 = lean_alloc_ctor(1, 2, 0); +} else { + x_2445 = x_2442; +} +lean_ctor_set(x_2445, 0, x_2440); +lean_ctor_set(x_2445, 1, x_2441); +return x_2445; +} +else +{ +lean_object* x_2446; lean_object* x_2447; lean_object* x_2448; lean_object* x_2449; +lean_dec(x_2442); +lean_dec(x_2441); +x_2446 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_2447 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2447, 0, x_2446); +x_2448 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_2449 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2448, x_2447, x_2440); +if (lean_obj_tag(x_2449) == 0) +{ +lean_object* x_2450; lean_object* x_2451; lean_object* x_2452; lean_object* x_2453; +lean_dec(x_2443); +x_2450 = lean_ctor_get(x_2449, 0); +lean_inc(x_2450); +x_2451 = lean_ctor_get(x_2449, 1); +lean_inc(x_2451); +if (lean_is_exclusive(x_2449)) { + lean_ctor_release(x_2449, 0); + lean_ctor_release(x_2449, 1); + x_2452 = x_2449; +} else { + lean_dec_ref(x_2449); + x_2452 = lean_box(0); +} +if (lean_is_scalar(x_2452)) { + x_2453 = lean_alloc_ctor(0, 2, 0); +} else { + x_2453 = x_2452; +} +lean_ctor_set(x_2453, 0, x_2450); +lean_ctor_set(x_2453, 1, x_2451); +return x_2453; +} +else +{ +lean_object* x_2454; lean_object* x_2455; lean_object* x_2456; lean_object* x_2457; uint8_t x_2458; +x_2454 = lean_ctor_get(x_2449, 0); +lean_inc(x_2454); +x_2455 = lean_ctor_get(x_2449, 1); +lean_inc(x_2455); +if (lean_is_exclusive(x_2449)) { + lean_ctor_release(x_2449, 0); + lean_ctor_release(x_2449, 1); + x_2456 = x_2449; +} else { + lean_dec_ref(x_2449); + x_2456 = lean_box(0); +} +x_2457 = lean_ctor_get(x_2454, 1); +lean_inc(x_2457); +x_2458 = lean_nat_dec_eq(x_2443, x_2457); +lean_dec(x_2457); +lean_dec(x_2443); +if (x_2458 == 0) +{ +lean_object* x_2459; +if (lean_is_scalar(x_2456)) { + x_2459 = lean_alloc_ctor(1, 2, 0); +} else { + x_2459 = x_2456; +} +lean_ctor_set(x_2459, 0, x_2454); +lean_ctor_set(x_2459, 1, x_2455); +return x_2459; +} +else +{ +lean_object* x_2460; lean_object* x_2461; lean_object* x_2462; lean_object* x_2463; +lean_dec(x_2456); +lean_dec(x_2455); +x_2460 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_2461 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2461, 0, x_2460); +x_2462 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_2463 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2462, x_2461, x_2454); +return x_2463; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_2464; lean_object* x_2465; lean_object* x_2466; uint8_t x_2467; +x_2464 = lean_ctor_get(x_172, 0); +x_2465 = lean_ctor_get(x_172, 1); +lean_inc(x_2465); +lean_inc(x_2464); +lean_dec(x_172); +x_2466 = lean_ctor_get(x_2464, 1); +lean_inc(x_2466); +x_2467 = lean_nat_dec_eq(x_168, x_2466); +lean_dec(x_168); +if (x_2467 == 0) +{ +lean_object* x_2468; +lean_dec(x_2466); +x_2468 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2468, 0, x_2464); +lean_ctor_set(x_2468, 1, x_2465); +return x_2468; +} +else +{ +lean_object* x_2469; lean_object* x_2470; lean_object* x_2471; lean_object* x_2472; +lean_dec(x_2465); +x_2469 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_2470 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2470, 0, x_2469); +x_2471 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_2472 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2471, x_2470, x_2464); +if (lean_obj_tag(x_2472) == 0) +{ +lean_object* x_2473; lean_object* x_2474; lean_object* x_2475; lean_object* x_2476; +lean_dec(x_2466); +x_2473 = lean_ctor_get(x_2472, 0); +lean_inc(x_2473); +x_2474 = lean_ctor_get(x_2472, 1); +lean_inc(x_2474); +if (lean_is_exclusive(x_2472)) { + lean_ctor_release(x_2472, 0); + lean_ctor_release(x_2472, 1); + x_2475 = x_2472; +} else { + lean_dec_ref(x_2472); + x_2475 = lean_box(0); +} +if (lean_is_scalar(x_2475)) { + x_2476 = lean_alloc_ctor(0, 2, 0); +} else { + x_2476 = x_2475; +} +lean_ctor_set(x_2476, 0, x_2473); +lean_ctor_set(x_2476, 1, x_2474); +return x_2476; +} +else +{ +lean_object* x_2477; lean_object* x_2478; lean_object* x_2479; lean_object* x_2480; uint8_t x_2481; +x_2477 = lean_ctor_get(x_2472, 0); +lean_inc(x_2477); +x_2478 = lean_ctor_get(x_2472, 1); +lean_inc(x_2478); +if (lean_is_exclusive(x_2472)) { + lean_ctor_release(x_2472, 0); + lean_ctor_release(x_2472, 1); + x_2479 = x_2472; +} else { + lean_dec_ref(x_2472); + x_2479 = lean_box(0); +} +x_2480 = lean_ctor_get(x_2477, 1); +lean_inc(x_2480); +x_2481 = lean_nat_dec_eq(x_2466, x_2480); +lean_dec(x_2466); +if (x_2481 == 0) +{ +lean_object* x_2482; +lean_dec(x_2480); +if (lean_is_scalar(x_2479)) { + x_2482 = lean_alloc_ctor(1, 2, 0); +} else { + x_2482 = x_2479; +} +lean_ctor_set(x_2482, 0, x_2477); +lean_ctor_set(x_2482, 1, x_2478); +return x_2482; +} +else +{ +lean_object* x_2483; lean_object* x_2484; lean_object* x_2485; lean_object* x_2486; +lean_dec(x_2479); +lean_dec(x_2478); +x_2483 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_2484 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2484, 0, x_2483); +x_2485 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_2486 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2485, x_2484, x_2477); +if (lean_obj_tag(x_2486) == 0) +{ +lean_object* x_2487; lean_object* x_2488; lean_object* x_2489; lean_object* x_2490; +lean_dec(x_2480); +x_2487 = lean_ctor_get(x_2486, 0); +lean_inc(x_2487); +x_2488 = lean_ctor_get(x_2486, 1); +lean_inc(x_2488); +if (lean_is_exclusive(x_2486)) { + lean_ctor_release(x_2486, 0); + lean_ctor_release(x_2486, 1); + x_2489 = x_2486; +} else { + lean_dec_ref(x_2486); + x_2489 = lean_box(0); +} +if (lean_is_scalar(x_2489)) { + x_2490 = lean_alloc_ctor(0, 2, 0); +} else { + x_2490 = x_2489; +} +lean_ctor_set(x_2490, 0, x_2487); +lean_ctor_set(x_2490, 1, x_2488); +return x_2490; +} +else +{ +lean_object* x_2491; lean_object* x_2492; lean_object* x_2493; lean_object* x_2494; uint8_t x_2495; +x_2491 = lean_ctor_get(x_2486, 0); +lean_inc(x_2491); +x_2492 = lean_ctor_get(x_2486, 1); +lean_inc(x_2492); +if (lean_is_exclusive(x_2486)) { + lean_ctor_release(x_2486, 0); + lean_ctor_release(x_2486, 1); + x_2493 = x_2486; +} else { + lean_dec_ref(x_2486); + x_2493 = lean_box(0); +} +x_2494 = lean_ctor_get(x_2491, 1); +lean_inc(x_2494); +x_2495 = lean_nat_dec_eq(x_2480, x_2494); +lean_dec(x_2480); +if (x_2495 == 0) +{ +lean_object* x_2496; +lean_dec(x_2494); +if (lean_is_scalar(x_2493)) { + x_2496 = lean_alloc_ctor(1, 2, 0); +} else { + x_2496 = x_2493; +} +lean_ctor_set(x_2496, 0, x_2491); +lean_ctor_set(x_2496, 1, x_2492); +return x_2496; +} +else +{ +lean_object* x_2497; lean_object* x_2498; lean_object* x_2499; lean_object* x_2500; +lean_dec(x_2493); +lean_dec(x_2492); +x_2497 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_2498 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2498, 0, x_2497); +x_2499 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_2500 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2499, x_2498, x_2491); +if (lean_obj_tag(x_2500) == 0) +{ +lean_object* x_2501; lean_object* x_2502; lean_object* x_2503; lean_object* x_2504; +lean_dec(x_2494); +x_2501 = lean_ctor_get(x_2500, 0); +lean_inc(x_2501); +x_2502 = lean_ctor_get(x_2500, 1); +lean_inc(x_2502); +if (lean_is_exclusive(x_2500)) { + lean_ctor_release(x_2500, 0); + lean_ctor_release(x_2500, 1); + x_2503 = x_2500; +} else { + lean_dec_ref(x_2500); + x_2503 = lean_box(0); +} +if (lean_is_scalar(x_2503)) { + x_2504 = lean_alloc_ctor(0, 2, 0); +} else { + x_2504 = x_2503; +} +lean_ctor_set(x_2504, 0, x_2501); +lean_ctor_set(x_2504, 1, x_2502); +return x_2504; +} +else +{ +lean_object* x_2505; lean_object* x_2506; lean_object* x_2507; lean_object* x_2508; uint8_t x_2509; +x_2505 = lean_ctor_get(x_2500, 0); +lean_inc(x_2505); +x_2506 = lean_ctor_get(x_2500, 1); +lean_inc(x_2506); +if (lean_is_exclusive(x_2500)) { + lean_ctor_release(x_2500, 0); + lean_ctor_release(x_2500, 1); + x_2507 = x_2500; +} else { + lean_dec_ref(x_2500); + x_2507 = lean_box(0); +} +x_2508 = lean_ctor_get(x_2505, 1); +lean_inc(x_2508); +x_2509 = lean_nat_dec_eq(x_2494, x_2508); +lean_dec(x_2494); +if (x_2509 == 0) +{ +lean_object* x_2510; +lean_dec(x_2508); +if (lean_is_scalar(x_2507)) { + x_2510 = lean_alloc_ctor(1, 2, 0); +} else { + x_2510 = x_2507; +} +lean_ctor_set(x_2510, 0, x_2505); +lean_ctor_set(x_2510, 1, x_2506); +return x_2510; +} +else +{ +lean_object* x_2511; lean_object* x_2512; lean_object* x_2513; lean_object* x_2514; +lean_dec(x_2507); +lean_dec(x_2506); +x_2511 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_2512 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2512, 0, x_2511); +x_2513 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_2514 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2513, x_2512, x_2505); +if (lean_obj_tag(x_2514) == 0) +{ +lean_object* x_2515; lean_object* x_2516; lean_object* x_2517; lean_object* x_2518; +lean_dec(x_2508); +x_2515 = lean_ctor_get(x_2514, 0); +lean_inc(x_2515); +x_2516 = lean_ctor_get(x_2514, 1); +lean_inc(x_2516); +if (lean_is_exclusive(x_2514)) { + lean_ctor_release(x_2514, 0); + lean_ctor_release(x_2514, 1); + x_2517 = x_2514; +} else { + lean_dec_ref(x_2514); + x_2517 = lean_box(0); +} +if (lean_is_scalar(x_2517)) { + x_2518 = lean_alloc_ctor(0, 2, 0); +} else { + x_2518 = x_2517; +} +lean_ctor_set(x_2518, 0, x_2515); +lean_ctor_set(x_2518, 1, x_2516); +return x_2518; +} +else +{ +lean_object* x_2519; lean_object* x_2520; lean_object* x_2521; lean_object* x_2522; uint8_t x_2523; +x_2519 = lean_ctor_get(x_2514, 0); +lean_inc(x_2519); +x_2520 = lean_ctor_get(x_2514, 1); +lean_inc(x_2520); +if (lean_is_exclusive(x_2514)) { + lean_ctor_release(x_2514, 0); + lean_ctor_release(x_2514, 1); + x_2521 = x_2514; +} else { + lean_dec_ref(x_2514); + x_2521 = lean_box(0); +} +x_2522 = lean_ctor_get(x_2519, 1); +lean_inc(x_2522); +x_2523 = lean_nat_dec_eq(x_2508, x_2522); +lean_dec(x_2508); +if (x_2523 == 0) +{ +lean_object* x_2524; +lean_dec(x_2522); +if (lean_is_scalar(x_2521)) { + x_2524 = lean_alloc_ctor(1, 2, 0); +} else { + x_2524 = x_2521; +} +lean_ctor_set(x_2524, 0, x_2519); +lean_ctor_set(x_2524, 1, x_2520); +return x_2524; +} +else +{ +lean_object* x_2525; lean_object* x_2526; lean_object* x_2527; lean_object* x_2528; +lean_dec(x_2521); +lean_dec(x_2520); +x_2525 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_2526 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2526, 0, x_2525); +x_2527 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_2528 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2527, x_2526, x_2519); +if (lean_obj_tag(x_2528) == 0) +{ +lean_object* x_2529; lean_object* x_2530; lean_object* x_2531; lean_object* x_2532; +lean_dec(x_2522); +x_2529 = lean_ctor_get(x_2528, 0); +lean_inc(x_2529); +x_2530 = lean_ctor_get(x_2528, 1); +lean_inc(x_2530); +if (lean_is_exclusive(x_2528)) { + lean_ctor_release(x_2528, 0); + lean_ctor_release(x_2528, 1); + x_2531 = x_2528; +} else { + lean_dec_ref(x_2528); + x_2531 = lean_box(0); +} +if (lean_is_scalar(x_2531)) { + x_2532 = lean_alloc_ctor(0, 2, 0); +} else { + x_2532 = x_2531; +} +lean_ctor_set(x_2532, 0, x_2529); +lean_ctor_set(x_2532, 1, x_2530); +return x_2532; +} +else +{ +lean_object* x_2533; lean_object* x_2534; lean_object* x_2535; lean_object* x_2536; uint8_t x_2537; +x_2533 = lean_ctor_get(x_2528, 0); +lean_inc(x_2533); +x_2534 = lean_ctor_get(x_2528, 1); +lean_inc(x_2534); +if (lean_is_exclusive(x_2528)) { + lean_ctor_release(x_2528, 0); + lean_ctor_release(x_2528, 1); + x_2535 = x_2528; +} else { + lean_dec_ref(x_2528); + x_2535 = lean_box(0); +} +x_2536 = lean_ctor_get(x_2533, 1); +lean_inc(x_2536); +x_2537 = lean_nat_dec_eq(x_2522, x_2536); +lean_dec(x_2522); +if (x_2537 == 0) +{ +lean_object* x_2538; +lean_dec(x_2536); +if (lean_is_scalar(x_2535)) { + x_2538 = lean_alloc_ctor(1, 2, 0); +} else { + x_2538 = x_2535; +} +lean_ctor_set(x_2538, 0, x_2533); +lean_ctor_set(x_2538, 1, x_2534); +return x_2538; +} +else +{ +lean_object* x_2539; lean_object* x_2540; lean_object* x_2541; lean_object* x_2542; +lean_dec(x_2535); +lean_dec(x_2534); +x_2539 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_2540 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2540, 0, x_2539); +x_2541 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_2542 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2541, x_2540, x_2533); +if (lean_obj_tag(x_2542) == 0) +{ +lean_object* x_2543; lean_object* x_2544; lean_object* x_2545; lean_object* x_2546; +lean_dec(x_2536); +x_2543 = lean_ctor_get(x_2542, 0); +lean_inc(x_2543); +x_2544 = lean_ctor_get(x_2542, 1); +lean_inc(x_2544); +if (lean_is_exclusive(x_2542)) { + lean_ctor_release(x_2542, 0); + lean_ctor_release(x_2542, 1); + x_2545 = x_2542; +} else { + lean_dec_ref(x_2542); + x_2545 = lean_box(0); +} +if (lean_is_scalar(x_2545)) { + x_2546 = lean_alloc_ctor(0, 2, 0); +} else { + x_2546 = x_2545; +} +lean_ctor_set(x_2546, 0, x_2543); +lean_ctor_set(x_2546, 1, x_2544); +return x_2546; +} +else +{ +lean_object* x_2547; lean_object* x_2548; lean_object* x_2549; lean_object* x_2550; uint8_t x_2551; +x_2547 = lean_ctor_get(x_2542, 0); +lean_inc(x_2547); +x_2548 = lean_ctor_get(x_2542, 1); +lean_inc(x_2548); +if (lean_is_exclusive(x_2542)) { + lean_ctor_release(x_2542, 0); + lean_ctor_release(x_2542, 1); + x_2549 = x_2542; +} else { + lean_dec_ref(x_2542); + x_2549 = lean_box(0); +} +x_2550 = lean_ctor_get(x_2547, 1); +lean_inc(x_2550); +x_2551 = lean_nat_dec_eq(x_2536, x_2550); +lean_dec(x_2536); +if (x_2551 == 0) +{ +lean_object* x_2552; +lean_dec(x_2550); +if (lean_is_scalar(x_2549)) { + x_2552 = lean_alloc_ctor(1, 2, 0); +} else { + x_2552 = x_2549; +} +lean_ctor_set(x_2552, 0, x_2547); +lean_ctor_set(x_2552, 1, x_2548); +return x_2552; +} +else +{ +lean_object* x_2553; lean_object* x_2554; lean_object* x_2555; lean_object* x_2556; +lean_dec(x_2549); +lean_dec(x_2548); +x_2553 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_2554 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2554, 0, x_2553); +x_2555 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_2556 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2555, x_2554, x_2547); +if (lean_obj_tag(x_2556) == 0) +{ +lean_object* x_2557; lean_object* x_2558; lean_object* x_2559; lean_object* x_2560; +lean_dec(x_2550); +x_2557 = lean_ctor_get(x_2556, 0); +lean_inc(x_2557); +x_2558 = lean_ctor_get(x_2556, 1); +lean_inc(x_2558); +if (lean_is_exclusive(x_2556)) { + lean_ctor_release(x_2556, 0); + lean_ctor_release(x_2556, 1); + x_2559 = x_2556; +} else { + lean_dec_ref(x_2556); + x_2559 = lean_box(0); +} +if (lean_is_scalar(x_2559)) { + x_2560 = lean_alloc_ctor(0, 2, 0); +} else { + x_2560 = x_2559; +} +lean_ctor_set(x_2560, 0, x_2557); +lean_ctor_set(x_2560, 1, x_2558); +return x_2560; +} +else +{ +lean_object* x_2561; lean_object* x_2562; lean_object* x_2563; lean_object* x_2564; uint8_t x_2565; +x_2561 = lean_ctor_get(x_2556, 0); +lean_inc(x_2561); +x_2562 = lean_ctor_get(x_2556, 1); +lean_inc(x_2562); +if (lean_is_exclusive(x_2556)) { + lean_ctor_release(x_2556, 0); + lean_ctor_release(x_2556, 1); + x_2563 = x_2556; +} else { + lean_dec_ref(x_2556); + x_2563 = lean_box(0); +} +x_2564 = lean_ctor_get(x_2561, 1); +lean_inc(x_2564); +x_2565 = lean_nat_dec_eq(x_2550, x_2564); +lean_dec(x_2550); +if (x_2565 == 0) +{ +lean_object* x_2566; +lean_dec(x_2564); +if (lean_is_scalar(x_2563)) { + x_2566 = lean_alloc_ctor(1, 2, 0); +} else { + x_2566 = x_2563; +} +lean_ctor_set(x_2566, 0, x_2561); +lean_ctor_set(x_2566, 1, x_2562); +return x_2566; +} +else +{ +lean_object* x_2567; lean_object* x_2568; lean_object* x_2569; lean_object* x_2570; +lean_dec(x_2563); +lean_dec(x_2562); +x_2567 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_2568 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2568, 0, x_2567); +x_2569 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_2570 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2569, x_2568, x_2561); +if (lean_obj_tag(x_2570) == 0) +{ +lean_object* x_2571; lean_object* x_2572; lean_object* x_2573; lean_object* x_2574; +lean_dec(x_2564); +x_2571 = lean_ctor_get(x_2570, 0); +lean_inc(x_2571); +x_2572 = lean_ctor_get(x_2570, 1); +lean_inc(x_2572); +if (lean_is_exclusive(x_2570)) { + lean_ctor_release(x_2570, 0); + lean_ctor_release(x_2570, 1); + x_2573 = x_2570; +} else { + lean_dec_ref(x_2570); + x_2573 = lean_box(0); +} +if (lean_is_scalar(x_2573)) { + x_2574 = lean_alloc_ctor(0, 2, 0); +} else { + x_2574 = x_2573; +} +lean_ctor_set(x_2574, 0, x_2571); +lean_ctor_set(x_2574, 1, x_2572); +return x_2574; +} +else +{ +lean_object* x_2575; lean_object* x_2576; lean_object* x_2577; lean_object* x_2578; uint8_t x_2579; +x_2575 = lean_ctor_get(x_2570, 0); +lean_inc(x_2575); +x_2576 = lean_ctor_get(x_2570, 1); +lean_inc(x_2576); +if (lean_is_exclusive(x_2570)) { + lean_ctor_release(x_2570, 0); + lean_ctor_release(x_2570, 1); + x_2577 = x_2570; +} else { + lean_dec_ref(x_2570); + x_2577 = lean_box(0); +} +x_2578 = lean_ctor_get(x_2575, 1); +lean_inc(x_2578); +x_2579 = lean_nat_dec_eq(x_2564, x_2578); +lean_dec(x_2564); +if (x_2579 == 0) +{ +lean_object* x_2580; +lean_dec(x_2578); +if (lean_is_scalar(x_2577)) { + x_2580 = lean_alloc_ctor(1, 2, 0); +} else { + x_2580 = x_2577; +} +lean_ctor_set(x_2580, 0, x_2575); +lean_ctor_set(x_2580, 1, x_2576); +return x_2580; +} +else +{ +lean_object* x_2581; lean_object* x_2582; lean_object* x_2583; lean_object* x_2584; +lean_dec(x_2577); +lean_dec(x_2576); +x_2581 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_2582 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2582, 0, x_2581); +x_2583 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_2584 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2583, x_2582, x_2575); +if (lean_obj_tag(x_2584) == 0) +{ +lean_object* x_2585; lean_object* x_2586; lean_object* x_2587; lean_object* x_2588; +lean_dec(x_2578); +x_2585 = lean_ctor_get(x_2584, 0); +lean_inc(x_2585); +x_2586 = lean_ctor_get(x_2584, 1); +lean_inc(x_2586); +if (lean_is_exclusive(x_2584)) { + lean_ctor_release(x_2584, 0); + lean_ctor_release(x_2584, 1); + x_2587 = x_2584; +} else { + lean_dec_ref(x_2584); + x_2587 = lean_box(0); +} +if (lean_is_scalar(x_2587)) { + x_2588 = lean_alloc_ctor(0, 2, 0); +} else { + x_2588 = x_2587; +} +lean_ctor_set(x_2588, 0, x_2585); +lean_ctor_set(x_2588, 1, x_2586); +return x_2588; +} +else +{ +lean_object* x_2589; lean_object* x_2590; lean_object* x_2591; lean_object* x_2592; uint8_t x_2593; +x_2589 = lean_ctor_get(x_2584, 0); +lean_inc(x_2589); +x_2590 = lean_ctor_get(x_2584, 1); +lean_inc(x_2590); +if (lean_is_exclusive(x_2584)) { + lean_ctor_release(x_2584, 0); + lean_ctor_release(x_2584, 1); + x_2591 = x_2584; +} else { + lean_dec_ref(x_2584); + x_2591 = lean_box(0); +} +x_2592 = lean_ctor_get(x_2589, 1); +lean_inc(x_2592); +x_2593 = lean_nat_dec_eq(x_2578, x_2592); +lean_dec(x_2578); +if (x_2593 == 0) +{ +lean_object* x_2594; +lean_dec(x_2592); +if (lean_is_scalar(x_2591)) { + x_2594 = lean_alloc_ctor(1, 2, 0); +} else { + x_2594 = x_2591; +} +lean_ctor_set(x_2594, 0, x_2589); +lean_ctor_set(x_2594, 1, x_2590); +return x_2594; +} +else +{ +lean_object* x_2595; lean_object* x_2596; lean_object* x_2597; lean_object* x_2598; +lean_dec(x_2591); +lean_dec(x_2590); +x_2595 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_2596 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2596, 0, x_2595); +x_2597 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_2598 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2597, x_2596, x_2589); +if (lean_obj_tag(x_2598) == 0) +{ +lean_object* x_2599; lean_object* x_2600; lean_object* x_2601; lean_object* x_2602; +lean_dec(x_2592); +x_2599 = lean_ctor_get(x_2598, 0); +lean_inc(x_2599); +x_2600 = lean_ctor_get(x_2598, 1); +lean_inc(x_2600); +if (lean_is_exclusive(x_2598)) { + lean_ctor_release(x_2598, 0); + lean_ctor_release(x_2598, 1); + x_2601 = x_2598; +} else { + lean_dec_ref(x_2598); + x_2601 = lean_box(0); +} +if (lean_is_scalar(x_2601)) { + x_2602 = lean_alloc_ctor(0, 2, 0); +} else { + x_2602 = x_2601; +} +lean_ctor_set(x_2602, 0, x_2599); +lean_ctor_set(x_2602, 1, x_2600); +return x_2602; +} +else +{ +lean_object* x_2603; lean_object* x_2604; lean_object* x_2605; lean_object* x_2606; uint8_t x_2607; +x_2603 = lean_ctor_get(x_2598, 0); +lean_inc(x_2603); +x_2604 = lean_ctor_get(x_2598, 1); +lean_inc(x_2604); +if (lean_is_exclusive(x_2598)) { + lean_ctor_release(x_2598, 0); + lean_ctor_release(x_2598, 1); + x_2605 = x_2598; +} else { + lean_dec_ref(x_2598); + x_2605 = lean_box(0); +} +x_2606 = lean_ctor_get(x_2603, 1); +lean_inc(x_2606); +x_2607 = lean_nat_dec_eq(x_2592, x_2606); +lean_dec(x_2592); +if (x_2607 == 0) +{ +lean_object* x_2608; +lean_dec(x_2606); +if (lean_is_scalar(x_2605)) { + x_2608 = lean_alloc_ctor(1, 2, 0); +} else { + x_2608 = x_2605; +} +lean_ctor_set(x_2608, 0, x_2603); +lean_ctor_set(x_2608, 1, x_2604); +return x_2608; +} +else +{ +lean_object* x_2609; lean_object* x_2610; lean_object* x_2611; lean_object* x_2612; +lean_dec(x_2605); +lean_dec(x_2604); +x_2609 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_2610 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2610, 0, x_2609); +x_2611 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_2612 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2611, x_2610, x_2603); +if (lean_obj_tag(x_2612) == 0) +{ +lean_object* x_2613; lean_object* x_2614; lean_object* x_2615; lean_object* x_2616; +lean_dec(x_2606); +x_2613 = lean_ctor_get(x_2612, 0); +lean_inc(x_2613); +x_2614 = lean_ctor_get(x_2612, 1); +lean_inc(x_2614); +if (lean_is_exclusive(x_2612)) { + lean_ctor_release(x_2612, 0); + lean_ctor_release(x_2612, 1); + x_2615 = x_2612; +} else { + lean_dec_ref(x_2612); + x_2615 = lean_box(0); +} +if (lean_is_scalar(x_2615)) { + x_2616 = lean_alloc_ctor(0, 2, 0); +} else { + x_2616 = x_2615; +} +lean_ctor_set(x_2616, 0, x_2613); +lean_ctor_set(x_2616, 1, x_2614); +return x_2616; +} +else +{ +lean_object* x_2617; lean_object* x_2618; lean_object* x_2619; lean_object* x_2620; uint8_t x_2621; +x_2617 = lean_ctor_get(x_2612, 0); +lean_inc(x_2617); +x_2618 = lean_ctor_get(x_2612, 1); +lean_inc(x_2618); +if (lean_is_exclusive(x_2612)) { + lean_ctor_release(x_2612, 0); + lean_ctor_release(x_2612, 1); + x_2619 = x_2612; +} else { + lean_dec_ref(x_2612); + x_2619 = lean_box(0); +} +x_2620 = lean_ctor_get(x_2617, 1); +lean_inc(x_2620); +x_2621 = lean_nat_dec_eq(x_2606, x_2620); +lean_dec(x_2606); +if (x_2621 == 0) +{ +lean_object* x_2622; +lean_dec(x_2620); +if (lean_is_scalar(x_2619)) { + x_2622 = lean_alloc_ctor(1, 2, 0); +} else { + x_2622 = x_2619; +} +lean_ctor_set(x_2622, 0, x_2617); +lean_ctor_set(x_2622, 1, x_2618); +return x_2622; +} +else +{ +lean_object* x_2623; lean_object* x_2624; lean_object* x_2625; lean_object* x_2626; +lean_dec(x_2619); +lean_dec(x_2618); +x_2623 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_2624 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2624, 0, x_2623); +x_2625 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_2626 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2625, x_2624, x_2617); +if (lean_obj_tag(x_2626) == 0) +{ +lean_object* x_2627; lean_object* x_2628; lean_object* x_2629; lean_object* x_2630; +lean_dec(x_2620); +x_2627 = lean_ctor_get(x_2626, 0); +lean_inc(x_2627); +x_2628 = lean_ctor_get(x_2626, 1); +lean_inc(x_2628); +if (lean_is_exclusive(x_2626)) { + lean_ctor_release(x_2626, 0); + lean_ctor_release(x_2626, 1); + x_2629 = x_2626; +} else { + lean_dec_ref(x_2626); + x_2629 = lean_box(0); +} +if (lean_is_scalar(x_2629)) { + x_2630 = lean_alloc_ctor(0, 2, 0); +} else { + x_2630 = x_2629; +} +lean_ctor_set(x_2630, 0, x_2627); +lean_ctor_set(x_2630, 1, x_2628); +return x_2630; +} +else +{ +lean_object* x_2631; lean_object* x_2632; lean_object* x_2633; lean_object* x_2634; uint8_t x_2635; +x_2631 = lean_ctor_get(x_2626, 0); +lean_inc(x_2631); +x_2632 = lean_ctor_get(x_2626, 1); +lean_inc(x_2632); +if (lean_is_exclusive(x_2626)) { + lean_ctor_release(x_2626, 0); + lean_ctor_release(x_2626, 1); + x_2633 = x_2626; +} else { + lean_dec_ref(x_2626); + x_2633 = lean_box(0); +} +x_2634 = lean_ctor_get(x_2631, 1); +lean_inc(x_2634); +x_2635 = lean_nat_dec_eq(x_2620, x_2634); +lean_dec(x_2620); +if (x_2635 == 0) +{ +lean_object* x_2636; +lean_dec(x_2634); +if (lean_is_scalar(x_2633)) { + x_2636 = lean_alloc_ctor(1, 2, 0); +} else { + x_2636 = x_2633; +} +lean_ctor_set(x_2636, 0, x_2631); +lean_ctor_set(x_2636, 1, x_2632); +return x_2636; +} +else +{ +lean_object* x_2637; lean_object* x_2638; lean_object* x_2639; lean_object* x_2640; +lean_dec(x_2633); +lean_dec(x_2632); +x_2637 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_2638 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2638, 0, x_2637); +x_2639 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_2640 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2639, x_2638, x_2631); +if (lean_obj_tag(x_2640) == 0) +{ +lean_object* x_2641; lean_object* x_2642; lean_object* x_2643; lean_object* x_2644; +lean_dec(x_2634); +x_2641 = lean_ctor_get(x_2640, 0); +lean_inc(x_2641); +x_2642 = lean_ctor_get(x_2640, 1); +lean_inc(x_2642); +if (lean_is_exclusive(x_2640)) { + lean_ctor_release(x_2640, 0); + lean_ctor_release(x_2640, 1); + x_2643 = x_2640; +} else { + lean_dec_ref(x_2640); + x_2643 = lean_box(0); +} +if (lean_is_scalar(x_2643)) { + x_2644 = lean_alloc_ctor(0, 2, 0); +} else { + x_2644 = x_2643; +} +lean_ctor_set(x_2644, 0, x_2641); +lean_ctor_set(x_2644, 1, x_2642); +return x_2644; +} +else +{ +lean_object* x_2645; lean_object* x_2646; lean_object* x_2647; lean_object* x_2648; uint8_t x_2649; +x_2645 = lean_ctor_get(x_2640, 0); +lean_inc(x_2645); +x_2646 = lean_ctor_get(x_2640, 1); +lean_inc(x_2646); +if (lean_is_exclusive(x_2640)) { + lean_ctor_release(x_2640, 0); + lean_ctor_release(x_2640, 1); + x_2647 = x_2640; +} else { + lean_dec_ref(x_2640); + x_2647 = lean_box(0); +} +x_2648 = lean_ctor_get(x_2645, 1); +lean_inc(x_2648); +x_2649 = lean_nat_dec_eq(x_2634, x_2648); +lean_dec(x_2634); +if (x_2649 == 0) +{ +lean_object* x_2650; +lean_dec(x_2648); +if (lean_is_scalar(x_2647)) { + x_2650 = lean_alloc_ctor(1, 2, 0); +} else { + x_2650 = x_2647; +} +lean_ctor_set(x_2650, 0, x_2645); +lean_ctor_set(x_2650, 1, x_2646); +return x_2650; +} +else +{ +lean_object* x_2651; lean_object* x_2652; lean_object* x_2653; lean_object* x_2654; +lean_dec(x_2647); +lean_dec(x_2646); +x_2651 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_2652 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2652, 0, x_2651); +x_2653 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_2654 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2653, x_2652, x_2645); +if (lean_obj_tag(x_2654) == 0) +{ +lean_object* x_2655; lean_object* x_2656; lean_object* x_2657; lean_object* x_2658; +lean_dec(x_2648); +x_2655 = lean_ctor_get(x_2654, 0); +lean_inc(x_2655); +x_2656 = lean_ctor_get(x_2654, 1); +lean_inc(x_2656); +if (lean_is_exclusive(x_2654)) { + lean_ctor_release(x_2654, 0); + lean_ctor_release(x_2654, 1); + x_2657 = x_2654; +} else { + lean_dec_ref(x_2654); + x_2657 = lean_box(0); +} +if (lean_is_scalar(x_2657)) { + x_2658 = lean_alloc_ctor(0, 2, 0); +} else { + x_2658 = x_2657; +} +lean_ctor_set(x_2658, 0, x_2655); +lean_ctor_set(x_2658, 1, x_2656); +return x_2658; +} +else +{ +lean_object* x_2659; lean_object* x_2660; lean_object* x_2661; lean_object* x_2662; uint8_t x_2663; +x_2659 = lean_ctor_get(x_2654, 0); +lean_inc(x_2659); +x_2660 = lean_ctor_get(x_2654, 1); +lean_inc(x_2660); +if (lean_is_exclusive(x_2654)) { + lean_ctor_release(x_2654, 0); + lean_ctor_release(x_2654, 1); + x_2661 = x_2654; +} else { + lean_dec_ref(x_2654); + x_2661 = lean_box(0); +} +x_2662 = lean_ctor_get(x_2659, 1); +lean_inc(x_2662); +x_2663 = lean_nat_dec_eq(x_2648, x_2662); +lean_dec(x_2648); +if (x_2663 == 0) +{ +lean_object* x_2664; +lean_dec(x_2662); +if (lean_is_scalar(x_2661)) { + x_2664 = lean_alloc_ctor(1, 2, 0); +} else { + x_2664 = x_2661; +} +lean_ctor_set(x_2664, 0, x_2659); +lean_ctor_set(x_2664, 1, x_2660); +return x_2664; +} +else +{ +lean_object* x_2665; lean_object* x_2666; lean_object* x_2667; lean_object* x_2668; +lean_dec(x_2661); +lean_dec(x_2660); +x_2665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_2666 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2666, 0, x_2665); +x_2667 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_2668 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2667, x_2666, x_2659); +if (lean_obj_tag(x_2668) == 0) +{ +lean_object* x_2669; lean_object* x_2670; lean_object* x_2671; lean_object* x_2672; +lean_dec(x_2662); +x_2669 = lean_ctor_get(x_2668, 0); +lean_inc(x_2669); +x_2670 = lean_ctor_get(x_2668, 1); +lean_inc(x_2670); +if (lean_is_exclusive(x_2668)) { + lean_ctor_release(x_2668, 0); + lean_ctor_release(x_2668, 1); + x_2671 = x_2668; +} else { + lean_dec_ref(x_2668); + x_2671 = lean_box(0); +} +if (lean_is_scalar(x_2671)) { + x_2672 = lean_alloc_ctor(0, 2, 0); +} else { + x_2672 = x_2671; +} +lean_ctor_set(x_2672, 0, x_2669); +lean_ctor_set(x_2672, 1, x_2670); +return x_2672; +} +else +{ +lean_object* x_2673; lean_object* x_2674; lean_object* x_2675; lean_object* x_2676; uint8_t x_2677; +x_2673 = lean_ctor_get(x_2668, 0); +lean_inc(x_2673); +x_2674 = lean_ctor_get(x_2668, 1); +lean_inc(x_2674); +if (lean_is_exclusive(x_2668)) { + lean_ctor_release(x_2668, 0); + lean_ctor_release(x_2668, 1); + x_2675 = x_2668; +} else { + lean_dec_ref(x_2668); + x_2675 = lean_box(0); +} +x_2676 = lean_ctor_get(x_2673, 1); +lean_inc(x_2676); +x_2677 = lean_nat_dec_eq(x_2662, x_2676); +lean_dec(x_2662); +if (x_2677 == 0) +{ +lean_object* x_2678; +lean_dec(x_2676); +if (lean_is_scalar(x_2675)) { + x_2678 = lean_alloc_ctor(1, 2, 0); +} else { + x_2678 = x_2675; +} +lean_ctor_set(x_2678, 0, x_2673); +lean_ctor_set(x_2678, 1, x_2674); +return x_2678; +} +else +{ +lean_object* x_2679; lean_object* x_2680; lean_object* x_2681; lean_object* x_2682; +lean_dec(x_2675); +lean_dec(x_2674); +x_2679 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_2680 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2680, 0, x_2679); +x_2681 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_2682 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2681, x_2680, x_2673); +if (lean_obj_tag(x_2682) == 0) +{ +lean_object* x_2683; lean_object* x_2684; lean_object* x_2685; lean_object* x_2686; +lean_dec(x_2676); +x_2683 = lean_ctor_get(x_2682, 0); +lean_inc(x_2683); +x_2684 = lean_ctor_get(x_2682, 1); +lean_inc(x_2684); +if (lean_is_exclusive(x_2682)) { + lean_ctor_release(x_2682, 0); + lean_ctor_release(x_2682, 1); + x_2685 = x_2682; +} else { + lean_dec_ref(x_2682); + x_2685 = lean_box(0); +} +if (lean_is_scalar(x_2685)) { + x_2686 = lean_alloc_ctor(0, 2, 0); +} else { + x_2686 = x_2685; +} +lean_ctor_set(x_2686, 0, x_2683); +lean_ctor_set(x_2686, 1, x_2684); +return x_2686; +} +else +{ +lean_object* x_2687; lean_object* x_2688; lean_object* x_2689; lean_object* x_2690; uint8_t x_2691; +x_2687 = lean_ctor_get(x_2682, 0); +lean_inc(x_2687); +x_2688 = lean_ctor_get(x_2682, 1); +lean_inc(x_2688); +if (lean_is_exclusive(x_2682)) { + lean_ctor_release(x_2682, 0); + lean_ctor_release(x_2682, 1); + x_2689 = x_2682; +} else { + lean_dec_ref(x_2682); + x_2689 = lean_box(0); +} +x_2690 = lean_ctor_get(x_2687, 1); +lean_inc(x_2690); +x_2691 = lean_nat_dec_eq(x_2676, x_2690); +lean_dec(x_2676); +if (x_2691 == 0) +{ +lean_object* x_2692; +lean_dec(x_2690); +if (lean_is_scalar(x_2689)) { + x_2692 = lean_alloc_ctor(1, 2, 0); +} else { + x_2692 = x_2689; +} +lean_ctor_set(x_2692, 0, x_2687); +lean_ctor_set(x_2692, 1, x_2688); +return x_2692; +} +else +{ +lean_object* x_2693; lean_object* x_2694; lean_object* x_2695; lean_object* x_2696; +lean_dec(x_2689); +lean_dec(x_2688); +x_2693 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_2694 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2694, 0, x_2693); +x_2695 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_2696 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2695, x_2694, x_2687); +if (lean_obj_tag(x_2696) == 0) +{ +lean_object* x_2697; lean_object* x_2698; lean_object* x_2699; lean_object* x_2700; +lean_dec(x_2690); +x_2697 = lean_ctor_get(x_2696, 0); +lean_inc(x_2697); +x_2698 = lean_ctor_get(x_2696, 1); +lean_inc(x_2698); +if (lean_is_exclusive(x_2696)) { + lean_ctor_release(x_2696, 0); + lean_ctor_release(x_2696, 1); + x_2699 = x_2696; +} else { + lean_dec_ref(x_2696); + x_2699 = lean_box(0); +} +if (lean_is_scalar(x_2699)) { + x_2700 = lean_alloc_ctor(0, 2, 0); +} else { + x_2700 = x_2699; +} +lean_ctor_set(x_2700, 0, x_2697); +lean_ctor_set(x_2700, 1, x_2698); +return x_2700; +} +else +{ +lean_object* x_2701; lean_object* x_2702; lean_object* x_2703; lean_object* x_2704; uint8_t x_2705; +x_2701 = lean_ctor_get(x_2696, 0); +lean_inc(x_2701); +x_2702 = lean_ctor_get(x_2696, 1); +lean_inc(x_2702); +if (lean_is_exclusive(x_2696)) { + lean_ctor_release(x_2696, 0); + lean_ctor_release(x_2696, 1); + x_2703 = x_2696; +} else { + lean_dec_ref(x_2696); + x_2703 = lean_box(0); +} +x_2704 = lean_ctor_get(x_2701, 1); +lean_inc(x_2704); +x_2705 = lean_nat_dec_eq(x_2690, x_2704); +lean_dec(x_2704); +lean_dec(x_2690); +if (x_2705 == 0) +{ +lean_object* x_2706; +if (lean_is_scalar(x_2703)) { + x_2706 = lean_alloc_ctor(1, 2, 0); +} else { + x_2706 = x_2703; +} +lean_ctor_set(x_2706, 0, x_2701); +lean_ctor_set(x_2706, 1, x_2702); +return x_2706; +} +else +{ +lean_object* x_2707; lean_object* x_2708; lean_object* x_2709; lean_object* x_2710; +lean_dec(x_2703); +lean_dec(x_2702); +x_2707 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_2708 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2708, 0, x_2707); +x_2709 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_2710 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2709, x_2708, x_2701); +return x_2710; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_2711; lean_object* x_2712; lean_object* x_2713; uint8_t x_2714; +x_2711 = lean_ctor_get(x_160, 0); +x_2712 = lean_ctor_get(x_160, 1); +lean_inc(x_2712); +lean_inc(x_2711); +lean_dec(x_160); +x_2713 = lean_ctor_get(x_2711, 1); +lean_inc(x_2713); +x_2714 = lean_nat_dec_eq(x_155, x_2713); +lean_dec(x_155); +if (x_2714 == 0) +{ +lean_object* x_2715; +lean_dec(x_2713); +x_2715 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2715, 0, x_2711); +lean_ctor_set(x_2715, 1, x_2712); +return x_2715; +} +else +{ +lean_object* x_2716; lean_object* x_2717; lean_object* x_2718; +lean_dec(x_2712); +x_2716 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_2717 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2717, 0, x_2716); +x_2718 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_159, x_2717, x_2711); +if (lean_obj_tag(x_2718) == 0) +{ +lean_object* x_2719; lean_object* x_2720; lean_object* x_2721; lean_object* x_2722; +lean_dec(x_2713); +x_2719 = lean_ctor_get(x_2718, 0); +lean_inc(x_2719); +x_2720 = lean_ctor_get(x_2718, 1); +lean_inc(x_2720); +if (lean_is_exclusive(x_2718)) { + lean_ctor_release(x_2718, 0); + lean_ctor_release(x_2718, 1); + x_2721 = x_2718; +} else { + lean_dec_ref(x_2718); + x_2721 = lean_box(0); +} +if (lean_is_scalar(x_2721)) { + x_2722 = lean_alloc_ctor(0, 2, 0); +} else { + x_2722 = x_2721; +} +lean_ctor_set(x_2722, 0, x_2719); +lean_ctor_set(x_2722, 1, x_2720); +return x_2722; +} +else +{ +lean_object* x_2723; lean_object* x_2724; lean_object* x_2725; lean_object* x_2726; uint8_t x_2727; +x_2723 = lean_ctor_get(x_2718, 0); +lean_inc(x_2723); +x_2724 = lean_ctor_get(x_2718, 1); +lean_inc(x_2724); +if (lean_is_exclusive(x_2718)) { + lean_ctor_release(x_2718, 0); + lean_ctor_release(x_2718, 1); + x_2725 = x_2718; +} else { + lean_dec_ref(x_2718); + x_2725 = lean_box(0); +} +x_2726 = lean_ctor_get(x_2723, 1); +lean_inc(x_2726); +x_2727 = lean_nat_dec_eq(x_2713, x_2726); +lean_dec(x_2713); +if (x_2727 == 0) +{ +lean_object* x_2728; +lean_dec(x_2726); +if (lean_is_scalar(x_2725)) { + x_2728 = lean_alloc_ctor(1, 2, 0); +} else { + x_2728 = x_2725; +} +lean_ctor_set(x_2728, 0, x_2723); +lean_ctor_set(x_2728, 1, x_2724); +return x_2728; +} +else +{ +lean_object* x_2729; lean_object* x_2730; lean_object* x_2731; lean_object* x_2732; +lean_dec(x_2725); +lean_dec(x_2724); +x_2729 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_2730 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2730, 0, x_2729); +x_2731 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_2732 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2731, x_2730, x_2723); +if (lean_obj_tag(x_2732) == 0) +{ +lean_object* x_2733; lean_object* x_2734; lean_object* x_2735; lean_object* x_2736; +lean_dec(x_2726); +x_2733 = lean_ctor_get(x_2732, 0); +lean_inc(x_2733); +x_2734 = lean_ctor_get(x_2732, 1); +lean_inc(x_2734); +if (lean_is_exclusive(x_2732)) { + lean_ctor_release(x_2732, 0); + lean_ctor_release(x_2732, 1); + x_2735 = x_2732; +} else { + lean_dec_ref(x_2732); + x_2735 = lean_box(0); +} +if (lean_is_scalar(x_2735)) { + x_2736 = lean_alloc_ctor(0, 2, 0); +} else { + x_2736 = x_2735; +} +lean_ctor_set(x_2736, 0, x_2733); +lean_ctor_set(x_2736, 1, x_2734); +return x_2736; +} +else +{ +lean_object* x_2737; lean_object* x_2738; lean_object* x_2739; lean_object* x_2740; uint8_t x_2741; +x_2737 = lean_ctor_get(x_2732, 0); +lean_inc(x_2737); +x_2738 = lean_ctor_get(x_2732, 1); +lean_inc(x_2738); +if (lean_is_exclusive(x_2732)) { + lean_ctor_release(x_2732, 0); + lean_ctor_release(x_2732, 1); + x_2739 = x_2732; +} else { + lean_dec_ref(x_2732); + x_2739 = lean_box(0); +} +x_2740 = lean_ctor_get(x_2737, 1); +lean_inc(x_2740); +x_2741 = lean_nat_dec_eq(x_2726, x_2740); +lean_dec(x_2726); +if (x_2741 == 0) +{ +lean_object* x_2742; +lean_dec(x_2740); +if (lean_is_scalar(x_2739)) { + x_2742 = lean_alloc_ctor(1, 2, 0); +} else { + x_2742 = x_2739; +} +lean_ctor_set(x_2742, 0, x_2737); +lean_ctor_set(x_2742, 1, x_2738); +return x_2742; +} +else +{ +lean_object* x_2743; lean_object* x_2744; lean_object* x_2745; lean_object* x_2746; +lean_dec(x_2739); +lean_dec(x_2738); +x_2743 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_2744 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2744, 0, x_2743); +x_2745 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_2746 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2745, x_2744, x_2737); +if (lean_obj_tag(x_2746) == 0) +{ +lean_object* x_2747; lean_object* x_2748; lean_object* x_2749; lean_object* x_2750; +lean_dec(x_2740); +x_2747 = lean_ctor_get(x_2746, 0); +lean_inc(x_2747); +x_2748 = lean_ctor_get(x_2746, 1); +lean_inc(x_2748); +if (lean_is_exclusive(x_2746)) { + lean_ctor_release(x_2746, 0); + lean_ctor_release(x_2746, 1); + x_2749 = x_2746; +} else { + lean_dec_ref(x_2746); + x_2749 = lean_box(0); +} +if (lean_is_scalar(x_2749)) { + x_2750 = lean_alloc_ctor(0, 2, 0); +} else { + x_2750 = x_2749; +} +lean_ctor_set(x_2750, 0, x_2747); +lean_ctor_set(x_2750, 1, x_2748); +return x_2750; +} +else +{ +lean_object* x_2751; lean_object* x_2752; lean_object* x_2753; lean_object* x_2754; uint8_t x_2755; +x_2751 = lean_ctor_get(x_2746, 0); +lean_inc(x_2751); +x_2752 = lean_ctor_get(x_2746, 1); +lean_inc(x_2752); +if (lean_is_exclusive(x_2746)) { + lean_ctor_release(x_2746, 0); + lean_ctor_release(x_2746, 1); + x_2753 = x_2746; +} else { + lean_dec_ref(x_2746); + x_2753 = lean_box(0); +} +x_2754 = lean_ctor_get(x_2751, 1); +lean_inc(x_2754); +x_2755 = lean_nat_dec_eq(x_2740, x_2754); +lean_dec(x_2740); +if (x_2755 == 0) +{ +lean_object* x_2756; +lean_dec(x_2754); +if (lean_is_scalar(x_2753)) { + x_2756 = lean_alloc_ctor(1, 2, 0); +} else { + x_2756 = x_2753; +} +lean_ctor_set(x_2756, 0, x_2751); +lean_ctor_set(x_2756, 1, x_2752); +return x_2756; +} +else +{ +lean_object* x_2757; lean_object* x_2758; lean_object* x_2759; lean_object* x_2760; +lean_dec(x_2753); +lean_dec(x_2752); +x_2757 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_2758 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2758, 0, x_2757); +x_2759 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_2760 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2759, x_2758, x_2751); +if (lean_obj_tag(x_2760) == 0) +{ +lean_object* x_2761; lean_object* x_2762; lean_object* x_2763; lean_object* x_2764; +lean_dec(x_2754); +x_2761 = lean_ctor_get(x_2760, 0); +lean_inc(x_2761); +x_2762 = lean_ctor_get(x_2760, 1); +lean_inc(x_2762); +if (lean_is_exclusive(x_2760)) { + lean_ctor_release(x_2760, 0); + lean_ctor_release(x_2760, 1); + x_2763 = x_2760; +} else { + lean_dec_ref(x_2760); + x_2763 = lean_box(0); +} +if (lean_is_scalar(x_2763)) { + x_2764 = lean_alloc_ctor(0, 2, 0); +} else { + x_2764 = x_2763; +} +lean_ctor_set(x_2764, 0, x_2761); +lean_ctor_set(x_2764, 1, x_2762); +return x_2764; +} +else +{ +lean_object* x_2765; lean_object* x_2766; lean_object* x_2767; lean_object* x_2768; uint8_t x_2769; +x_2765 = lean_ctor_get(x_2760, 0); +lean_inc(x_2765); +x_2766 = lean_ctor_get(x_2760, 1); +lean_inc(x_2766); +if (lean_is_exclusive(x_2760)) { + lean_ctor_release(x_2760, 0); + lean_ctor_release(x_2760, 1); + x_2767 = x_2760; +} else { + lean_dec_ref(x_2760); + x_2767 = lean_box(0); +} +x_2768 = lean_ctor_get(x_2765, 1); +lean_inc(x_2768); +x_2769 = lean_nat_dec_eq(x_2754, x_2768); +lean_dec(x_2754); +if (x_2769 == 0) +{ +lean_object* x_2770; +lean_dec(x_2768); +if (lean_is_scalar(x_2767)) { + x_2770 = lean_alloc_ctor(1, 2, 0); +} else { + x_2770 = x_2767; +} +lean_ctor_set(x_2770, 0, x_2765); +lean_ctor_set(x_2770, 1, x_2766); +return x_2770; +} +else +{ +lean_object* x_2771; lean_object* x_2772; lean_object* x_2773; lean_object* x_2774; +lean_dec(x_2767); +lean_dec(x_2766); +x_2771 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_2772 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2772, 0, x_2771); +x_2773 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_2774 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2773, x_2772, x_2765); +if (lean_obj_tag(x_2774) == 0) +{ +lean_object* x_2775; lean_object* x_2776; lean_object* x_2777; lean_object* x_2778; +lean_dec(x_2768); +x_2775 = lean_ctor_get(x_2774, 0); +lean_inc(x_2775); +x_2776 = lean_ctor_get(x_2774, 1); +lean_inc(x_2776); +if (lean_is_exclusive(x_2774)) { + lean_ctor_release(x_2774, 0); + lean_ctor_release(x_2774, 1); + x_2777 = x_2774; +} else { + lean_dec_ref(x_2774); + x_2777 = lean_box(0); +} +if (lean_is_scalar(x_2777)) { + x_2778 = lean_alloc_ctor(0, 2, 0); +} else { + x_2778 = x_2777; +} +lean_ctor_set(x_2778, 0, x_2775); +lean_ctor_set(x_2778, 1, x_2776); +return x_2778; +} +else +{ +lean_object* x_2779; lean_object* x_2780; lean_object* x_2781; lean_object* x_2782; uint8_t x_2783; +x_2779 = lean_ctor_get(x_2774, 0); +lean_inc(x_2779); +x_2780 = lean_ctor_get(x_2774, 1); +lean_inc(x_2780); +if (lean_is_exclusive(x_2774)) { + lean_ctor_release(x_2774, 0); + lean_ctor_release(x_2774, 1); + x_2781 = x_2774; +} else { + lean_dec_ref(x_2774); + x_2781 = lean_box(0); +} +x_2782 = lean_ctor_get(x_2779, 1); +lean_inc(x_2782); +x_2783 = lean_nat_dec_eq(x_2768, x_2782); +lean_dec(x_2768); +if (x_2783 == 0) +{ +lean_object* x_2784; +lean_dec(x_2782); +if (lean_is_scalar(x_2781)) { + x_2784 = lean_alloc_ctor(1, 2, 0); +} else { + x_2784 = x_2781; +} +lean_ctor_set(x_2784, 0, x_2779); +lean_ctor_set(x_2784, 1, x_2780); +return x_2784; +} +else +{ +lean_object* x_2785; lean_object* x_2786; lean_object* x_2787; lean_object* x_2788; +lean_dec(x_2781); +lean_dec(x_2780); +x_2785 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_2786 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2786, 0, x_2785); +x_2787 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_2788 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2787, x_2786, x_2779); +if (lean_obj_tag(x_2788) == 0) +{ +lean_object* x_2789; lean_object* x_2790; lean_object* x_2791; lean_object* x_2792; +lean_dec(x_2782); +x_2789 = lean_ctor_get(x_2788, 0); +lean_inc(x_2789); +x_2790 = lean_ctor_get(x_2788, 1); +lean_inc(x_2790); +if (lean_is_exclusive(x_2788)) { + lean_ctor_release(x_2788, 0); + lean_ctor_release(x_2788, 1); + x_2791 = x_2788; +} else { + lean_dec_ref(x_2788); + x_2791 = lean_box(0); +} +if (lean_is_scalar(x_2791)) { + x_2792 = lean_alloc_ctor(0, 2, 0); +} else { + x_2792 = x_2791; +} +lean_ctor_set(x_2792, 0, x_2789); +lean_ctor_set(x_2792, 1, x_2790); +return x_2792; +} +else +{ +lean_object* x_2793; lean_object* x_2794; lean_object* x_2795; lean_object* x_2796; uint8_t x_2797; +x_2793 = lean_ctor_get(x_2788, 0); +lean_inc(x_2793); +x_2794 = lean_ctor_get(x_2788, 1); +lean_inc(x_2794); +if (lean_is_exclusive(x_2788)) { + lean_ctor_release(x_2788, 0); + lean_ctor_release(x_2788, 1); + x_2795 = x_2788; +} else { + lean_dec_ref(x_2788); + x_2795 = lean_box(0); +} +x_2796 = lean_ctor_get(x_2793, 1); +lean_inc(x_2796); +x_2797 = lean_nat_dec_eq(x_2782, x_2796); +lean_dec(x_2782); +if (x_2797 == 0) +{ +lean_object* x_2798; +lean_dec(x_2796); +if (lean_is_scalar(x_2795)) { + x_2798 = lean_alloc_ctor(1, 2, 0); +} else { + x_2798 = x_2795; +} +lean_ctor_set(x_2798, 0, x_2793); +lean_ctor_set(x_2798, 1, x_2794); +return x_2798; +} +else +{ +lean_object* x_2799; lean_object* x_2800; lean_object* x_2801; lean_object* x_2802; +lean_dec(x_2795); +lean_dec(x_2794); +x_2799 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_2800 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2800, 0, x_2799); +x_2801 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_2802 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2801, x_2800, x_2793); +if (lean_obj_tag(x_2802) == 0) +{ +lean_object* x_2803; lean_object* x_2804; lean_object* x_2805; lean_object* x_2806; +lean_dec(x_2796); +x_2803 = lean_ctor_get(x_2802, 0); +lean_inc(x_2803); +x_2804 = lean_ctor_get(x_2802, 1); +lean_inc(x_2804); +if (lean_is_exclusive(x_2802)) { + lean_ctor_release(x_2802, 0); + lean_ctor_release(x_2802, 1); + x_2805 = x_2802; +} else { + lean_dec_ref(x_2802); + x_2805 = lean_box(0); +} +if (lean_is_scalar(x_2805)) { + x_2806 = lean_alloc_ctor(0, 2, 0); +} else { + x_2806 = x_2805; +} +lean_ctor_set(x_2806, 0, x_2803); +lean_ctor_set(x_2806, 1, x_2804); +return x_2806; +} +else +{ +lean_object* x_2807; lean_object* x_2808; lean_object* x_2809; lean_object* x_2810; uint8_t x_2811; +x_2807 = lean_ctor_get(x_2802, 0); +lean_inc(x_2807); +x_2808 = lean_ctor_get(x_2802, 1); +lean_inc(x_2808); +if (lean_is_exclusive(x_2802)) { + lean_ctor_release(x_2802, 0); + lean_ctor_release(x_2802, 1); + x_2809 = x_2802; +} else { + lean_dec_ref(x_2802); + x_2809 = lean_box(0); +} +x_2810 = lean_ctor_get(x_2807, 1); +lean_inc(x_2810); +x_2811 = lean_nat_dec_eq(x_2796, x_2810); +lean_dec(x_2796); +if (x_2811 == 0) +{ +lean_object* x_2812; +lean_dec(x_2810); +if (lean_is_scalar(x_2809)) { + x_2812 = lean_alloc_ctor(1, 2, 0); +} else { + x_2812 = x_2809; +} +lean_ctor_set(x_2812, 0, x_2807); +lean_ctor_set(x_2812, 1, x_2808); +return x_2812; +} +else +{ +lean_object* x_2813; lean_object* x_2814; lean_object* x_2815; lean_object* x_2816; +lean_dec(x_2809); +lean_dec(x_2808); +x_2813 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_2814 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2814, 0, x_2813); +x_2815 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_2816 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2815, x_2814, x_2807); +if (lean_obj_tag(x_2816) == 0) +{ +lean_object* x_2817; lean_object* x_2818; lean_object* x_2819; lean_object* x_2820; +lean_dec(x_2810); +x_2817 = lean_ctor_get(x_2816, 0); +lean_inc(x_2817); +x_2818 = lean_ctor_get(x_2816, 1); +lean_inc(x_2818); +if (lean_is_exclusive(x_2816)) { + lean_ctor_release(x_2816, 0); + lean_ctor_release(x_2816, 1); + x_2819 = x_2816; +} else { + lean_dec_ref(x_2816); + x_2819 = lean_box(0); +} +if (lean_is_scalar(x_2819)) { + x_2820 = lean_alloc_ctor(0, 2, 0); +} else { + x_2820 = x_2819; +} +lean_ctor_set(x_2820, 0, x_2817); +lean_ctor_set(x_2820, 1, x_2818); +return x_2820; +} +else +{ +lean_object* x_2821; lean_object* x_2822; lean_object* x_2823; lean_object* x_2824; uint8_t x_2825; +x_2821 = lean_ctor_get(x_2816, 0); +lean_inc(x_2821); +x_2822 = lean_ctor_get(x_2816, 1); +lean_inc(x_2822); +if (lean_is_exclusive(x_2816)) { + lean_ctor_release(x_2816, 0); + lean_ctor_release(x_2816, 1); + x_2823 = x_2816; +} else { + lean_dec_ref(x_2816); + x_2823 = lean_box(0); +} +x_2824 = lean_ctor_get(x_2821, 1); +lean_inc(x_2824); +x_2825 = lean_nat_dec_eq(x_2810, x_2824); +lean_dec(x_2810); +if (x_2825 == 0) +{ +lean_object* x_2826; +lean_dec(x_2824); +if (lean_is_scalar(x_2823)) { + x_2826 = lean_alloc_ctor(1, 2, 0); +} else { + x_2826 = x_2823; +} +lean_ctor_set(x_2826, 0, x_2821); +lean_ctor_set(x_2826, 1, x_2822); +return x_2826; +} +else +{ +lean_object* x_2827; lean_object* x_2828; lean_object* x_2829; lean_object* x_2830; +lean_dec(x_2823); +lean_dec(x_2822); +x_2827 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_2828 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2828, 0, x_2827); +x_2829 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_2830 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2829, x_2828, x_2821); +if (lean_obj_tag(x_2830) == 0) +{ +lean_object* x_2831; lean_object* x_2832; lean_object* x_2833; lean_object* x_2834; +lean_dec(x_2824); +x_2831 = lean_ctor_get(x_2830, 0); +lean_inc(x_2831); +x_2832 = lean_ctor_get(x_2830, 1); +lean_inc(x_2832); +if (lean_is_exclusive(x_2830)) { + lean_ctor_release(x_2830, 0); + lean_ctor_release(x_2830, 1); + x_2833 = x_2830; +} else { + lean_dec_ref(x_2830); + x_2833 = lean_box(0); +} +if (lean_is_scalar(x_2833)) { + x_2834 = lean_alloc_ctor(0, 2, 0); +} else { + x_2834 = x_2833; +} +lean_ctor_set(x_2834, 0, x_2831); +lean_ctor_set(x_2834, 1, x_2832); +return x_2834; +} +else +{ +lean_object* x_2835; lean_object* x_2836; lean_object* x_2837; lean_object* x_2838; uint8_t x_2839; +x_2835 = lean_ctor_get(x_2830, 0); +lean_inc(x_2835); +x_2836 = lean_ctor_get(x_2830, 1); +lean_inc(x_2836); +if (lean_is_exclusive(x_2830)) { + lean_ctor_release(x_2830, 0); + lean_ctor_release(x_2830, 1); + x_2837 = x_2830; +} else { + lean_dec_ref(x_2830); + x_2837 = lean_box(0); +} +x_2838 = lean_ctor_get(x_2835, 1); +lean_inc(x_2838); +x_2839 = lean_nat_dec_eq(x_2824, x_2838); +lean_dec(x_2824); +if (x_2839 == 0) +{ +lean_object* x_2840; +lean_dec(x_2838); +if (lean_is_scalar(x_2837)) { + x_2840 = lean_alloc_ctor(1, 2, 0); +} else { + x_2840 = x_2837; +} +lean_ctor_set(x_2840, 0, x_2835); +lean_ctor_set(x_2840, 1, x_2836); +return x_2840; +} +else +{ +lean_object* x_2841; lean_object* x_2842; lean_object* x_2843; lean_object* x_2844; +lean_dec(x_2837); +lean_dec(x_2836); +x_2841 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_2842 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2842, 0, x_2841); +x_2843 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_2844 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2843, x_2842, x_2835); +if (lean_obj_tag(x_2844) == 0) +{ +lean_object* x_2845; lean_object* x_2846; lean_object* x_2847; lean_object* x_2848; +lean_dec(x_2838); +x_2845 = lean_ctor_get(x_2844, 0); +lean_inc(x_2845); +x_2846 = lean_ctor_get(x_2844, 1); +lean_inc(x_2846); +if (lean_is_exclusive(x_2844)) { + lean_ctor_release(x_2844, 0); + lean_ctor_release(x_2844, 1); + x_2847 = x_2844; +} else { + lean_dec_ref(x_2844); + x_2847 = lean_box(0); +} +if (lean_is_scalar(x_2847)) { + x_2848 = lean_alloc_ctor(0, 2, 0); +} else { + x_2848 = x_2847; +} +lean_ctor_set(x_2848, 0, x_2845); +lean_ctor_set(x_2848, 1, x_2846); +return x_2848; +} +else +{ +lean_object* x_2849; lean_object* x_2850; lean_object* x_2851; lean_object* x_2852; uint8_t x_2853; +x_2849 = lean_ctor_get(x_2844, 0); +lean_inc(x_2849); +x_2850 = lean_ctor_get(x_2844, 1); +lean_inc(x_2850); +if (lean_is_exclusive(x_2844)) { + lean_ctor_release(x_2844, 0); + lean_ctor_release(x_2844, 1); + x_2851 = x_2844; +} else { + lean_dec_ref(x_2844); + x_2851 = lean_box(0); +} +x_2852 = lean_ctor_get(x_2849, 1); +lean_inc(x_2852); +x_2853 = lean_nat_dec_eq(x_2838, x_2852); +lean_dec(x_2838); +if (x_2853 == 0) +{ +lean_object* x_2854; +lean_dec(x_2852); +if (lean_is_scalar(x_2851)) { + x_2854 = lean_alloc_ctor(1, 2, 0); +} else { + x_2854 = x_2851; +} +lean_ctor_set(x_2854, 0, x_2849); +lean_ctor_set(x_2854, 1, x_2850); +return x_2854; +} +else +{ +lean_object* x_2855; lean_object* x_2856; lean_object* x_2857; lean_object* x_2858; +lean_dec(x_2851); +lean_dec(x_2850); +x_2855 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_2856 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2856, 0, x_2855); +x_2857 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_2858 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2857, x_2856, x_2849); +if (lean_obj_tag(x_2858) == 0) +{ +lean_object* x_2859; lean_object* x_2860; lean_object* x_2861; lean_object* x_2862; +lean_dec(x_2852); +x_2859 = lean_ctor_get(x_2858, 0); +lean_inc(x_2859); +x_2860 = lean_ctor_get(x_2858, 1); +lean_inc(x_2860); +if (lean_is_exclusive(x_2858)) { + lean_ctor_release(x_2858, 0); + lean_ctor_release(x_2858, 1); + x_2861 = x_2858; +} else { + lean_dec_ref(x_2858); + x_2861 = lean_box(0); +} +if (lean_is_scalar(x_2861)) { + x_2862 = lean_alloc_ctor(0, 2, 0); +} else { + x_2862 = x_2861; +} +lean_ctor_set(x_2862, 0, x_2859); +lean_ctor_set(x_2862, 1, x_2860); +return x_2862; +} +else +{ +lean_object* x_2863; lean_object* x_2864; lean_object* x_2865; lean_object* x_2866; uint8_t x_2867; +x_2863 = lean_ctor_get(x_2858, 0); +lean_inc(x_2863); +x_2864 = lean_ctor_get(x_2858, 1); +lean_inc(x_2864); +if (lean_is_exclusive(x_2858)) { + lean_ctor_release(x_2858, 0); + lean_ctor_release(x_2858, 1); + x_2865 = x_2858; +} else { + lean_dec_ref(x_2858); + x_2865 = lean_box(0); +} +x_2866 = lean_ctor_get(x_2863, 1); +lean_inc(x_2866); +x_2867 = lean_nat_dec_eq(x_2852, x_2866); +lean_dec(x_2852); +if (x_2867 == 0) +{ +lean_object* x_2868; +lean_dec(x_2866); +if (lean_is_scalar(x_2865)) { + x_2868 = lean_alloc_ctor(1, 2, 0); +} else { + x_2868 = x_2865; +} +lean_ctor_set(x_2868, 0, x_2863); +lean_ctor_set(x_2868, 1, x_2864); +return x_2868; +} +else +{ +lean_object* x_2869; lean_object* x_2870; lean_object* x_2871; lean_object* x_2872; +lean_dec(x_2865); +lean_dec(x_2864); +x_2869 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_2870 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2870, 0, x_2869); +x_2871 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_2872 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2871, x_2870, x_2863); +if (lean_obj_tag(x_2872) == 0) +{ +lean_object* x_2873; lean_object* x_2874; lean_object* x_2875; lean_object* x_2876; +lean_dec(x_2866); +x_2873 = lean_ctor_get(x_2872, 0); +lean_inc(x_2873); +x_2874 = lean_ctor_get(x_2872, 1); +lean_inc(x_2874); +if (lean_is_exclusive(x_2872)) { + lean_ctor_release(x_2872, 0); + lean_ctor_release(x_2872, 1); + x_2875 = x_2872; +} else { + lean_dec_ref(x_2872); + x_2875 = lean_box(0); +} +if (lean_is_scalar(x_2875)) { + x_2876 = lean_alloc_ctor(0, 2, 0); +} else { + x_2876 = x_2875; +} +lean_ctor_set(x_2876, 0, x_2873); +lean_ctor_set(x_2876, 1, x_2874); +return x_2876; +} +else +{ +lean_object* x_2877; lean_object* x_2878; lean_object* x_2879; lean_object* x_2880; uint8_t x_2881; +x_2877 = lean_ctor_get(x_2872, 0); +lean_inc(x_2877); +x_2878 = lean_ctor_get(x_2872, 1); +lean_inc(x_2878); +if (lean_is_exclusive(x_2872)) { + lean_ctor_release(x_2872, 0); + lean_ctor_release(x_2872, 1); + x_2879 = x_2872; +} else { + lean_dec_ref(x_2872); + x_2879 = lean_box(0); +} +x_2880 = lean_ctor_get(x_2877, 1); +lean_inc(x_2880); +x_2881 = lean_nat_dec_eq(x_2866, x_2880); +lean_dec(x_2866); +if (x_2881 == 0) +{ +lean_object* x_2882; +lean_dec(x_2880); +if (lean_is_scalar(x_2879)) { + x_2882 = lean_alloc_ctor(1, 2, 0); +} else { + x_2882 = x_2879; +} +lean_ctor_set(x_2882, 0, x_2877); +lean_ctor_set(x_2882, 1, x_2878); +return x_2882; +} +else +{ +lean_object* x_2883; lean_object* x_2884; lean_object* x_2885; lean_object* x_2886; +lean_dec(x_2879); +lean_dec(x_2878); +x_2883 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_2884 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2884, 0, x_2883); +x_2885 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_2886 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2885, x_2884, x_2877); +if (lean_obj_tag(x_2886) == 0) +{ +lean_object* x_2887; lean_object* x_2888; lean_object* x_2889; lean_object* x_2890; +lean_dec(x_2880); +x_2887 = lean_ctor_get(x_2886, 0); +lean_inc(x_2887); +x_2888 = lean_ctor_get(x_2886, 1); +lean_inc(x_2888); +if (lean_is_exclusive(x_2886)) { + lean_ctor_release(x_2886, 0); + lean_ctor_release(x_2886, 1); + x_2889 = x_2886; +} else { + lean_dec_ref(x_2886); + x_2889 = lean_box(0); +} +if (lean_is_scalar(x_2889)) { + x_2890 = lean_alloc_ctor(0, 2, 0); +} else { + x_2890 = x_2889; +} +lean_ctor_set(x_2890, 0, x_2887); +lean_ctor_set(x_2890, 1, x_2888); +return x_2890; +} +else +{ +lean_object* x_2891; lean_object* x_2892; lean_object* x_2893; lean_object* x_2894; uint8_t x_2895; +x_2891 = lean_ctor_get(x_2886, 0); +lean_inc(x_2891); +x_2892 = lean_ctor_get(x_2886, 1); +lean_inc(x_2892); +if (lean_is_exclusive(x_2886)) { + lean_ctor_release(x_2886, 0); + lean_ctor_release(x_2886, 1); + x_2893 = x_2886; +} else { + lean_dec_ref(x_2886); + x_2893 = lean_box(0); +} +x_2894 = lean_ctor_get(x_2891, 1); +lean_inc(x_2894); +x_2895 = lean_nat_dec_eq(x_2880, x_2894); +lean_dec(x_2880); +if (x_2895 == 0) +{ +lean_object* x_2896; +lean_dec(x_2894); +if (lean_is_scalar(x_2893)) { + x_2896 = lean_alloc_ctor(1, 2, 0); +} else { + x_2896 = x_2893; +} +lean_ctor_set(x_2896, 0, x_2891); +lean_ctor_set(x_2896, 1, x_2892); +return x_2896; +} +else +{ +lean_object* x_2897; lean_object* x_2898; lean_object* x_2899; lean_object* x_2900; +lean_dec(x_2893); +lean_dec(x_2892); +x_2897 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_2898 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2898, 0, x_2897); +x_2899 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_2900 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2899, x_2898, x_2891); +if (lean_obj_tag(x_2900) == 0) +{ +lean_object* x_2901; lean_object* x_2902; lean_object* x_2903; lean_object* x_2904; +lean_dec(x_2894); +x_2901 = lean_ctor_get(x_2900, 0); +lean_inc(x_2901); +x_2902 = lean_ctor_get(x_2900, 1); +lean_inc(x_2902); +if (lean_is_exclusive(x_2900)) { + lean_ctor_release(x_2900, 0); + lean_ctor_release(x_2900, 1); + x_2903 = x_2900; +} else { + lean_dec_ref(x_2900); + x_2903 = lean_box(0); +} +if (lean_is_scalar(x_2903)) { + x_2904 = lean_alloc_ctor(0, 2, 0); +} else { + x_2904 = x_2903; +} +lean_ctor_set(x_2904, 0, x_2901); +lean_ctor_set(x_2904, 1, x_2902); +return x_2904; +} +else +{ +lean_object* x_2905; lean_object* x_2906; lean_object* x_2907; lean_object* x_2908; uint8_t x_2909; +x_2905 = lean_ctor_get(x_2900, 0); +lean_inc(x_2905); +x_2906 = lean_ctor_get(x_2900, 1); +lean_inc(x_2906); +if (lean_is_exclusive(x_2900)) { + lean_ctor_release(x_2900, 0); + lean_ctor_release(x_2900, 1); + x_2907 = x_2900; +} else { + lean_dec_ref(x_2900); + x_2907 = lean_box(0); +} +x_2908 = lean_ctor_get(x_2905, 1); +lean_inc(x_2908); +x_2909 = lean_nat_dec_eq(x_2894, x_2908); +lean_dec(x_2894); +if (x_2909 == 0) +{ +lean_object* x_2910; +lean_dec(x_2908); +if (lean_is_scalar(x_2907)) { + x_2910 = lean_alloc_ctor(1, 2, 0); +} else { + x_2910 = x_2907; +} +lean_ctor_set(x_2910, 0, x_2905); +lean_ctor_set(x_2910, 1, x_2906); +return x_2910; +} +else +{ +lean_object* x_2911; lean_object* x_2912; lean_object* x_2913; lean_object* x_2914; +lean_dec(x_2907); +lean_dec(x_2906); +x_2911 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_2912 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2912, 0, x_2911); +x_2913 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_2914 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2913, x_2912, x_2905); +if (lean_obj_tag(x_2914) == 0) +{ +lean_object* x_2915; lean_object* x_2916; lean_object* x_2917; lean_object* x_2918; +lean_dec(x_2908); +x_2915 = lean_ctor_get(x_2914, 0); +lean_inc(x_2915); +x_2916 = lean_ctor_get(x_2914, 1); +lean_inc(x_2916); +if (lean_is_exclusive(x_2914)) { + lean_ctor_release(x_2914, 0); + lean_ctor_release(x_2914, 1); + x_2917 = x_2914; +} else { + lean_dec_ref(x_2914); + x_2917 = lean_box(0); +} +if (lean_is_scalar(x_2917)) { + x_2918 = lean_alloc_ctor(0, 2, 0); +} else { + x_2918 = x_2917; +} +lean_ctor_set(x_2918, 0, x_2915); +lean_ctor_set(x_2918, 1, x_2916); +return x_2918; +} +else +{ +lean_object* x_2919; lean_object* x_2920; lean_object* x_2921; lean_object* x_2922; uint8_t x_2923; +x_2919 = lean_ctor_get(x_2914, 0); +lean_inc(x_2919); +x_2920 = lean_ctor_get(x_2914, 1); +lean_inc(x_2920); +if (lean_is_exclusive(x_2914)) { + lean_ctor_release(x_2914, 0); + lean_ctor_release(x_2914, 1); + x_2921 = x_2914; +} else { + lean_dec_ref(x_2914); + x_2921 = lean_box(0); +} +x_2922 = lean_ctor_get(x_2919, 1); +lean_inc(x_2922); +x_2923 = lean_nat_dec_eq(x_2908, x_2922); +lean_dec(x_2908); +if (x_2923 == 0) +{ +lean_object* x_2924; +lean_dec(x_2922); +if (lean_is_scalar(x_2921)) { + x_2924 = lean_alloc_ctor(1, 2, 0); +} else { + x_2924 = x_2921; +} +lean_ctor_set(x_2924, 0, x_2919); +lean_ctor_set(x_2924, 1, x_2920); +return x_2924; +} +else +{ +lean_object* x_2925; lean_object* x_2926; lean_object* x_2927; lean_object* x_2928; +lean_dec(x_2921); +lean_dec(x_2920); +x_2925 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_2926 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2926, 0, x_2925); +x_2927 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_2928 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2927, x_2926, x_2919); +if (lean_obj_tag(x_2928) == 0) +{ +lean_object* x_2929; lean_object* x_2930; lean_object* x_2931; lean_object* x_2932; +lean_dec(x_2922); +x_2929 = lean_ctor_get(x_2928, 0); +lean_inc(x_2929); +x_2930 = lean_ctor_get(x_2928, 1); +lean_inc(x_2930); +if (lean_is_exclusive(x_2928)) { + lean_ctor_release(x_2928, 0); + lean_ctor_release(x_2928, 1); + x_2931 = x_2928; +} else { + lean_dec_ref(x_2928); + x_2931 = lean_box(0); +} +if (lean_is_scalar(x_2931)) { + x_2932 = lean_alloc_ctor(0, 2, 0); +} else { + x_2932 = x_2931; +} +lean_ctor_set(x_2932, 0, x_2929); +lean_ctor_set(x_2932, 1, x_2930); +return x_2932; +} +else +{ +lean_object* x_2933; lean_object* x_2934; lean_object* x_2935; lean_object* x_2936; uint8_t x_2937; +x_2933 = lean_ctor_get(x_2928, 0); +lean_inc(x_2933); +x_2934 = lean_ctor_get(x_2928, 1); +lean_inc(x_2934); +if (lean_is_exclusive(x_2928)) { + lean_ctor_release(x_2928, 0); + lean_ctor_release(x_2928, 1); + x_2935 = x_2928; +} else { + lean_dec_ref(x_2928); + x_2935 = lean_box(0); +} +x_2936 = lean_ctor_get(x_2933, 1); +lean_inc(x_2936); +x_2937 = lean_nat_dec_eq(x_2922, x_2936); +lean_dec(x_2922); +if (x_2937 == 0) +{ +lean_object* x_2938; +lean_dec(x_2936); +if (lean_is_scalar(x_2935)) { + x_2938 = lean_alloc_ctor(1, 2, 0); +} else { + x_2938 = x_2935; +} +lean_ctor_set(x_2938, 0, x_2933); +lean_ctor_set(x_2938, 1, x_2934); +return x_2938; +} +else +{ +lean_object* x_2939; lean_object* x_2940; lean_object* x_2941; lean_object* x_2942; +lean_dec(x_2935); +lean_dec(x_2934); +x_2939 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_2940 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2940, 0, x_2939); +x_2941 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_2942 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2941, x_2940, x_2933); +if (lean_obj_tag(x_2942) == 0) +{ +lean_object* x_2943; lean_object* x_2944; lean_object* x_2945; lean_object* x_2946; +lean_dec(x_2936); +x_2943 = lean_ctor_get(x_2942, 0); +lean_inc(x_2943); +x_2944 = lean_ctor_get(x_2942, 1); +lean_inc(x_2944); +if (lean_is_exclusive(x_2942)) { + lean_ctor_release(x_2942, 0); + lean_ctor_release(x_2942, 1); + x_2945 = x_2942; +} else { + lean_dec_ref(x_2942); + x_2945 = lean_box(0); +} +if (lean_is_scalar(x_2945)) { + x_2946 = lean_alloc_ctor(0, 2, 0); +} else { + x_2946 = x_2945; +} +lean_ctor_set(x_2946, 0, x_2943); +lean_ctor_set(x_2946, 1, x_2944); +return x_2946; +} +else +{ +lean_object* x_2947; lean_object* x_2948; lean_object* x_2949; lean_object* x_2950; uint8_t x_2951; +x_2947 = lean_ctor_get(x_2942, 0); +lean_inc(x_2947); +x_2948 = lean_ctor_get(x_2942, 1); +lean_inc(x_2948); +if (lean_is_exclusive(x_2942)) { + lean_ctor_release(x_2942, 0); + lean_ctor_release(x_2942, 1); + x_2949 = x_2942; +} else { + lean_dec_ref(x_2942); + x_2949 = lean_box(0); +} +x_2950 = lean_ctor_get(x_2947, 1); +lean_inc(x_2950); +x_2951 = lean_nat_dec_eq(x_2936, x_2950); +lean_dec(x_2936); +if (x_2951 == 0) +{ +lean_object* x_2952; +lean_dec(x_2950); +if (lean_is_scalar(x_2949)) { + x_2952 = lean_alloc_ctor(1, 2, 0); +} else { + x_2952 = x_2949; +} +lean_ctor_set(x_2952, 0, x_2947); +lean_ctor_set(x_2952, 1, x_2948); +return x_2952; +} +else +{ +lean_object* x_2953; lean_object* x_2954; lean_object* x_2955; lean_object* x_2956; +lean_dec(x_2949); +lean_dec(x_2948); +x_2953 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_2954 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2954, 0, x_2953); +x_2955 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_2956 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2955, x_2954, x_2947); +if (lean_obj_tag(x_2956) == 0) +{ +lean_object* x_2957; lean_object* x_2958; lean_object* x_2959; lean_object* x_2960; +lean_dec(x_2950); +x_2957 = lean_ctor_get(x_2956, 0); +lean_inc(x_2957); +x_2958 = lean_ctor_get(x_2956, 1); +lean_inc(x_2958); +if (lean_is_exclusive(x_2956)) { + lean_ctor_release(x_2956, 0); + lean_ctor_release(x_2956, 1); + x_2959 = x_2956; +} else { + lean_dec_ref(x_2956); + x_2959 = lean_box(0); +} +if (lean_is_scalar(x_2959)) { + x_2960 = lean_alloc_ctor(0, 2, 0); +} else { + x_2960 = x_2959; +} +lean_ctor_set(x_2960, 0, x_2957); +lean_ctor_set(x_2960, 1, x_2958); +return x_2960; +} +else +{ +lean_object* x_2961; lean_object* x_2962; lean_object* x_2963; lean_object* x_2964; uint8_t x_2965; +x_2961 = lean_ctor_get(x_2956, 0); +lean_inc(x_2961); +x_2962 = lean_ctor_get(x_2956, 1); +lean_inc(x_2962); +if (lean_is_exclusive(x_2956)) { + lean_ctor_release(x_2956, 0); + lean_ctor_release(x_2956, 1); + x_2963 = x_2956; +} else { + lean_dec_ref(x_2956); + x_2963 = lean_box(0); +} +x_2964 = lean_ctor_get(x_2961, 1); +lean_inc(x_2964); +x_2965 = lean_nat_dec_eq(x_2950, x_2964); +lean_dec(x_2964); +lean_dec(x_2950); +if (x_2965 == 0) +{ +lean_object* x_2966; +if (lean_is_scalar(x_2963)) { + x_2966 = lean_alloc_ctor(1, 2, 0); +} else { + x_2966 = x_2963; +} +lean_ctor_set(x_2966, 0, x_2961); +lean_ctor_set(x_2966, 1, x_2962); +return x_2966; +} +else +{ +lean_object* x_2967; lean_object* x_2968; lean_object* x_2969; lean_object* x_2970; +lean_dec(x_2963); +lean_dec(x_2962); +x_2967 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_2968 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2968, 0, x_2967); +x_2969 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_2970 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2969, x_2968, x_2961); +return x_2970; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_2971; lean_object* x_2972; lean_object* x_2973; uint8_t x_2974; +x_2971 = lean_ctor_get(x_147, 0); +x_2972 = lean_ctor_get(x_147, 1); +lean_inc(x_2972); +lean_inc(x_2971); +lean_dec(x_147); +x_2973 = lean_ctor_get(x_2971, 1); +lean_inc(x_2973); +x_2974 = lean_nat_dec_eq(x_142, x_2973); +lean_dec(x_142); +if (x_2974 == 0) +{ +lean_object* x_2975; +lean_dec(x_2973); +x_2975 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_2975, 0, x_2971); +lean_ctor_set(x_2975, 1, x_2972); +return x_2975; +} +else +{ +lean_object* x_2976; lean_object* x_2977; lean_object* x_2978; lean_object* x_2979; +lean_dec(x_2972); +x_2976 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_2977 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2977, 0, x_2976); +x_2978 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_2979 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2978, x_2977, x_2971); +if (lean_obj_tag(x_2979) == 0) +{ +lean_object* x_2980; lean_object* x_2981; lean_object* x_2982; lean_object* x_2983; +lean_dec(x_2973); +x_2980 = lean_ctor_get(x_2979, 0); +lean_inc(x_2980); +x_2981 = lean_ctor_get(x_2979, 1); +lean_inc(x_2981); +if (lean_is_exclusive(x_2979)) { + lean_ctor_release(x_2979, 0); + lean_ctor_release(x_2979, 1); + x_2982 = x_2979; +} else { + lean_dec_ref(x_2979); + x_2982 = lean_box(0); +} +if (lean_is_scalar(x_2982)) { + x_2983 = lean_alloc_ctor(0, 2, 0); +} else { + x_2983 = x_2982; +} +lean_ctor_set(x_2983, 0, x_2980); +lean_ctor_set(x_2983, 1, x_2981); +return x_2983; +} +else +{ +lean_object* x_2984; lean_object* x_2985; lean_object* x_2986; lean_object* x_2987; uint8_t x_2988; +x_2984 = lean_ctor_get(x_2979, 0); +lean_inc(x_2984); +x_2985 = lean_ctor_get(x_2979, 1); +lean_inc(x_2985); +if (lean_is_exclusive(x_2979)) { + lean_ctor_release(x_2979, 0); + lean_ctor_release(x_2979, 1); + x_2986 = x_2979; +} else { + lean_dec_ref(x_2979); + x_2986 = lean_box(0); +} +x_2987 = lean_ctor_get(x_2984, 1); +lean_inc(x_2987); +x_2988 = lean_nat_dec_eq(x_2973, x_2987); +lean_dec(x_2973); +if (x_2988 == 0) +{ +lean_object* x_2989; +lean_dec(x_2987); +if (lean_is_scalar(x_2986)) { + x_2989 = lean_alloc_ctor(1, 2, 0); +} else { + x_2989 = x_2986; +} +lean_ctor_set(x_2989, 0, x_2984); +lean_ctor_set(x_2989, 1, x_2985); +return x_2989; +} +else +{ +lean_object* x_2990; lean_object* x_2991; lean_object* x_2992; +lean_dec(x_2986); +lean_dec(x_2985); +x_2990 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_2991 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_2991, 0, x_2990); +x_2992 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_2978, x_2991, x_2984); +if (lean_obj_tag(x_2992) == 0) +{ +lean_object* x_2993; lean_object* x_2994; lean_object* x_2995; lean_object* x_2996; +lean_dec(x_2987); +x_2993 = lean_ctor_get(x_2992, 0); +lean_inc(x_2993); +x_2994 = lean_ctor_get(x_2992, 1); +lean_inc(x_2994); +if (lean_is_exclusive(x_2992)) { + lean_ctor_release(x_2992, 0); + lean_ctor_release(x_2992, 1); + x_2995 = x_2992; +} else { + lean_dec_ref(x_2992); + x_2995 = lean_box(0); +} +if (lean_is_scalar(x_2995)) { + x_2996 = lean_alloc_ctor(0, 2, 0); +} else { + x_2996 = x_2995; +} +lean_ctor_set(x_2996, 0, x_2993); +lean_ctor_set(x_2996, 1, x_2994); +return x_2996; +} +else +{ +lean_object* x_2997; lean_object* x_2998; lean_object* x_2999; lean_object* x_3000; uint8_t x_3001; +x_2997 = lean_ctor_get(x_2992, 0); +lean_inc(x_2997); +x_2998 = lean_ctor_get(x_2992, 1); +lean_inc(x_2998); +if (lean_is_exclusive(x_2992)) { + lean_ctor_release(x_2992, 0); + lean_ctor_release(x_2992, 1); + x_2999 = x_2992; +} else { + lean_dec_ref(x_2992); + x_2999 = lean_box(0); +} +x_3000 = lean_ctor_get(x_2997, 1); +lean_inc(x_3000); +x_3001 = lean_nat_dec_eq(x_2987, x_3000); +lean_dec(x_2987); +if (x_3001 == 0) +{ +lean_object* x_3002; +lean_dec(x_3000); +if (lean_is_scalar(x_2999)) { + x_3002 = lean_alloc_ctor(1, 2, 0); +} else { + x_3002 = x_2999; +} +lean_ctor_set(x_3002, 0, x_2997); +lean_ctor_set(x_3002, 1, x_2998); +return x_3002; +} +else +{ +lean_object* x_3003; lean_object* x_3004; lean_object* x_3005; lean_object* x_3006; +lean_dec(x_2999); +lean_dec(x_2998); +x_3003 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_3004 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3004, 0, x_3003); +x_3005 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_3006 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3005, x_3004, x_2997); +if (lean_obj_tag(x_3006) == 0) +{ +lean_object* x_3007; lean_object* x_3008; lean_object* x_3009; lean_object* x_3010; +lean_dec(x_3000); +x_3007 = lean_ctor_get(x_3006, 0); +lean_inc(x_3007); +x_3008 = lean_ctor_get(x_3006, 1); +lean_inc(x_3008); +if (lean_is_exclusive(x_3006)) { + lean_ctor_release(x_3006, 0); + lean_ctor_release(x_3006, 1); + x_3009 = x_3006; +} else { + lean_dec_ref(x_3006); + x_3009 = lean_box(0); +} +if (lean_is_scalar(x_3009)) { + x_3010 = lean_alloc_ctor(0, 2, 0); +} else { + x_3010 = x_3009; +} +lean_ctor_set(x_3010, 0, x_3007); +lean_ctor_set(x_3010, 1, x_3008); +return x_3010; +} +else +{ +lean_object* x_3011; lean_object* x_3012; lean_object* x_3013; lean_object* x_3014; uint8_t x_3015; +x_3011 = lean_ctor_get(x_3006, 0); +lean_inc(x_3011); +x_3012 = lean_ctor_get(x_3006, 1); +lean_inc(x_3012); +if (lean_is_exclusive(x_3006)) { + lean_ctor_release(x_3006, 0); + lean_ctor_release(x_3006, 1); + x_3013 = x_3006; +} else { + lean_dec_ref(x_3006); + x_3013 = lean_box(0); +} +x_3014 = lean_ctor_get(x_3011, 1); +lean_inc(x_3014); +x_3015 = lean_nat_dec_eq(x_3000, x_3014); +lean_dec(x_3000); +if (x_3015 == 0) +{ +lean_object* x_3016; +lean_dec(x_3014); +if (lean_is_scalar(x_3013)) { + x_3016 = lean_alloc_ctor(1, 2, 0); +} else { + x_3016 = x_3013; +} +lean_ctor_set(x_3016, 0, x_3011); +lean_ctor_set(x_3016, 1, x_3012); +return x_3016; +} +else +{ +lean_object* x_3017; lean_object* x_3018; lean_object* x_3019; lean_object* x_3020; +lean_dec(x_3013); +lean_dec(x_3012); +x_3017 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_3018 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3018, 0, x_3017); +x_3019 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_3020 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3019, x_3018, x_3011); +if (lean_obj_tag(x_3020) == 0) +{ +lean_object* x_3021; lean_object* x_3022; lean_object* x_3023; lean_object* x_3024; +lean_dec(x_3014); +x_3021 = lean_ctor_get(x_3020, 0); +lean_inc(x_3021); +x_3022 = lean_ctor_get(x_3020, 1); +lean_inc(x_3022); +if (lean_is_exclusive(x_3020)) { + lean_ctor_release(x_3020, 0); + lean_ctor_release(x_3020, 1); + x_3023 = x_3020; +} else { + lean_dec_ref(x_3020); + x_3023 = lean_box(0); +} +if (lean_is_scalar(x_3023)) { + x_3024 = lean_alloc_ctor(0, 2, 0); +} else { + x_3024 = x_3023; +} +lean_ctor_set(x_3024, 0, x_3021); +lean_ctor_set(x_3024, 1, x_3022); +return x_3024; +} +else +{ +lean_object* x_3025; lean_object* x_3026; lean_object* x_3027; lean_object* x_3028; uint8_t x_3029; +x_3025 = lean_ctor_get(x_3020, 0); +lean_inc(x_3025); +x_3026 = lean_ctor_get(x_3020, 1); +lean_inc(x_3026); +if (lean_is_exclusive(x_3020)) { + lean_ctor_release(x_3020, 0); + lean_ctor_release(x_3020, 1); + x_3027 = x_3020; +} else { + lean_dec_ref(x_3020); + x_3027 = lean_box(0); +} +x_3028 = lean_ctor_get(x_3025, 1); +lean_inc(x_3028); +x_3029 = lean_nat_dec_eq(x_3014, x_3028); +lean_dec(x_3014); +if (x_3029 == 0) +{ +lean_object* x_3030; +lean_dec(x_3028); +if (lean_is_scalar(x_3027)) { + x_3030 = lean_alloc_ctor(1, 2, 0); +} else { + x_3030 = x_3027; +} +lean_ctor_set(x_3030, 0, x_3025); +lean_ctor_set(x_3030, 1, x_3026); +return x_3030; +} +else +{ +lean_object* x_3031; lean_object* x_3032; lean_object* x_3033; lean_object* x_3034; +lean_dec(x_3027); +lean_dec(x_3026); +x_3031 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_3032 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3032, 0, x_3031); +x_3033 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_3034 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3033, x_3032, x_3025); +if (lean_obj_tag(x_3034) == 0) +{ +lean_object* x_3035; lean_object* x_3036; lean_object* x_3037; lean_object* x_3038; +lean_dec(x_3028); +x_3035 = lean_ctor_get(x_3034, 0); +lean_inc(x_3035); +x_3036 = lean_ctor_get(x_3034, 1); +lean_inc(x_3036); +if (lean_is_exclusive(x_3034)) { + lean_ctor_release(x_3034, 0); + lean_ctor_release(x_3034, 1); + x_3037 = x_3034; +} else { + lean_dec_ref(x_3034); + x_3037 = lean_box(0); +} +if (lean_is_scalar(x_3037)) { + x_3038 = lean_alloc_ctor(0, 2, 0); +} else { + x_3038 = x_3037; +} +lean_ctor_set(x_3038, 0, x_3035); +lean_ctor_set(x_3038, 1, x_3036); +return x_3038; +} +else +{ +lean_object* x_3039; lean_object* x_3040; lean_object* x_3041; lean_object* x_3042; uint8_t x_3043; +x_3039 = lean_ctor_get(x_3034, 0); +lean_inc(x_3039); +x_3040 = lean_ctor_get(x_3034, 1); +lean_inc(x_3040); +if (lean_is_exclusive(x_3034)) { + lean_ctor_release(x_3034, 0); + lean_ctor_release(x_3034, 1); + x_3041 = x_3034; +} else { + lean_dec_ref(x_3034); + x_3041 = lean_box(0); +} +x_3042 = lean_ctor_get(x_3039, 1); +lean_inc(x_3042); +x_3043 = lean_nat_dec_eq(x_3028, x_3042); +lean_dec(x_3028); +if (x_3043 == 0) +{ +lean_object* x_3044; +lean_dec(x_3042); +if (lean_is_scalar(x_3041)) { + x_3044 = lean_alloc_ctor(1, 2, 0); +} else { + x_3044 = x_3041; +} +lean_ctor_set(x_3044, 0, x_3039); +lean_ctor_set(x_3044, 1, x_3040); +return x_3044; +} +else +{ +lean_object* x_3045; lean_object* x_3046; lean_object* x_3047; lean_object* x_3048; +lean_dec(x_3041); +lean_dec(x_3040); +x_3045 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_3046 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3046, 0, x_3045); +x_3047 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_3048 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3047, x_3046, x_3039); +if (lean_obj_tag(x_3048) == 0) +{ +lean_object* x_3049; lean_object* x_3050; lean_object* x_3051; lean_object* x_3052; +lean_dec(x_3042); +x_3049 = lean_ctor_get(x_3048, 0); +lean_inc(x_3049); +x_3050 = lean_ctor_get(x_3048, 1); +lean_inc(x_3050); +if (lean_is_exclusive(x_3048)) { + lean_ctor_release(x_3048, 0); + lean_ctor_release(x_3048, 1); + x_3051 = x_3048; +} else { + lean_dec_ref(x_3048); + x_3051 = lean_box(0); +} +if (lean_is_scalar(x_3051)) { + x_3052 = lean_alloc_ctor(0, 2, 0); +} else { + x_3052 = x_3051; +} +lean_ctor_set(x_3052, 0, x_3049); +lean_ctor_set(x_3052, 1, x_3050); +return x_3052; +} +else +{ +lean_object* x_3053; lean_object* x_3054; lean_object* x_3055; lean_object* x_3056; uint8_t x_3057; +x_3053 = lean_ctor_get(x_3048, 0); +lean_inc(x_3053); +x_3054 = lean_ctor_get(x_3048, 1); +lean_inc(x_3054); +if (lean_is_exclusive(x_3048)) { + lean_ctor_release(x_3048, 0); + lean_ctor_release(x_3048, 1); + x_3055 = x_3048; +} else { + lean_dec_ref(x_3048); + x_3055 = lean_box(0); +} +x_3056 = lean_ctor_get(x_3053, 1); +lean_inc(x_3056); +x_3057 = lean_nat_dec_eq(x_3042, x_3056); +lean_dec(x_3042); +if (x_3057 == 0) +{ +lean_object* x_3058; +lean_dec(x_3056); +if (lean_is_scalar(x_3055)) { + x_3058 = lean_alloc_ctor(1, 2, 0); +} else { + x_3058 = x_3055; +} +lean_ctor_set(x_3058, 0, x_3053); +lean_ctor_set(x_3058, 1, x_3054); +return x_3058; +} +else +{ +lean_object* x_3059; lean_object* x_3060; lean_object* x_3061; lean_object* x_3062; +lean_dec(x_3055); +lean_dec(x_3054); +x_3059 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_3060 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3060, 0, x_3059); +x_3061 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_3062 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3061, x_3060, x_3053); +if (lean_obj_tag(x_3062) == 0) +{ +lean_object* x_3063; lean_object* x_3064; lean_object* x_3065; lean_object* x_3066; +lean_dec(x_3056); +x_3063 = lean_ctor_get(x_3062, 0); +lean_inc(x_3063); +x_3064 = lean_ctor_get(x_3062, 1); +lean_inc(x_3064); +if (lean_is_exclusive(x_3062)) { + lean_ctor_release(x_3062, 0); + lean_ctor_release(x_3062, 1); + x_3065 = x_3062; +} else { + lean_dec_ref(x_3062); + x_3065 = lean_box(0); +} +if (lean_is_scalar(x_3065)) { + x_3066 = lean_alloc_ctor(0, 2, 0); +} else { + x_3066 = x_3065; +} +lean_ctor_set(x_3066, 0, x_3063); +lean_ctor_set(x_3066, 1, x_3064); +return x_3066; +} +else +{ +lean_object* x_3067; lean_object* x_3068; lean_object* x_3069; lean_object* x_3070; uint8_t x_3071; +x_3067 = lean_ctor_get(x_3062, 0); +lean_inc(x_3067); +x_3068 = lean_ctor_get(x_3062, 1); +lean_inc(x_3068); +if (lean_is_exclusive(x_3062)) { + lean_ctor_release(x_3062, 0); + lean_ctor_release(x_3062, 1); + x_3069 = x_3062; +} else { + lean_dec_ref(x_3062); + x_3069 = lean_box(0); +} +x_3070 = lean_ctor_get(x_3067, 1); +lean_inc(x_3070); +x_3071 = lean_nat_dec_eq(x_3056, x_3070); +lean_dec(x_3056); +if (x_3071 == 0) +{ +lean_object* x_3072; +lean_dec(x_3070); +if (lean_is_scalar(x_3069)) { + x_3072 = lean_alloc_ctor(1, 2, 0); +} else { + x_3072 = x_3069; +} +lean_ctor_set(x_3072, 0, x_3067); +lean_ctor_set(x_3072, 1, x_3068); +return x_3072; +} +else +{ +lean_object* x_3073; lean_object* x_3074; lean_object* x_3075; lean_object* x_3076; +lean_dec(x_3069); +lean_dec(x_3068); +x_3073 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_3074 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3074, 0, x_3073); +x_3075 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_3076 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3075, x_3074, x_3067); +if (lean_obj_tag(x_3076) == 0) +{ +lean_object* x_3077; lean_object* x_3078; lean_object* x_3079; lean_object* x_3080; +lean_dec(x_3070); +x_3077 = lean_ctor_get(x_3076, 0); +lean_inc(x_3077); +x_3078 = lean_ctor_get(x_3076, 1); +lean_inc(x_3078); +if (lean_is_exclusive(x_3076)) { + lean_ctor_release(x_3076, 0); + lean_ctor_release(x_3076, 1); + x_3079 = x_3076; +} else { + lean_dec_ref(x_3076); + x_3079 = lean_box(0); +} +if (lean_is_scalar(x_3079)) { + x_3080 = lean_alloc_ctor(0, 2, 0); +} else { + x_3080 = x_3079; +} +lean_ctor_set(x_3080, 0, x_3077); +lean_ctor_set(x_3080, 1, x_3078); +return x_3080; +} +else +{ +lean_object* x_3081; lean_object* x_3082; lean_object* x_3083; lean_object* x_3084; uint8_t x_3085; +x_3081 = lean_ctor_get(x_3076, 0); +lean_inc(x_3081); +x_3082 = lean_ctor_get(x_3076, 1); +lean_inc(x_3082); +if (lean_is_exclusive(x_3076)) { + lean_ctor_release(x_3076, 0); + lean_ctor_release(x_3076, 1); + x_3083 = x_3076; +} else { + lean_dec_ref(x_3076); + x_3083 = lean_box(0); +} +x_3084 = lean_ctor_get(x_3081, 1); +lean_inc(x_3084); +x_3085 = lean_nat_dec_eq(x_3070, x_3084); +lean_dec(x_3070); +if (x_3085 == 0) +{ +lean_object* x_3086; +lean_dec(x_3084); +if (lean_is_scalar(x_3083)) { + x_3086 = lean_alloc_ctor(1, 2, 0); +} else { + x_3086 = x_3083; +} +lean_ctor_set(x_3086, 0, x_3081); +lean_ctor_set(x_3086, 1, x_3082); +return x_3086; +} +else +{ +lean_object* x_3087; lean_object* x_3088; lean_object* x_3089; lean_object* x_3090; +lean_dec(x_3083); +lean_dec(x_3082); +x_3087 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_3088 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3088, 0, x_3087); +x_3089 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_3090 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3089, x_3088, x_3081); +if (lean_obj_tag(x_3090) == 0) +{ +lean_object* x_3091; lean_object* x_3092; lean_object* x_3093; lean_object* x_3094; +lean_dec(x_3084); +x_3091 = lean_ctor_get(x_3090, 0); +lean_inc(x_3091); +x_3092 = lean_ctor_get(x_3090, 1); +lean_inc(x_3092); +if (lean_is_exclusive(x_3090)) { + lean_ctor_release(x_3090, 0); + lean_ctor_release(x_3090, 1); + x_3093 = x_3090; +} else { + lean_dec_ref(x_3090); + x_3093 = lean_box(0); +} +if (lean_is_scalar(x_3093)) { + x_3094 = lean_alloc_ctor(0, 2, 0); +} else { + x_3094 = x_3093; +} +lean_ctor_set(x_3094, 0, x_3091); +lean_ctor_set(x_3094, 1, x_3092); +return x_3094; +} +else +{ +lean_object* x_3095; lean_object* x_3096; lean_object* x_3097; lean_object* x_3098; uint8_t x_3099; +x_3095 = lean_ctor_get(x_3090, 0); +lean_inc(x_3095); +x_3096 = lean_ctor_get(x_3090, 1); +lean_inc(x_3096); +if (lean_is_exclusive(x_3090)) { + lean_ctor_release(x_3090, 0); + lean_ctor_release(x_3090, 1); + x_3097 = x_3090; +} else { + lean_dec_ref(x_3090); + x_3097 = lean_box(0); +} +x_3098 = lean_ctor_get(x_3095, 1); +lean_inc(x_3098); +x_3099 = lean_nat_dec_eq(x_3084, x_3098); +lean_dec(x_3084); +if (x_3099 == 0) +{ +lean_object* x_3100; +lean_dec(x_3098); +if (lean_is_scalar(x_3097)) { + x_3100 = lean_alloc_ctor(1, 2, 0); +} else { + x_3100 = x_3097; +} +lean_ctor_set(x_3100, 0, x_3095); +lean_ctor_set(x_3100, 1, x_3096); +return x_3100; +} +else +{ +lean_object* x_3101; lean_object* x_3102; lean_object* x_3103; lean_object* x_3104; +lean_dec(x_3097); +lean_dec(x_3096); +x_3101 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_3102 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3102, 0, x_3101); +x_3103 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_3104 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3103, x_3102, x_3095); +if (lean_obj_tag(x_3104) == 0) +{ +lean_object* x_3105; lean_object* x_3106; lean_object* x_3107; lean_object* x_3108; +lean_dec(x_3098); +x_3105 = lean_ctor_get(x_3104, 0); +lean_inc(x_3105); +x_3106 = lean_ctor_get(x_3104, 1); +lean_inc(x_3106); +if (lean_is_exclusive(x_3104)) { + lean_ctor_release(x_3104, 0); + lean_ctor_release(x_3104, 1); + x_3107 = x_3104; +} else { + lean_dec_ref(x_3104); + x_3107 = lean_box(0); +} +if (lean_is_scalar(x_3107)) { + x_3108 = lean_alloc_ctor(0, 2, 0); +} else { + x_3108 = x_3107; +} +lean_ctor_set(x_3108, 0, x_3105); +lean_ctor_set(x_3108, 1, x_3106); +return x_3108; +} +else +{ +lean_object* x_3109; lean_object* x_3110; lean_object* x_3111; lean_object* x_3112; uint8_t x_3113; +x_3109 = lean_ctor_get(x_3104, 0); +lean_inc(x_3109); +x_3110 = lean_ctor_get(x_3104, 1); +lean_inc(x_3110); +if (lean_is_exclusive(x_3104)) { + lean_ctor_release(x_3104, 0); + lean_ctor_release(x_3104, 1); + x_3111 = x_3104; +} else { + lean_dec_ref(x_3104); + x_3111 = lean_box(0); +} +x_3112 = lean_ctor_get(x_3109, 1); +lean_inc(x_3112); +x_3113 = lean_nat_dec_eq(x_3098, x_3112); +lean_dec(x_3098); +if (x_3113 == 0) +{ +lean_object* x_3114; +lean_dec(x_3112); +if (lean_is_scalar(x_3111)) { + x_3114 = lean_alloc_ctor(1, 2, 0); +} else { + x_3114 = x_3111; +} +lean_ctor_set(x_3114, 0, x_3109); +lean_ctor_set(x_3114, 1, x_3110); +return x_3114; +} +else +{ +lean_object* x_3115; lean_object* x_3116; lean_object* x_3117; lean_object* x_3118; +lean_dec(x_3111); +lean_dec(x_3110); +x_3115 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_3116 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3116, 0, x_3115); +x_3117 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_3118 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3117, x_3116, x_3109); +if (lean_obj_tag(x_3118) == 0) +{ +lean_object* x_3119; lean_object* x_3120; lean_object* x_3121; lean_object* x_3122; +lean_dec(x_3112); +x_3119 = lean_ctor_get(x_3118, 0); +lean_inc(x_3119); +x_3120 = lean_ctor_get(x_3118, 1); +lean_inc(x_3120); +if (lean_is_exclusive(x_3118)) { + lean_ctor_release(x_3118, 0); + lean_ctor_release(x_3118, 1); + x_3121 = x_3118; +} else { + lean_dec_ref(x_3118); + x_3121 = lean_box(0); +} +if (lean_is_scalar(x_3121)) { + x_3122 = lean_alloc_ctor(0, 2, 0); +} else { + x_3122 = x_3121; +} +lean_ctor_set(x_3122, 0, x_3119); +lean_ctor_set(x_3122, 1, x_3120); +return x_3122; +} +else +{ +lean_object* x_3123; lean_object* x_3124; lean_object* x_3125; lean_object* x_3126; uint8_t x_3127; +x_3123 = lean_ctor_get(x_3118, 0); +lean_inc(x_3123); +x_3124 = lean_ctor_get(x_3118, 1); +lean_inc(x_3124); +if (lean_is_exclusive(x_3118)) { + lean_ctor_release(x_3118, 0); + lean_ctor_release(x_3118, 1); + x_3125 = x_3118; +} else { + lean_dec_ref(x_3118); + x_3125 = lean_box(0); +} +x_3126 = lean_ctor_get(x_3123, 1); +lean_inc(x_3126); +x_3127 = lean_nat_dec_eq(x_3112, x_3126); +lean_dec(x_3112); +if (x_3127 == 0) +{ +lean_object* x_3128; +lean_dec(x_3126); +if (lean_is_scalar(x_3125)) { + x_3128 = lean_alloc_ctor(1, 2, 0); +} else { + x_3128 = x_3125; +} +lean_ctor_set(x_3128, 0, x_3123); +lean_ctor_set(x_3128, 1, x_3124); +return x_3128; +} +else +{ +lean_object* x_3129; lean_object* x_3130; lean_object* x_3131; lean_object* x_3132; +lean_dec(x_3125); +lean_dec(x_3124); +x_3129 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_3130 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3130, 0, x_3129); +x_3131 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_3132 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3131, x_3130, x_3123); +if (lean_obj_tag(x_3132) == 0) +{ +lean_object* x_3133; lean_object* x_3134; lean_object* x_3135; lean_object* x_3136; +lean_dec(x_3126); +x_3133 = lean_ctor_get(x_3132, 0); +lean_inc(x_3133); +x_3134 = lean_ctor_get(x_3132, 1); +lean_inc(x_3134); +if (lean_is_exclusive(x_3132)) { + lean_ctor_release(x_3132, 0); + lean_ctor_release(x_3132, 1); + x_3135 = x_3132; +} else { + lean_dec_ref(x_3132); + x_3135 = lean_box(0); +} +if (lean_is_scalar(x_3135)) { + x_3136 = lean_alloc_ctor(0, 2, 0); +} else { + x_3136 = x_3135; +} +lean_ctor_set(x_3136, 0, x_3133); +lean_ctor_set(x_3136, 1, x_3134); +return x_3136; +} +else +{ +lean_object* x_3137; lean_object* x_3138; lean_object* x_3139; lean_object* x_3140; uint8_t x_3141; +x_3137 = lean_ctor_get(x_3132, 0); +lean_inc(x_3137); +x_3138 = lean_ctor_get(x_3132, 1); +lean_inc(x_3138); +if (lean_is_exclusive(x_3132)) { + lean_ctor_release(x_3132, 0); + lean_ctor_release(x_3132, 1); + x_3139 = x_3132; +} else { + lean_dec_ref(x_3132); + x_3139 = lean_box(0); +} +x_3140 = lean_ctor_get(x_3137, 1); +lean_inc(x_3140); +x_3141 = lean_nat_dec_eq(x_3126, x_3140); +lean_dec(x_3126); +if (x_3141 == 0) +{ +lean_object* x_3142; +lean_dec(x_3140); +if (lean_is_scalar(x_3139)) { + x_3142 = lean_alloc_ctor(1, 2, 0); +} else { + x_3142 = x_3139; +} +lean_ctor_set(x_3142, 0, x_3137); +lean_ctor_set(x_3142, 1, x_3138); +return x_3142; +} +else +{ +lean_object* x_3143; lean_object* x_3144; lean_object* x_3145; lean_object* x_3146; +lean_dec(x_3139); +lean_dec(x_3138); +x_3143 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_3144 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3144, 0, x_3143); +x_3145 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_3146 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3145, x_3144, x_3137); +if (lean_obj_tag(x_3146) == 0) +{ +lean_object* x_3147; lean_object* x_3148; lean_object* x_3149; lean_object* x_3150; +lean_dec(x_3140); +x_3147 = lean_ctor_get(x_3146, 0); +lean_inc(x_3147); +x_3148 = lean_ctor_get(x_3146, 1); +lean_inc(x_3148); +if (lean_is_exclusive(x_3146)) { + lean_ctor_release(x_3146, 0); + lean_ctor_release(x_3146, 1); + x_3149 = x_3146; +} else { + lean_dec_ref(x_3146); + x_3149 = lean_box(0); +} +if (lean_is_scalar(x_3149)) { + x_3150 = lean_alloc_ctor(0, 2, 0); +} else { + x_3150 = x_3149; +} +lean_ctor_set(x_3150, 0, x_3147); +lean_ctor_set(x_3150, 1, x_3148); +return x_3150; +} +else +{ +lean_object* x_3151; lean_object* x_3152; lean_object* x_3153; lean_object* x_3154; uint8_t x_3155; +x_3151 = lean_ctor_get(x_3146, 0); +lean_inc(x_3151); +x_3152 = lean_ctor_get(x_3146, 1); +lean_inc(x_3152); +if (lean_is_exclusive(x_3146)) { + lean_ctor_release(x_3146, 0); + lean_ctor_release(x_3146, 1); + x_3153 = x_3146; +} else { + lean_dec_ref(x_3146); + x_3153 = lean_box(0); +} +x_3154 = lean_ctor_get(x_3151, 1); +lean_inc(x_3154); +x_3155 = lean_nat_dec_eq(x_3140, x_3154); +lean_dec(x_3140); +if (x_3155 == 0) +{ +lean_object* x_3156; +lean_dec(x_3154); +if (lean_is_scalar(x_3153)) { + x_3156 = lean_alloc_ctor(1, 2, 0); +} else { + x_3156 = x_3153; +} +lean_ctor_set(x_3156, 0, x_3151); +lean_ctor_set(x_3156, 1, x_3152); +return x_3156; +} +else +{ +lean_object* x_3157; lean_object* x_3158; lean_object* x_3159; lean_object* x_3160; +lean_dec(x_3153); +lean_dec(x_3152); +x_3157 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_3158 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3158, 0, x_3157); +x_3159 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_3160 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3159, x_3158, x_3151); +if (lean_obj_tag(x_3160) == 0) +{ +lean_object* x_3161; lean_object* x_3162; lean_object* x_3163; lean_object* x_3164; +lean_dec(x_3154); +x_3161 = lean_ctor_get(x_3160, 0); +lean_inc(x_3161); +x_3162 = lean_ctor_get(x_3160, 1); +lean_inc(x_3162); +if (lean_is_exclusive(x_3160)) { + lean_ctor_release(x_3160, 0); + lean_ctor_release(x_3160, 1); + x_3163 = x_3160; +} else { + lean_dec_ref(x_3160); + x_3163 = lean_box(0); +} +if (lean_is_scalar(x_3163)) { + x_3164 = lean_alloc_ctor(0, 2, 0); +} else { + x_3164 = x_3163; +} +lean_ctor_set(x_3164, 0, x_3161); +lean_ctor_set(x_3164, 1, x_3162); +return x_3164; +} +else +{ +lean_object* x_3165; lean_object* x_3166; lean_object* x_3167; lean_object* x_3168; uint8_t x_3169; +x_3165 = lean_ctor_get(x_3160, 0); +lean_inc(x_3165); +x_3166 = lean_ctor_get(x_3160, 1); +lean_inc(x_3166); +if (lean_is_exclusive(x_3160)) { + lean_ctor_release(x_3160, 0); + lean_ctor_release(x_3160, 1); + x_3167 = x_3160; +} else { + lean_dec_ref(x_3160); + x_3167 = lean_box(0); +} +x_3168 = lean_ctor_get(x_3165, 1); +lean_inc(x_3168); +x_3169 = lean_nat_dec_eq(x_3154, x_3168); +lean_dec(x_3154); +if (x_3169 == 0) +{ +lean_object* x_3170; +lean_dec(x_3168); +if (lean_is_scalar(x_3167)) { + x_3170 = lean_alloc_ctor(1, 2, 0); +} else { + x_3170 = x_3167; +} +lean_ctor_set(x_3170, 0, x_3165); +lean_ctor_set(x_3170, 1, x_3166); +return x_3170; +} +else +{ +lean_object* x_3171; lean_object* x_3172; lean_object* x_3173; lean_object* x_3174; +lean_dec(x_3167); +lean_dec(x_3166); +x_3171 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_3172 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3172, 0, x_3171); +x_3173 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_3174 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3173, x_3172, x_3165); +if (lean_obj_tag(x_3174) == 0) +{ +lean_object* x_3175; lean_object* x_3176; lean_object* x_3177; lean_object* x_3178; +lean_dec(x_3168); +x_3175 = lean_ctor_get(x_3174, 0); +lean_inc(x_3175); +x_3176 = lean_ctor_get(x_3174, 1); +lean_inc(x_3176); +if (lean_is_exclusive(x_3174)) { + lean_ctor_release(x_3174, 0); + lean_ctor_release(x_3174, 1); + x_3177 = x_3174; +} else { + lean_dec_ref(x_3174); + x_3177 = lean_box(0); +} +if (lean_is_scalar(x_3177)) { + x_3178 = lean_alloc_ctor(0, 2, 0); +} else { + x_3178 = x_3177; +} +lean_ctor_set(x_3178, 0, x_3175); +lean_ctor_set(x_3178, 1, x_3176); +return x_3178; +} +else +{ +lean_object* x_3179; lean_object* x_3180; lean_object* x_3181; lean_object* x_3182; uint8_t x_3183; +x_3179 = lean_ctor_get(x_3174, 0); +lean_inc(x_3179); +x_3180 = lean_ctor_get(x_3174, 1); +lean_inc(x_3180); +if (lean_is_exclusive(x_3174)) { + lean_ctor_release(x_3174, 0); + lean_ctor_release(x_3174, 1); + x_3181 = x_3174; +} else { + lean_dec_ref(x_3174); + x_3181 = lean_box(0); +} +x_3182 = lean_ctor_get(x_3179, 1); +lean_inc(x_3182); +x_3183 = lean_nat_dec_eq(x_3168, x_3182); +lean_dec(x_3168); +if (x_3183 == 0) +{ +lean_object* x_3184; +lean_dec(x_3182); +if (lean_is_scalar(x_3181)) { + x_3184 = lean_alloc_ctor(1, 2, 0); +} else { + x_3184 = x_3181; +} +lean_ctor_set(x_3184, 0, x_3179); +lean_ctor_set(x_3184, 1, x_3180); +return x_3184; +} +else +{ +lean_object* x_3185; lean_object* x_3186; lean_object* x_3187; lean_object* x_3188; +lean_dec(x_3181); +lean_dec(x_3180); +x_3185 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_3186 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3186, 0, x_3185); +x_3187 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_3188 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3187, x_3186, x_3179); +if (lean_obj_tag(x_3188) == 0) +{ +lean_object* x_3189; lean_object* x_3190; lean_object* x_3191; lean_object* x_3192; +lean_dec(x_3182); +x_3189 = lean_ctor_get(x_3188, 0); +lean_inc(x_3189); +x_3190 = lean_ctor_get(x_3188, 1); +lean_inc(x_3190); +if (lean_is_exclusive(x_3188)) { + lean_ctor_release(x_3188, 0); + lean_ctor_release(x_3188, 1); + x_3191 = x_3188; +} else { + lean_dec_ref(x_3188); + x_3191 = lean_box(0); +} +if (lean_is_scalar(x_3191)) { + x_3192 = lean_alloc_ctor(0, 2, 0); +} else { + x_3192 = x_3191; +} +lean_ctor_set(x_3192, 0, x_3189); +lean_ctor_set(x_3192, 1, x_3190); +return x_3192; +} +else +{ +lean_object* x_3193; lean_object* x_3194; lean_object* x_3195; lean_object* x_3196; uint8_t x_3197; +x_3193 = lean_ctor_get(x_3188, 0); +lean_inc(x_3193); +x_3194 = lean_ctor_get(x_3188, 1); +lean_inc(x_3194); +if (lean_is_exclusive(x_3188)) { + lean_ctor_release(x_3188, 0); + lean_ctor_release(x_3188, 1); + x_3195 = x_3188; +} else { + lean_dec_ref(x_3188); + x_3195 = lean_box(0); +} +x_3196 = lean_ctor_get(x_3193, 1); +lean_inc(x_3196); +x_3197 = lean_nat_dec_eq(x_3182, x_3196); +lean_dec(x_3182); +if (x_3197 == 0) +{ +lean_object* x_3198; +lean_dec(x_3196); +if (lean_is_scalar(x_3195)) { + x_3198 = lean_alloc_ctor(1, 2, 0); +} else { + x_3198 = x_3195; +} +lean_ctor_set(x_3198, 0, x_3193); +lean_ctor_set(x_3198, 1, x_3194); +return x_3198; +} +else +{ +lean_object* x_3199; lean_object* x_3200; lean_object* x_3201; lean_object* x_3202; +lean_dec(x_3195); +lean_dec(x_3194); +x_3199 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_3200 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3200, 0, x_3199); +x_3201 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_3202 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3201, x_3200, x_3193); +if (lean_obj_tag(x_3202) == 0) +{ +lean_object* x_3203; lean_object* x_3204; lean_object* x_3205; lean_object* x_3206; +lean_dec(x_3196); +x_3203 = lean_ctor_get(x_3202, 0); +lean_inc(x_3203); +x_3204 = lean_ctor_get(x_3202, 1); +lean_inc(x_3204); +if (lean_is_exclusive(x_3202)) { + lean_ctor_release(x_3202, 0); + lean_ctor_release(x_3202, 1); + x_3205 = x_3202; +} else { + lean_dec_ref(x_3202); + x_3205 = lean_box(0); +} +if (lean_is_scalar(x_3205)) { + x_3206 = lean_alloc_ctor(0, 2, 0); +} else { + x_3206 = x_3205; +} +lean_ctor_set(x_3206, 0, x_3203); +lean_ctor_set(x_3206, 1, x_3204); +return x_3206; +} +else +{ +lean_object* x_3207; lean_object* x_3208; lean_object* x_3209; lean_object* x_3210; uint8_t x_3211; +x_3207 = lean_ctor_get(x_3202, 0); +lean_inc(x_3207); +x_3208 = lean_ctor_get(x_3202, 1); +lean_inc(x_3208); +if (lean_is_exclusive(x_3202)) { + lean_ctor_release(x_3202, 0); + lean_ctor_release(x_3202, 1); + x_3209 = x_3202; +} else { + lean_dec_ref(x_3202); + x_3209 = lean_box(0); +} +x_3210 = lean_ctor_get(x_3207, 1); +lean_inc(x_3210); +x_3211 = lean_nat_dec_eq(x_3196, x_3210); +lean_dec(x_3196); +if (x_3211 == 0) +{ +lean_object* x_3212; +lean_dec(x_3210); +if (lean_is_scalar(x_3209)) { + x_3212 = lean_alloc_ctor(1, 2, 0); +} else { + x_3212 = x_3209; +} +lean_ctor_set(x_3212, 0, x_3207); +lean_ctor_set(x_3212, 1, x_3208); +return x_3212; +} +else +{ +lean_object* x_3213; lean_object* x_3214; lean_object* x_3215; lean_object* x_3216; +lean_dec(x_3209); +lean_dec(x_3208); +x_3213 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_3214 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3214, 0, x_3213); +x_3215 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_3216 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3215, x_3214, x_3207); +if (lean_obj_tag(x_3216) == 0) +{ +lean_object* x_3217; lean_object* x_3218; lean_object* x_3219; lean_object* x_3220; +lean_dec(x_3210); +x_3217 = lean_ctor_get(x_3216, 0); +lean_inc(x_3217); +x_3218 = lean_ctor_get(x_3216, 1); +lean_inc(x_3218); +if (lean_is_exclusive(x_3216)) { + lean_ctor_release(x_3216, 0); + lean_ctor_release(x_3216, 1); + x_3219 = x_3216; +} else { + lean_dec_ref(x_3216); + x_3219 = lean_box(0); +} +if (lean_is_scalar(x_3219)) { + x_3220 = lean_alloc_ctor(0, 2, 0); +} else { + x_3220 = x_3219; +} +lean_ctor_set(x_3220, 0, x_3217); +lean_ctor_set(x_3220, 1, x_3218); +return x_3220; +} +else +{ +lean_object* x_3221; lean_object* x_3222; lean_object* x_3223; lean_object* x_3224; uint8_t x_3225; +x_3221 = lean_ctor_get(x_3216, 0); +lean_inc(x_3221); +x_3222 = lean_ctor_get(x_3216, 1); +lean_inc(x_3222); +if (lean_is_exclusive(x_3216)) { + lean_ctor_release(x_3216, 0); + lean_ctor_release(x_3216, 1); + x_3223 = x_3216; +} else { + lean_dec_ref(x_3216); + x_3223 = lean_box(0); +} +x_3224 = lean_ctor_get(x_3221, 1); +lean_inc(x_3224); +x_3225 = lean_nat_dec_eq(x_3210, x_3224); +lean_dec(x_3210); +if (x_3225 == 0) +{ +lean_object* x_3226; +lean_dec(x_3224); +if (lean_is_scalar(x_3223)) { + x_3226 = lean_alloc_ctor(1, 2, 0); +} else { + x_3226 = x_3223; +} +lean_ctor_set(x_3226, 0, x_3221); +lean_ctor_set(x_3226, 1, x_3222); +return x_3226; +} +else +{ +lean_object* x_3227; lean_object* x_3228; lean_object* x_3229; lean_object* x_3230; +lean_dec(x_3223); +lean_dec(x_3222); +x_3227 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_3228 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3228, 0, x_3227); +x_3229 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_3230 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3229, x_3228, x_3221); +if (lean_obj_tag(x_3230) == 0) +{ +lean_object* x_3231; lean_object* x_3232; lean_object* x_3233; lean_object* x_3234; +lean_dec(x_3224); +x_3231 = lean_ctor_get(x_3230, 0); +lean_inc(x_3231); +x_3232 = lean_ctor_get(x_3230, 1); +lean_inc(x_3232); +if (lean_is_exclusive(x_3230)) { + lean_ctor_release(x_3230, 0); + lean_ctor_release(x_3230, 1); + x_3233 = x_3230; +} else { + lean_dec_ref(x_3230); + x_3233 = lean_box(0); +} +if (lean_is_scalar(x_3233)) { + x_3234 = lean_alloc_ctor(0, 2, 0); +} else { + x_3234 = x_3233; +} +lean_ctor_set(x_3234, 0, x_3231); +lean_ctor_set(x_3234, 1, x_3232); +return x_3234; +} +else +{ +lean_object* x_3235; lean_object* x_3236; lean_object* x_3237; lean_object* x_3238; uint8_t x_3239; +x_3235 = lean_ctor_get(x_3230, 0); +lean_inc(x_3235); +x_3236 = lean_ctor_get(x_3230, 1); +lean_inc(x_3236); +if (lean_is_exclusive(x_3230)) { + lean_ctor_release(x_3230, 0); + lean_ctor_release(x_3230, 1); + x_3237 = x_3230; +} else { + lean_dec_ref(x_3230); + x_3237 = lean_box(0); +} +x_3238 = lean_ctor_get(x_3235, 1); +lean_inc(x_3238); +x_3239 = lean_nat_dec_eq(x_3224, x_3238); +lean_dec(x_3238); +lean_dec(x_3224); +if (x_3239 == 0) +{ +lean_object* x_3240; +if (lean_is_scalar(x_3237)) { + x_3240 = lean_alloc_ctor(1, 2, 0); +} else { + x_3240 = x_3237; +} +lean_ctor_set(x_3240, 0, x_3235); +lean_ctor_set(x_3240, 1, x_3236); +return x_3240; +} +else +{ +lean_object* x_3241; lean_object* x_3242; lean_object* x_3243; lean_object* x_3244; +lean_dec(x_3237); +lean_dec(x_3236); +x_3241 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_3242 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3242, 0, x_3241); +x_3243 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_3244 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3243, x_3242, x_3235); +return x_3244; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_3245; lean_object* x_3246; lean_object* x_3247; uint8_t x_3248; +x_3245 = lean_ctor_get(x_134, 0); +x_3246 = lean_ctor_get(x_134, 1); +lean_inc(x_3246); +lean_inc(x_3245); +lean_dec(x_134); +x_3247 = lean_ctor_get(x_3245, 1); +lean_inc(x_3247); +x_3248 = lean_nat_dec_eq(x_129, x_3247); +lean_dec(x_129); +if (x_3248 == 0) +{ +lean_object* x_3249; +lean_dec(x_3247); +x_3249 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3249, 0, x_3245); +lean_ctor_set(x_3249, 1, x_3246); +return x_3249; +} +else +{ +lean_object* x_3250; lean_object* x_3251; lean_object* x_3252; lean_object* x_3253; +lean_dec(x_3246); +x_3250 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_3251 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3251, 0, x_3250); +x_3252 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_3253 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3252, x_3251, x_3245); +if (lean_obj_tag(x_3253) == 0) +{ +lean_object* x_3254; lean_object* x_3255; lean_object* x_3256; lean_object* x_3257; +lean_dec(x_3247); +x_3254 = lean_ctor_get(x_3253, 0); +lean_inc(x_3254); +x_3255 = lean_ctor_get(x_3253, 1); +lean_inc(x_3255); +if (lean_is_exclusive(x_3253)) { + lean_ctor_release(x_3253, 0); + lean_ctor_release(x_3253, 1); + x_3256 = x_3253; +} else { + lean_dec_ref(x_3253); + x_3256 = lean_box(0); +} +if (lean_is_scalar(x_3256)) { + x_3257 = lean_alloc_ctor(0, 2, 0); +} else { + x_3257 = x_3256; +} +lean_ctor_set(x_3257, 0, x_3254); +lean_ctor_set(x_3257, 1, x_3255); +return x_3257; +} +else +{ +lean_object* x_3258; lean_object* x_3259; lean_object* x_3260; lean_object* x_3261; uint8_t x_3262; +x_3258 = lean_ctor_get(x_3253, 0); +lean_inc(x_3258); +x_3259 = lean_ctor_get(x_3253, 1); +lean_inc(x_3259); +if (lean_is_exclusive(x_3253)) { + lean_ctor_release(x_3253, 0); + lean_ctor_release(x_3253, 1); + x_3260 = x_3253; +} else { + lean_dec_ref(x_3253); + x_3260 = lean_box(0); +} +x_3261 = lean_ctor_get(x_3258, 1); +lean_inc(x_3261); +x_3262 = lean_nat_dec_eq(x_3247, x_3261); +lean_dec(x_3247); +if (x_3262 == 0) +{ +lean_object* x_3263; +lean_dec(x_3261); +if (lean_is_scalar(x_3260)) { + x_3263 = lean_alloc_ctor(1, 2, 0); +} else { + x_3263 = x_3260; +} +lean_ctor_set(x_3263, 0, x_3258); +lean_ctor_set(x_3263, 1, x_3259); +return x_3263; +} +else +{ +lean_object* x_3264; lean_object* x_3265; lean_object* x_3266; lean_object* x_3267; +lean_dec(x_3260); +lean_dec(x_3259); +x_3264 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_3265 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3265, 0, x_3264); +x_3266 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_3267 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3266, x_3265, x_3258); +if (lean_obj_tag(x_3267) == 0) +{ +lean_object* x_3268; lean_object* x_3269; lean_object* x_3270; lean_object* x_3271; +lean_dec(x_3261); +x_3268 = lean_ctor_get(x_3267, 0); +lean_inc(x_3268); +x_3269 = lean_ctor_get(x_3267, 1); +lean_inc(x_3269); +if (lean_is_exclusive(x_3267)) { + lean_ctor_release(x_3267, 0); + lean_ctor_release(x_3267, 1); + x_3270 = x_3267; +} else { + lean_dec_ref(x_3267); + x_3270 = lean_box(0); +} +if (lean_is_scalar(x_3270)) { + x_3271 = lean_alloc_ctor(0, 2, 0); +} else { + x_3271 = x_3270; +} +lean_ctor_set(x_3271, 0, x_3268); +lean_ctor_set(x_3271, 1, x_3269); +return x_3271; +} +else +{ +lean_object* x_3272; lean_object* x_3273; lean_object* x_3274; lean_object* x_3275; uint8_t x_3276; +x_3272 = lean_ctor_get(x_3267, 0); +lean_inc(x_3272); +x_3273 = lean_ctor_get(x_3267, 1); +lean_inc(x_3273); +if (lean_is_exclusive(x_3267)) { + lean_ctor_release(x_3267, 0); + lean_ctor_release(x_3267, 1); + x_3274 = x_3267; +} else { + lean_dec_ref(x_3267); + x_3274 = lean_box(0); +} +x_3275 = lean_ctor_get(x_3272, 1); +lean_inc(x_3275); +x_3276 = lean_nat_dec_eq(x_3261, x_3275); +lean_dec(x_3261); +if (x_3276 == 0) +{ +lean_object* x_3277; +lean_dec(x_3275); +if (lean_is_scalar(x_3274)) { + x_3277 = lean_alloc_ctor(1, 2, 0); +} else { + x_3277 = x_3274; +} +lean_ctor_set(x_3277, 0, x_3272); +lean_ctor_set(x_3277, 1, x_3273); +return x_3277; +} +else +{ +lean_object* x_3278; lean_object* x_3279; lean_object* x_3280; +lean_dec(x_3274); +lean_dec(x_3273); +x_3278 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_3279 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3279, 0, x_3278); +x_3280 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3266, x_3279, x_3272); +if (lean_obj_tag(x_3280) == 0) +{ +lean_object* x_3281; lean_object* x_3282; lean_object* x_3283; lean_object* x_3284; +lean_dec(x_3275); +x_3281 = lean_ctor_get(x_3280, 0); +lean_inc(x_3281); +x_3282 = lean_ctor_get(x_3280, 1); +lean_inc(x_3282); +if (lean_is_exclusive(x_3280)) { + lean_ctor_release(x_3280, 0); + lean_ctor_release(x_3280, 1); + x_3283 = x_3280; +} else { + lean_dec_ref(x_3280); + x_3283 = lean_box(0); +} +if (lean_is_scalar(x_3283)) { + x_3284 = lean_alloc_ctor(0, 2, 0); +} else { + x_3284 = x_3283; +} +lean_ctor_set(x_3284, 0, x_3281); +lean_ctor_set(x_3284, 1, x_3282); +return x_3284; +} +else +{ +lean_object* x_3285; lean_object* x_3286; lean_object* x_3287; lean_object* x_3288; uint8_t x_3289; +x_3285 = lean_ctor_get(x_3280, 0); +lean_inc(x_3285); +x_3286 = lean_ctor_get(x_3280, 1); +lean_inc(x_3286); +if (lean_is_exclusive(x_3280)) { + lean_ctor_release(x_3280, 0); + lean_ctor_release(x_3280, 1); + x_3287 = x_3280; +} else { + lean_dec_ref(x_3280); + x_3287 = lean_box(0); +} +x_3288 = lean_ctor_get(x_3285, 1); +lean_inc(x_3288); +x_3289 = lean_nat_dec_eq(x_3275, x_3288); +lean_dec(x_3275); +if (x_3289 == 0) +{ +lean_object* x_3290; +lean_dec(x_3288); +if (lean_is_scalar(x_3287)) { + x_3290 = lean_alloc_ctor(1, 2, 0); +} else { + x_3290 = x_3287; +} +lean_ctor_set(x_3290, 0, x_3285); +lean_ctor_set(x_3290, 1, x_3286); +return x_3290; +} +else +{ +lean_object* x_3291; lean_object* x_3292; lean_object* x_3293; lean_object* x_3294; +lean_dec(x_3287); +lean_dec(x_3286); +x_3291 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_3292 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3292, 0, x_3291); +x_3293 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_3294 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3293, x_3292, x_3285); +if (lean_obj_tag(x_3294) == 0) +{ +lean_object* x_3295; lean_object* x_3296; lean_object* x_3297; lean_object* x_3298; +lean_dec(x_3288); +x_3295 = lean_ctor_get(x_3294, 0); +lean_inc(x_3295); +x_3296 = lean_ctor_get(x_3294, 1); +lean_inc(x_3296); +if (lean_is_exclusive(x_3294)) { + lean_ctor_release(x_3294, 0); + lean_ctor_release(x_3294, 1); + x_3297 = x_3294; +} else { + lean_dec_ref(x_3294); + x_3297 = lean_box(0); +} +if (lean_is_scalar(x_3297)) { + x_3298 = lean_alloc_ctor(0, 2, 0); +} else { + x_3298 = x_3297; +} +lean_ctor_set(x_3298, 0, x_3295); +lean_ctor_set(x_3298, 1, x_3296); +return x_3298; +} +else +{ +lean_object* x_3299; lean_object* x_3300; lean_object* x_3301; lean_object* x_3302; uint8_t x_3303; +x_3299 = lean_ctor_get(x_3294, 0); +lean_inc(x_3299); +x_3300 = lean_ctor_get(x_3294, 1); +lean_inc(x_3300); +if (lean_is_exclusive(x_3294)) { + lean_ctor_release(x_3294, 0); + lean_ctor_release(x_3294, 1); + x_3301 = x_3294; +} else { + lean_dec_ref(x_3294); + x_3301 = lean_box(0); +} +x_3302 = lean_ctor_get(x_3299, 1); +lean_inc(x_3302); +x_3303 = lean_nat_dec_eq(x_3288, x_3302); +lean_dec(x_3288); +if (x_3303 == 0) +{ +lean_object* x_3304; +lean_dec(x_3302); +if (lean_is_scalar(x_3301)) { + x_3304 = lean_alloc_ctor(1, 2, 0); +} else { + x_3304 = x_3301; +} +lean_ctor_set(x_3304, 0, x_3299); +lean_ctor_set(x_3304, 1, x_3300); +return x_3304; +} +else +{ +lean_object* x_3305; lean_object* x_3306; lean_object* x_3307; lean_object* x_3308; +lean_dec(x_3301); +lean_dec(x_3300); +x_3305 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_3306 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3306, 0, x_3305); +x_3307 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_3308 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3307, x_3306, x_3299); +if (lean_obj_tag(x_3308) == 0) +{ +lean_object* x_3309; lean_object* x_3310; lean_object* x_3311; lean_object* x_3312; +lean_dec(x_3302); +x_3309 = lean_ctor_get(x_3308, 0); +lean_inc(x_3309); +x_3310 = lean_ctor_get(x_3308, 1); +lean_inc(x_3310); +if (lean_is_exclusive(x_3308)) { + lean_ctor_release(x_3308, 0); + lean_ctor_release(x_3308, 1); + x_3311 = x_3308; +} else { + lean_dec_ref(x_3308); + x_3311 = lean_box(0); +} +if (lean_is_scalar(x_3311)) { + x_3312 = lean_alloc_ctor(0, 2, 0); +} else { + x_3312 = x_3311; +} +lean_ctor_set(x_3312, 0, x_3309); +lean_ctor_set(x_3312, 1, x_3310); +return x_3312; +} +else +{ +lean_object* x_3313; lean_object* x_3314; lean_object* x_3315; lean_object* x_3316; uint8_t x_3317; +x_3313 = lean_ctor_get(x_3308, 0); +lean_inc(x_3313); +x_3314 = lean_ctor_get(x_3308, 1); +lean_inc(x_3314); +if (lean_is_exclusive(x_3308)) { + lean_ctor_release(x_3308, 0); + lean_ctor_release(x_3308, 1); + x_3315 = x_3308; +} else { + lean_dec_ref(x_3308); + x_3315 = lean_box(0); +} +x_3316 = lean_ctor_get(x_3313, 1); +lean_inc(x_3316); +x_3317 = lean_nat_dec_eq(x_3302, x_3316); +lean_dec(x_3302); +if (x_3317 == 0) +{ +lean_object* x_3318; +lean_dec(x_3316); +if (lean_is_scalar(x_3315)) { + x_3318 = lean_alloc_ctor(1, 2, 0); +} else { + x_3318 = x_3315; +} +lean_ctor_set(x_3318, 0, x_3313); +lean_ctor_set(x_3318, 1, x_3314); +return x_3318; +} +else +{ +lean_object* x_3319; lean_object* x_3320; lean_object* x_3321; lean_object* x_3322; +lean_dec(x_3315); +lean_dec(x_3314); +x_3319 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_3320 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3320, 0, x_3319); +x_3321 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_3322 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3321, x_3320, x_3313); +if (lean_obj_tag(x_3322) == 0) +{ +lean_object* x_3323; lean_object* x_3324; lean_object* x_3325; lean_object* x_3326; +lean_dec(x_3316); +x_3323 = lean_ctor_get(x_3322, 0); +lean_inc(x_3323); +x_3324 = lean_ctor_get(x_3322, 1); +lean_inc(x_3324); +if (lean_is_exclusive(x_3322)) { + lean_ctor_release(x_3322, 0); + lean_ctor_release(x_3322, 1); + x_3325 = x_3322; +} else { + lean_dec_ref(x_3322); + x_3325 = lean_box(0); +} +if (lean_is_scalar(x_3325)) { + x_3326 = lean_alloc_ctor(0, 2, 0); +} else { + x_3326 = x_3325; +} +lean_ctor_set(x_3326, 0, x_3323); +lean_ctor_set(x_3326, 1, x_3324); +return x_3326; +} +else +{ +lean_object* x_3327; lean_object* x_3328; lean_object* x_3329; lean_object* x_3330; uint8_t x_3331; +x_3327 = lean_ctor_get(x_3322, 0); +lean_inc(x_3327); +x_3328 = lean_ctor_get(x_3322, 1); +lean_inc(x_3328); +if (lean_is_exclusive(x_3322)) { + lean_ctor_release(x_3322, 0); + lean_ctor_release(x_3322, 1); + x_3329 = x_3322; +} else { + lean_dec_ref(x_3322); + x_3329 = lean_box(0); +} +x_3330 = lean_ctor_get(x_3327, 1); +lean_inc(x_3330); +x_3331 = lean_nat_dec_eq(x_3316, x_3330); +lean_dec(x_3316); +if (x_3331 == 0) +{ +lean_object* x_3332; +lean_dec(x_3330); +if (lean_is_scalar(x_3329)) { + x_3332 = lean_alloc_ctor(1, 2, 0); +} else { + x_3332 = x_3329; +} +lean_ctor_set(x_3332, 0, x_3327); +lean_ctor_set(x_3332, 1, x_3328); +return x_3332; +} +else +{ +lean_object* x_3333; lean_object* x_3334; lean_object* x_3335; lean_object* x_3336; +lean_dec(x_3329); +lean_dec(x_3328); +x_3333 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_3334 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3334, 0, x_3333); +x_3335 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_3336 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3335, x_3334, x_3327); +if (lean_obj_tag(x_3336) == 0) +{ +lean_object* x_3337; lean_object* x_3338; lean_object* x_3339; lean_object* x_3340; +lean_dec(x_3330); +x_3337 = lean_ctor_get(x_3336, 0); +lean_inc(x_3337); +x_3338 = lean_ctor_get(x_3336, 1); +lean_inc(x_3338); +if (lean_is_exclusive(x_3336)) { + lean_ctor_release(x_3336, 0); + lean_ctor_release(x_3336, 1); + x_3339 = x_3336; +} else { + lean_dec_ref(x_3336); + x_3339 = lean_box(0); +} +if (lean_is_scalar(x_3339)) { + x_3340 = lean_alloc_ctor(0, 2, 0); +} else { + x_3340 = x_3339; +} +lean_ctor_set(x_3340, 0, x_3337); +lean_ctor_set(x_3340, 1, x_3338); +return x_3340; +} +else +{ +lean_object* x_3341; lean_object* x_3342; lean_object* x_3343; lean_object* x_3344; uint8_t x_3345; +x_3341 = lean_ctor_get(x_3336, 0); +lean_inc(x_3341); +x_3342 = lean_ctor_get(x_3336, 1); +lean_inc(x_3342); +if (lean_is_exclusive(x_3336)) { + lean_ctor_release(x_3336, 0); + lean_ctor_release(x_3336, 1); + x_3343 = x_3336; +} else { + lean_dec_ref(x_3336); + x_3343 = lean_box(0); +} +x_3344 = lean_ctor_get(x_3341, 1); +lean_inc(x_3344); +x_3345 = lean_nat_dec_eq(x_3330, x_3344); +lean_dec(x_3330); +if (x_3345 == 0) +{ +lean_object* x_3346; +lean_dec(x_3344); +if (lean_is_scalar(x_3343)) { + x_3346 = lean_alloc_ctor(1, 2, 0); +} else { + x_3346 = x_3343; +} +lean_ctor_set(x_3346, 0, x_3341); +lean_ctor_set(x_3346, 1, x_3342); +return x_3346; +} +else +{ +lean_object* x_3347; lean_object* x_3348; lean_object* x_3349; lean_object* x_3350; +lean_dec(x_3343); +lean_dec(x_3342); +x_3347 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_3348 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3348, 0, x_3347); +x_3349 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_3350 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3349, x_3348, x_3341); +if (lean_obj_tag(x_3350) == 0) +{ +lean_object* x_3351; lean_object* x_3352; lean_object* x_3353; lean_object* x_3354; +lean_dec(x_3344); +x_3351 = lean_ctor_get(x_3350, 0); +lean_inc(x_3351); +x_3352 = lean_ctor_get(x_3350, 1); +lean_inc(x_3352); +if (lean_is_exclusive(x_3350)) { + lean_ctor_release(x_3350, 0); + lean_ctor_release(x_3350, 1); + x_3353 = x_3350; +} else { + lean_dec_ref(x_3350); + x_3353 = lean_box(0); +} +if (lean_is_scalar(x_3353)) { + x_3354 = lean_alloc_ctor(0, 2, 0); +} else { + x_3354 = x_3353; +} +lean_ctor_set(x_3354, 0, x_3351); +lean_ctor_set(x_3354, 1, x_3352); +return x_3354; +} +else +{ +lean_object* x_3355; lean_object* x_3356; lean_object* x_3357; lean_object* x_3358; uint8_t x_3359; +x_3355 = lean_ctor_get(x_3350, 0); +lean_inc(x_3355); +x_3356 = lean_ctor_get(x_3350, 1); +lean_inc(x_3356); +if (lean_is_exclusive(x_3350)) { + lean_ctor_release(x_3350, 0); + lean_ctor_release(x_3350, 1); + x_3357 = x_3350; +} else { + lean_dec_ref(x_3350); + x_3357 = lean_box(0); +} +x_3358 = lean_ctor_get(x_3355, 1); +lean_inc(x_3358); +x_3359 = lean_nat_dec_eq(x_3344, x_3358); +lean_dec(x_3344); +if (x_3359 == 0) +{ +lean_object* x_3360; +lean_dec(x_3358); +if (lean_is_scalar(x_3357)) { + x_3360 = lean_alloc_ctor(1, 2, 0); +} else { + x_3360 = x_3357; +} +lean_ctor_set(x_3360, 0, x_3355); +lean_ctor_set(x_3360, 1, x_3356); +return x_3360; +} +else +{ +lean_object* x_3361; lean_object* x_3362; lean_object* x_3363; lean_object* x_3364; +lean_dec(x_3357); +lean_dec(x_3356); +x_3361 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_3362 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3362, 0, x_3361); +x_3363 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_3364 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3363, x_3362, x_3355); +if (lean_obj_tag(x_3364) == 0) +{ +lean_object* x_3365; lean_object* x_3366; lean_object* x_3367; lean_object* x_3368; +lean_dec(x_3358); +x_3365 = lean_ctor_get(x_3364, 0); +lean_inc(x_3365); +x_3366 = lean_ctor_get(x_3364, 1); +lean_inc(x_3366); +if (lean_is_exclusive(x_3364)) { + lean_ctor_release(x_3364, 0); + lean_ctor_release(x_3364, 1); + x_3367 = x_3364; +} else { + lean_dec_ref(x_3364); + x_3367 = lean_box(0); +} +if (lean_is_scalar(x_3367)) { + x_3368 = lean_alloc_ctor(0, 2, 0); +} else { + x_3368 = x_3367; +} +lean_ctor_set(x_3368, 0, x_3365); +lean_ctor_set(x_3368, 1, x_3366); +return x_3368; +} +else +{ +lean_object* x_3369; lean_object* x_3370; lean_object* x_3371; lean_object* x_3372; uint8_t x_3373; +x_3369 = lean_ctor_get(x_3364, 0); +lean_inc(x_3369); +x_3370 = lean_ctor_get(x_3364, 1); +lean_inc(x_3370); +if (lean_is_exclusive(x_3364)) { + lean_ctor_release(x_3364, 0); + lean_ctor_release(x_3364, 1); + x_3371 = x_3364; +} else { + lean_dec_ref(x_3364); + x_3371 = lean_box(0); +} +x_3372 = lean_ctor_get(x_3369, 1); +lean_inc(x_3372); +x_3373 = lean_nat_dec_eq(x_3358, x_3372); +lean_dec(x_3358); +if (x_3373 == 0) +{ +lean_object* x_3374; +lean_dec(x_3372); +if (lean_is_scalar(x_3371)) { + x_3374 = lean_alloc_ctor(1, 2, 0); +} else { + x_3374 = x_3371; +} +lean_ctor_set(x_3374, 0, x_3369); +lean_ctor_set(x_3374, 1, x_3370); +return x_3374; +} +else +{ +lean_object* x_3375; lean_object* x_3376; lean_object* x_3377; lean_object* x_3378; +lean_dec(x_3371); +lean_dec(x_3370); +x_3375 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_3376 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3376, 0, x_3375); +x_3377 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_3378 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3377, x_3376, x_3369); +if (lean_obj_tag(x_3378) == 0) +{ +lean_object* x_3379; lean_object* x_3380; lean_object* x_3381; lean_object* x_3382; +lean_dec(x_3372); +x_3379 = lean_ctor_get(x_3378, 0); +lean_inc(x_3379); +x_3380 = lean_ctor_get(x_3378, 1); +lean_inc(x_3380); +if (lean_is_exclusive(x_3378)) { + lean_ctor_release(x_3378, 0); + lean_ctor_release(x_3378, 1); + x_3381 = x_3378; +} else { + lean_dec_ref(x_3378); + x_3381 = lean_box(0); +} +if (lean_is_scalar(x_3381)) { + x_3382 = lean_alloc_ctor(0, 2, 0); +} else { + x_3382 = x_3381; +} +lean_ctor_set(x_3382, 0, x_3379); +lean_ctor_set(x_3382, 1, x_3380); +return x_3382; +} +else +{ +lean_object* x_3383; lean_object* x_3384; lean_object* x_3385; lean_object* x_3386; uint8_t x_3387; +x_3383 = lean_ctor_get(x_3378, 0); +lean_inc(x_3383); +x_3384 = lean_ctor_get(x_3378, 1); +lean_inc(x_3384); +if (lean_is_exclusive(x_3378)) { + lean_ctor_release(x_3378, 0); + lean_ctor_release(x_3378, 1); + x_3385 = x_3378; +} else { + lean_dec_ref(x_3378); + x_3385 = lean_box(0); +} +x_3386 = lean_ctor_get(x_3383, 1); +lean_inc(x_3386); +x_3387 = lean_nat_dec_eq(x_3372, x_3386); +lean_dec(x_3372); +if (x_3387 == 0) +{ +lean_object* x_3388; +lean_dec(x_3386); +if (lean_is_scalar(x_3385)) { + x_3388 = lean_alloc_ctor(1, 2, 0); +} else { + x_3388 = x_3385; +} +lean_ctor_set(x_3388, 0, x_3383); +lean_ctor_set(x_3388, 1, x_3384); +return x_3388; +} +else +{ +lean_object* x_3389; lean_object* x_3390; lean_object* x_3391; lean_object* x_3392; +lean_dec(x_3385); +lean_dec(x_3384); +x_3389 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_3390 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3390, 0, x_3389); +x_3391 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_3392 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3391, x_3390, x_3383); +if (lean_obj_tag(x_3392) == 0) +{ +lean_object* x_3393; lean_object* x_3394; lean_object* x_3395; lean_object* x_3396; +lean_dec(x_3386); +x_3393 = lean_ctor_get(x_3392, 0); +lean_inc(x_3393); +x_3394 = lean_ctor_get(x_3392, 1); +lean_inc(x_3394); +if (lean_is_exclusive(x_3392)) { + lean_ctor_release(x_3392, 0); + lean_ctor_release(x_3392, 1); + x_3395 = x_3392; +} else { + lean_dec_ref(x_3392); + x_3395 = lean_box(0); +} +if (lean_is_scalar(x_3395)) { + x_3396 = lean_alloc_ctor(0, 2, 0); +} else { + x_3396 = x_3395; +} +lean_ctor_set(x_3396, 0, x_3393); +lean_ctor_set(x_3396, 1, x_3394); +return x_3396; +} +else +{ +lean_object* x_3397; lean_object* x_3398; lean_object* x_3399; lean_object* x_3400; uint8_t x_3401; +x_3397 = lean_ctor_get(x_3392, 0); +lean_inc(x_3397); +x_3398 = lean_ctor_get(x_3392, 1); +lean_inc(x_3398); +if (lean_is_exclusive(x_3392)) { + lean_ctor_release(x_3392, 0); + lean_ctor_release(x_3392, 1); + x_3399 = x_3392; +} else { + lean_dec_ref(x_3392); + x_3399 = lean_box(0); +} +x_3400 = lean_ctor_get(x_3397, 1); +lean_inc(x_3400); +x_3401 = lean_nat_dec_eq(x_3386, x_3400); +lean_dec(x_3386); +if (x_3401 == 0) +{ +lean_object* x_3402; +lean_dec(x_3400); +if (lean_is_scalar(x_3399)) { + x_3402 = lean_alloc_ctor(1, 2, 0); +} else { + x_3402 = x_3399; +} +lean_ctor_set(x_3402, 0, x_3397); +lean_ctor_set(x_3402, 1, x_3398); +return x_3402; +} +else +{ +lean_object* x_3403; lean_object* x_3404; lean_object* x_3405; lean_object* x_3406; +lean_dec(x_3399); +lean_dec(x_3398); +x_3403 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_3404 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3404, 0, x_3403); +x_3405 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_3406 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3405, x_3404, x_3397); +if (lean_obj_tag(x_3406) == 0) +{ +lean_object* x_3407; lean_object* x_3408; lean_object* x_3409; lean_object* x_3410; +lean_dec(x_3400); +x_3407 = lean_ctor_get(x_3406, 0); +lean_inc(x_3407); +x_3408 = lean_ctor_get(x_3406, 1); +lean_inc(x_3408); +if (lean_is_exclusive(x_3406)) { + lean_ctor_release(x_3406, 0); + lean_ctor_release(x_3406, 1); + x_3409 = x_3406; +} else { + lean_dec_ref(x_3406); + x_3409 = lean_box(0); +} +if (lean_is_scalar(x_3409)) { + x_3410 = lean_alloc_ctor(0, 2, 0); +} else { + x_3410 = x_3409; +} +lean_ctor_set(x_3410, 0, x_3407); +lean_ctor_set(x_3410, 1, x_3408); +return x_3410; +} +else +{ +lean_object* x_3411; lean_object* x_3412; lean_object* x_3413; lean_object* x_3414; uint8_t x_3415; +x_3411 = lean_ctor_get(x_3406, 0); +lean_inc(x_3411); +x_3412 = lean_ctor_get(x_3406, 1); +lean_inc(x_3412); +if (lean_is_exclusive(x_3406)) { + lean_ctor_release(x_3406, 0); + lean_ctor_release(x_3406, 1); + x_3413 = x_3406; +} else { + lean_dec_ref(x_3406); + x_3413 = lean_box(0); +} +x_3414 = lean_ctor_get(x_3411, 1); +lean_inc(x_3414); +x_3415 = lean_nat_dec_eq(x_3400, x_3414); +lean_dec(x_3400); +if (x_3415 == 0) +{ +lean_object* x_3416; +lean_dec(x_3414); +if (lean_is_scalar(x_3413)) { + x_3416 = lean_alloc_ctor(1, 2, 0); +} else { + x_3416 = x_3413; +} +lean_ctor_set(x_3416, 0, x_3411); +lean_ctor_set(x_3416, 1, x_3412); +return x_3416; +} +else +{ +lean_object* x_3417; lean_object* x_3418; lean_object* x_3419; lean_object* x_3420; +lean_dec(x_3413); +lean_dec(x_3412); +x_3417 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_3418 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3418, 0, x_3417); +x_3419 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_3420 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3419, x_3418, x_3411); +if (lean_obj_tag(x_3420) == 0) +{ +lean_object* x_3421; lean_object* x_3422; lean_object* x_3423; lean_object* x_3424; +lean_dec(x_3414); +x_3421 = lean_ctor_get(x_3420, 0); +lean_inc(x_3421); +x_3422 = lean_ctor_get(x_3420, 1); +lean_inc(x_3422); +if (lean_is_exclusive(x_3420)) { + lean_ctor_release(x_3420, 0); + lean_ctor_release(x_3420, 1); + x_3423 = x_3420; +} else { + lean_dec_ref(x_3420); + x_3423 = lean_box(0); +} +if (lean_is_scalar(x_3423)) { + x_3424 = lean_alloc_ctor(0, 2, 0); +} else { + x_3424 = x_3423; +} +lean_ctor_set(x_3424, 0, x_3421); +lean_ctor_set(x_3424, 1, x_3422); +return x_3424; +} +else +{ +lean_object* x_3425; lean_object* x_3426; lean_object* x_3427; lean_object* x_3428; uint8_t x_3429; +x_3425 = lean_ctor_get(x_3420, 0); +lean_inc(x_3425); +x_3426 = lean_ctor_get(x_3420, 1); +lean_inc(x_3426); +if (lean_is_exclusive(x_3420)) { + lean_ctor_release(x_3420, 0); + lean_ctor_release(x_3420, 1); + x_3427 = x_3420; +} else { + lean_dec_ref(x_3420); + x_3427 = lean_box(0); +} +x_3428 = lean_ctor_get(x_3425, 1); +lean_inc(x_3428); +x_3429 = lean_nat_dec_eq(x_3414, x_3428); +lean_dec(x_3414); +if (x_3429 == 0) +{ +lean_object* x_3430; +lean_dec(x_3428); +if (lean_is_scalar(x_3427)) { + x_3430 = lean_alloc_ctor(1, 2, 0); +} else { + x_3430 = x_3427; +} +lean_ctor_set(x_3430, 0, x_3425); +lean_ctor_set(x_3430, 1, x_3426); +return x_3430; +} +else +{ +lean_object* x_3431; lean_object* x_3432; lean_object* x_3433; lean_object* x_3434; +lean_dec(x_3427); +lean_dec(x_3426); +x_3431 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_3432 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3432, 0, x_3431); +x_3433 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_3434 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3433, x_3432, x_3425); +if (lean_obj_tag(x_3434) == 0) +{ +lean_object* x_3435; lean_object* x_3436; lean_object* x_3437; lean_object* x_3438; +lean_dec(x_3428); +x_3435 = lean_ctor_get(x_3434, 0); +lean_inc(x_3435); +x_3436 = lean_ctor_get(x_3434, 1); +lean_inc(x_3436); +if (lean_is_exclusive(x_3434)) { + lean_ctor_release(x_3434, 0); + lean_ctor_release(x_3434, 1); + x_3437 = x_3434; +} else { + lean_dec_ref(x_3434); + x_3437 = lean_box(0); +} +if (lean_is_scalar(x_3437)) { + x_3438 = lean_alloc_ctor(0, 2, 0); +} else { + x_3438 = x_3437; +} +lean_ctor_set(x_3438, 0, x_3435); +lean_ctor_set(x_3438, 1, x_3436); +return x_3438; +} +else +{ +lean_object* x_3439; lean_object* x_3440; lean_object* x_3441; lean_object* x_3442; uint8_t x_3443; +x_3439 = lean_ctor_get(x_3434, 0); +lean_inc(x_3439); +x_3440 = lean_ctor_get(x_3434, 1); +lean_inc(x_3440); +if (lean_is_exclusive(x_3434)) { + lean_ctor_release(x_3434, 0); + lean_ctor_release(x_3434, 1); + x_3441 = x_3434; +} else { + lean_dec_ref(x_3434); + x_3441 = lean_box(0); +} +x_3442 = lean_ctor_get(x_3439, 1); +lean_inc(x_3442); +x_3443 = lean_nat_dec_eq(x_3428, x_3442); +lean_dec(x_3428); +if (x_3443 == 0) +{ +lean_object* x_3444; +lean_dec(x_3442); +if (lean_is_scalar(x_3441)) { + x_3444 = lean_alloc_ctor(1, 2, 0); +} else { + x_3444 = x_3441; +} +lean_ctor_set(x_3444, 0, x_3439); +lean_ctor_set(x_3444, 1, x_3440); +return x_3444; +} +else +{ +lean_object* x_3445; lean_object* x_3446; lean_object* x_3447; lean_object* x_3448; +lean_dec(x_3441); +lean_dec(x_3440); +x_3445 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_3446 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3446, 0, x_3445); +x_3447 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_3448 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3447, x_3446, x_3439); +if (lean_obj_tag(x_3448) == 0) +{ +lean_object* x_3449; lean_object* x_3450; lean_object* x_3451; lean_object* x_3452; +lean_dec(x_3442); +x_3449 = lean_ctor_get(x_3448, 0); +lean_inc(x_3449); +x_3450 = lean_ctor_get(x_3448, 1); +lean_inc(x_3450); +if (lean_is_exclusive(x_3448)) { + lean_ctor_release(x_3448, 0); + lean_ctor_release(x_3448, 1); + x_3451 = x_3448; +} else { + lean_dec_ref(x_3448); + x_3451 = lean_box(0); +} +if (lean_is_scalar(x_3451)) { + x_3452 = lean_alloc_ctor(0, 2, 0); +} else { + x_3452 = x_3451; +} +lean_ctor_set(x_3452, 0, x_3449); +lean_ctor_set(x_3452, 1, x_3450); +return x_3452; +} +else +{ +lean_object* x_3453; lean_object* x_3454; lean_object* x_3455; lean_object* x_3456; uint8_t x_3457; +x_3453 = lean_ctor_get(x_3448, 0); +lean_inc(x_3453); +x_3454 = lean_ctor_get(x_3448, 1); +lean_inc(x_3454); +if (lean_is_exclusive(x_3448)) { + lean_ctor_release(x_3448, 0); + lean_ctor_release(x_3448, 1); + x_3455 = x_3448; +} else { + lean_dec_ref(x_3448); + x_3455 = lean_box(0); +} +x_3456 = lean_ctor_get(x_3453, 1); +lean_inc(x_3456); +x_3457 = lean_nat_dec_eq(x_3442, x_3456); +lean_dec(x_3442); +if (x_3457 == 0) +{ +lean_object* x_3458; +lean_dec(x_3456); +if (lean_is_scalar(x_3455)) { + x_3458 = lean_alloc_ctor(1, 2, 0); +} else { + x_3458 = x_3455; +} +lean_ctor_set(x_3458, 0, x_3453); +lean_ctor_set(x_3458, 1, x_3454); +return x_3458; +} +else +{ +lean_object* x_3459; lean_object* x_3460; lean_object* x_3461; lean_object* x_3462; +lean_dec(x_3455); +lean_dec(x_3454); +x_3459 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_3460 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3460, 0, x_3459); +x_3461 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_3462 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3461, x_3460, x_3453); +if (lean_obj_tag(x_3462) == 0) +{ +lean_object* x_3463; lean_object* x_3464; lean_object* x_3465; lean_object* x_3466; +lean_dec(x_3456); +x_3463 = lean_ctor_get(x_3462, 0); +lean_inc(x_3463); +x_3464 = lean_ctor_get(x_3462, 1); +lean_inc(x_3464); +if (lean_is_exclusive(x_3462)) { + lean_ctor_release(x_3462, 0); + lean_ctor_release(x_3462, 1); + x_3465 = x_3462; +} else { + lean_dec_ref(x_3462); + x_3465 = lean_box(0); +} +if (lean_is_scalar(x_3465)) { + x_3466 = lean_alloc_ctor(0, 2, 0); +} else { + x_3466 = x_3465; +} +lean_ctor_set(x_3466, 0, x_3463); +lean_ctor_set(x_3466, 1, x_3464); +return x_3466; +} +else +{ +lean_object* x_3467; lean_object* x_3468; lean_object* x_3469; lean_object* x_3470; uint8_t x_3471; +x_3467 = lean_ctor_get(x_3462, 0); +lean_inc(x_3467); +x_3468 = lean_ctor_get(x_3462, 1); +lean_inc(x_3468); +if (lean_is_exclusive(x_3462)) { + lean_ctor_release(x_3462, 0); + lean_ctor_release(x_3462, 1); + x_3469 = x_3462; +} else { + lean_dec_ref(x_3462); + x_3469 = lean_box(0); +} +x_3470 = lean_ctor_get(x_3467, 1); +lean_inc(x_3470); +x_3471 = lean_nat_dec_eq(x_3456, x_3470); +lean_dec(x_3456); +if (x_3471 == 0) +{ +lean_object* x_3472; +lean_dec(x_3470); +if (lean_is_scalar(x_3469)) { + x_3472 = lean_alloc_ctor(1, 2, 0); +} else { + x_3472 = x_3469; +} +lean_ctor_set(x_3472, 0, x_3467); +lean_ctor_set(x_3472, 1, x_3468); +return x_3472; +} +else +{ +lean_object* x_3473; lean_object* x_3474; lean_object* x_3475; lean_object* x_3476; +lean_dec(x_3469); +lean_dec(x_3468); +x_3473 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_3474 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3474, 0, x_3473); +x_3475 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_3476 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3475, x_3474, x_3467); +if (lean_obj_tag(x_3476) == 0) +{ +lean_object* x_3477; lean_object* x_3478; lean_object* x_3479; lean_object* x_3480; +lean_dec(x_3470); +x_3477 = lean_ctor_get(x_3476, 0); +lean_inc(x_3477); +x_3478 = lean_ctor_get(x_3476, 1); +lean_inc(x_3478); +if (lean_is_exclusive(x_3476)) { + lean_ctor_release(x_3476, 0); + lean_ctor_release(x_3476, 1); + x_3479 = x_3476; +} else { + lean_dec_ref(x_3476); + x_3479 = lean_box(0); +} +if (lean_is_scalar(x_3479)) { + x_3480 = lean_alloc_ctor(0, 2, 0); +} else { + x_3480 = x_3479; +} +lean_ctor_set(x_3480, 0, x_3477); +lean_ctor_set(x_3480, 1, x_3478); +return x_3480; +} +else +{ +lean_object* x_3481; lean_object* x_3482; lean_object* x_3483; lean_object* x_3484; uint8_t x_3485; +x_3481 = lean_ctor_get(x_3476, 0); +lean_inc(x_3481); +x_3482 = lean_ctor_get(x_3476, 1); +lean_inc(x_3482); +if (lean_is_exclusive(x_3476)) { + lean_ctor_release(x_3476, 0); + lean_ctor_release(x_3476, 1); + x_3483 = x_3476; +} else { + lean_dec_ref(x_3476); + x_3483 = lean_box(0); +} +x_3484 = lean_ctor_get(x_3481, 1); +lean_inc(x_3484); +x_3485 = lean_nat_dec_eq(x_3470, x_3484); +lean_dec(x_3470); +if (x_3485 == 0) +{ +lean_object* x_3486; +lean_dec(x_3484); +if (lean_is_scalar(x_3483)) { + x_3486 = lean_alloc_ctor(1, 2, 0); +} else { + x_3486 = x_3483; +} +lean_ctor_set(x_3486, 0, x_3481); +lean_ctor_set(x_3486, 1, x_3482); +return x_3486; +} +else +{ +lean_object* x_3487; lean_object* x_3488; lean_object* x_3489; lean_object* x_3490; +lean_dec(x_3483); +lean_dec(x_3482); +x_3487 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_3488 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3488, 0, x_3487); +x_3489 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_3490 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3489, x_3488, x_3481); +if (lean_obj_tag(x_3490) == 0) +{ +lean_object* x_3491; lean_object* x_3492; lean_object* x_3493; lean_object* x_3494; +lean_dec(x_3484); +x_3491 = lean_ctor_get(x_3490, 0); +lean_inc(x_3491); +x_3492 = lean_ctor_get(x_3490, 1); +lean_inc(x_3492); +if (lean_is_exclusive(x_3490)) { + lean_ctor_release(x_3490, 0); + lean_ctor_release(x_3490, 1); + x_3493 = x_3490; +} else { + lean_dec_ref(x_3490); + x_3493 = lean_box(0); +} +if (lean_is_scalar(x_3493)) { + x_3494 = lean_alloc_ctor(0, 2, 0); +} else { + x_3494 = x_3493; +} +lean_ctor_set(x_3494, 0, x_3491); +lean_ctor_set(x_3494, 1, x_3492); +return x_3494; +} +else +{ +lean_object* x_3495; lean_object* x_3496; lean_object* x_3497; lean_object* x_3498; uint8_t x_3499; +x_3495 = lean_ctor_get(x_3490, 0); +lean_inc(x_3495); +x_3496 = lean_ctor_get(x_3490, 1); +lean_inc(x_3496); +if (lean_is_exclusive(x_3490)) { + lean_ctor_release(x_3490, 0); + lean_ctor_release(x_3490, 1); + x_3497 = x_3490; +} else { + lean_dec_ref(x_3490); + x_3497 = lean_box(0); +} +x_3498 = lean_ctor_get(x_3495, 1); +lean_inc(x_3498); +x_3499 = lean_nat_dec_eq(x_3484, x_3498); +lean_dec(x_3484); +if (x_3499 == 0) +{ +lean_object* x_3500; +lean_dec(x_3498); +if (lean_is_scalar(x_3497)) { + x_3500 = lean_alloc_ctor(1, 2, 0); +} else { + x_3500 = x_3497; +} +lean_ctor_set(x_3500, 0, x_3495); +lean_ctor_set(x_3500, 1, x_3496); +return x_3500; +} +else +{ +lean_object* x_3501; lean_object* x_3502; lean_object* x_3503; lean_object* x_3504; +lean_dec(x_3497); +lean_dec(x_3496); +x_3501 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_3502 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3502, 0, x_3501); +x_3503 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_3504 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3503, x_3502, x_3495); +if (lean_obj_tag(x_3504) == 0) +{ +lean_object* x_3505; lean_object* x_3506; lean_object* x_3507; lean_object* x_3508; +lean_dec(x_3498); +x_3505 = lean_ctor_get(x_3504, 0); +lean_inc(x_3505); +x_3506 = lean_ctor_get(x_3504, 1); +lean_inc(x_3506); +if (lean_is_exclusive(x_3504)) { + lean_ctor_release(x_3504, 0); + lean_ctor_release(x_3504, 1); + x_3507 = x_3504; +} else { + lean_dec_ref(x_3504); + x_3507 = lean_box(0); +} +if (lean_is_scalar(x_3507)) { + x_3508 = lean_alloc_ctor(0, 2, 0); +} else { + x_3508 = x_3507; +} +lean_ctor_set(x_3508, 0, x_3505); +lean_ctor_set(x_3508, 1, x_3506); +return x_3508; +} +else +{ +lean_object* x_3509; lean_object* x_3510; lean_object* x_3511; lean_object* x_3512; uint8_t x_3513; +x_3509 = lean_ctor_get(x_3504, 0); +lean_inc(x_3509); +x_3510 = lean_ctor_get(x_3504, 1); +lean_inc(x_3510); +if (lean_is_exclusive(x_3504)) { + lean_ctor_release(x_3504, 0); + lean_ctor_release(x_3504, 1); + x_3511 = x_3504; +} else { + lean_dec_ref(x_3504); + x_3511 = lean_box(0); +} +x_3512 = lean_ctor_get(x_3509, 1); +lean_inc(x_3512); +x_3513 = lean_nat_dec_eq(x_3498, x_3512); +lean_dec(x_3498); +if (x_3513 == 0) +{ +lean_object* x_3514; +lean_dec(x_3512); +if (lean_is_scalar(x_3511)) { + x_3514 = lean_alloc_ctor(1, 2, 0); +} else { + x_3514 = x_3511; +} +lean_ctor_set(x_3514, 0, x_3509); +lean_ctor_set(x_3514, 1, x_3510); +return x_3514; +} +else +{ +lean_object* x_3515; lean_object* x_3516; lean_object* x_3517; lean_object* x_3518; +lean_dec(x_3511); +lean_dec(x_3510); +x_3515 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_3516 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3516, 0, x_3515); +x_3517 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_3518 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3517, x_3516, x_3509); +if (lean_obj_tag(x_3518) == 0) +{ +lean_object* x_3519; lean_object* x_3520; lean_object* x_3521; lean_object* x_3522; +lean_dec(x_3512); +x_3519 = lean_ctor_get(x_3518, 0); +lean_inc(x_3519); +x_3520 = lean_ctor_get(x_3518, 1); +lean_inc(x_3520); +if (lean_is_exclusive(x_3518)) { + lean_ctor_release(x_3518, 0); + lean_ctor_release(x_3518, 1); + x_3521 = x_3518; +} else { + lean_dec_ref(x_3518); + x_3521 = lean_box(0); +} +if (lean_is_scalar(x_3521)) { + x_3522 = lean_alloc_ctor(0, 2, 0); +} else { + x_3522 = x_3521; +} +lean_ctor_set(x_3522, 0, x_3519); +lean_ctor_set(x_3522, 1, x_3520); +return x_3522; +} +else +{ +lean_object* x_3523; lean_object* x_3524; lean_object* x_3525; lean_object* x_3526; uint8_t x_3527; +x_3523 = lean_ctor_get(x_3518, 0); +lean_inc(x_3523); +x_3524 = lean_ctor_get(x_3518, 1); +lean_inc(x_3524); +if (lean_is_exclusive(x_3518)) { + lean_ctor_release(x_3518, 0); + lean_ctor_release(x_3518, 1); + x_3525 = x_3518; +} else { + lean_dec_ref(x_3518); + x_3525 = lean_box(0); +} +x_3526 = lean_ctor_get(x_3523, 1); +lean_inc(x_3526); +x_3527 = lean_nat_dec_eq(x_3512, x_3526); +lean_dec(x_3526); +lean_dec(x_3512); +if (x_3527 == 0) +{ +lean_object* x_3528; +if (lean_is_scalar(x_3525)) { + x_3528 = lean_alloc_ctor(1, 2, 0); +} else { + x_3528 = x_3525; +} +lean_ctor_set(x_3528, 0, x_3523); +lean_ctor_set(x_3528, 1, x_3524); +return x_3528; +} +else +{ +lean_object* x_3529; lean_object* x_3530; lean_object* x_3531; lean_object* x_3532; +lean_dec(x_3525); +lean_dec(x_3524); +x_3529 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_3530 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3530, 0, x_3529); +x_3531 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_3532 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3531, x_3530, x_3523); +return x_3532; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_3533; lean_object* x_3534; lean_object* x_3535; uint8_t x_3536; +x_3533 = lean_ctor_get(x_121, 0); +x_3534 = lean_ctor_get(x_121, 1); +lean_inc(x_3534); +lean_inc(x_3533); +lean_dec(x_121); +x_3535 = lean_ctor_get(x_3533, 1); +lean_inc(x_3535); +x_3536 = lean_nat_dec_eq(x_116, x_3535); +lean_dec(x_116); +if (x_3536 == 0) +{ +lean_object* x_3537; +lean_dec(x_3535); +x_3537 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3537, 0, x_3533); +lean_ctor_set(x_3537, 1, x_3534); +return x_3537; +} +else +{ +lean_object* x_3538; lean_object* x_3539; lean_object* x_3540; lean_object* x_3541; +lean_dec(x_3534); +x_3538 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_3539 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3539, 0, x_3538); +x_3540 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_3541 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3540, x_3539, x_3533); +if (lean_obj_tag(x_3541) == 0) +{ +lean_object* x_3542; lean_object* x_3543; lean_object* x_3544; lean_object* x_3545; +lean_dec(x_3535); +x_3542 = lean_ctor_get(x_3541, 0); +lean_inc(x_3542); +x_3543 = lean_ctor_get(x_3541, 1); +lean_inc(x_3543); +if (lean_is_exclusive(x_3541)) { + lean_ctor_release(x_3541, 0); + lean_ctor_release(x_3541, 1); + x_3544 = x_3541; +} else { + lean_dec_ref(x_3541); + x_3544 = lean_box(0); +} +if (lean_is_scalar(x_3544)) { + x_3545 = lean_alloc_ctor(0, 2, 0); +} else { + x_3545 = x_3544; +} +lean_ctor_set(x_3545, 0, x_3542); +lean_ctor_set(x_3545, 1, x_3543); +return x_3545; +} +else +{ +lean_object* x_3546; lean_object* x_3547; lean_object* x_3548; lean_object* x_3549; uint8_t x_3550; +x_3546 = lean_ctor_get(x_3541, 0); +lean_inc(x_3546); +x_3547 = lean_ctor_get(x_3541, 1); +lean_inc(x_3547); +if (lean_is_exclusive(x_3541)) { + lean_ctor_release(x_3541, 0); + lean_ctor_release(x_3541, 1); + x_3548 = x_3541; +} else { + lean_dec_ref(x_3541); + x_3548 = lean_box(0); +} +x_3549 = lean_ctor_get(x_3546, 1); +lean_inc(x_3549); +x_3550 = lean_nat_dec_eq(x_3535, x_3549); +lean_dec(x_3535); +if (x_3550 == 0) +{ +lean_object* x_3551; +lean_dec(x_3549); +if (lean_is_scalar(x_3548)) { + x_3551 = lean_alloc_ctor(1, 2, 0); +} else { + x_3551 = x_3548; +} +lean_ctor_set(x_3551, 0, x_3546); +lean_ctor_set(x_3551, 1, x_3547); +return x_3551; +} +else +{ +lean_object* x_3552; lean_object* x_3553; lean_object* x_3554; lean_object* x_3555; +lean_dec(x_3548); +lean_dec(x_3547); +x_3552 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_3553 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3553, 0, x_3552); +x_3554 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_3555 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3554, x_3553, x_3546); +if (lean_obj_tag(x_3555) == 0) +{ +lean_object* x_3556; lean_object* x_3557; lean_object* x_3558; lean_object* x_3559; +lean_dec(x_3549); +x_3556 = lean_ctor_get(x_3555, 0); +lean_inc(x_3556); +x_3557 = lean_ctor_get(x_3555, 1); +lean_inc(x_3557); +if (lean_is_exclusive(x_3555)) { + lean_ctor_release(x_3555, 0); + lean_ctor_release(x_3555, 1); + x_3558 = x_3555; +} else { + lean_dec_ref(x_3555); + x_3558 = lean_box(0); +} +if (lean_is_scalar(x_3558)) { + x_3559 = lean_alloc_ctor(0, 2, 0); +} else { + x_3559 = x_3558; +} +lean_ctor_set(x_3559, 0, x_3556); +lean_ctor_set(x_3559, 1, x_3557); +return x_3559; +} +else +{ +lean_object* x_3560; lean_object* x_3561; lean_object* x_3562; lean_object* x_3563; uint8_t x_3564; +x_3560 = lean_ctor_get(x_3555, 0); +lean_inc(x_3560); +x_3561 = lean_ctor_get(x_3555, 1); +lean_inc(x_3561); +if (lean_is_exclusive(x_3555)) { + lean_ctor_release(x_3555, 0); + lean_ctor_release(x_3555, 1); + x_3562 = x_3555; +} else { + lean_dec_ref(x_3555); + x_3562 = lean_box(0); +} +x_3563 = lean_ctor_get(x_3560, 1); +lean_inc(x_3563); +x_3564 = lean_nat_dec_eq(x_3549, x_3563); +lean_dec(x_3549); +if (x_3564 == 0) +{ +lean_object* x_3565; +lean_dec(x_3563); +if (lean_is_scalar(x_3562)) { + x_3565 = lean_alloc_ctor(1, 2, 0); +} else { + x_3565 = x_3562; +} +lean_ctor_set(x_3565, 0, x_3560); +lean_ctor_set(x_3565, 1, x_3561); +return x_3565; +} +else +{ +lean_object* x_3566; lean_object* x_3567; lean_object* x_3568; lean_object* x_3569; +lean_dec(x_3562); +lean_dec(x_3561); +x_3566 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_3567 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3567, 0, x_3566); +x_3568 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_3569 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3568, x_3567, x_3560); +if (lean_obj_tag(x_3569) == 0) +{ +lean_object* x_3570; lean_object* x_3571; lean_object* x_3572; lean_object* x_3573; +lean_dec(x_3563); +x_3570 = lean_ctor_get(x_3569, 0); +lean_inc(x_3570); +x_3571 = lean_ctor_get(x_3569, 1); +lean_inc(x_3571); +if (lean_is_exclusive(x_3569)) { + lean_ctor_release(x_3569, 0); + lean_ctor_release(x_3569, 1); + x_3572 = x_3569; +} else { + lean_dec_ref(x_3569); + x_3572 = lean_box(0); +} +if (lean_is_scalar(x_3572)) { + x_3573 = lean_alloc_ctor(0, 2, 0); +} else { + x_3573 = x_3572; +} +lean_ctor_set(x_3573, 0, x_3570); +lean_ctor_set(x_3573, 1, x_3571); +return x_3573; +} +else +{ +lean_object* x_3574; lean_object* x_3575; lean_object* x_3576; lean_object* x_3577; uint8_t x_3578; +x_3574 = lean_ctor_get(x_3569, 0); +lean_inc(x_3574); +x_3575 = lean_ctor_get(x_3569, 1); +lean_inc(x_3575); +if (lean_is_exclusive(x_3569)) { + lean_ctor_release(x_3569, 0); + lean_ctor_release(x_3569, 1); + x_3576 = x_3569; +} else { + lean_dec_ref(x_3569); + x_3576 = lean_box(0); +} +x_3577 = lean_ctor_get(x_3574, 1); +lean_inc(x_3577); +x_3578 = lean_nat_dec_eq(x_3563, x_3577); +lean_dec(x_3563); +if (x_3578 == 0) +{ +lean_object* x_3579; +lean_dec(x_3577); +if (lean_is_scalar(x_3576)) { + x_3579 = lean_alloc_ctor(1, 2, 0); +} else { + x_3579 = x_3576; +} +lean_ctor_set(x_3579, 0, x_3574); +lean_ctor_set(x_3579, 1, x_3575); +return x_3579; +} +else +{ +lean_object* x_3580; lean_object* x_3581; lean_object* x_3582; +lean_dec(x_3576); +lean_dec(x_3575); +x_3580 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_3581 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3581, 0, x_3580); +x_3582 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3568, x_3581, x_3574); +if (lean_obj_tag(x_3582) == 0) +{ +lean_object* x_3583; lean_object* x_3584; lean_object* x_3585; lean_object* x_3586; +lean_dec(x_3577); +x_3583 = lean_ctor_get(x_3582, 0); +lean_inc(x_3583); +x_3584 = lean_ctor_get(x_3582, 1); +lean_inc(x_3584); +if (lean_is_exclusive(x_3582)) { + lean_ctor_release(x_3582, 0); + lean_ctor_release(x_3582, 1); + x_3585 = x_3582; +} else { + lean_dec_ref(x_3582); + x_3585 = lean_box(0); +} +if (lean_is_scalar(x_3585)) { + x_3586 = lean_alloc_ctor(0, 2, 0); +} else { + x_3586 = x_3585; +} +lean_ctor_set(x_3586, 0, x_3583); +lean_ctor_set(x_3586, 1, x_3584); +return x_3586; +} +else +{ +lean_object* x_3587; lean_object* x_3588; lean_object* x_3589; lean_object* x_3590; uint8_t x_3591; +x_3587 = lean_ctor_get(x_3582, 0); +lean_inc(x_3587); +x_3588 = lean_ctor_get(x_3582, 1); +lean_inc(x_3588); +if (lean_is_exclusive(x_3582)) { + lean_ctor_release(x_3582, 0); + lean_ctor_release(x_3582, 1); + x_3589 = x_3582; +} else { + lean_dec_ref(x_3582); + x_3589 = lean_box(0); +} +x_3590 = lean_ctor_get(x_3587, 1); +lean_inc(x_3590); +x_3591 = lean_nat_dec_eq(x_3577, x_3590); +lean_dec(x_3577); +if (x_3591 == 0) +{ +lean_object* x_3592; +lean_dec(x_3590); +if (lean_is_scalar(x_3589)) { + x_3592 = lean_alloc_ctor(1, 2, 0); +} else { + x_3592 = x_3589; +} +lean_ctor_set(x_3592, 0, x_3587); +lean_ctor_set(x_3592, 1, x_3588); +return x_3592; +} +else +{ +lean_object* x_3593; lean_object* x_3594; lean_object* x_3595; lean_object* x_3596; +lean_dec(x_3589); +lean_dec(x_3588); +x_3593 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_3594 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3594, 0, x_3593); +x_3595 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_3596 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3595, x_3594, x_3587); +if (lean_obj_tag(x_3596) == 0) +{ +lean_object* x_3597; lean_object* x_3598; lean_object* x_3599; lean_object* x_3600; +lean_dec(x_3590); +x_3597 = lean_ctor_get(x_3596, 0); +lean_inc(x_3597); +x_3598 = lean_ctor_get(x_3596, 1); +lean_inc(x_3598); +if (lean_is_exclusive(x_3596)) { + lean_ctor_release(x_3596, 0); + lean_ctor_release(x_3596, 1); + x_3599 = x_3596; +} else { + lean_dec_ref(x_3596); + x_3599 = lean_box(0); +} +if (lean_is_scalar(x_3599)) { + x_3600 = lean_alloc_ctor(0, 2, 0); +} else { + x_3600 = x_3599; +} +lean_ctor_set(x_3600, 0, x_3597); +lean_ctor_set(x_3600, 1, x_3598); +return x_3600; +} +else +{ +lean_object* x_3601; lean_object* x_3602; lean_object* x_3603; lean_object* x_3604; uint8_t x_3605; +x_3601 = lean_ctor_get(x_3596, 0); +lean_inc(x_3601); +x_3602 = lean_ctor_get(x_3596, 1); +lean_inc(x_3602); +if (lean_is_exclusive(x_3596)) { + lean_ctor_release(x_3596, 0); + lean_ctor_release(x_3596, 1); + x_3603 = x_3596; +} else { + lean_dec_ref(x_3596); + x_3603 = lean_box(0); +} +x_3604 = lean_ctor_get(x_3601, 1); +lean_inc(x_3604); +x_3605 = lean_nat_dec_eq(x_3590, x_3604); +lean_dec(x_3590); +if (x_3605 == 0) +{ +lean_object* x_3606; +lean_dec(x_3604); +if (lean_is_scalar(x_3603)) { + x_3606 = lean_alloc_ctor(1, 2, 0); +} else { + x_3606 = x_3603; +} +lean_ctor_set(x_3606, 0, x_3601); +lean_ctor_set(x_3606, 1, x_3602); +return x_3606; +} +else +{ +lean_object* x_3607; lean_object* x_3608; lean_object* x_3609; lean_object* x_3610; +lean_dec(x_3603); +lean_dec(x_3602); +x_3607 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_3608 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3608, 0, x_3607); +x_3609 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_3610 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3609, x_3608, x_3601); +if (lean_obj_tag(x_3610) == 0) +{ +lean_object* x_3611; lean_object* x_3612; lean_object* x_3613; lean_object* x_3614; +lean_dec(x_3604); +x_3611 = lean_ctor_get(x_3610, 0); +lean_inc(x_3611); +x_3612 = lean_ctor_get(x_3610, 1); +lean_inc(x_3612); +if (lean_is_exclusive(x_3610)) { + lean_ctor_release(x_3610, 0); + lean_ctor_release(x_3610, 1); + x_3613 = x_3610; +} else { + lean_dec_ref(x_3610); + x_3613 = lean_box(0); +} +if (lean_is_scalar(x_3613)) { + x_3614 = lean_alloc_ctor(0, 2, 0); +} else { + x_3614 = x_3613; +} +lean_ctor_set(x_3614, 0, x_3611); +lean_ctor_set(x_3614, 1, x_3612); +return x_3614; +} +else +{ +lean_object* x_3615; lean_object* x_3616; lean_object* x_3617; lean_object* x_3618; uint8_t x_3619; +x_3615 = lean_ctor_get(x_3610, 0); +lean_inc(x_3615); +x_3616 = lean_ctor_get(x_3610, 1); +lean_inc(x_3616); +if (lean_is_exclusive(x_3610)) { + lean_ctor_release(x_3610, 0); + lean_ctor_release(x_3610, 1); + x_3617 = x_3610; +} else { + lean_dec_ref(x_3610); + x_3617 = lean_box(0); +} +x_3618 = lean_ctor_get(x_3615, 1); +lean_inc(x_3618); +x_3619 = lean_nat_dec_eq(x_3604, x_3618); +lean_dec(x_3604); +if (x_3619 == 0) +{ +lean_object* x_3620; +lean_dec(x_3618); +if (lean_is_scalar(x_3617)) { + x_3620 = lean_alloc_ctor(1, 2, 0); +} else { + x_3620 = x_3617; +} +lean_ctor_set(x_3620, 0, x_3615); +lean_ctor_set(x_3620, 1, x_3616); +return x_3620; +} +else +{ +lean_object* x_3621; lean_object* x_3622; lean_object* x_3623; lean_object* x_3624; +lean_dec(x_3617); +lean_dec(x_3616); +x_3621 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_3622 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3622, 0, x_3621); +x_3623 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_3624 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3623, x_3622, x_3615); +if (lean_obj_tag(x_3624) == 0) +{ +lean_object* x_3625; lean_object* x_3626; lean_object* x_3627; lean_object* x_3628; +lean_dec(x_3618); +x_3625 = lean_ctor_get(x_3624, 0); +lean_inc(x_3625); +x_3626 = lean_ctor_get(x_3624, 1); +lean_inc(x_3626); +if (lean_is_exclusive(x_3624)) { + lean_ctor_release(x_3624, 0); + lean_ctor_release(x_3624, 1); + x_3627 = x_3624; +} else { + lean_dec_ref(x_3624); + x_3627 = lean_box(0); +} +if (lean_is_scalar(x_3627)) { + x_3628 = lean_alloc_ctor(0, 2, 0); +} else { + x_3628 = x_3627; +} +lean_ctor_set(x_3628, 0, x_3625); +lean_ctor_set(x_3628, 1, x_3626); +return x_3628; +} +else +{ +lean_object* x_3629; lean_object* x_3630; lean_object* x_3631; lean_object* x_3632; uint8_t x_3633; +x_3629 = lean_ctor_get(x_3624, 0); +lean_inc(x_3629); +x_3630 = lean_ctor_get(x_3624, 1); +lean_inc(x_3630); +if (lean_is_exclusive(x_3624)) { + lean_ctor_release(x_3624, 0); + lean_ctor_release(x_3624, 1); + x_3631 = x_3624; +} else { + lean_dec_ref(x_3624); + x_3631 = lean_box(0); +} +x_3632 = lean_ctor_get(x_3629, 1); +lean_inc(x_3632); +x_3633 = lean_nat_dec_eq(x_3618, x_3632); +lean_dec(x_3618); +if (x_3633 == 0) +{ +lean_object* x_3634; +lean_dec(x_3632); +if (lean_is_scalar(x_3631)) { + x_3634 = lean_alloc_ctor(1, 2, 0); +} else { + x_3634 = x_3631; +} +lean_ctor_set(x_3634, 0, x_3629); +lean_ctor_set(x_3634, 1, x_3630); +return x_3634; +} +else +{ +lean_object* x_3635; lean_object* x_3636; lean_object* x_3637; lean_object* x_3638; +lean_dec(x_3631); +lean_dec(x_3630); +x_3635 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_3636 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3636, 0, x_3635); +x_3637 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_3638 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3637, x_3636, x_3629); +if (lean_obj_tag(x_3638) == 0) +{ +lean_object* x_3639; lean_object* x_3640; lean_object* x_3641; lean_object* x_3642; +lean_dec(x_3632); +x_3639 = lean_ctor_get(x_3638, 0); +lean_inc(x_3639); +x_3640 = lean_ctor_get(x_3638, 1); +lean_inc(x_3640); +if (lean_is_exclusive(x_3638)) { + lean_ctor_release(x_3638, 0); + lean_ctor_release(x_3638, 1); + x_3641 = x_3638; +} else { + lean_dec_ref(x_3638); + x_3641 = lean_box(0); +} +if (lean_is_scalar(x_3641)) { + x_3642 = lean_alloc_ctor(0, 2, 0); +} else { + x_3642 = x_3641; +} +lean_ctor_set(x_3642, 0, x_3639); +lean_ctor_set(x_3642, 1, x_3640); +return x_3642; +} +else +{ +lean_object* x_3643; lean_object* x_3644; lean_object* x_3645; lean_object* x_3646; uint8_t x_3647; +x_3643 = lean_ctor_get(x_3638, 0); +lean_inc(x_3643); +x_3644 = lean_ctor_get(x_3638, 1); +lean_inc(x_3644); +if (lean_is_exclusive(x_3638)) { + lean_ctor_release(x_3638, 0); + lean_ctor_release(x_3638, 1); + x_3645 = x_3638; +} else { + lean_dec_ref(x_3638); + x_3645 = lean_box(0); +} +x_3646 = lean_ctor_get(x_3643, 1); +lean_inc(x_3646); +x_3647 = lean_nat_dec_eq(x_3632, x_3646); +lean_dec(x_3632); +if (x_3647 == 0) +{ +lean_object* x_3648; +lean_dec(x_3646); +if (lean_is_scalar(x_3645)) { + x_3648 = lean_alloc_ctor(1, 2, 0); +} else { + x_3648 = x_3645; +} +lean_ctor_set(x_3648, 0, x_3643); +lean_ctor_set(x_3648, 1, x_3644); +return x_3648; +} +else +{ +lean_object* x_3649; lean_object* x_3650; lean_object* x_3651; lean_object* x_3652; +lean_dec(x_3645); +lean_dec(x_3644); +x_3649 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_3650 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3650, 0, x_3649); +x_3651 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_3652 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3651, x_3650, x_3643); +if (lean_obj_tag(x_3652) == 0) +{ +lean_object* x_3653; lean_object* x_3654; lean_object* x_3655; lean_object* x_3656; +lean_dec(x_3646); +x_3653 = lean_ctor_get(x_3652, 0); +lean_inc(x_3653); +x_3654 = lean_ctor_get(x_3652, 1); +lean_inc(x_3654); +if (lean_is_exclusive(x_3652)) { + lean_ctor_release(x_3652, 0); + lean_ctor_release(x_3652, 1); + x_3655 = x_3652; +} else { + lean_dec_ref(x_3652); + x_3655 = lean_box(0); +} +if (lean_is_scalar(x_3655)) { + x_3656 = lean_alloc_ctor(0, 2, 0); +} else { + x_3656 = x_3655; +} +lean_ctor_set(x_3656, 0, x_3653); +lean_ctor_set(x_3656, 1, x_3654); +return x_3656; +} +else +{ +lean_object* x_3657; lean_object* x_3658; lean_object* x_3659; lean_object* x_3660; uint8_t x_3661; +x_3657 = lean_ctor_get(x_3652, 0); +lean_inc(x_3657); +x_3658 = lean_ctor_get(x_3652, 1); +lean_inc(x_3658); +if (lean_is_exclusive(x_3652)) { + lean_ctor_release(x_3652, 0); + lean_ctor_release(x_3652, 1); + x_3659 = x_3652; +} else { + lean_dec_ref(x_3652); + x_3659 = lean_box(0); +} +x_3660 = lean_ctor_get(x_3657, 1); +lean_inc(x_3660); +x_3661 = lean_nat_dec_eq(x_3646, x_3660); +lean_dec(x_3646); +if (x_3661 == 0) +{ +lean_object* x_3662; +lean_dec(x_3660); +if (lean_is_scalar(x_3659)) { + x_3662 = lean_alloc_ctor(1, 2, 0); +} else { + x_3662 = x_3659; +} +lean_ctor_set(x_3662, 0, x_3657); +lean_ctor_set(x_3662, 1, x_3658); +return x_3662; +} +else +{ +lean_object* x_3663; lean_object* x_3664; lean_object* x_3665; lean_object* x_3666; +lean_dec(x_3659); +lean_dec(x_3658); +x_3663 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_3664 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3664, 0, x_3663); +x_3665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_3666 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3665, x_3664, x_3657); +if (lean_obj_tag(x_3666) == 0) +{ +lean_object* x_3667; lean_object* x_3668; lean_object* x_3669; lean_object* x_3670; +lean_dec(x_3660); +x_3667 = lean_ctor_get(x_3666, 0); +lean_inc(x_3667); +x_3668 = lean_ctor_get(x_3666, 1); +lean_inc(x_3668); +if (lean_is_exclusive(x_3666)) { + lean_ctor_release(x_3666, 0); + lean_ctor_release(x_3666, 1); + x_3669 = x_3666; +} else { + lean_dec_ref(x_3666); + x_3669 = lean_box(0); +} +if (lean_is_scalar(x_3669)) { + x_3670 = lean_alloc_ctor(0, 2, 0); +} else { + x_3670 = x_3669; +} +lean_ctor_set(x_3670, 0, x_3667); +lean_ctor_set(x_3670, 1, x_3668); +return x_3670; +} +else +{ +lean_object* x_3671; lean_object* x_3672; lean_object* x_3673; lean_object* x_3674; uint8_t x_3675; +x_3671 = lean_ctor_get(x_3666, 0); +lean_inc(x_3671); +x_3672 = lean_ctor_get(x_3666, 1); +lean_inc(x_3672); +if (lean_is_exclusive(x_3666)) { + lean_ctor_release(x_3666, 0); + lean_ctor_release(x_3666, 1); + x_3673 = x_3666; +} else { + lean_dec_ref(x_3666); + x_3673 = lean_box(0); +} +x_3674 = lean_ctor_get(x_3671, 1); +lean_inc(x_3674); +x_3675 = lean_nat_dec_eq(x_3660, x_3674); +lean_dec(x_3660); +if (x_3675 == 0) +{ +lean_object* x_3676; +lean_dec(x_3674); +if (lean_is_scalar(x_3673)) { + x_3676 = lean_alloc_ctor(1, 2, 0); +} else { + x_3676 = x_3673; +} +lean_ctor_set(x_3676, 0, x_3671); +lean_ctor_set(x_3676, 1, x_3672); +return x_3676; +} +else +{ +lean_object* x_3677; lean_object* x_3678; lean_object* x_3679; lean_object* x_3680; +lean_dec(x_3673); +lean_dec(x_3672); +x_3677 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_3678 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3678, 0, x_3677); +x_3679 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_3680 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3679, x_3678, x_3671); +if (lean_obj_tag(x_3680) == 0) +{ +lean_object* x_3681; lean_object* x_3682; lean_object* x_3683; lean_object* x_3684; +lean_dec(x_3674); +x_3681 = lean_ctor_get(x_3680, 0); +lean_inc(x_3681); +x_3682 = lean_ctor_get(x_3680, 1); +lean_inc(x_3682); +if (lean_is_exclusive(x_3680)) { + lean_ctor_release(x_3680, 0); + lean_ctor_release(x_3680, 1); + x_3683 = x_3680; +} else { + lean_dec_ref(x_3680); + x_3683 = lean_box(0); +} +if (lean_is_scalar(x_3683)) { + x_3684 = lean_alloc_ctor(0, 2, 0); +} else { + x_3684 = x_3683; +} +lean_ctor_set(x_3684, 0, x_3681); +lean_ctor_set(x_3684, 1, x_3682); +return x_3684; +} +else +{ +lean_object* x_3685; lean_object* x_3686; lean_object* x_3687; lean_object* x_3688; uint8_t x_3689; +x_3685 = lean_ctor_get(x_3680, 0); +lean_inc(x_3685); +x_3686 = lean_ctor_get(x_3680, 1); +lean_inc(x_3686); +if (lean_is_exclusive(x_3680)) { + lean_ctor_release(x_3680, 0); + lean_ctor_release(x_3680, 1); + x_3687 = x_3680; +} else { + lean_dec_ref(x_3680); + x_3687 = lean_box(0); +} +x_3688 = lean_ctor_get(x_3685, 1); +lean_inc(x_3688); +x_3689 = lean_nat_dec_eq(x_3674, x_3688); +lean_dec(x_3674); +if (x_3689 == 0) +{ +lean_object* x_3690; +lean_dec(x_3688); +if (lean_is_scalar(x_3687)) { + x_3690 = lean_alloc_ctor(1, 2, 0); +} else { + x_3690 = x_3687; +} +lean_ctor_set(x_3690, 0, x_3685); +lean_ctor_set(x_3690, 1, x_3686); +return x_3690; +} +else +{ +lean_object* x_3691; lean_object* x_3692; lean_object* x_3693; lean_object* x_3694; +lean_dec(x_3687); +lean_dec(x_3686); +x_3691 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_3692 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3692, 0, x_3691); +x_3693 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_3694 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3693, x_3692, x_3685); +if (lean_obj_tag(x_3694) == 0) +{ +lean_object* x_3695; lean_object* x_3696; lean_object* x_3697; lean_object* x_3698; +lean_dec(x_3688); +x_3695 = lean_ctor_get(x_3694, 0); +lean_inc(x_3695); +x_3696 = lean_ctor_get(x_3694, 1); +lean_inc(x_3696); +if (lean_is_exclusive(x_3694)) { + lean_ctor_release(x_3694, 0); + lean_ctor_release(x_3694, 1); + x_3697 = x_3694; +} else { + lean_dec_ref(x_3694); + x_3697 = lean_box(0); +} +if (lean_is_scalar(x_3697)) { + x_3698 = lean_alloc_ctor(0, 2, 0); +} else { + x_3698 = x_3697; +} +lean_ctor_set(x_3698, 0, x_3695); +lean_ctor_set(x_3698, 1, x_3696); +return x_3698; +} +else +{ +lean_object* x_3699; lean_object* x_3700; lean_object* x_3701; lean_object* x_3702; uint8_t x_3703; +x_3699 = lean_ctor_get(x_3694, 0); +lean_inc(x_3699); +x_3700 = lean_ctor_get(x_3694, 1); +lean_inc(x_3700); +if (lean_is_exclusive(x_3694)) { + lean_ctor_release(x_3694, 0); + lean_ctor_release(x_3694, 1); + x_3701 = x_3694; +} else { + lean_dec_ref(x_3694); + x_3701 = lean_box(0); +} +x_3702 = lean_ctor_get(x_3699, 1); +lean_inc(x_3702); +x_3703 = lean_nat_dec_eq(x_3688, x_3702); +lean_dec(x_3688); +if (x_3703 == 0) +{ +lean_object* x_3704; +lean_dec(x_3702); +if (lean_is_scalar(x_3701)) { + x_3704 = lean_alloc_ctor(1, 2, 0); +} else { + x_3704 = x_3701; +} +lean_ctor_set(x_3704, 0, x_3699); +lean_ctor_set(x_3704, 1, x_3700); +return x_3704; +} +else +{ +lean_object* x_3705; lean_object* x_3706; lean_object* x_3707; lean_object* x_3708; +lean_dec(x_3701); +lean_dec(x_3700); +x_3705 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_3706 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3706, 0, x_3705); +x_3707 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_3708 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3707, x_3706, x_3699); +if (lean_obj_tag(x_3708) == 0) +{ +lean_object* x_3709; lean_object* x_3710; lean_object* x_3711; lean_object* x_3712; +lean_dec(x_3702); +x_3709 = lean_ctor_get(x_3708, 0); +lean_inc(x_3709); +x_3710 = lean_ctor_get(x_3708, 1); +lean_inc(x_3710); +if (lean_is_exclusive(x_3708)) { + lean_ctor_release(x_3708, 0); + lean_ctor_release(x_3708, 1); + x_3711 = x_3708; +} else { + lean_dec_ref(x_3708); + x_3711 = lean_box(0); +} +if (lean_is_scalar(x_3711)) { + x_3712 = lean_alloc_ctor(0, 2, 0); +} else { + x_3712 = x_3711; +} +lean_ctor_set(x_3712, 0, x_3709); +lean_ctor_set(x_3712, 1, x_3710); +return x_3712; +} +else +{ +lean_object* x_3713; lean_object* x_3714; lean_object* x_3715; lean_object* x_3716; uint8_t x_3717; +x_3713 = lean_ctor_get(x_3708, 0); +lean_inc(x_3713); +x_3714 = lean_ctor_get(x_3708, 1); +lean_inc(x_3714); +if (lean_is_exclusive(x_3708)) { + lean_ctor_release(x_3708, 0); + lean_ctor_release(x_3708, 1); + x_3715 = x_3708; +} else { + lean_dec_ref(x_3708); + x_3715 = lean_box(0); +} +x_3716 = lean_ctor_get(x_3713, 1); +lean_inc(x_3716); +x_3717 = lean_nat_dec_eq(x_3702, x_3716); +lean_dec(x_3702); +if (x_3717 == 0) +{ +lean_object* x_3718; +lean_dec(x_3716); +if (lean_is_scalar(x_3715)) { + x_3718 = lean_alloc_ctor(1, 2, 0); +} else { + x_3718 = x_3715; +} +lean_ctor_set(x_3718, 0, x_3713); +lean_ctor_set(x_3718, 1, x_3714); +return x_3718; +} +else +{ +lean_object* x_3719; lean_object* x_3720; lean_object* x_3721; lean_object* x_3722; +lean_dec(x_3715); +lean_dec(x_3714); +x_3719 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_3720 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3720, 0, x_3719); +x_3721 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_3722 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3721, x_3720, x_3713); +if (lean_obj_tag(x_3722) == 0) +{ +lean_object* x_3723; lean_object* x_3724; lean_object* x_3725; lean_object* x_3726; +lean_dec(x_3716); +x_3723 = lean_ctor_get(x_3722, 0); +lean_inc(x_3723); +x_3724 = lean_ctor_get(x_3722, 1); +lean_inc(x_3724); +if (lean_is_exclusive(x_3722)) { + lean_ctor_release(x_3722, 0); + lean_ctor_release(x_3722, 1); + x_3725 = x_3722; +} else { + lean_dec_ref(x_3722); + x_3725 = lean_box(0); +} +if (lean_is_scalar(x_3725)) { + x_3726 = lean_alloc_ctor(0, 2, 0); +} else { + x_3726 = x_3725; +} +lean_ctor_set(x_3726, 0, x_3723); +lean_ctor_set(x_3726, 1, x_3724); +return x_3726; +} +else +{ +lean_object* x_3727; lean_object* x_3728; lean_object* x_3729; lean_object* x_3730; uint8_t x_3731; +x_3727 = lean_ctor_get(x_3722, 0); +lean_inc(x_3727); +x_3728 = lean_ctor_get(x_3722, 1); +lean_inc(x_3728); +if (lean_is_exclusive(x_3722)) { + lean_ctor_release(x_3722, 0); + lean_ctor_release(x_3722, 1); + x_3729 = x_3722; +} else { + lean_dec_ref(x_3722); + x_3729 = lean_box(0); +} +x_3730 = lean_ctor_get(x_3727, 1); +lean_inc(x_3730); +x_3731 = lean_nat_dec_eq(x_3716, x_3730); +lean_dec(x_3716); +if (x_3731 == 0) +{ +lean_object* x_3732; +lean_dec(x_3730); +if (lean_is_scalar(x_3729)) { + x_3732 = lean_alloc_ctor(1, 2, 0); +} else { + x_3732 = x_3729; +} +lean_ctor_set(x_3732, 0, x_3727); +lean_ctor_set(x_3732, 1, x_3728); +return x_3732; +} +else +{ +lean_object* x_3733; lean_object* x_3734; lean_object* x_3735; lean_object* x_3736; +lean_dec(x_3729); +lean_dec(x_3728); +x_3733 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_3734 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3734, 0, x_3733); +x_3735 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_3736 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3735, x_3734, x_3727); +if (lean_obj_tag(x_3736) == 0) +{ +lean_object* x_3737; lean_object* x_3738; lean_object* x_3739; lean_object* x_3740; +lean_dec(x_3730); +x_3737 = lean_ctor_get(x_3736, 0); +lean_inc(x_3737); +x_3738 = lean_ctor_get(x_3736, 1); +lean_inc(x_3738); +if (lean_is_exclusive(x_3736)) { + lean_ctor_release(x_3736, 0); + lean_ctor_release(x_3736, 1); + x_3739 = x_3736; +} else { + lean_dec_ref(x_3736); + x_3739 = lean_box(0); +} +if (lean_is_scalar(x_3739)) { + x_3740 = lean_alloc_ctor(0, 2, 0); +} else { + x_3740 = x_3739; +} +lean_ctor_set(x_3740, 0, x_3737); +lean_ctor_set(x_3740, 1, x_3738); +return x_3740; +} +else +{ +lean_object* x_3741; lean_object* x_3742; lean_object* x_3743; lean_object* x_3744; uint8_t x_3745; +x_3741 = lean_ctor_get(x_3736, 0); +lean_inc(x_3741); +x_3742 = lean_ctor_get(x_3736, 1); +lean_inc(x_3742); +if (lean_is_exclusive(x_3736)) { + lean_ctor_release(x_3736, 0); + lean_ctor_release(x_3736, 1); + x_3743 = x_3736; +} else { + lean_dec_ref(x_3736); + x_3743 = lean_box(0); +} +x_3744 = lean_ctor_get(x_3741, 1); +lean_inc(x_3744); +x_3745 = lean_nat_dec_eq(x_3730, x_3744); +lean_dec(x_3730); +if (x_3745 == 0) +{ +lean_object* x_3746; +lean_dec(x_3744); +if (lean_is_scalar(x_3743)) { + x_3746 = lean_alloc_ctor(1, 2, 0); +} else { + x_3746 = x_3743; +} +lean_ctor_set(x_3746, 0, x_3741); +lean_ctor_set(x_3746, 1, x_3742); +return x_3746; +} +else +{ +lean_object* x_3747; lean_object* x_3748; lean_object* x_3749; lean_object* x_3750; +lean_dec(x_3743); +lean_dec(x_3742); +x_3747 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_3748 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3748, 0, x_3747); +x_3749 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_3750 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3749, x_3748, x_3741); +if (lean_obj_tag(x_3750) == 0) +{ +lean_object* x_3751; lean_object* x_3752; lean_object* x_3753; lean_object* x_3754; +lean_dec(x_3744); +x_3751 = lean_ctor_get(x_3750, 0); +lean_inc(x_3751); +x_3752 = lean_ctor_get(x_3750, 1); +lean_inc(x_3752); +if (lean_is_exclusive(x_3750)) { + lean_ctor_release(x_3750, 0); + lean_ctor_release(x_3750, 1); + x_3753 = x_3750; +} else { + lean_dec_ref(x_3750); + x_3753 = lean_box(0); +} +if (lean_is_scalar(x_3753)) { + x_3754 = lean_alloc_ctor(0, 2, 0); +} else { + x_3754 = x_3753; +} +lean_ctor_set(x_3754, 0, x_3751); +lean_ctor_set(x_3754, 1, x_3752); +return x_3754; +} +else +{ +lean_object* x_3755; lean_object* x_3756; lean_object* x_3757; lean_object* x_3758; uint8_t x_3759; +x_3755 = lean_ctor_get(x_3750, 0); +lean_inc(x_3755); +x_3756 = lean_ctor_get(x_3750, 1); +lean_inc(x_3756); +if (lean_is_exclusive(x_3750)) { + lean_ctor_release(x_3750, 0); + lean_ctor_release(x_3750, 1); + x_3757 = x_3750; +} else { + lean_dec_ref(x_3750); + x_3757 = lean_box(0); +} +x_3758 = lean_ctor_get(x_3755, 1); +lean_inc(x_3758); +x_3759 = lean_nat_dec_eq(x_3744, x_3758); +lean_dec(x_3744); +if (x_3759 == 0) +{ +lean_object* x_3760; +lean_dec(x_3758); +if (lean_is_scalar(x_3757)) { + x_3760 = lean_alloc_ctor(1, 2, 0); +} else { + x_3760 = x_3757; +} +lean_ctor_set(x_3760, 0, x_3755); +lean_ctor_set(x_3760, 1, x_3756); +return x_3760; +} +else +{ +lean_object* x_3761; lean_object* x_3762; lean_object* x_3763; lean_object* x_3764; +lean_dec(x_3757); +lean_dec(x_3756); +x_3761 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_3762 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3762, 0, x_3761); +x_3763 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_3764 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3763, x_3762, x_3755); +if (lean_obj_tag(x_3764) == 0) +{ +lean_object* x_3765; lean_object* x_3766; lean_object* x_3767; lean_object* x_3768; +lean_dec(x_3758); +x_3765 = lean_ctor_get(x_3764, 0); +lean_inc(x_3765); +x_3766 = lean_ctor_get(x_3764, 1); +lean_inc(x_3766); +if (lean_is_exclusive(x_3764)) { + lean_ctor_release(x_3764, 0); + lean_ctor_release(x_3764, 1); + x_3767 = x_3764; +} else { + lean_dec_ref(x_3764); + x_3767 = lean_box(0); +} +if (lean_is_scalar(x_3767)) { + x_3768 = lean_alloc_ctor(0, 2, 0); +} else { + x_3768 = x_3767; +} +lean_ctor_set(x_3768, 0, x_3765); +lean_ctor_set(x_3768, 1, x_3766); +return x_3768; +} +else +{ +lean_object* x_3769; lean_object* x_3770; lean_object* x_3771; lean_object* x_3772; uint8_t x_3773; +x_3769 = lean_ctor_get(x_3764, 0); +lean_inc(x_3769); +x_3770 = lean_ctor_get(x_3764, 1); +lean_inc(x_3770); +if (lean_is_exclusive(x_3764)) { + lean_ctor_release(x_3764, 0); + lean_ctor_release(x_3764, 1); + x_3771 = x_3764; +} else { + lean_dec_ref(x_3764); + x_3771 = lean_box(0); +} +x_3772 = lean_ctor_get(x_3769, 1); +lean_inc(x_3772); +x_3773 = lean_nat_dec_eq(x_3758, x_3772); +lean_dec(x_3758); +if (x_3773 == 0) +{ +lean_object* x_3774; +lean_dec(x_3772); +if (lean_is_scalar(x_3771)) { + x_3774 = lean_alloc_ctor(1, 2, 0); +} else { + x_3774 = x_3771; +} +lean_ctor_set(x_3774, 0, x_3769); +lean_ctor_set(x_3774, 1, x_3770); +return x_3774; +} +else +{ +lean_object* x_3775; lean_object* x_3776; lean_object* x_3777; lean_object* x_3778; +lean_dec(x_3771); +lean_dec(x_3770); +x_3775 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_3776 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3776, 0, x_3775); +x_3777 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_3778 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3777, x_3776, x_3769); +if (lean_obj_tag(x_3778) == 0) +{ +lean_object* x_3779; lean_object* x_3780; lean_object* x_3781; lean_object* x_3782; +lean_dec(x_3772); +x_3779 = lean_ctor_get(x_3778, 0); +lean_inc(x_3779); +x_3780 = lean_ctor_get(x_3778, 1); +lean_inc(x_3780); +if (lean_is_exclusive(x_3778)) { + lean_ctor_release(x_3778, 0); + lean_ctor_release(x_3778, 1); + x_3781 = x_3778; +} else { + lean_dec_ref(x_3778); + x_3781 = lean_box(0); +} +if (lean_is_scalar(x_3781)) { + x_3782 = lean_alloc_ctor(0, 2, 0); +} else { + x_3782 = x_3781; +} +lean_ctor_set(x_3782, 0, x_3779); +lean_ctor_set(x_3782, 1, x_3780); +return x_3782; +} +else +{ +lean_object* x_3783; lean_object* x_3784; lean_object* x_3785; lean_object* x_3786; uint8_t x_3787; +x_3783 = lean_ctor_get(x_3778, 0); +lean_inc(x_3783); +x_3784 = lean_ctor_get(x_3778, 1); +lean_inc(x_3784); +if (lean_is_exclusive(x_3778)) { + lean_ctor_release(x_3778, 0); + lean_ctor_release(x_3778, 1); + x_3785 = x_3778; +} else { + lean_dec_ref(x_3778); + x_3785 = lean_box(0); +} +x_3786 = lean_ctor_get(x_3783, 1); +lean_inc(x_3786); +x_3787 = lean_nat_dec_eq(x_3772, x_3786); +lean_dec(x_3772); +if (x_3787 == 0) +{ +lean_object* x_3788; +lean_dec(x_3786); +if (lean_is_scalar(x_3785)) { + x_3788 = lean_alloc_ctor(1, 2, 0); +} else { + x_3788 = x_3785; +} +lean_ctor_set(x_3788, 0, x_3783); +lean_ctor_set(x_3788, 1, x_3784); +return x_3788; +} +else +{ +lean_object* x_3789; lean_object* x_3790; lean_object* x_3791; lean_object* x_3792; +lean_dec(x_3785); +lean_dec(x_3784); +x_3789 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_3790 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3790, 0, x_3789); +x_3791 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_3792 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3791, x_3790, x_3783); +if (lean_obj_tag(x_3792) == 0) +{ +lean_object* x_3793; lean_object* x_3794; lean_object* x_3795; lean_object* x_3796; +lean_dec(x_3786); +x_3793 = lean_ctor_get(x_3792, 0); +lean_inc(x_3793); +x_3794 = lean_ctor_get(x_3792, 1); +lean_inc(x_3794); +if (lean_is_exclusive(x_3792)) { + lean_ctor_release(x_3792, 0); + lean_ctor_release(x_3792, 1); + x_3795 = x_3792; +} else { + lean_dec_ref(x_3792); + x_3795 = lean_box(0); +} +if (lean_is_scalar(x_3795)) { + x_3796 = lean_alloc_ctor(0, 2, 0); +} else { + x_3796 = x_3795; +} +lean_ctor_set(x_3796, 0, x_3793); +lean_ctor_set(x_3796, 1, x_3794); +return x_3796; +} +else +{ +lean_object* x_3797; lean_object* x_3798; lean_object* x_3799; lean_object* x_3800; uint8_t x_3801; +x_3797 = lean_ctor_get(x_3792, 0); +lean_inc(x_3797); +x_3798 = lean_ctor_get(x_3792, 1); +lean_inc(x_3798); +if (lean_is_exclusive(x_3792)) { + lean_ctor_release(x_3792, 0); + lean_ctor_release(x_3792, 1); + x_3799 = x_3792; +} else { + lean_dec_ref(x_3792); + x_3799 = lean_box(0); +} +x_3800 = lean_ctor_get(x_3797, 1); +lean_inc(x_3800); +x_3801 = lean_nat_dec_eq(x_3786, x_3800); +lean_dec(x_3786); +if (x_3801 == 0) +{ +lean_object* x_3802; +lean_dec(x_3800); +if (lean_is_scalar(x_3799)) { + x_3802 = lean_alloc_ctor(1, 2, 0); +} else { + x_3802 = x_3799; +} +lean_ctor_set(x_3802, 0, x_3797); +lean_ctor_set(x_3802, 1, x_3798); +return x_3802; +} +else +{ +lean_object* x_3803; lean_object* x_3804; lean_object* x_3805; lean_object* x_3806; +lean_dec(x_3799); +lean_dec(x_3798); +x_3803 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_3804 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3804, 0, x_3803); +x_3805 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_3806 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3805, x_3804, x_3797); +if (lean_obj_tag(x_3806) == 0) +{ +lean_object* x_3807; lean_object* x_3808; lean_object* x_3809; lean_object* x_3810; +lean_dec(x_3800); +x_3807 = lean_ctor_get(x_3806, 0); +lean_inc(x_3807); +x_3808 = lean_ctor_get(x_3806, 1); +lean_inc(x_3808); +if (lean_is_exclusive(x_3806)) { + lean_ctor_release(x_3806, 0); + lean_ctor_release(x_3806, 1); + x_3809 = x_3806; +} else { + lean_dec_ref(x_3806); + x_3809 = lean_box(0); +} +if (lean_is_scalar(x_3809)) { + x_3810 = lean_alloc_ctor(0, 2, 0); +} else { + x_3810 = x_3809; +} +lean_ctor_set(x_3810, 0, x_3807); +lean_ctor_set(x_3810, 1, x_3808); +return x_3810; +} +else +{ +lean_object* x_3811; lean_object* x_3812; lean_object* x_3813; lean_object* x_3814; uint8_t x_3815; +x_3811 = lean_ctor_get(x_3806, 0); +lean_inc(x_3811); +x_3812 = lean_ctor_get(x_3806, 1); +lean_inc(x_3812); +if (lean_is_exclusive(x_3806)) { + lean_ctor_release(x_3806, 0); + lean_ctor_release(x_3806, 1); + x_3813 = x_3806; +} else { + lean_dec_ref(x_3806); + x_3813 = lean_box(0); +} +x_3814 = lean_ctor_get(x_3811, 1); +lean_inc(x_3814); +x_3815 = lean_nat_dec_eq(x_3800, x_3814); +lean_dec(x_3800); +if (x_3815 == 0) +{ +lean_object* x_3816; +lean_dec(x_3814); +if (lean_is_scalar(x_3813)) { + x_3816 = lean_alloc_ctor(1, 2, 0); +} else { + x_3816 = x_3813; +} +lean_ctor_set(x_3816, 0, x_3811); +lean_ctor_set(x_3816, 1, x_3812); +return x_3816; +} +else +{ +lean_object* x_3817; lean_object* x_3818; lean_object* x_3819; lean_object* x_3820; +lean_dec(x_3813); +lean_dec(x_3812); +x_3817 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_3818 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3818, 0, x_3817); +x_3819 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_3820 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3819, x_3818, x_3811); +if (lean_obj_tag(x_3820) == 0) +{ +lean_object* x_3821; lean_object* x_3822; lean_object* x_3823; lean_object* x_3824; +lean_dec(x_3814); +x_3821 = lean_ctor_get(x_3820, 0); +lean_inc(x_3821); +x_3822 = lean_ctor_get(x_3820, 1); +lean_inc(x_3822); +if (lean_is_exclusive(x_3820)) { + lean_ctor_release(x_3820, 0); + lean_ctor_release(x_3820, 1); + x_3823 = x_3820; +} else { + lean_dec_ref(x_3820); + x_3823 = lean_box(0); +} +if (lean_is_scalar(x_3823)) { + x_3824 = lean_alloc_ctor(0, 2, 0); +} else { + x_3824 = x_3823; +} +lean_ctor_set(x_3824, 0, x_3821); +lean_ctor_set(x_3824, 1, x_3822); +return x_3824; +} +else +{ +lean_object* x_3825; lean_object* x_3826; lean_object* x_3827; lean_object* x_3828; uint8_t x_3829; +x_3825 = lean_ctor_get(x_3820, 0); +lean_inc(x_3825); +x_3826 = lean_ctor_get(x_3820, 1); +lean_inc(x_3826); +if (lean_is_exclusive(x_3820)) { + lean_ctor_release(x_3820, 0); + lean_ctor_release(x_3820, 1); + x_3827 = x_3820; +} else { + lean_dec_ref(x_3820); + x_3827 = lean_box(0); +} +x_3828 = lean_ctor_get(x_3825, 1); +lean_inc(x_3828); +x_3829 = lean_nat_dec_eq(x_3814, x_3828); +lean_dec(x_3828); +lean_dec(x_3814); +if (x_3829 == 0) +{ +lean_object* x_3830; +if (lean_is_scalar(x_3827)) { + x_3830 = lean_alloc_ctor(1, 2, 0); +} else { + x_3830 = x_3827; +} +lean_ctor_set(x_3830, 0, x_3825); +lean_ctor_set(x_3830, 1, x_3826); +return x_3830; +} +else +{ +lean_object* x_3831; lean_object* x_3832; lean_object* x_3833; lean_object* x_3834; +lean_dec(x_3827); +lean_dec(x_3826); +x_3831 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_3832 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3832, 0, x_3831); +x_3833 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_3834 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3833, x_3832, x_3825); +return x_3834; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_3835; lean_object* x_3836; lean_object* x_3837; uint8_t x_3838; +x_3835 = lean_ctor_get(x_108, 0); +x_3836 = lean_ctor_get(x_108, 1); +lean_inc(x_3836); +lean_inc(x_3835); +lean_dec(x_108); +x_3837 = lean_ctor_get(x_3835, 1); +lean_inc(x_3837); +x_3838 = lean_nat_dec_eq(x_104, x_3837); +lean_dec(x_104); +if (x_3838 == 0) +{ +lean_object* x_3839; +lean_dec(x_3837); +x_3839 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3839, 0, x_3835); +lean_ctor_set(x_3839, 1, x_3836); +return x_3839; +} +else +{ +lean_object* x_3840; lean_object* x_3841; lean_object* x_3842; lean_object* x_3843; +lean_dec(x_3836); +x_3840 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_3841 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3841, 0, x_3840); +x_3842 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_3843 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3842, x_3841, x_3835); +if (lean_obj_tag(x_3843) == 0) +{ +lean_object* x_3844; lean_object* x_3845; lean_object* x_3846; lean_object* x_3847; +lean_dec(x_3837); +x_3844 = lean_ctor_get(x_3843, 0); +lean_inc(x_3844); +x_3845 = lean_ctor_get(x_3843, 1); +lean_inc(x_3845); +if (lean_is_exclusive(x_3843)) { + lean_ctor_release(x_3843, 0); + lean_ctor_release(x_3843, 1); + x_3846 = x_3843; +} else { + lean_dec_ref(x_3843); + x_3846 = lean_box(0); +} +if (lean_is_scalar(x_3846)) { + x_3847 = lean_alloc_ctor(0, 2, 0); +} else { + x_3847 = x_3846; +} +lean_ctor_set(x_3847, 0, x_3844); +lean_ctor_set(x_3847, 1, x_3845); +return x_3847; +} +else +{ +lean_object* x_3848; lean_object* x_3849; lean_object* x_3850; lean_object* x_3851; uint8_t x_3852; +x_3848 = lean_ctor_get(x_3843, 0); +lean_inc(x_3848); +x_3849 = lean_ctor_get(x_3843, 1); +lean_inc(x_3849); +if (lean_is_exclusive(x_3843)) { + lean_ctor_release(x_3843, 0); + lean_ctor_release(x_3843, 1); + x_3850 = x_3843; +} else { + lean_dec_ref(x_3843); + x_3850 = lean_box(0); +} +x_3851 = lean_ctor_get(x_3848, 1); +lean_inc(x_3851); +x_3852 = lean_nat_dec_eq(x_3837, x_3851); +lean_dec(x_3837); +if (x_3852 == 0) +{ +lean_object* x_3853; +lean_dec(x_3851); +if (lean_is_scalar(x_3850)) { + x_3853 = lean_alloc_ctor(1, 2, 0); +} else { + x_3853 = x_3850; +} +lean_ctor_set(x_3853, 0, x_3848); +lean_ctor_set(x_3853, 1, x_3849); +return x_3853; +} +else +{ +lean_object* x_3854; lean_object* x_3855; lean_object* x_3856; lean_object* x_3857; +lean_dec(x_3850); +lean_dec(x_3849); +x_3854 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_3855 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3855, 0, x_3854); +x_3856 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_3857 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3856, x_3855, x_3848); +if (lean_obj_tag(x_3857) == 0) +{ +lean_object* x_3858; lean_object* x_3859; lean_object* x_3860; lean_object* x_3861; +lean_dec(x_3851); +x_3858 = lean_ctor_get(x_3857, 0); +lean_inc(x_3858); +x_3859 = lean_ctor_get(x_3857, 1); +lean_inc(x_3859); +if (lean_is_exclusive(x_3857)) { + lean_ctor_release(x_3857, 0); + lean_ctor_release(x_3857, 1); + x_3860 = x_3857; +} else { + lean_dec_ref(x_3857); + x_3860 = lean_box(0); +} +if (lean_is_scalar(x_3860)) { + x_3861 = lean_alloc_ctor(0, 2, 0); +} else { + x_3861 = x_3860; +} +lean_ctor_set(x_3861, 0, x_3858); +lean_ctor_set(x_3861, 1, x_3859); +return x_3861; +} +else +{ +lean_object* x_3862; lean_object* x_3863; lean_object* x_3864; lean_object* x_3865; uint8_t x_3866; +x_3862 = lean_ctor_get(x_3857, 0); +lean_inc(x_3862); +x_3863 = lean_ctor_get(x_3857, 1); +lean_inc(x_3863); +if (lean_is_exclusive(x_3857)) { + lean_ctor_release(x_3857, 0); + lean_ctor_release(x_3857, 1); + x_3864 = x_3857; +} else { + lean_dec_ref(x_3857); + x_3864 = lean_box(0); +} +x_3865 = lean_ctor_get(x_3862, 1); +lean_inc(x_3865); +x_3866 = lean_nat_dec_eq(x_3851, x_3865); +lean_dec(x_3851); +if (x_3866 == 0) +{ +lean_object* x_3867; +lean_dec(x_3865); +if (lean_is_scalar(x_3864)) { + x_3867 = lean_alloc_ctor(1, 2, 0); +} else { + x_3867 = x_3864; +} +lean_ctor_set(x_3867, 0, x_3862); +lean_ctor_set(x_3867, 1, x_3863); +return x_3867; +} +else +{ +lean_object* x_3868; lean_object* x_3869; lean_object* x_3870; lean_object* x_3871; +lean_dec(x_3864); +lean_dec(x_3863); +x_3868 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_3869 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3869, 0, x_3868); +x_3870 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_3871 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3870, x_3869, x_3862); +if (lean_obj_tag(x_3871) == 0) +{ +lean_object* x_3872; lean_object* x_3873; lean_object* x_3874; lean_object* x_3875; +lean_dec(x_3865); +x_3872 = lean_ctor_get(x_3871, 0); +lean_inc(x_3872); +x_3873 = lean_ctor_get(x_3871, 1); +lean_inc(x_3873); +if (lean_is_exclusive(x_3871)) { + lean_ctor_release(x_3871, 0); + lean_ctor_release(x_3871, 1); + x_3874 = x_3871; +} else { + lean_dec_ref(x_3871); + x_3874 = lean_box(0); +} +if (lean_is_scalar(x_3874)) { + x_3875 = lean_alloc_ctor(0, 2, 0); +} else { + x_3875 = x_3874; +} +lean_ctor_set(x_3875, 0, x_3872); +lean_ctor_set(x_3875, 1, x_3873); +return x_3875; +} +else +{ +lean_object* x_3876; lean_object* x_3877; lean_object* x_3878; lean_object* x_3879; uint8_t x_3880; +x_3876 = lean_ctor_get(x_3871, 0); +lean_inc(x_3876); +x_3877 = lean_ctor_get(x_3871, 1); +lean_inc(x_3877); +if (lean_is_exclusive(x_3871)) { + lean_ctor_release(x_3871, 0); + lean_ctor_release(x_3871, 1); + x_3878 = x_3871; +} else { + lean_dec_ref(x_3871); + x_3878 = lean_box(0); +} +x_3879 = lean_ctor_get(x_3876, 1); +lean_inc(x_3879); +x_3880 = lean_nat_dec_eq(x_3865, x_3879); +lean_dec(x_3865); +if (x_3880 == 0) +{ +lean_object* x_3881; +lean_dec(x_3879); +if (lean_is_scalar(x_3878)) { + x_3881 = lean_alloc_ctor(1, 2, 0); +} else { + x_3881 = x_3878; +} +lean_ctor_set(x_3881, 0, x_3876); +lean_ctor_set(x_3881, 1, x_3877); +return x_3881; +} +else +{ +lean_object* x_3882; lean_object* x_3883; lean_object* x_3884; lean_object* x_3885; +lean_dec(x_3878); +lean_dec(x_3877); +x_3882 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_3883 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3883, 0, x_3882); +x_3884 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_3885 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3884, x_3883, x_3876); +if (lean_obj_tag(x_3885) == 0) +{ +lean_object* x_3886; lean_object* x_3887; lean_object* x_3888; lean_object* x_3889; +lean_dec(x_3879); +x_3886 = lean_ctor_get(x_3885, 0); +lean_inc(x_3886); +x_3887 = lean_ctor_get(x_3885, 1); +lean_inc(x_3887); +if (lean_is_exclusive(x_3885)) { + lean_ctor_release(x_3885, 0); + lean_ctor_release(x_3885, 1); + x_3888 = x_3885; +} else { + lean_dec_ref(x_3885); + x_3888 = lean_box(0); +} +if (lean_is_scalar(x_3888)) { + x_3889 = lean_alloc_ctor(0, 2, 0); +} else { + x_3889 = x_3888; +} +lean_ctor_set(x_3889, 0, x_3886); +lean_ctor_set(x_3889, 1, x_3887); +return x_3889; +} +else +{ +lean_object* x_3890; lean_object* x_3891; lean_object* x_3892; lean_object* x_3893; uint8_t x_3894; +x_3890 = lean_ctor_get(x_3885, 0); +lean_inc(x_3890); +x_3891 = lean_ctor_get(x_3885, 1); +lean_inc(x_3891); +if (lean_is_exclusive(x_3885)) { + lean_ctor_release(x_3885, 0); + lean_ctor_release(x_3885, 1); + x_3892 = x_3885; +} else { + lean_dec_ref(x_3885); + x_3892 = lean_box(0); +} +x_3893 = lean_ctor_get(x_3890, 1); +lean_inc(x_3893); +x_3894 = lean_nat_dec_eq(x_3879, x_3893); +lean_dec(x_3879); +if (x_3894 == 0) +{ +lean_object* x_3895; +lean_dec(x_3893); +if (lean_is_scalar(x_3892)) { + x_3895 = lean_alloc_ctor(1, 2, 0); +} else { + x_3895 = x_3892; +} +lean_ctor_set(x_3895, 0, x_3890); +lean_ctor_set(x_3895, 1, x_3891); +return x_3895; +} +else +{ +lean_object* x_3896; lean_object* x_3897; lean_object* x_3898; +lean_dec(x_3892); +lean_dec(x_3891); +x_3896 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_3897 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3897, 0, x_3896); +x_3898 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3884, x_3897, x_3890); +if (lean_obj_tag(x_3898) == 0) +{ +lean_object* x_3899; lean_object* x_3900; lean_object* x_3901; lean_object* x_3902; +lean_dec(x_3893); +x_3899 = lean_ctor_get(x_3898, 0); +lean_inc(x_3899); +x_3900 = lean_ctor_get(x_3898, 1); +lean_inc(x_3900); +if (lean_is_exclusive(x_3898)) { + lean_ctor_release(x_3898, 0); + lean_ctor_release(x_3898, 1); + x_3901 = x_3898; +} else { + lean_dec_ref(x_3898); + x_3901 = lean_box(0); +} +if (lean_is_scalar(x_3901)) { + x_3902 = lean_alloc_ctor(0, 2, 0); +} else { + x_3902 = x_3901; +} +lean_ctor_set(x_3902, 0, x_3899); +lean_ctor_set(x_3902, 1, x_3900); +return x_3902; +} +else +{ +lean_object* x_3903; lean_object* x_3904; lean_object* x_3905; lean_object* x_3906; uint8_t x_3907; +x_3903 = lean_ctor_get(x_3898, 0); +lean_inc(x_3903); +x_3904 = lean_ctor_get(x_3898, 1); +lean_inc(x_3904); +if (lean_is_exclusive(x_3898)) { + lean_ctor_release(x_3898, 0); + lean_ctor_release(x_3898, 1); + x_3905 = x_3898; +} else { + lean_dec_ref(x_3898); + x_3905 = lean_box(0); +} +x_3906 = lean_ctor_get(x_3903, 1); +lean_inc(x_3906); +x_3907 = lean_nat_dec_eq(x_3893, x_3906); +lean_dec(x_3893); +if (x_3907 == 0) +{ +lean_object* x_3908; +lean_dec(x_3906); +if (lean_is_scalar(x_3905)) { + x_3908 = lean_alloc_ctor(1, 2, 0); +} else { + x_3908 = x_3905; +} +lean_ctor_set(x_3908, 0, x_3903); +lean_ctor_set(x_3908, 1, x_3904); +return x_3908; +} +else +{ +lean_object* x_3909; lean_object* x_3910; lean_object* x_3911; lean_object* x_3912; +lean_dec(x_3905); +lean_dec(x_3904); +x_3909 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_3910 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3910, 0, x_3909); +x_3911 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_3912 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3911, x_3910, x_3903); +if (lean_obj_tag(x_3912) == 0) +{ +lean_object* x_3913; lean_object* x_3914; lean_object* x_3915; lean_object* x_3916; +lean_dec(x_3906); +x_3913 = lean_ctor_get(x_3912, 0); +lean_inc(x_3913); +x_3914 = lean_ctor_get(x_3912, 1); +lean_inc(x_3914); +if (lean_is_exclusive(x_3912)) { + lean_ctor_release(x_3912, 0); + lean_ctor_release(x_3912, 1); + x_3915 = x_3912; +} else { + lean_dec_ref(x_3912); + x_3915 = lean_box(0); +} +if (lean_is_scalar(x_3915)) { + x_3916 = lean_alloc_ctor(0, 2, 0); +} else { + x_3916 = x_3915; +} +lean_ctor_set(x_3916, 0, x_3913); +lean_ctor_set(x_3916, 1, x_3914); +return x_3916; +} +else +{ +lean_object* x_3917; lean_object* x_3918; lean_object* x_3919; lean_object* x_3920; uint8_t x_3921; +x_3917 = lean_ctor_get(x_3912, 0); +lean_inc(x_3917); +x_3918 = lean_ctor_get(x_3912, 1); +lean_inc(x_3918); +if (lean_is_exclusive(x_3912)) { + lean_ctor_release(x_3912, 0); + lean_ctor_release(x_3912, 1); + x_3919 = x_3912; +} else { + lean_dec_ref(x_3912); + x_3919 = lean_box(0); +} +x_3920 = lean_ctor_get(x_3917, 1); +lean_inc(x_3920); +x_3921 = lean_nat_dec_eq(x_3906, x_3920); +lean_dec(x_3906); +if (x_3921 == 0) +{ +lean_object* x_3922; +lean_dec(x_3920); +if (lean_is_scalar(x_3919)) { + x_3922 = lean_alloc_ctor(1, 2, 0); +} else { + x_3922 = x_3919; +} +lean_ctor_set(x_3922, 0, x_3917); +lean_ctor_set(x_3922, 1, x_3918); +return x_3922; +} +else +{ +lean_object* x_3923; lean_object* x_3924; lean_object* x_3925; lean_object* x_3926; +lean_dec(x_3919); +lean_dec(x_3918); +x_3923 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_3924 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3924, 0, x_3923); +x_3925 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_3926 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3925, x_3924, x_3917); +if (lean_obj_tag(x_3926) == 0) +{ +lean_object* x_3927; lean_object* x_3928; lean_object* x_3929; lean_object* x_3930; +lean_dec(x_3920); +x_3927 = lean_ctor_get(x_3926, 0); +lean_inc(x_3927); +x_3928 = lean_ctor_get(x_3926, 1); +lean_inc(x_3928); +if (lean_is_exclusive(x_3926)) { + lean_ctor_release(x_3926, 0); + lean_ctor_release(x_3926, 1); + x_3929 = x_3926; +} else { + lean_dec_ref(x_3926); + x_3929 = lean_box(0); +} +if (lean_is_scalar(x_3929)) { + x_3930 = lean_alloc_ctor(0, 2, 0); +} else { + x_3930 = x_3929; +} +lean_ctor_set(x_3930, 0, x_3927); +lean_ctor_set(x_3930, 1, x_3928); +return x_3930; +} +else +{ +lean_object* x_3931; lean_object* x_3932; lean_object* x_3933; lean_object* x_3934; uint8_t x_3935; +x_3931 = lean_ctor_get(x_3926, 0); +lean_inc(x_3931); +x_3932 = lean_ctor_get(x_3926, 1); +lean_inc(x_3932); +if (lean_is_exclusive(x_3926)) { + lean_ctor_release(x_3926, 0); + lean_ctor_release(x_3926, 1); + x_3933 = x_3926; +} else { + lean_dec_ref(x_3926); + x_3933 = lean_box(0); +} +x_3934 = lean_ctor_get(x_3931, 1); +lean_inc(x_3934); +x_3935 = lean_nat_dec_eq(x_3920, x_3934); +lean_dec(x_3920); +if (x_3935 == 0) +{ +lean_object* x_3936; +lean_dec(x_3934); +if (lean_is_scalar(x_3933)) { + x_3936 = lean_alloc_ctor(1, 2, 0); +} else { + x_3936 = x_3933; +} +lean_ctor_set(x_3936, 0, x_3931); +lean_ctor_set(x_3936, 1, x_3932); +return x_3936; +} +else +{ +lean_object* x_3937; lean_object* x_3938; lean_object* x_3939; lean_object* x_3940; +lean_dec(x_3933); +lean_dec(x_3932); +x_3937 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_3938 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3938, 0, x_3937); +x_3939 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_3940 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3939, x_3938, x_3931); +if (lean_obj_tag(x_3940) == 0) +{ +lean_object* x_3941; lean_object* x_3942; lean_object* x_3943; lean_object* x_3944; +lean_dec(x_3934); +x_3941 = lean_ctor_get(x_3940, 0); +lean_inc(x_3941); +x_3942 = lean_ctor_get(x_3940, 1); +lean_inc(x_3942); +if (lean_is_exclusive(x_3940)) { + lean_ctor_release(x_3940, 0); + lean_ctor_release(x_3940, 1); + x_3943 = x_3940; +} else { + lean_dec_ref(x_3940); + x_3943 = lean_box(0); +} +if (lean_is_scalar(x_3943)) { + x_3944 = lean_alloc_ctor(0, 2, 0); +} else { + x_3944 = x_3943; +} +lean_ctor_set(x_3944, 0, x_3941); +lean_ctor_set(x_3944, 1, x_3942); +return x_3944; +} +else +{ +lean_object* x_3945; lean_object* x_3946; lean_object* x_3947; lean_object* x_3948; uint8_t x_3949; +x_3945 = lean_ctor_get(x_3940, 0); +lean_inc(x_3945); +x_3946 = lean_ctor_get(x_3940, 1); +lean_inc(x_3946); +if (lean_is_exclusive(x_3940)) { + lean_ctor_release(x_3940, 0); + lean_ctor_release(x_3940, 1); + x_3947 = x_3940; +} else { + lean_dec_ref(x_3940); + x_3947 = lean_box(0); +} +x_3948 = lean_ctor_get(x_3945, 1); +lean_inc(x_3948); +x_3949 = lean_nat_dec_eq(x_3934, x_3948); +lean_dec(x_3934); +if (x_3949 == 0) +{ +lean_object* x_3950; +lean_dec(x_3948); +if (lean_is_scalar(x_3947)) { + x_3950 = lean_alloc_ctor(1, 2, 0); +} else { + x_3950 = x_3947; +} +lean_ctor_set(x_3950, 0, x_3945); +lean_ctor_set(x_3950, 1, x_3946); +return x_3950; +} +else +{ +lean_object* x_3951; lean_object* x_3952; lean_object* x_3953; lean_object* x_3954; +lean_dec(x_3947); +lean_dec(x_3946); +x_3951 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_3952 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3952, 0, x_3951); +x_3953 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_3954 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3953, x_3952, x_3945); +if (lean_obj_tag(x_3954) == 0) +{ +lean_object* x_3955; lean_object* x_3956; lean_object* x_3957; lean_object* x_3958; +lean_dec(x_3948); +x_3955 = lean_ctor_get(x_3954, 0); +lean_inc(x_3955); +x_3956 = lean_ctor_get(x_3954, 1); +lean_inc(x_3956); +if (lean_is_exclusive(x_3954)) { + lean_ctor_release(x_3954, 0); + lean_ctor_release(x_3954, 1); + x_3957 = x_3954; +} else { + lean_dec_ref(x_3954); + x_3957 = lean_box(0); +} +if (lean_is_scalar(x_3957)) { + x_3958 = lean_alloc_ctor(0, 2, 0); +} else { + x_3958 = x_3957; +} +lean_ctor_set(x_3958, 0, x_3955); +lean_ctor_set(x_3958, 1, x_3956); +return x_3958; +} +else +{ +lean_object* x_3959; lean_object* x_3960; lean_object* x_3961; lean_object* x_3962; uint8_t x_3963; +x_3959 = lean_ctor_get(x_3954, 0); +lean_inc(x_3959); +x_3960 = lean_ctor_get(x_3954, 1); +lean_inc(x_3960); +if (lean_is_exclusive(x_3954)) { + lean_ctor_release(x_3954, 0); + lean_ctor_release(x_3954, 1); + x_3961 = x_3954; +} else { + lean_dec_ref(x_3954); + x_3961 = lean_box(0); +} +x_3962 = lean_ctor_get(x_3959, 1); +lean_inc(x_3962); +x_3963 = lean_nat_dec_eq(x_3948, x_3962); +lean_dec(x_3948); +if (x_3963 == 0) +{ +lean_object* x_3964; +lean_dec(x_3962); +if (lean_is_scalar(x_3961)) { + x_3964 = lean_alloc_ctor(1, 2, 0); +} else { + x_3964 = x_3961; +} +lean_ctor_set(x_3964, 0, x_3959); +lean_ctor_set(x_3964, 1, x_3960); +return x_3964; +} +else +{ +lean_object* x_3965; lean_object* x_3966; lean_object* x_3967; lean_object* x_3968; +lean_dec(x_3961); +lean_dec(x_3960); +x_3965 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_3966 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3966, 0, x_3965); +x_3967 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_3968 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3967, x_3966, x_3959); +if (lean_obj_tag(x_3968) == 0) +{ +lean_object* x_3969; lean_object* x_3970; lean_object* x_3971; lean_object* x_3972; +lean_dec(x_3962); +x_3969 = lean_ctor_get(x_3968, 0); +lean_inc(x_3969); +x_3970 = lean_ctor_get(x_3968, 1); +lean_inc(x_3970); +if (lean_is_exclusive(x_3968)) { + lean_ctor_release(x_3968, 0); + lean_ctor_release(x_3968, 1); + x_3971 = x_3968; +} else { + lean_dec_ref(x_3968); + x_3971 = lean_box(0); +} +if (lean_is_scalar(x_3971)) { + x_3972 = lean_alloc_ctor(0, 2, 0); +} else { + x_3972 = x_3971; +} +lean_ctor_set(x_3972, 0, x_3969); +lean_ctor_set(x_3972, 1, x_3970); +return x_3972; +} +else +{ +lean_object* x_3973; lean_object* x_3974; lean_object* x_3975; lean_object* x_3976; uint8_t x_3977; +x_3973 = lean_ctor_get(x_3968, 0); +lean_inc(x_3973); +x_3974 = lean_ctor_get(x_3968, 1); +lean_inc(x_3974); +if (lean_is_exclusive(x_3968)) { + lean_ctor_release(x_3968, 0); + lean_ctor_release(x_3968, 1); + x_3975 = x_3968; +} else { + lean_dec_ref(x_3968); + x_3975 = lean_box(0); +} +x_3976 = lean_ctor_get(x_3973, 1); +lean_inc(x_3976); +x_3977 = lean_nat_dec_eq(x_3962, x_3976); +lean_dec(x_3962); +if (x_3977 == 0) +{ +lean_object* x_3978; +lean_dec(x_3976); +if (lean_is_scalar(x_3975)) { + x_3978 = lean_alloc_ctor(1, 2, 0); +} else { + x_3978 = x_3975; +} +lean_ctor_set(x_3978, 0, x_3973); +lean_ctor_set(x_3978, 1, x_3974); +return x_3978; +} +else +{ +lean_object* x_3979; lean_object* x_3980; lean_object* x_3981; lean_object* x_3982; +lean_dec(x_3975); +lean_dec(x_3974); +x_3979 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_3980 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3980, 0, x_3979); +x_3981 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_3982 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3981, x_3980, x_3973); +if (lean_obj_tag(x_3982) == 0) +{ +lean_object* x_3983; lean_object* x_3984; lean_object* x_3985; lean_object* x_3986; +lean_dec(x_3976); +x_3983 = lean_ctor_get(x_3982, 0); +lean_inc(x_3983); +x_3984 = lean_ctor_get(x_3982, 1); +lean_inc(x_3984); +if (lean_is_exclusive(x_3982)) { + lean_ctor_release(x_3982, 0); + lean_ctor_release(x_3982, 1); + x_3985 = x_3982; +} else { + lean_dec_ref(x_3982); + x_3985 = lean_box(0); +} +if (lean_is_scalar(x_3985)) { + x_3986 = lean_alloc_ctor(0, 2, 0); +} else { + x_3986 = x_3985; +} +lean_ctor_set(x_3986, 0, x_3983); +lean_ctor_set(x_3986, 1, x_3984); +return x_3986; +} +else +{ +lean_object* x_3987; lean_object* x_3988; lean_object* x_3989; lean_object* x_3990; uint8_t x_3991; +x_3987 = lean_ctor_get(x_3982, 0); +lean_inc(x_3987); +x_3988 = lean_ctor_get(x_3982, 1); +lean_inc(x_3988); +if (lean_is_exclusive(x_3982)) { + lean_ctor_release(x_3982, 0); + lean_ctor_release(x_3982, 1); + x_3989 = x_3982; +} else { + lean_dec_ref(x_3982); + x_3989 = lean_box(0); +} +x_3990 = lean_ctor_get(x_3987, 1); +lean_inc(x_3990); +x_3991 = lean_nat_dec_eq(x_3976, x_3990); +lean_dec(x_3976); +if (x_3991 == 0) +{ +lean_object* x_3992; +lean_dec(x_3990); +if (lean_is_scalar(x_3989)) { + x_3992 = lean_alloc_ctor(1, 2, 0); +} else { + x_3992 = x_3989; +} +lean_ctor_set(x_3992, 0, x_3987); +lean_ctor_set(x_3992, 1, x_3988); +return x_3992; +} +else +{ +lean_object* x_3993; lean_object* x_3994; lean_object* x_3995; lean_object* x_3996; +lean_dec(x_3989); +lean_dec(x_3988); +x_3993 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_3994 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_3994, 0, x_3993); +x_3995 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_3996 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_3995, x_3994, x_3987); +if (lean_obj_tag(x_3996) == 0) +{ +lean_object* x_3997; lean_object* x_3998; lean_object* x_3999; lean_object* x_4000; +lean_dec(x_3990); +x_3997 = lean_ctor_get(x_3996, 0); +lean_inc(x_3997); +x_3998 = lean_ctor_get(x_3996, 1); +lean_inc(x_3998); +if (lean_is_exclusive(x_3996)) { + lean_ctor_release(x_3996, 0); + lean_ctor_release(x_3996, 1); + x_3999 = x_3996; +} else { + lean_dec_ref(x_3996); + x_3999 = lean_box(0); +} +if (lean_is_scalar(x_3999)) { + x_4000 = lean_alloc_ctor(0, 2, 0); +} else { + x_4000 = x_3999; +} +lean_ctor_set(x_4000, 0, x_3997); +lean_ctor_set(x_4000, 1, x_3998); +return x_4000; +} +else +{ +lean_object* x_4001; lean_object* x_4002; lean_object* x_4003; lean_object* x_4004; uint8_t x_4005; +x_4001 = lean_ctor_get(x_3996, 0); +lean_inc(x_4001); +x_4002 = lean_ctor_get(x_3996, 1); +lean_inc(x_4002); +if (lean_is_exclusive(x_3996)) { + lean_ctor_release(x_3996, 0); + lean_ctor_release(x_3996, 1); + x_4003 = x_3996; +} else { + lean_dec_ref(x_3996); + x_4003 = lean_box(0); +} +x_4004 = lean_ctor_get(x_4001, 1); +lean_inc(x_4004); +x_4005 = lean_nat_dec_eq(x_3990, x_4004); +lean_dec(x_3990); +if (x_4005 == 0) +{ +lean_object* x_4006; +lean_dec(x_4004); +if (lean_is_scalar(x_4003)) { + x_4006 = lean_alloc_ctor(1, 2, 0); +} else { + x_4006 = x_4003; +} +lean_ctor_set(x_4006, 0, x_4001); +lean_ctor_set(x_4006, 1, x_4002); +return x_4006; +} +else +{ +lean_object* x_4007; lean_object* x_4008; lean_object* x_4009; lean_object* x_4010; +lean_dec(x_4003); +lean_dec(x_4002); +x_4007 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_4008 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4008, 0, x_4007); +x_4009 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_4010 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4009, x_4008, x_4001); +if (lean_obj_tag(x_4010) == 0) +{ +lean_object* x_4011; lean_object* x_4012; lean_object* x_4013; lean_object* x_4014; +lean_dec(x_4004); +x_4011 = lean_ctor_get(x_4010, 0); +lean_inc(x_4011); +x_4012 = lean_ctor_get(x_4010, 1); +lean_inc(x_4012); +if (lean_is_exclusive(x_4010)) { + lean_ctor_release(x_4010, 0); + lean_ctor_release(x_4010, 1); + x_4013 = x_4010; +} else { + lean_dec_ref(x_4010); + x_4013 = lean_box(0); +} +if (lean_is_scalar(x_4013)) { + x_4014 = lean_alloc_ctor(0, 2, 0); +} else { + x_4014 = x_4013; +} +lean_ctor_set(x_4014, 0, x_4011); +lean_ctor_set(x_4014, 1, x_4012); +return x_4014; +} +else +{ +lean_object* x_4015; lean_object* x_4016; lean_object* x_4017; lean_object* x_4018; uint8_t x_4019; +x_4015 = lean_ctor_get(x_4010, 0); +lean_inc(x_4015); +x_4016 = lean_ctor_get(x_4010, 1); +lean_inc(x_4016); +if (lean_is_exclusive(x_4010)) { + lean_ctor_release(x_4010, 0); + lean_ctor_release(x_4010, 1); + x_4017 = x_4010; +} else { + lean_dec_ref(x_4010); + x_4017 = lean_box(0); +} +x_4018 = lean_ctor_get(x_4015, 1); +lean_inc(x_4018); +x_4019 = lean_nat_dec_eq(x_4004, x_4018); +lean_dec(x_4004); +if (x_4019 == 0) +{ +lean_object* x_4020; +lean_dec(x_4018); +if (lean_is_scalar(x_4017)) { + x_4020 = lean_alloc_ctor(1, 2, 0); +} else { + x_4020 = x_4017; +} +lean_ctor_set(x_4020, 0, x_4015); +lean_ctor_set(x_4020, 1, x_4016); +return x_4020; +} +else +{ +lean_object* x_4021; lean_object* x_4022; lean_object* x_4023; lean_object* x_4024; +lean_dec(x_4017); +lean_dec(x_4016); +x_4021 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_4022 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4022, 0, x_4021); +x_4023 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_4024 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4023, x_4022, x_4015); +if (lean_obj_tag(x_4024) == 0) +{ +lean_object* x_4025; lean_object* x_4026; lean_object* x_4027; lean_object* x_4028; +lean_dec(x_4018); +x_4025 = lean_ctor_get(x_4024, 0); +lean_inc(x_4025); +x_4026 = lean_ctor_get(x_4024, 1); +lean_inc(x_4026); +if (lean_is_exclusive(x_4024)) { + lean_ctor_release(x_4024, 0); + lean_ctor_release(x_4024, 1); + x_4027 = x_4024; +} else { + lean_dec_ref(x_4024); + x_4027 = lean_box(0); +} +if (lean_is_scalar(x_4027)) { + x_4028 = lean_alloc_ctor(0, 2, 0); +} else { + x_4028 = x_4027; +} +lean_ctor_set(x_4028, 0, x_4025); +lean_ctor_set(x_4028, 1, x_4026); +return x_4028; +} +else +{ +lean_object* x_4029; lean_object* x_4030; lean_object* x_4031; lean_object* x_4032; uint8_t x_4033; +x_4029 = lean_ctor_get(x_4024, 0); +lean_inc(x_4029); +x_4030 = lean_ctor_get(x_4024, 1); +lean_inc(x_4030); +if (lean_is_exclusive(x_4024)) { + lean_ctor_release(x_4024, 0); + lean_ctor_release(x_4024, 1); + x_4031 = x_4024; +} else { + lean_dec_ref(x_4024); + x_4031 = lean_box(0); +} +x_4032 = lean_ctor_get(x_4029, 1); +lean_inc(x_4032); +x_4033 = lean_nat_dec_eq(x_4018, x_4032); +lean_dec(x_4018); +if (x_4033 == 0) +{ +lean_object* x_4034; +lean_dec(x_4032); +if (lean_is_scalar(x_4031)) { + x_4034 = lean_alloc_ctor(1, 2, 0); +} else { + x_4034 = x_4031; +} +lean_ctor_set(x_4034, 0, x_4029); +lean_ctor_set(x_4034, 1, x_4030); +return x_4034; +} +else +{ +lean_object* x_4035; lean_object* x_4036; lean_object* x_4037; lean_object* x_4038; +lean_dec(x_4031); +lean_dec(x_4030); +x_4035 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_4036 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4036, 0, x_4035); +x_4037 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_4038 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4037, x_4036, x_4029); +if (lean_obj_tag(x_4038) == 0) +{ +lean_object* x_4039; lean_object* x_4040; lean_object* x_4041; lean_object* x_4042; +lean_dec(x_4032); +x_4039 = lean_ctor_get(x_4038, 0); +lean_inc(x_4039); +x_4040 = lean_ctor_get(x_4038, 1); +lean_inc(x_4040); +if (lean_is_exclusive(x_4038)) { + lean_ctor_release(x_4038, 0); + lean_ctor_release(x_4038, 1); + x_4041 = x_4038; +} else { + lean_dec_ref(x_4038); + x_4041 = lean_box(0); +} +if (lean_is_scalar(x_4041)) { + x_4042 = lean_alloc_ctor(0, 2, 0); +} else { + x_4042 = x_4041; +} +lean_ctor_set(x_4042, 0, x_4039); +lean_ctor_set(x_4042, 1, x_4040); +return x_4042; +} +else +{ +lean_object* x_4043; lean_object* x_4044; lean_object* x_4045; lean_object* x_4046; uint8_t x_4047; +x_4043 = lean_ctor_get(x_4038, 0); +lean_inc(x_4043); +x_4044 = lean_ctor_get(x_4038, 1); +lean_inc(x_4044); +if (lean_is_exclusive(x_4038)) { + lean_ctor_release(x_4038, 0); + lean_ctor_release(x_4038, 1); + x_4045 = x_4038; +} else { + lean_dec_ref(x_4038); + x_4045 = lean_box(0); +} +x_4046 = lean_ctor_get(x_4043, 1); +lean_inc(x_4046); +x_4047 = lean_nat_dec_eq(x_4032, x_4046); +lean_dec(x_4032); +if (x_4047 == 0) +{ +lean_object* x_4048; +lean_dec(x_4046); +if (lean_is_scalar(x_4045)) { + x_4048 = lean_alloc_ctor(1, 2, 0); +} else { + x_4048 = x_4045; +} +lean_ctor_set(x_4048, 0, x_4043); +lean_ctor_set(x_4048, 1, x_4044); +return x_4048; +} +else +{ +lean_object* x_4049; lean_object* x_4050; lean_object* x_4051; lean_object* x_4052; +lean_dec(x_4045); +lean_dec(x_4044); +x_4049 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_4050 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4050, 0, x_4049); +x_4051 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_4052 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4051, x_4050, x_4043); +if (lean_obj_tag(x_4052) == 0) +{ +lean_object* x_4053; lean_object* x_4054; lean_object* x_4055; lean_object* x_4056; +lean_dec(x_4046); +x_4053 = lean_ctor_get(x_4052, 0); +lean_inc(x_4053); +x_4054 = lean_ctor_get(x_4052, 1); +lean_inc(x_4054); +if (lean_is_exclusive(x_4052)) { + lean_ctor_release(x_4052, 0); + lean_ctor_release(x_4052, 1); + x_4055 = x_4052; +} else { + lean_dec_ref(x_4052); + x_4055 = lean_box(0); +} +if (lean_is_scalar(x_4055)) { + x_4056 = lean_alloc_ctor(0, 2, 0); +} else { + x_4056 = x_4055; +} +lean_ctor_set(x_4056, 0, x_4053); +lean_ctor_set(x_4056, 1, x_4054); +return x_4056; +} +else +{ +lean_object* x_4057; lean_object* x_4058; lean_object* x_4059; lean_object* x_4060; uint8_t x_4061; +x_4057 = lean_ctor_get(x_4052, 0); +lean_inc(x_4057); +x_4058 = lean_ctor_get(x_4052, 1); +lean_inc(x_4058); +if (lean_is_exclusive(x_4052)) { + lean_ctor_release(x_4052, 0); + lean_ctor_release(x_4052, 1); + x_4059 = x_4052; +} else { + lean_dec_ref(x_4052); + x_4059 = lean_box(0); +} +x_4060 = lean_ctor_get(x_4057, 1); +lean_inc(x_4060); +x_4061 = lean_nat_dec_eq(x_4046, x_4060); +lean_dec(x_4046); +if (x_4061 == 0) +{ +lean_object* x_4062; +lean_dec(x_4060); +if (lean_is_scalar(x_4059)) { + x_4062 = lean_alloc_ctor(1, 2, 0); +} else { + x_4062 = x_4059; +} +lean_ctor_set(x_4062, 0, x_4057); +lean_ctor_set(x_4062, 1, x_4058); +return x_4062; +} +else +{ +lean_object* x_4063; lean_object* x_4064; lean_object* x_4065; lean_object* x_4066; +lean_dec(x_4059); +lean_dec(x_4058); +x_4063 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_4064 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4064, 0, x_4063); +x_4065 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_4066 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4065, x_4064, x_4057); +if (lean_obj_tag(x_4066) == 0) +{ +lean_object* x_4067; lean_object* x_4068; lean_object* x_4069; lean_object* x_4070; +lean_dec(x_4060); +x_4067 = lean_ctor_get(x_4066, 0); +lean_inc(x_4067); +x_4068 = lean_ctor_get(x_4066, 1); +lean_inc(x_4068); +if (lean_is_exclusive(x_4066)) { + lean_ctor_release(x_4066, 0); + lean_ctor_release(x_4066, 1); + x_4069 = x_4066; +} else { + lean_dec_ref(x_4066); + x_4069 = lean_box(0); +} +if (lean_is_scalar(x_4069)) { + x_4070 = lean_alloc_ctor(0, 2, 0); +} else { + x_4070 = x_4069; +} +lean_ctor_set(x_4070, 0, x_4067); +lean_ctor_set(x_4070, 1, x_4068); +return x_4070; +} +else +{ +lean_object* x_4071; lean_object* x_4072; lean_object* x_4073; lean_object* x_4074; uint8_t x_4075; +x_4071 = lean_ctor_get(x_4066, 0); +lean_inc(x_4071); +x_4072 = lean_ctor_get(x_4066, 1); +lean_inc(x_4072); +if (lean_is_exclusive(x_4066)) { + lean_ctor_release(x_4066, 0); + lean_ctor_release(x_4066, 1); + x_4073 = x_4066; +} else { + lean_dec_ref(x_4066); + x_4073 = lean_box(0); +} +x_4074 = lean_ctor_get(x_4071, 1); +lean_inc(x_4074); +x_4075 = lean_nat_dec_eq(x_4060, x_4074); +lean_dec(x_4060); +if (x_4075 == 0) +{ +lean_object* x_4076; +lean_dec(x_4074); +if (lean_is_scalar(x_4073)) { + x_4076 = lean_alloc_ctor(1, 2, 0); +} else { + x_4076 = x_4073; +} +lean_ctor_set(x_4076, 0, x_4071); +lean_ctor_set(x_4076, 1, x_4072); +return x_4076; +} +else +{ +lean_object* x_4077; lean_object* x_4078; lean_object* x_4079; lean_object* x_4080; +lean_dec(x_4073); +lean_dec(x_4072); +x_4077 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_4078 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4078, 0, x_4077); +x_4079 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_4080 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4079, x_4078, x_4071); +if (lean_obj_tag(x_4080) == 0) +{ +lean_object* x_4081; lean_object* x_4082; lean_object* x_4083; lean_object* x_4084; +lean_dec(x_4074); +x_4081 = lean_ctor_get(x_4080, 0); +lean_inc(x_4081); +x_4082 = lean_ctor_get(x_4080, 1); +lean_inc(x_4082); +if (lean_is_exclusive(x_4080)) { + lean_ctor_release(x_4080, 0); + lean_ctor_release(x_4080, 1); + x_4083 = x_4080; +} else { + lean_dec_ref(x_4080); + x_4083 = lean_box(0); +} +if (lean_is_scalar(x_4083)) { + x_4084 = lean_alloc_ctor(0, 2, 0); +} else { + x_4084 = x_4083; +} +lean_ctor_set(x_4084, 0, x_4081); +lean_ctor_set(x_4084, 1, x_4082); +return x_4084; +} +else +{ +lean_object* x_4085; lean_object* x_4086; lean_object* x_4087; lean_object* x_4088; uint8_t x_4089; +x_4085 = lean_ctor_get(x_4080, 0); +lean_inc(x_4085); +x_4086 = lean_ctor_get(x_4080, 1); +lean_inc(x_4086); +if (lean_is_exclusive(x_4080)) { + lean_ctor_release(x_4080, 0); + lean_ctor_release(x_4080, 1); + x_4087 = x_4080; +} else { + lean_dec_ref(x_4080); + x_4087 = lean_box(0); +} +x_4088 = lean_ctor_get(x_4085, 1); +lean_inc(x_4088); +x_4089 = lean_nat_dec_eq(x_4074, x_4088); +lean_dec(x_4074); +if (x_4089 == 0) +{ +lean_object* x_4090; +lean_dec(x_4088); +if (lean_is_scalar(x_4087)) { + x_4090 = lean_alloc_ctor(1, 2, 0); +} else { + x_4090 = x_4087; +} +lean_ctor_set(x_4090, 0, x_4085); +lean_ctor_set(x_4090, 1, x_4086); +return x_4090; +} +else +{ +lean_object* x_4091; lean_object* x_4092; lean_object* x_4093; lean_object* x_4094; +lean_dec(x_4087); +lean_dec(x_4086); +x_4091 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_4092 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4092, 0, x_4091); +x_4093 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_4094 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4093, x_4092, x_4085); +if (lean_obj_tag(x_4094) == 0) +{ +lean_object* x_4095; lean_object* x_4096; lean_object* x_4097; lean_object* x_4098; +lean_dec(x_4088); +x_4095 = lean_ctor_get(x_4094, 0); +lean_inc(x_4095); +x_4096 = lean_ctor_get(x_4094, 1); +lean_inc(x_4096); +if (lean_is_exclusive(x_4094)) { + lean_ctor_release(x_4094, 0); + lean_ctor_release(x_4094, 1); + x_4097 = x_4094; +} else { + lean_dec_ref(x_4094); + x_4097 = lean_box(0); +} +if (lean_is_scalar(x_4097)) { + x_4098 = lean_alloc_ctor(0, 2, 0); +} else { + x_4098 = x_4097; +} +lean_ctor_set(x_4098, 0, x_4095); +lean_ctor_set(x_4098, 1, x_4096); +return x_4098; +} +else +{ +lean_object* x_4099; lean_object* x_4100; lean_object* x_4101; lean_object* x_4102; uint8_t x_4103; +x_4099 = lean_ctor_get(x_4094, 0); +lean_inc(x_4099); +x_4100 = lean_ctor_get(x_4094, 1); +lean_inc(x_4100); +if (lean_is_exclusive(x_4094)) { + lean_ctor_release(x_4094, 0); + lean_ctor_release(x_4094, 1); + x_4101 = x_4094; +} else { + lean_dec_ref(x_4094); + x_4101 = lean_box(0); +} +x_4102 = lean_ctor_get(x_4099, 1); +lean_inc(x_4102); +x_4103 = lean_nat_dec_eq(x_4088, x_4102); +lean_dec(x_4088); +if (x_4103 == 0) +{ +lean_object* x_4104; +lean_dec(x_4102); +if (lean_is_scalar(x_4101)) { + x_4104 = lean_alloc_ctor(1, 2, 0); +} else { + x_4104 = x_4101; +} +lean_ctor_set(x_4104, 0, x_4099); +lean_ctor_set(x_4104, 1, x_4100); +return x_4104; +} +else +{ +lean_object* x_4105; lean_object* x_4106; lean_object* x_4107; lean_object* x_4108; +lean_dec(x_4101); +lean_dec(x_4100); +x_4105 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_4106 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4106, 0, x_4105); +x_4107 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_4108 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4107, x_4106, x_4099); +if (lean_obj_tag(x_4108) == 0) +{ +lean_object* x_4109; lean_object* x_4110; lean_object* x_4111; lean_object* x_4112; +lean_dec(x_4102); +x_4109 = lean_ctor_get(x_4108, 0); +lean_inc(x_4109); +x_4110 = lean_ctor_get(x_4108, 1); +lean_inc(x_4110); +if (lean_is_exclusive(x_4108)) { + lean_ctor_release(x_4108, 0); + lean_ctor_release(x_4108, 1); + x_4111 = x_4108; +} else { + lean_dec_ref(x_4108); + x_4111 = lean_box(0); +} +if (lean_is_scalar(x_4111)) { + x_4112 = lean_alloc_ctor(0, 2, 0); +} else { + x_4112 = x_4111; +} +lean_ctor_set(x_4112, 0, x_4109); +lean_ctor_set(x_4112, 1, x_4110); +return x_4112; +} +else +{ +lean_object* x_4113; lean_object* x_4114; lean_object* x_4115; lean_object* x_4116; uint8_t x_4117; +x_4113 = lean_ctor_get(x_4108, 0); +lean_inc(x_4113); +x_4114 = lean_ctor_get(x_4108, 1); +lean_inc(x_4114); +if (lean_is_exclusive(x_4108)) { + lean_ctor_release(x_4108, 0); + lean_ctor_release(x_4108, 1); + x_4115 = x_4108; +} else { + lean_dec_ref(x_4108); + x_4115 = lean_box(0); +} +x_4116 = lean_ctor_get(x_4113, 1); +lean_inc(x_4116); +x_4117 = lean_nat_dec_eq(x_4102, x_4116); +lean_dec(x_4102); +if (x_4117 == 0) +{ +lean_object* x_4118; +lean_dec(x_4116); +if (lean_is_scalar(x_4115)) { + x_4118 = lean_alloc_ctor(1, 2, 0); +} else { + x_4118 = x_4115; +} +lean_ctor_set(x_4118, 0, x_4113); +lean_ctor_set(x_4118, 1, x_4114); +return x_4118; +} +else +{ +lean_object* x_4119; lean_object* x_4120; lean_object* x_4121; lean_object* x_4122; +lean_dec(x_4115); +lean_dec(x_4114); +x_4119 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_4120 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4120, 0, x_4119); +x_4121 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_4122 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4121, x_4120, x_4113); +if (lean_obj_tag(x_4122) == 0) +{ +lean_object* x_4123; lean_object* x_4124; lean_object* x_4125; lean_object* x_4126; +lean_dec(x_4116); +x_4123 = lean_ctor_get(x_4122, 0); +lean_inc(x_4123); +x_4124 = lean_ctor_get(x_4122, 1); +lean_inc(x_4124); +if (lean_is_exclusive(x_4122)) { + lean_ctor_release(x_4122, 0); + lean_ctor_release(x_4122, 1); + x_4125 = x_4122; +} else { + lean_dec_ref(x_4122); + x_4125 = lean_box(0); +} +if (lean_is_scalar(x_4125)) { + x_4126 = lean_alloc_ctor(0, 2, 0); +} else { + x_4126 = x_4125; +} +lean_ctor_set(x_4126, 0, x_4123); +lean_ctor_set(x_4126, 1, x_4124); +return x_4126; +} +else +{ +lean_object* x_4127; lean_object* x_4128; lean_object* x_4129; lean_object* x_4130; uint8_t x_4131; +x_4127 = lean_ctor_get(x_4122, 0); +lean_inc(x_4127); +x_4128 = lean_ctor_get(x_4122, 1); +lean_inc(x_4128); +if (lean_is_exclusive(x_4122)) { + lean_ctor_release(x_4122, 0); + lean_ctor_release(x_4122, 1); + x_4129 = x_4122; +} else { + lean_dec_ref(x_4122); + x_4129 = lean_box(0); +} +x_4130 = lean_ctor_get(x_4127, 1); +lean_inc(x_4130); +x_4131 = lean_nat_dec_eq(x_4116, x_4130); +lean_dec(x_4116); +if (x_4131 == 0) +{ +lean_object* x_4132; +lean_dec(x_4130); +if (lean_is_scalar(x_4129)) { + x_4132 = lean_alloc_ctor(1, 2, 0); +} else { + x_4132 = x_4129; +} +lean_ctor_set(x_4132, 0, x_4127); +lean_ctor_set(x_4132, 1, x_4128); +return x_4132; +} +else +{ +lean_object* x_4133; lean_object* x_4134; lean_object* x_4135; lean_object* x_4136; +lean_dec(x_4129); +lean_dec(x_4128); +x_4133 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_4134 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4134, 0, x_4133); +x_4135 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_4136 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4135, x_4134, x_4127); +if (lean_obj_tag(x_4136) == 0) +{ +lean_object* x_4137; lean_object* x_4138; lean_object* x_4139; lean_object* x_4140; +lean_dec(x_4130); +x_4137 = lean_ctor_get(x_4136, 0); +lean_inc(x_4137); +x_4138 = lean_ctor_get(x_4136, 1); +lean_inc(x_4138); +if (lean_is_exclusive(x_4136)) { + lean_ctor_release(x_4136, 0); + lean_ctor_release(x_4136, 1); + x_4139 = x_4136; +} else { + lean_dec_ref(x_4136); + x_4139 = lean_box(0); +} +if (lean_is_scalar(x_4139)) { + x_4140 = lean_alloc_ctor(0, 2, 0); +} else { + x_4140 = x_4139; +} +lean_ctor_set(x_4140, 0, x_4137); +lean_ctor_set(x_4140, 1, x_4138); +return x_4140; +} +else +{ +lean_object* x_4141; lean_object* x_4142; lean_object* x_4143; lean_object* x_4144; uint8_t x_4145; +x_4141 = lean_ctor_get(x_4136, 0); +lean_inc(x_4141); +x_4142 = lean_ctor_get(x_4136, 1); +lean_inc(x_4142); +if (lean_is_exclusive(x_4136)) { + lean_ctor_release(x_4136, 0); + lean_ctor_release(x_4136, 1); + x_4143 = x_4136; +} else { + lean_dec_ref(x_4136); + x_4143 = lean_box(0); +} +x_4144 = lean_ctor_get(x_4141, 1); +lean_inc(x_4144); +x_4145 = lean_nat_dec_eq(x_4130, x_4144); +lean_dec(x_4144); +lean_dec(x_4130); +if (x_4145 == 0) +{ +lean_object* x_4146; +if (lean_is_scalar(x_4143)) { + x_4146 = lean_alloc_ctor(1, 2, 0); +} else { + x_4146 = x_4143; +} +lean_ctor_set(x_4146, 0, x_4141); +lean_ctor_set(x_4146, 1, x_4142); +return x_4146; +} +else +{ +lean_object* x_4147; lean_object* x_4148; lean_object* x_4149; lean_object* x_4150; +lean_dec(x_4143); +lean_dec(x_4142); +x_4147 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_4148 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4148, 0, x_4147); +x_4149 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_4150 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4149, x_4148, x_4141); +return x_4150; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_4151; lean_object* x_4152; lean_object* x_4153; uint8_t x_4154; +x_4151 = lean_ctor_get(x_96, 0); +x_4152 = lean_ctor_get(x_96, 1); +lean_inc(x_4152); +lean_inc(x_4151); +lean_dec(x_96); +x_4153 = lean_ctor_get(x_4151, 1); +lean_inc(x_4153); +x_4154 = lean_nat_dec_eq(x_91, x_4153); +lean_dec(x_91); +if (x_4154 == 0) +{ +lean_object* x_4155; +lean_dec(x_4153); +x_4155 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4155, 0, x_4151); +lean_ctor_set(x_4155, 1, x_4152); +return x_4155; +} +else +{ +lean_object* x_4156; lean_object* x_4157; lean_object* x_4158; +lean_dec(x_4152); +x_4156 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_4157 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4157, 0, x_4156); +x_4158 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_95, x_4157, x_4151); +if (lean_obj_tag(x_4158) == 0) +{ +lean_object* x_4159; lean_object* x_4160; lean_object* x_4161; lean_object* x_4162; +lean_dec(x_4153); +x_4159 = lean_ctor_get(x_4158, 0); +lean_inc(x_4159); +x_4160 = lean_ctor_get(x_4158, 1); +lean_inc(x_4160); +if (lean_is_exclusive(x_4158)) { + lean_ctor_release(x_4158, 0); + lean_ctor_release(x_4158, 1); + x_4161 = x_4158; +} else { + lean_dec_ref(x_4158); + x_4161 = lean_box(0); +} +if (lean_is_scalar(x_4161)) { + x_4162 = lean_alloc_ctor(0, 2, 0); +} else { + x_4162 = x_4161; +} +lean_ctor_set(x_4162, 0, x_4159); +lean_ctor_set(x_4162, 1, x_4160); +return x_4162; +} +else +{ +lean_object* x_4163; lean_object* x_4164; lean_object* x_4165; lean_object* x_4166; uint8_t x_4167; +x_4163 = lean_ctor_get(x_4158, 0); +lean_inc(x_4163); +x_4164 = lean_ctor_get(x_4158, 1); +lean_inc(x_4164); +if (lean_is_exclusive(x_4158)) { + lean_ctor_release(x_4158, 0); + lean_ctor_release(x_4158, 1); + x_4165 = x_4158; +} else { + lean_dec_ref(x_4158); + x_4165 = lean_box(0); +} +x_4166 = lean_ctor_get(x_4163, 1); +lean_inc(x_4166); +x_4167 = lean_nat_dec_eq(x_4153, x_4166); +lean_dec(x_4153); +if (x_4167 == 0) +{ +lean_object* x_4168; +lean_dec(x_4166); +if (lean_is_scalar(x_4165)) { + x_4168 = lean_alloc_ctor(1, 2, 0); +} else { + x_4168 = x_4165; +} +lean_ctor_set(x_4168, 0, x_4163); +lean_ctor_set(x_4168, 1, x_4164); +return x_4168; +} +else +{ +lean_object* x_4169; lean_object* x_4170; lean_object* x_4171; lean_object* x_4172; +lean_dec(x_4165); +lean_dec(x_4164); +x_4169 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_4170 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4170, 0, x_4169); +x_4171 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_4172 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4171, x_4170, x_4163); +if (lean_obj_tag(x_4172) == 0) +{ +lean_object* x_4173; lean_object* x_4174; lean_object* x_4175; lean_object* x_4176; +lean_dec(x_4166); +x_4173 = lean_ctor_get(x_4172, 0); +lean_inc(x_4173); +x_4174 = lean_ctor_get(x_4172, 1); +lean_inc(x_4174); +if (lean_is_exclusive(x_4172)) { + lean_ctor_release(x_4172, 0); + lean_ctor_release(x_4172, 1); + x_4175 = x_4172; +} else { + lean_dec_ref(x_4172); + x_4175 = lean_box(0); +} +if (lean_is_scalar(x_4175)) { + x_4176 = lean_alloc_ctor(0, 2, 0); +} else { + x_4176 = x_4175; +} +lean_ctor_set(x_4176, 0, x_4173); +lean_ctor_set(x_4176, 1, x_4174); +return x_4176; +} +else +{ +lean_object* x_4177; lean_object* x_4178; lean_object* x_4179; lean_object* x_4180; uint8_t x_4181; +x_4177 = lean_ctor_get(x_4172, 0); +lean_inc(x_4177); +x_4178 = lean_ctor_get(x_4172, 1); +lean_inc(x_4178); +if (lean_is_exclusive(x_4172)) { + lean_ctor_release(x_4172, 0); + lean_ctor_release(x_4172, 1); + x_4179 = x_4172; +} else { + lean_dec_ref(x_4172); + x_4179 = lean_box(0); +} +x_4180 = lean_ctor_get(x_4177, 1); +lean_inc(x_4180); +x_4181 = lean_nat_dec_eq(x_4166, x_4180); +lean_dec(x_4166); +if (x_4181 == 0) +{ +lean_object* x_4182; +lean_dec(x_4180); +if (lean_is_scalar(x_4179)) { + x_4182 = lean_alloc_ctor(1, 2, 0); +} else { + x_4182 = x_4179; +} +lean_ctor_set(x_4182, 0, x_4177); +lean_ctor_set(x_4182, 1, x_4178); +return x_4182; +} +else +{ +lean_object* x_4183; lean_object* x_4184; lean_object* x_4185; lean_object* x_4186; +lean_dec(x_4179); +lean_dec(x_4178); +x_4183 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_4184 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4184, 0, x_4183); +x_4185 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_4186 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4185, x_4184, x_4177); +if (lean_obj_tag(x_4186) == 0) +{ +lean_object* x_4187; lean_object* x_4188; lean_object* x_4189; lean_object* x_4190; +lean_dec(x_4180); +x_4187 = lean_ctor_get(x_4186, 0); +lean_inc(x_4187); +x_4188 = lean_ctor_get(x_4186, 1); +lean_inc(x_4188); +if (lean_is_exclusive(x_4186)) { + lean_ctor_release(x_4186, 0); + lean_ctor_release(x_4186, 1); + x_4189 = x_4186; +} else { + lean_dec_ref(x_4186); + x_4189 = lean_box(0); +} +if (lean_is_scalar(x_4189)) { + x_4190 = lean_alloc_ctor(0, 2, 0); +} else { + x_4190 = x_4189; +} +lean_ctor_set(x_4190, 0, x_4187); +lean_ctor_set(x_4190, 1, x_4188); +return x_4190; +} +else +{ +lean_object* x_4191; lean_object* x_4192; lean_object* x_4193; lean_object* x_4194; uint8_t x_4195; +x_4191 = lean_ctor_get(x_4186, 0); +lean_inc(x_4191); +x_4192 = lean_ctor_get(x_4186, 1); +lean_inc(x_4192); +if (lean_is_exclusive(x_4186)) { + lean_ctor_release(x_4186, 0); + lean_ctor_release(x_4186, 1); + x_4193 = x_4186; +} else { + lean_dec_ref(x_4186); + x_4193 = lean_box(0); +} +x_4194 = lean_ctor_get(x_4191, 1); +lean_inc(x_4194); +x_4195 = lean_nat_dec_eq(x_4180, x_4194); +lean_dec(x_4180); +if (x_4195 == 0) +{ +lean_object* x_4196; +lean_dec(x_4194); +if (lean_is_scalar(x_4193)) { + x_4196 = lean_alloc_ctor(1, 2, 0); +} else { + x_4196 = x_4193; +} +lean_ctor_set(x_4196, 0, x_4191); +lean_ctor_set(x_4196, 1, x_4192); +return x_4196; +} +else +{ +lean_object* x_4197; lean_object* x_4198; lean_object* x_4199; lean_object* x_4200; +lean_dec(x_4193); +lean_dec(x_4192); +x_4197 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_4198 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4198, 0, x_4197); +x_4199 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_4200 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4199, x_4198, x_4191); +if (lean_obj_tag(x_4200) == 0) +{ +lean_object* x_4201; lean_object* x_4202; lean_object* x_4203; lean_object* x_4204; +lean_dec(x_4194); +x_4201 = lean_ctor_get(x_4200, 0); +lean_inc(x_4201); +x_4202 = lean_ctor_get(x_4200, 1); +lean_inc(x_4202); +if (lean_is_exclusive(x_4200)) { + lean_ctor_release(x_4200, 0); + lean_ctor_release(x_4200, 1); + x_4203 = x_4200; +} else { + lean_dec_ref(x_4200); + x_4203 = lean_box(0); +} +if (lean_is_scalar(x_4203)) { + x_4204 = lean_alloc_ctor(0, 2, 0); +} else { + x_4204 = x_4203; +} +lean_ctor_set(x_4204, 0, x_4201); +lean_ctor_set(x_4204, 1, x_4202); +return x_4204; +} +else +{ +lean_object* x_4205; lean_object* x_4206; lean_object* x_4207; lean_object* x_4208; uint8_t x_4209; +x_4205 = lean_ctor_get(x_4200, 0); +lean_inc(x_4205); +x_4206 = lean_ctor_get(x_4200, 1); +lean_inc(x_4206); +if (lean_is_exclusive(x_4200)) { + lean_ctor_release(x_4200, 0); + lean_ctor_release(x_4200, 1); + x_4207 = x_4200; +} else { + lean_dec_ref(x_4200); + x_4207 = lean_box(0); +} +x_4208 = lean_ctor_get(x_4205, 1); +lean_inc(x_4208); +x_4209 = lean_nat_dec_eq(x_4194, x_4208); +lean_dec(x_4194); +if (x_4209 == 0) +{ +lean_object* x_4210; +lean_dec(x_4208); +if (lean_is_scalar(x_4207)) { + x_4210 = lean_alloc_ctor(1, 2, 0); +} else { + x_4210 = x_4207; +} +lean_ctor_set(x_4210, 0, x_4205); +lean_ctor_set(x_4210, 1, x_4206); +return x_4210; +} +else +{ +lean_object* x_4211; lean_object* x_4212; lean_object* x_4213; lean_object* x_4214; +lean_dec(x_4207); +lean_dec(x_4206); +x_4211 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_4212 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4212, 0, x_4211); +x_4213 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_4214 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4213, x_4212, x_4205); +if (lean_obj_tag(x_4214) == 0) +{ +lean_object* x_4215; lean_object* x_4216; lean_object* x_4217; lean_object* x_4218; +lean_dec(x_4208); +x_4215 = lean_ctor_get(x_4214, 0); +lean_inc(x_4215); +x_4216 = lean_ctor_get(x_4214, 1); +lean_inc(x_4216); +if (lean_is_exclusive(x_4214)) { + lean_ctor_release(x_4214, 0); + lean_ctor_release(x_4214, 1); + x_4217 = x_4214; +} else { + lean_dec_ref(x_4214); + x_4217 = lean_box(0); +} +if (lean_is_scalar(x_4217)) { + x_4218 = lean_alloc_ctor(0, 2, 0); +} else { + x_4218 = x_4217; +} +lean_ctor_set(x_4218, 0, x_4215); +lean_ctor_set(x_4218, 1, x_4216); +return x_4218; +} +else +{ +lean_object* x_4219; lean_object* x_4220; lean_object* x_4221; lean_object* x_4222; uint8_t x_4223; +x_4219 = lean_ctor_get(x_4214, 0); +lean_inc(x_4219); +x_4220 = lean_ctor_get(x_4214, 1); +lean_inc(x_4220); +if (lean_is_exclusive(x_4214)) { + lean_ctor_release(x_4214, 0); + lean_ctor_release(x_4214, 1); + x_4221 = x_4214; +} else { + lean_dec_ref(x_4214); + x_4221 = lean_box(0); +} +x_4222 = lean_ctor_get(x_4219, 1); +lean_inc(x_4222); +x_4223 = lean_nat_dec_eq(x_4208, x_4222); +lean_dec(x_4208); +if (x_4223 == 0) +{ +lean_object* x_4224; +lean_dec(x_4222); +if (lean_is_scalar(x_4221)) { + x_4224 = lean_alloc_ctor(1, 2, 0); +} else { + x_4224 = x_4221; +} +lean_ctor_set(x_4224, 0, x_4219); +lean_ctor_set(x_4224, 1, x_4220); +return x_4224; +} +else +{ +lean_object* x_4225; lean_object* x_4226; lean_object* x_4227; +lean_dec(x_4221); +lean_dec(x_4220); +x_4225 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_4226 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4226, 0, x_4225); +x_4227 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4213, x_4226, x_4219); +if (lean_obj_tag(x_4227) == 0) +{ +lean_object* x_4228; lean_object* x_4229; lean_object* x_4230; lean_object* x_4231; +lean_dec(x_4222); +x_4228 = lean_ctor_get(x_4227, 0); +lean_inc(x_4228); +x_4229 = lean_ctor_get(x_4227, 1); +lean_inc(x_4229); +if (lean_is_exclusive(x_4227)) { + lean_ctor_release(x_4227, 0); + lean_ctor_release(x_4227, 1); + x_4230 = x_4227; +} else { + lean_dec_ref(x_4227); + x_4230 = lean_box(0); +} +if (lean_is_scalar(x_4230)) { + x_4231 = lean_alloc_ctor(0, 2, 0); +} else { + x_4231 = x_4230; +} +lean_ctor_set(x_4231, 0, x_4228); +lean_ctor_set(x_4231, 1, x_4229); +return x_4231; +} +else +{ +lean_object* x_4232; lean_object* x_4233; lean_object* x_4234; lean_object* x_4235; uint8_t x_4236; +x_4232 = lean_ctor_get(x_4227, 0); +lean_inc(x_4232); +x_4233 = lean_ctor_get(x_4227, 1); +lean_inc(x_4233); +if (lean_is_exclusive(x_4227)) { + lean_ctor_release(x_4227, 0); + lean_ctor_release(x_4227, 1); + x_4234 = x_4227; +} else { + lean_dec_ref(x_4227); + x_4234 = lean_box(0); +} +x_4235 = lean_ctor_get(x_4232, 1); +lean_inc(x_4235); +x_4236 = lean_nat_dec_eq(x_4222, x_4235); +lean_dec(x_4222); +if (x_4236 == 0) +{ +lean_object* x_4237; +lean_dec(x_4235); +if (lean_is_scalar(x_4234)) { + x_4237 = lean_alloc_ctor(1, 2, 0); +} else { + x_4237 = x_4234; +} +lean_ctor_set(x_4237, 0, x_4232); +lean_ctor_set(x_4237, 1, x_4233); +return x_4237; +} +else +{ +lean_object* x_4238; lean_object* x_4239; lean_object* x_4240; lean_object* x_4241; +lean_dec(x_4234); +lean_dec(x_4233); +x_4238 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_4239 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4239, 0, x_4238); +x_4240 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_4241 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4240, x_4239, x_4232); +if (lean_obj_tag(x_4241) == 0) +{ +lean_object* x_4242; lean_object* x_4243; lean_object* x_4244; lean_object* x_4245; +lean_dec(x_4235); +x_4242 = lean_ctor_get(x_4241, 0); +lean_inc(x_4242); +x_4243 = lean_ctor_get(x_4241, 1); +lean_inc(x_4243); +if (lean_is_exclusive(x_4241)) { + lean_ctor_release(x_4241, 0); + lean_ctor_release(x_4241, 1); + x_4244 = x_4241; +} else { + lean_dec_ref(x_4241); + x_4244 = lean_box(0); +} +if (lean_is_scalar(x_4244)) { + x_4245 = lean_alloc_ctor(0, 2, 0); +} else { + x_4245 = x_4244; +} +lean_ctor_set(x_4245, 0, x_4242); +lean_ctor_set(x_4245, 1, x_4243); +return x_4245; +} +else +{ +lean_object* x_4246; lean_object* x_4247; lean_object* x_4248; lean_object* x_4249; uint8_t x_4250; +x_4246 = lean_ctor_get(x_4241, 0); +lean_inc(x_4246); +x_4247 = lean_ctor_get(x_4241, 1); +lean_inc(x_4247); +if (lean_is_exclusive(x_4241)) { + lean_ctor_release(x_4241, 0); + lean_ctor_release(x_4241, 1); + x_4248 = x_4241; +} else { + lean_dec_ref(x_4241); + x_4248 = lean_box(0); +} +x_4249 = lean_ctor_get(x_4246, 1); +lean_inc(x_4249); +x_4250 = lean_nat_dec_eq(x_4235, x_4249); +lean_dec(x_4235); +if (x_4250 == 0) +{ +lean_object* x_4251; +lean_dec(x_4249); +if (lean_is_scalar(x_4248)) { + x_4251 = lean_alloc_ctor(1, 2, 0); +} else { + x_4251 = x_4248; +} +lean_ctor_set(x_4251, 0, x_4246); +lean_ctor_set(x_4251, 1, x_4247); +return x_4251; +} +else +{ +lean_object* x_4252; lean_object* x_4253; lean_object* x_4254; lean_object* x_4255; +lean_dec(x_4248); +lean_dec(x_4247); +x_4252 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_4253 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4253, 0, x_4252); +x_4254 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_4255 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4254, x_4253, x_4246); +if (lean_obj_tag(x_4255) == 0) +{ +lean_object* x_4256; lean_object* x_4257; lean_object* x_4258; lean_object* x_4259; +lean_dec(x_4249); +x_4256 = lean_ctor_get(x_4255, 0); +lean_inc(x_4256); +x_4257 = lean_ctor_get(x_4255, 1); +lean_inc(x_4257); +if (lean_is_exclusive(x_4255)) { + lean_ctor_release(x_4255, 0); + lean_ctor_release(x_4255, 1); + x_4258 = x_4255; +} else { + lean_dec_ref(x_4255); + x_4258 = lean_box(0); +} +if (lean_is_scalar(x_4258)) { + x_4259 = lean_alloc_ctor(0, 2, 0); +} else { + x_4259 = x_4258; +} +lean_ctor_set(x_4259, 0, x_4256); +lean_ctor_set(x_4259, 1, x_4257); +return x_4259; +} +else +{ +lean_object* x_4260; lean_object* x_4261; lean_object* x_4262; lean_object* x_4263; uint8_t x_4264; +x_4260 = lean_ctor_get(x_4255, 0); +lean_inc(x_4260); +x_4261 = lean_ctor_get(x_4255, 1); +lean_inc(x_4261); +if (lean_is_exclusive(x_4255)) { + lean_ctor_release(x_4255, 0); + lean_ctor_release(x_4255, 1); + x_4262 = x_4255; +} else { + lean_dec_ref(x_4255); + x_4262 = lean_box(0); +} +x_4263 = lean_ctor_get(x_4260, 1); +lean_inc(x_4263); +x_4264 = lean_nat_dec_eq(x_4249, x_4263); +lean_dec(x_4249); +if (x_4264 == 0) +{ +lean_object* x_4265; +lean_dec(x_4263); +if (lean_is_scalar(x_4262)) { + x_4265 = lean_alloc_ctor(1, 2, 0); +} else { + x_4265 = x_4262; +} +lean_ctor_set(x_4265, 0, x_4260); +lean_ctor_set(x_4265, 1, x_4261); +return x_4265; +} +else +{ +lean_object* x_4266; lean_object* x_4267; lean_object* x_4268; lean_object* x_4269; +lean_dec(x_4262); +lean_dec(x_4261); +x_4266 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_4267 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4267, 0, x_4266); +x_4268 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_4269 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4268, x_4267, x_4260); +if (lean_obj_tag(x_4269) == 0) +{ +lean_object* x_4270; lean_object* x_4271; lean_object* x_4272; lean_object* x_4273; +lean_dec(x_4263); +x_4270 = lean_ctor_get(x_4269, 0); +lean_inc(x_4270); +x_4271 = lean_ctor_get(x_4269, 1); +lean_inc(x_4271); +if (lean_is_exclusive(x_4269)) { + lean_ctor_release(x_4269, 0); + lean_ctor_release(x_4269, 1); + x_4272 = x_4269; +} else { + lean_dec_ref(x_4269); + x_4272 = lean_box(0); +} +if (lean_is_scalar(x_4272)) { + x_4273 = lean_alloc_ctor(0, 2, 0); +} else { + x_4273 = x_4272; +} +lean_ctor_set(x_4273, 0, x_4270); +lean_ctor_set(x_4273, 1, x_4271); +return x_4273; +} +else +{ +lean_object* x_4274; lean_object* x_4275; lean_object* x_4276; lean_object* x_4277; uint8_t x_4278; +x_4274 = lean_ctor_get(x_4269, 0); +lean_inc(x_4274); +x_4275 = lean_ctor_get(x_4269, 1); +lean_inc(x_4275); +if (lean_is_exclusive(x_4269)) { + lean_ctor_release(x_4269, 0); + lean_ctor_release(x_4269, 1); + x_4276 = x_4269; +} else { + lean_dec_ref(x_4269); + x_4276 = lean_box(0); +} +x_4277 = lean_ctor_get(x_4274, 1); +lean_inc(x_4277); +x_4278 = lean_nat_dec_eq(x_4263, x_4277); +lean_dec(x_4263); +if (x_4278 == 0) +{ +lean_object* x_4279; +lean_dec(x_4277); +if (lean_is_scalar(x_4276)) { + x_4279 = lean_alloc_ctor(1, 2, 0); +} else { + x_4279 = x_4276; +} +lean_ctor_set(x_4279, 0, x_4274); +lean_ctor_set(x_4279, 1, x_4275); +return x_4279; +} +else +{ +lean_object* x_4280; lean_object* x_4281; lean_object* x_4282; lean_object* x_4283; +lean_dec(x_4276); +lean_dec(x_4275); +x_4280 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_4281 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4281, 0, x_4280); +x_4282 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_4283 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4282, x_4281, x_4274); +if (lean_obj_tag(x_4283) == 0) +{ +lean_object* x_4284; lean_object* x_4285; lean_object* x_4286; lean_object* x_4287; +lean_dec(x_4277); +x_4284 = lean_ctor_get(x_4283, 0); +lean_inc(x_4284); +x_4285 = lean_ctor_get(x_4283, 1); +lean_inc(x_4285); +if (lean_is_exclusive(x_4283)) { + lean_ctor_release(x_4283, 0); + lean_ctor_release(x_4283, 1); + x_4286 = x_4283; +} else { + lean_dec_ref(x_4283); + x_4286 = lean_box(0); +} +if (lean_is_scalar(x_4286)) { + x_4287 = lean_alloc_ctor(0, 2, 0); +} else { + x_4287 = x_4286; +} +lean_ctor_set(x_4287, 0, x_4284); +lean_ctor_set(x_4287, 1, x_4285); +return x_4287; +} +else +{ +lean_object* x_4288; lean_object* x_4289; lean_object* x_4290; lean_object* x_4291; uint8_t x_4292; +x_4288 = lean_ctor_get(x_4283, 0); +lean_inc(x_4288); +x_4289 = lean_ctor_get(x_4283, 1); +lean_inc(x_4289); +if (lean_is_exclusive(x_4283)) { + lean_ctor_release(x_4283, 0); + lean_ctor_release(x_4283, 1); + x_4290 = x_4283; +} else { + lean_dec_ref(x_4283); + x_4290 = lean_box(0); +} +x_4291 = lean_ctor_get(x_4288, 1); +lean_inc(x_4291); +x_4292 = lean_nat_dec_eq(x_4277, x_4291); +lean_dec(x_4277); +if (x_4292 == 0) +{ +lean_object* x_4293; +lean_dec(x_4291); +if (lean_is_scalar(x_4290)) { + x_4293 = lean_alloc_ctor(1, 2, 0); +} else { + x_4293 = x_4290; +} +lean_ctor_set(x_4293, 0, x_4288); +lean_ctor_set(x_4293, 1, x_4289); +return x_4293; +} +else +{ +lean_object* x_4294; lean_object* x_4295; lean_object* x_4296; lean_object* x_4297; +lean_dec(x_4290); +lean_dec(x_4289); +x_4294 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_4295 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4295, 0, x_4294); +x_4296 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_4297 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4296, x_4295, x_4288); +if (lean_obj_tag(x_4297) == 0) +{ +lean_object* x_4298; lean_object* x_4299; lean_object* x_4300; lean_object* x_4301; +lean_dec(x_4291); +x_4298 = lean_ctor_get(x_4297, 0); +lean_inc(x_4298); +x_4299 = lean_ctor_get(x_4297, 1); +lean_inc(x_4299); +if (lean_is_exclusive(x_4297)) { + lean_ctor_release(x_4297, 0); + lean_ctor_release(x_4297, 1); + x_4300 = x_4297; +} else { + lean_dec_ref(x_4297); + x_4300 = lean_box(0); +} +if (lean_is_scalar(x_4300)) { + x_4301 = lean_alloc_ctor(0, 2, 0); +} else { + x_4301 = x_4300; +} +lean_ctor_set(x_4301, 0, x_4298); +lean_ctor_set(x_4301, 1, x_4299); +return x_4301; +} +else +{ +lean_object* x_4302; lean_object* x_4303; lean_object* x_4304; lean_object* x_4305; uint8_t x_4306; +x_4302 = lean_ctor_get(x_4297, 0); +lean_inc(x_4302); +x_4303 = lean_ctor_get(x_4297, 1); +lean_inc(x_4303); +if (lean_is_exclusive(x_4297)) { + lean_ctor_release(x_4297, 0); + lean_ctor_release(x_4297, 1); + x_4304 = x_4297; +} else { + lean_dec_ref(x_4297); + x_4304 = lean_box(0); +} +x_4305 = lean_ctor_get(x_4302, 1); +lean_inc(x_4305); +x_4306 = lean_nat_dec_eq(x_4291, x_4305); +lean_dec(x_4291); +if (x_4306 == 0) +{ +lean_object* x_4307; +lean_dec(x_4305); +if (lean_is_scalar(x_4304)) { + x_4307 = lean_alloc_ctor(1, 2, 0); +} else { + x_4307 = x_4304; +} +lean_ctor_set(x_4307, 0, x_4302); +lean_ctor_set(x_4307, 1, x_4303); +return x_4307; +} +else +{ +lean_object* x_4308; lean_object* x_4309; lean_object* x_4310; lean_object* x_4311; +lean_dec(x_4304); +lean_dec(x_4303); +x_4308 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_4309 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4309, 0, x_4308); +x_4310 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_4311 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4310, x_4309, x_4302); +if (lean_obj_tag(x_4311) == 0) +{ +lean_object* x_4312; lean_object* x_4313; lean_object* x_4314; lean_object* x_4315; +lean_dec(x_4305); +x_4312 = lean_ctor_get(x_4311, 0); +lean_inc(x_4312); +x_4313 = lean_ctor_get(x_4311, 1); +lean_inc(x_4313); +if (lean_is_exclusive(x_4311)) { + lean_ctor_release(x_4311, 0); + lean_ctor_release(x_4311, 1); + x_4314 = x_4311; +} else { + lean_dec_ref(x_4311); + x_4314 = lean_box(0); +} +if (lean_is_scalar(x_4314)) { + x_4315 = lean_alloc_ctor(0, 2, 0); +} else { + x_4315 = x_4314; +} +lean_ctor_set(x_4315, 0, x_4312); +lean_ctor_set(x_4315, 1, x_4313); +return x_4315; +} +else +{ +lean_object* x_4316; lean_object* x_4317; lean_object* x_4318; lean_object* x_4319; uint8_t x_4320; +x_4316 = lean_ctor_get(x_4311, 0); +lean_inc(x_4316); +x_4317 = lean_ctor_get(x_4311, 1); +lean_inc(x_4317); +if (lean_is_exclusive(x_4311)) { + lean_ctor_release(x_4311, 0); + lean_ctor_release(x_4311, 1); + x_4318 = x_4311; +} else { + lean_dec_ref(x_4311); + x_4318 = lean_box(0); +} +x_4319 = lean_ctor_get(x_4316, 1); +lean_inc(x_4319); +x_4320 = lean_nat_dec_eq(x_4305, x_4319); +lean_dec(x_4305); +if (x_4320 == 0) +{ +lean_object* x_4321; +lean_dec(x_4319); +if (lean_is_scalar(x_4318)) { + x_4321 = lean_alloc_ctor(1, 2, 0); +} else { + x_4321 = x_4318; +} +lean_ctor_set(x_4321, 0, x_4316); +lean_ctor_set(x_4321, 1, x_4317); +return x_4321; +} +else +{ +lean_object* x_4322; lean_object* x_4323; lean_object* x_4324; lean_object* x_4325; +lean_dec(x_4318); +lean_dec(x_4317); +x_4322 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_4323 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4323, 0, x_4322); +x_4324 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_4325 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4324, x_4323, x_4316); +if (lean_obj_tag(x_4325) == 0) +{ +lean_object* x_4326; lean_object* x_4327; lean_object* x_4328; lean_object* x_4329; +lean_dec(x_4319); +x_4326 = lean_ctor_get(x_4325, 0); +lean_inc(x_4326); +x_4327 = lean_ctor_get(x_4325, 1); +lean_inc(x_4327); +if (lean_is_exclusive(x_4325)) { + lean_ctor_release(x_4325, 0); + lean_ctor_release(x_4325, 1); + x_4328 = x_4325; +} else { + lean_dec_ref(x_4325); + x_4328 = lean_box(0); +} +if (lean_is_scalar(x_4328)) { + x_4329 = lean_alloc_ctor(0, 2, 0); +} else { + x_4329 = x_4328; +} +lean_ctor_set(x_4329, 0, x_4326); +lean_ctor_set(x_4329, 1, x_4327); +return x_4329; +} +else +{ +lean_object* x_4330; lean_object* x_4331; lean_object* x_4332; lean_object* x_4333; uint8_t x_4334; +x_4330 = lean_ctor_get(x_4325, 0); +lean_inc(x_4330); +x_4331 = lean_ctor_get(x_4325, 1); +lean_inc(x_4331); +if (lean_is_exclusive(x_4325)) { + lean_ctor_release(x_4325, 0); + lean_ctor_release(x_4325, 1); + x_4332 = x_4325; +} else { + lean_dec_ref(x_4325); + x_4332 = lean_box(0); +} +x_4333 = lean_ctor_get(x_4330, 1); +lean_inc(x_4333); +x_4334 = lean_nat_dec_eq(x_4319, x_4333); +lean_dec(x_4319); +if (x_4334 == 0) +{ +lean_object* x_4335; +lean_dec(x_4333); +if (lean_is_scalar(x_4332)) { + x_4335 = lean_alloc_ctor(1, 2, 0); +} else { + x_4335 = x_4332; +} +lean_ctor_set(x_4335, 0, x_4330); +lean_ctor_set(x_4335, 1, x_4331); +return x_4335; +} +else +{ +lean_object* x_4336; lean_object* x_4337; lean_object* x_4338; lean_object* x_4339; +lean_dec(x_4332); +lean_dec(x_4331); +x_4336 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_4337 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4337, 0, x_4336); +x_4338 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_4339 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4338, x_4337, x_4330); +if (lean_obj_tag(x_4339) == 0) +{ +lean_object* x_4340; lean_object* x_4341; lean_object* x_4342; lean_object* x_4343; +lean_dec(x_4333); +x_4340 = lean_ctor_get(x_4339, 0); +lean_inc(x_4340); +x_4341 = lean_ctor_get(x_4339, 1); +lean_inc(x_4341); +if (lean_is_exclusive(x_4339)) { + lean_ctor_release(x_4339, 0); + lean_ctor_release(x_4339, 1); + x_4342 = x_4339; +} else { + lean_dec_ref(x_4339); + x_4342 = lean_box(0); +} +if (lean_is_scalar(x_4342)) { + x_4343 = lean_alloc_ctor(0, 2, 0); +} else { + x_4343 = x_4342; +} +lean_ctor_set(x_4343, 0, x_4340); +lean_ctor_set(x_4343, 1, x_4341); +return x_4343; +} +else +{ +lean_object* x_4344; lean_object* x_4345; lean_object* x_4346; lean_object* x_4347; uint8_t x_4348; +x_4344 = lean_ctor_get(x_4339, 0); +lean_inc(x_4344); +x_4345 = lean_ctor_get(x_4339, 1); +lean_inc(x_4345); +if (lean_is_exclusive(x_4339)) { + lean_ctor_release(x_4339, 0); + lean_ctor_release(x_4339, 1); + x_4346 = x_4339; +} else { + lean_dec_ref(x_4339); + x_4346 = lean_box(0); +} +x_4347 = lean_ctor_get(x_4344, 1); +lean_inc(x_4347); +x_4348 = lean_nat_dec_eq(x_4333, x_4347); +lean_dec(x_4333); +if (x_4348 == 0) +{ +lean_object* x_4349; +lean_dec(x_4347); +if (lean_is_scalar(x_4346)) { + x_4349 = lean_alloc_ctor(1, 2, 0); +} else { + x_4349 = x_4346; +} +lean_ctor_set(x_4349, 0, x_4344); +lean_ctor_set(x_4349, 1, x_4345); +return x_4349; +} +else +{ +lean_object* x_4350; lean_object* x_4351; lean_object* x_4352; lean_object* x_4353; +lean_dec(x_4346); +lean_dec(x_4345); +x_4350 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_4351 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4351, 0, x_4350); +x_4352 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_4353 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4352, x_4351, x_4344); +if (lean_obj_tag(x_4353) == 0) +{ +lean_object* x_4354; lean_object* x_4355; lean_object* x_4356; lean_object* x_4357; +lean_dec(x_4347); +x_4354 = lean_ctor_get(x_4353, 0); +lean_inc(x_4354); +x_4355 = lean_ctor_get(x_4353, 1); +lean_inc(x_4355); +if (lean_is_exclusive(x_4353)) { + lean_ctor_release(x_4353, 0); + lean_ctor_release(x_4353, 1); + x_4356 = x_4353; +} else { + lean_dec_ref(x_4353); + x_4356 = lean_box(0); +} +if (lean_is_scalar(x_4356)) { + x_4357 = lean_alloc_ctor(0, 2, 0); +} else { + x_4357 = x_4356; +} +lean_ctor_set(x_4357, 0, x_4354); +lean_ctor_set(x_4357, 1, x_4355); +return x_4357; +} +else +{ +lean_object* x_4358; lean_object* x_4359; lean_object* x_4360; lean_object* x_4361; uint8_t x_4362; +x_4358 = lean_ctor_get(x_4353, 0); +lean_inc(x_4358); +x_4359 = lean_ctor_get(x_4353, 1); +lean_inc(x_4359); +if (lean_is_exclusive(x_4353)) { + lean_ctor_release(x_4353, 0); + lean_ctor_release(x_4353, 1); + x_4360 = x_4353; +} else { + lean_dec_ref(x_4353); + x_4360 = lean_box(0); +} +x_4361 = lean_ctor_get(x_4358, 1); +lean_inc(x_4361); +x_4362 = lean_nat_dec_eq(x_4347, x_4361); +lean_dec(x_4347); +if (x_4362 == 0) +{ +lean_object* x_4363; +lean_dec(x_4361); +if (lean_is_scalar(x_4360)) { + x_4363 = lean_alloc_ctor(1, 2, 0); +} else { + x_4363 = x_4360; +} +lean_ctor_set(x_4363, 0, x_4358); +lean_ctor_set(x_4363, 1, x_4359); +return x_4363; +} +else +{ +lean_object* x_4364; lean_object* x_4365; lean_object* x_4366; lean_object* x_4367; +lean_dec(x_4360); +lean_dec(x_4359); +x_4364 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_4365 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4365, 0, x_4364); +x_4366 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_4367 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4366, x_4365, x_4358); +if (lean_obj_tag(x_4367) == 0) +{ +lean_object* x_4368; lean_object* x_4369; lean_object* x_4370; lean_object* x_4371; +lean_dec(x_4361); +x_4368 = lean_ctor_get(x_4367, 0); +lean_inc(x_4368); +x_4369 = lean_ctor_get(x_4367, 1); +lean_inc(x_4369); +if (lean_is_exclusive(x_4367)) { + lean_ctor_release(x_4367, 0); + lean_ctor_release(x_4367, 1); + x_4370 = x_4367; +} else { + lean_dec_ref(x_4367); + x_4370 = lean_box(0); +} +if (lean_is_scalar(x_4370)) { + x_4371 = lean_alloc_ctor(0, 2, 0); +} else { + x_4371 = x_4370; +} +lean_ctor_set(x_4371, 0, x_4368); +lean_ctor_set(x_4371, 1, x_4369); +return x_4371; +} +else +{ +lean_object* x_4372; lean_object* x_4373; lean_object* x_4374; lean_object* x_4375; uint8_t x_4376; +x_4372 = lean_ctor_get(x_4367, 0); +lean_inc(x_4372); +x_4373 = lean_ctor_get(x_4367, 1); +lean_inc(x_4373); +if (lean_is_exclusive(x_4367)) { + lean_ctor_release(x_4367, 0); + lean_ctor_release(x_4367, 1); + x_4374 = x_4367; +} else { + lean_dec_ref(x_4367); + x_4374 = lean_box(0); +} +x_4375 = lean_ctor_get(x_4372, 1); +lean_inc(x_4375); +x_4376 = lean_nat_dec_eq(x_4361, x_4375); +lean_dec(x_4361); +if (x_4376 == 0) +{ +lean_object* x_4377; +lean_dec(x_4375); +if (lean_is_scalar(x_4374)) { + x_4377 = lean_alloc_ctor(1, 2, 0); +} else { + x_4377 = x_4374; +} +lean_ctor_set(x_4377, 0, x_4372); +lean_ctor_set(x_4377, 1, x_4373); +return x_4377; +} +else +{ +lean_object* x_4378; lean_object* x_4379; lean_object* x_4380; lean_object* x_4381; +lean_dec(x_4374); +lean_dec(x_4373); +x_4378 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_4379 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4379, 0, x_4378); +x_4380 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_4381 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4380, x_4379, x_4372); +if (lean_obj_tag(x_4381) == 0) +{ +lean_object* x_4382; lean_object* x_4383; lean_object* x_4384; lean_object* x_4385; +lean_dec(x_4375); +x_4382 = lean_ctor_get(x_4381, 0); +lean_inc(x_4382); +x_4383 = lean_ctor_get(x_4381, 1); +lean_inc(x_4383); +if (lean_is_exclusive(x_4381)) { + lean_ctor_release(x_4381, 0); + lean_ctor_release(x_4381, 1); + x_4384 = x_4381; +} else { + lean_dec_ref(x_4381); + x_4384 = lean_box(0); +} +if (lean_is_scalar(x_4384)) { + x_4385 = lean_alloc_ctor(0, 2, 0); +} else { + x_4385 = x_4384; +} +lean_ctor_set(x_4385, 0, x_4382); +lean_ctor_set(x_4385, 1, x_4383); +return x_4385; +} +else +{ +lean_object* x_4386; lean_object* x_4387; lean_object* x_4388; lean_object* x_4389; uint8_t x_4390; +x_4386 = lean_ctor_get(x_4381, 0); +lean_inc(x_4386); +x_4387 = lean_ctor_get(x_4381, 1); +lean_inc(x_4387); +if (lean_is_exclusive(x_4381)) { + lean_ctor_release(x_4381, 0); + lean_ctor_release(x_4381, 1); + x_4388 = x_4381; +} else { + lean_dec_ref(x_4381); + x_4388 = lean_box(0); +} +x_4389 = lean_ctor_get(x_4386, 1); +lean_inc(x_4389); +x_4390 = lean_nat_dec_eq(x_4375, x_4389); +lean_dec(x_4375); +if (x_4390 == 0) +{ +lean_object* x_4391; +lean_dec(x_4389); +if (lean_is_scalar(x_4388)) { + x_4391 = lean_alloc_ctor(1, 2, 0); +} else { + x_4391 = x_4388; +} +lean_ctor_set(x_4391, 0, x_4386); +lean_ctor_set(x_4391, 1, x_4387); +return x_4391; +} +else +{ +lean_object* x_4392; lean_object* x_4393; lean_object* x_4394; lean_object* x_4395; +lean_dec(x_4388); +lean_dec(x_4387); +x_4392 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_4393 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4393, 0, x_4392); +x_4394 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_4395 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4394, x_4393, x_4386); +if (lean_obj_tag(x_4395) == 0) +{ +lean_object* x_4396; lean_object* x_4397; lean_object* x_4398; lean_object* x_4399; +lean_dec(x_4389); +x_4396 = lean_ctor_get(x_4395, 0); +lean_inc(x_4396); +x_4397 = lean_ctor_get(x_4395, 1); +lean_inc(x_4397); +if (lean_is_exclusive(x_4395)) { + lean_ctor_release(x_4395, 0); + lean_ctor_release(x_4395, 1); + x_4398 = x_4395; +} else { + lean_dec_ref(x_4395); + x_4398 = lean_box(0); +} +if (lean_is_scalar(x_4398)) { + x_4399 = lean_alloc_ctor(0, 2, 0); +} else { + x_4399 = x_4398; +} +lean_ctor_set(x_4399, 0, x_4396); +lean_ctor_set(x_4399, 1, x_4397); +return x_4399; +} +else +{ +lean_object* x_4400; lean_object* x_4401; lean_object* x_4402; lean_object* x_4403; uint8_t x_4404; +x_4400 = lean_ctor_get(x_4395, 0); +lean_inc(x_4400); +x_4401 = lean_ctor_get(x_4395, 1); +lean_inc(x_4401); +if (lean_is_exclusive(x_4395)) { + lean_ctor_release(x_4395, 0); + lean_ctor_release(x_4395, 1); + x_4402 = x_4395; +} else { + lean_dec_ref(x_4395); + x_4402 = lean_box(0); +} +x_4403 = lean_ctor_get(x_4400, 1); +lean_inc(x_4403); +x_4404 = lean_nat_dec_eq(x_4389, x_4403); +lean_dec(x_4389); +if (x_4404 == 0) +{ +lean_object* x_4405; +lean_dec(x_4403); +if (lean_is_scalar(x_4402)) { + x_4405 = lean_alloc_ctor(1, 2, 0); +} else { + x_4405 = x_4402; +} +lean_ctor_set(x_4405, 0, x_4400); +lean_ctor_set(x_4405, 1, x_4401); +return x_4405; +} +else +{ +lean_object* x_4406; lean_object* x_4407; lean_object* x_4408; lean_object* x_4409; +lean_dec(x_4402); +lean_dec(x_4401); +x_4406 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_4407 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4407, 0, x_4406); +x_4408 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_4409 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4408, x_4407, x_4400); +if (lean_obj_tag(x_4409) == 0) +{ +lean_object* x_4410; lean_object* x_4411; lean_object* x_4412; lean_object* x_4413; +lean_dec(x_4403); +x_4410 = lean_ctor_get(x_4409, 0); +lean_inc(x_4410); +x_4411 = lean_ctor_get(x_4409, 1); +lean_inc(x_4411); +if (lean_is_exclusive(x_4409)) { + lean_ctor_release(x_4409, 0); + lean_ctor_release(x_4409, 1); + x_4412 = x_4409; +} else { + lean_dec_ref(x_4409); + x_4412 = lean_box(0); +} +if (lean_is_scalar(x_4412)) { + x_4413 = lean_alloc_ctor(0, 2, 0); +} else { + x_4413 = x_4412; +} +lean_ctor_set(x_4413, 0, x_4410); +lean_ctor_set(x_4413, 1, x_4411); +return x_4413; +} +else +{ +lean_object* x_4414; lean_object* x_4415; lean_object* x_4416; lean_object* x_4417; uint8_t x_4418; +x_4414 = lean_ctor_get(x_4409, 0); +lean_inc(x_4414); +x_4415 = lean_ctor_get(x_4409, 1); +lean_inc(x_4415); +if (lean_is_exclusive(x_4409)) { + lean_ctor_release(x_4409, 0); + lean_ctor_release(x_4409, 1); + x_4416 = x_4409; +} else { + lean_dec_ref(x_4409); + x_4416 = lean_box(0); +} +x_4417 = lean_ctor_get(x_4414, 1); +lean_inc(x_4417); +x_4418 = lean_nat_dec_eq(x_4403, x_4417); +lean_dec(x_4403); +if (x_4418 == 0) +{ +lean_object* x_4419; +lean_dec(x_4417); +if (lean_is_scalar(x_4416)) { + x_4419 = lean_alloc_ctor(1, 2, 0); +} else { + x_4419 = x_4416; +} +lean_ctor_set(x_4419, 0, x_4414); +lean_ctor_set(x_4419, 1, x_4415); +return x_4419; +} +else +{ +lean_object* x_4420; lean_object* x_4421; lean_object* x_4422; lean_object* x_4423; +lean_dec(x_4416); +lean_dec(x_4415); +x_4420 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_4421 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4421, 0, x_4420); +x_4422 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_4423 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4422, x_4421, x_4414); +if (lean_obj_tag(x_4423) == 0) +{ +lean_object* x_4424; lean_object* x_4425; lean_object* x_4426; lean_object* x_4427; +lean_dec(x_4417); +x_4424 = lean_ctor_get(x_4423, 0); +lean_inc(x_4424); +x_4425 = lean_ctor_get(x_4423, 1); +lean_inc(x_4425); +if (lean_is_exclusive(x_4423)) { + lean_ctor_release(x_4423, 0); + lean_ctor_release(x_4423, 1); + x_4426 = x_4423; +} else { + lean_dec_ref(x_4423); + x_4426 = lean_box(0); +} +if (lean_is_scalar(x_4426)) { + x_4427 = lean_alloc_ctor(0, 2, 0); +} else { + x_4427 = x_4426; +} +lean_ctor_set(x_4427, 0, x_4424); +lean_ctor_set(x_4427, 1, x_4425); +return x_4427; +} +else +{ +lean_object* x_4428; lean_object* x_4429; lean_object* x_4430; lean_object* x_4431; uint8_t x_4432; +x_4428 = lean_ctor_get(x_4423, 0); +lean_inc(x_4428); +x_4429 = lean_ctor_get(x_4423, 1); +lean_inc(x_4429); +if (lean_is_exclusive(x_4423)) { + lean_ctor_release(x_4423, 0); + lean_ctor_release(x_4423, 1); + x_4430 = x_4423; +} else { + lean_dec_ref(x_4423); + x_4430 = lean_box(0); +} +x_4431 = lean_ctor_get(x_4428, 1); +lean_inc(x_4431); +x_4432 = lean_nat_dec_eq(x_4417, x_4431); +lean_dec(x_4417); +if (x_4432 == 0) +{ +lean_object* x_4433; +lean_dec(x_4431); +if (lean_is_scalar(x_4430)) { + x_4433 = lean_alloc_ctor(1, 2, 0); +} else { + x_4433 = x_4430; +} +lean_ctor_set(x_4433, 0, x_4428); +lean_ctor_set(x_4433, 1, x_4429); +return x_4433; +} +else +{ +lean_object* x_4434; lean_object* x_4435; lean_object* x_4436; lean_object* x_4437; +lean_dec(x_4430); +lean_dec(x_4429); +x_4434 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_4435 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4435, 0, x_4434); +x_4436 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_4437 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4436, x_4435, x_4428); +if (lean_obj_tag(x_4437) == 0) +{ +lean_object* x_4438; lean_object* x_4439; lean_object* x_4440; lean_object* x_4441; +lean_dec(x_4431); +x_4438 = lean_ctor_get(x_4437, 0); +lean_inc(x_4438); +x_4439 = lean_ctor_get(x_4437, 1); +lean_inc(x_4439); +if (lean_is_exclusive(x_4437)) { + lean_ctor_release(x_4437, 0); + lean_ctor_release(x_4437, 1); + x_4440 = x_4437; +} else { + lean_dec_ref(x_4437); + x_4440 = lean_box(0); +} +if (lean_is_scalar(x_4440)) { + x_4441 = lean_alloc_ctor(0, 2, 0); +} else { + x_4441 = x_4440; +} +lean_ctor_set(x_4441, 0, x_4438); +lean_ctor_set(x_4441, 1, x_4439); +return x_4441; +} +else +{ +lean_object* x_4442; lean_object* x_4443; lean_object* x_4444; lean_object* x_4445; uint8_t x_4446; +x_4442 = lean_ctor_get(x_4437, 0); +lean_inc(x_4442); +x_4443 = lean_ctor_get(x_4437, 1); +lean_inc(x_4443); +if (lean_is_exclusive(x_4437)) { + lean_ctor_release(x_4437, 0); + lean_ctor_release(x_4437, 1); + x_4444 = x_4437; +} else { + lean_dec_ref(x_4437); + x_4444 = lean_box(0); +} +x_4445 = lean_ctor_get(x_4442, 1); +lean_inc(x_4445); +x_4446 = lean_nat_dec_eq(x_4431, x_4445); +lean_dec(x_4431); +if (x_4446 == 0) +{ +lean_object* x_4447; +lean_dec(x_4445); +if (lean_is_scalar(x_4444)) { + x_4447 = lean_alloc_ctor(1, 2, 0); +} else { + x_4447 = x_4444; +} +lean_ctor_set(x_4447, 0, x_4442); +lean_ctor_set(x_4447, 1, x_4443); +return x_4447; +} +else +{ +lean_object* x_4448; lean_object* x_4449; lean_object* x_4450; lean_object* x_4451; +lean_dec(x_4444); +lean_dec(x_4443); +x_4448 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_4449 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4449, 0, x_4448); +x_4450 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_4451 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4450, x_4449, x_4442); +if (lean_obj_tag(x_4451) == 0) +{ +lean_object* x_4452; lean_object* x_4453; lean_object* x_4454; lean_object* x_4455; +lean_dec(x_4445); +x_4452 = lean_ctor_get(x_4451, 0); +lean_inc(x_4452); +x_4453 = lean_ctor_get(x_4451, 1); +lean_inc(x_4453); +if (lean_is_exclusive(x_4451)) { + lean_ctor_release(x_4451, 0); + lean_ctor_release(x_4451, 1); + x_4454 = x_4451; +} else { + lean_dec_ref(x_4451); + x_4454 = lean_box(0); +} +if (lean_is_scalar(x_4454)) { + x_4455 = lean_alloc_ctor(0, 2, 0); +} else { + x_4455 = x_4454; +} +lean_ctor_set(x_4455, 0, x_4452); +lean_ctor_set(x_4455, 1, x_4453); +return x_4455; +} +else +{ +lean_object* x_4456; lean_object* x_4457; lean_object* x_4458; lean_object* x_4459; uint8_t x_4460; +x_4456 = lean_ctor_get(x_4451, 0); +lean_inc(x_4456); +x_4457 = lean_ctor_get(x_4451, 1); +lean_inc(x_4457); +if (lean_is_exclusive(x_4451)) { + lean_ctor_release(x_4451, 0); + lean_ctor_release(x_4451, 1); + x_4458 = x_4451; +} else { + lean_dec_ref(x_4451); + x_4458 = lean_box(0); +} +x_4459 = lean_ctor_get(x_4456, 1); +lean_inc(x_4459); +x_4460 = lean_nat_dec_eq(x_4445, x_4459); +lean_dec(x_4445); +if (x_4460 == 0) +{ +lean_object* x_4461; +lean_dec(x_4459); +if (lean_is_scalar(x_4458)) { + x_4461 = lean_alloc_ctor(1, 2, 0); +} else { + x_4461 = x_4458; +} +lean_ctor_set(x_4461, 0, x_4456); +lean_ctor_set(x_4461, 1, x_4457); +return x_4461; +} +else +{ +lean_object* x_4462; lean_object* x_4463; lean_object* x_4464; lean_object* x_4465; +lean_dec(x_4458); +lean_dec(x_4457); +x_4462 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_4463 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4463, 0, x_4462); +x_4464 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_4465 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4464, x_4463, x_4456); +if (lean_obj_tag(x_4465) == 0) +{ +lean_object* x_4466; lean_object* x_4467; lean_object* x_4468; lean_object* x_4469; +lean_dec(x_4459); +x_4466 = lean_ctor_get(x_4465, 0); +lean_inc(x_4466); +x_4467 = lean_ctor_get(x_4465, 1); +lean_inc(x_4467); +if (lean_is_exclusive(x_4465)) { + lean_ctor_release(x_4465, 0); + lean_ctor_release(x_4465, 1); + x_4468 = x_4465; +} else { + lean_dec_ref(x_4465); + x_4468 = lean_box(0); +} +if (lean_is_scalar(x_4468)) { + x_4469 = lean_alloc_ctor(0, 2, 0); +} else { + x_4469 = x_4468; +} +lean_ctor_set(x_4469, 0, x_4466); +lean_ctor_set(x_4469, 1, x_4467); +return x_4469; +} +else +{ +lean_object* x_4470; lean_object* x_4471; lean_object* x_4472; lean_object* x_4473; uint8_t x_4474; +x_4470 = lean_ctor_get(x_4465, 0); +lean_inc(x_4470); +x_4471 = lean_ctor_get(x_4465, 1); +lean_inc(x_4471); +if (lean_is_exclusive(x_4465)) { + lean_ctor_release(x_4465, 0); + lean_ctor_release(x_4465, 1); + x_4472 = x_4465; +} else { + lean_dec_ref(x_4465); + x_4472 = lean_box(0); +} +x_4473 = lean_ctor_get(x_4470, 1); +lean_inc(x_4473); +x_4474 = lean_nat_dec_eq(x_4459, x_4473); +lean_dec(x_4473); +lean_dec(x_4459); +if (x_4474 == 0) +{ +lean_object* x_4475; +if (lean_is_scalar(x_4472)) { + x_4475 = lean_alloc_ctor(1, 2, 0); +} else { + x_4475 = x_4472; +} +lean_ctor_set(x_4475, 0, x_4470); +lean_ctor_set(x_4475, 1, x_4471); +return x_4475; +} +else +{ +lean_object* x_4476; lean_object* x_4477; lean_object* x_4478; lean_object* x_4479; +lean_dec(x_4472); +lean_dec(x_4471); +x_4476 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_4477 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4477, 0, x_4476); +x_4478 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_4479 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4478, x_4477, x_4470); +return x_4479; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_4480; lean_object* x_4481; lean_object* x_4482; uint8_t x_4483; +x_4480 = lean_ctor_get(x_83, 0); +x_4481 = lean_ctor_get(x_83, 1); +lean_inc(x_4481); +lean_inc(x_4480); +lean_dec(x_83); +x_4482 = lean_ctor_get(x_4480, 1); +lean_inc(x_4482); +x_4483 = lean_nat_dec_eq(x_78, x_4482); +lean_dec(x_78); +if (x_4483 == 0) +{ +lean_object* x_4484; +lean_dec(x_4482); +x_4484 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4484, 0, x_4480); +lean_ctor_set(x_4484, 1, x_4481); +return x_4484; +} +else +{ +lean_object* x_4485; lean_object* x_4486; lean_object* x_4487; lean_object* x_4488; +lean_dec(x_4481); +x_4485 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_4486 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4486, 0, x_4485); +x_4487 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_4488 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4487, x_4486, x_4480); +if (lean_obj_tag(x_4488) == 0) +{ +lean_object* x_4489; lean_object* x_4490; lean_object* x_4491; lean_object* x_4492; +lean_dec(x_4482); +x_4489 = lean_ctor_get(x_4488, 0); +lean_inc(x_4489); +x_4490 = lean_ctor_get(x_4488, 1); +lean_inc(x_4490); +if (lean_is_exclusive(x_4488)) { + lean_ctor_release(x_4488, 0); + lean_ctor_release(x_4488, 1); + x_4491 = x_4488; +} else { + lean_dec_ref(x_4488); + x_4491 = lean_box(0); +} +if (lean_is_scalar(x_4491)) { + x_4492 = lean_alloc_ctor(0, 2, 0); +} else { + x_4492 = x_4491; +} +lean_ctor_set(x_4492, 0, x_4489); +lean_ctor_set(x_4492, 1, x_4490); +return x_4492; +} +else +{ +lean_object* x_4493; lean_object* x_4494; lean_object* x_4495; lean_object* x_4496; uint8_t x_4497; +x_4493 = lean_ctor_get(x_4488, 0); +lean_inc(x_4493); +x_4494 = lean_ctor_get(x_4488, 1); +lean_inc(x_4494); +if (lean_is_exclusive(x_4488)) { + lean_ctor_release(x_4488, 0); + lean_ctor_release(x_4488, 1); + x_4495 = x_4488; +} else { + lean_dec_ref(x_4488); + x_4495 = lean_box(0); +} +x_4496 = lean_ctor_get(x_4493, 1); +lean_inc(x_4496); +x_4497 = lean_nat_dec_eq(x_4482, x_4496); +lean_dec(x_4482); +if (x_4497 == 0) +{ +lean_object* x_4498; +lean_dec(x_4496); +if (lean_is_scalar(x_4495)) { + x_4498 = lean_alloc_ctor(1, 2, 0); +} else { + x_4498 = x_4495; +} +lean_ctor_set(x_4498, 0, x_4493); +lean_ctor_set(x_4498, 1, x_4494); +return x_4498; +} +else +{ +lean_object* x_4499; lean_object* x_4500; lean_object* x_4501; +lean_dec(x_4495); +lean_dec(x_4494); +x_4499 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_4500 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4500, 0, x_4499); +x_4501 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4487, x_4500, x_4493); +if (lean_obj_tag(x_4501) == 0) +{ +lean_object* x_4502; lean_object* x_4503; lean_object* x_4504; lean_object* x_4505; +lean_dec(x_4496); +x_4502 = lean_ctor_get(x_4501, 0); +lean_inc(x_4502); +x_4503 = lean_ctor_get(x_4501, 1); +lean_inc(x_4503); +if (lean_is_exclusive(x_4501)) { + lean_ctor_release(x_4501, 0); + lean_ctor_release(x_4501, 1); + x_4504 = x_4501; +} else { + lean_dec_ref(x_4501); + x_4504 = lean_box(0); +} +if (lean_is_scalar(x_4504)) { + x_4505 = lean_alloc_ctor(0, 2, 0); +} else { + x_4505 = x_4504; +} +lean_ctor_set(x_4505, 0, x_4502); +lean_ctor_set(x_4505, 1, x_4503); +return x_4505; +} +else +{ +lean_object* x_4506; lean_object* x_4507; lean_object* x_4508; lean_object* x_4509; uint8_t x_4510; +x_4506 = lean_ctor_get(x_4501, 0); +lean_inc(x_4506); +x_4507 = lean_ctor_get(x_4501, 1); +lean_inc(x_4507); +if (lean_is_exclusive(x_4501)) { + lean_ctor_release(x_4501, 0); + lean_ctor_release(x_4501, 1); + x_4508 = x_4501; +} else { + lean_dec_ref(x_4501); + x_4508 = lean_box(0); +} +x_4509 = lean_ctor_get(x_4506, 1); +lean_inc(x_4509); +x_4510 = lean_nat_dec_eq(x_4496, x_4509); +lean_dec(x_4496); +if (x_4510 == 0) +{ +lean_object* x_4511; +lean_dec(x_4509); +if (lean_is_scalar(x_4508)) { + x_4511 = lean_alloc_ctor(1, 2, 0); +} else { + x_4511 = x_4508; +} +lean_ctor_set(x_4511, 0, x_4506); +lean_ctor_set(x_4511, 1, x_4507); +return x_4511; +} +else +{ +lean_object* x_4512; lean_object* x_4513; lean_object* x_4514; lean_object* x_4515; +lean_dec(x_4508); +lean_dec(x_4507); +x_4512 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_4513 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4513, 0, x_4512); +x_4514 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_4515 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4514, x_4513, x_4506); +if (lean_obj_tag(x_4515) == 0) +{ +lean_object* x_4516; lean_object* x_4517; lean_object* x_4518; lean_object* x_4519; +lean_dec(x_4509); +x_4516 = lean_ctor_get(x_4515, 0); +lean_inc(x_4516); +x_4517 = lean_ctor_get(x_4515, 1); +lean_inc(x_4517); +if (lean_is_exclusive(x_4515)) { + lean_ctor_release(x_4515, 0); + lean_ctor_release(x_4515, 1); + x_4518 = x_4515; +} else { + lean_dec_ref(x_4515); + x_4518 = lean_box(0); +} +if (lean_is_scalar(x_4518)) { + x_4519 = lean_alloc_ctor(0, 2, 0); +} else { + x_4519 = x_4518; +} +lean_ctor_set(x_4519, 0, x_4516); +lean_ctor_set(x_4519, 1, x_4517); +return x_4519; +} +else +{ +lean_object* x_4520; lean_object* x_4521; lean_object* x_4522; lean_object* x_4523; uint8_t x_4524; +x_4520 = lean_ctor_get(x_4515, 0); +lean_inc(x_4520); +x_4521 = lean_ctor_get(x_4515, 1); +lean_inc(x_4521); +if (lean_is_exclusive(x_4515)) { + lean_ctor_release(x_4515, 0); + lean_ctor_release(x_4515, 1); + x_4522 = x_4515; +} else { + lean_dec_ref(x_4515); + x_4522 = lean_box(0); +} +x_4523 = lean_ctor_get(x_4520, 1); +lean_inc(x_4523); +x_4524 = lean_nat_dec_eq(x_4509, x_4523); +lean_dec(x_4509); +if (x_4524 == 0) +{ +lean_object* x_4525; +lean_dec(x_4523); +if (lean_is_scalar(x_4522)) { + x_4525 = lean_alloc_ctor(1, 2, 0); +} else { + x_4525 = x_4522; +} +lean_ctor_set(x_4525, 0, x_4520); +lean_ctor_set(x_4525, 1, x_4521); +return x_4525; +} +else +{ +lean_object* x_4526; lean_object* x_4527; lean_object* x_4528; lean_object* x_4529; +lean_dec(x_4522); +lean_dec(x_4521); +x_4526 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_4527 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4527, 0, x_4526); +x_4528 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_4529 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4528, x_4527, x_4520); +if (lean_obj_tag(x_4529) == 0) +{ +lean_object* x_4530; lean_object* x_4531; lean_object* x_4532; lean_object* x_4533; +lean_dec(x_4523); +x_4530 = lean_ctor_get(x_4529, 0); +lean_inc(x_4530); +x_4531 = lean_ctor_get(x_4529, 1); +lean_inc(x_4531); +if (lean_is_exclusive(x_4529)) { + lean_ctor_release(x_4529, 0); + lean_ctor_release(x_4529, 1); + x_4532 = x_4529; +} else { + lean_dec_ref(x_4529); + x_4532 = lean_box(0); +} +if (lean_is_scalar(x_4532)) { + x_4533 = lean_alloc_ctor(0, 2, 0); +} else { + x_4533 = x_4532; +} +lean_ctor_set(x_4533, 0, x_4530); +lean_ctor_set(x_4533, 1, x_4531); +return x_4533; +} +else +{ +lean_object* x_4534; lean_object* x_4535; lean_object* x_4536; lean_object* x_4537; uint8_t x_4538; +x_4534 = lean_ctor_get(x_4529, 0); +lean_inc(x_4534); +x_4535 = lean_ctor_get(x_4529, 1); +lean_inc(x_4535); +if (lean_is_exclusive(x_4529)) { + lean_ctor_release(x_4529, 0); + lean_ctor_release(x_4529, 1); + x_4536 = x_4529; +} else { + lean_dec_ref(x_4529); + x_4536 = lean_box(0); +} +x_4537 = lean_ctor_get(x_4534, 1); +lean_inc(x_4537); +x_4538 = lean_nat_dec_eq(x_4523, x_4537); +lean_dec(x_4523); +if (x_4538 == 0) +{ +lean_object* x_4539; +lean_dec(x_4537); +if (lean_is_scalar(x_4536)) { + x_4539 = lean_alloc_ctor(1, 2, 0); +} else { + x_4539 = x_4536; +} +lean_ctor_set(x_4539, 0, x_4534); +lean_ctor_set(x_4539, 1, x_4535); +return x_4539; +} +else +{ +lean_object* x_4540; lean_object* x_4541; lean_object* x_4542; lean_object* x_4543; +lean_dec(x_4536); +lean_dec(x_4535); +x_4540 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_4541 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4541, 0, x_4540); +x_4542 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_4543 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4542, x_4541, x_4534); +if (lean_obj_tag(x_4543) == 0) +{ +lean_object* x_4544; lean_object* x_4545; lean_object* x_4546; lean_object* x_4547; +lean_dec(x_4537); +x_4544 = lean_ctor_get(x_4543, 0); +lean_inc(x_4544); +x_4545 = lean_ctor_get(x_4543, 1); +lean_inc(x_4545); +if (lean_is_exclusive(x_4543)) { + lean_ctor_release(x_4543, 0); + lean_ctor_release(x_4543, 1); + x_4546 = x_4543; +} else { + lean_dec_ref(x_4543); + x_4546 = lean_box(0); +} +if (lean_is_scalar(x_4546)) { + x_4547 = lean_alloc_ctor(0, 2, 0); +} else { + x_4547 = x_4546; +} +lean_ctor_set(x_4547, 0, x_4544); +lean_ctor_set(x_4547, 1, x_4545); +return x_4547; +} +else +{ +lean_object* x_4548; lean_object* x_4549; lean_object* x_4550; lean_object* x_4551; uint8_t x_4552; +x_4548 = lean_ctor_get(x_4543, 0); +lean_inc(x_4548); +x_4549 = lean_ctor_get(x_4543, 1); +lean_inc(x_4549); +if (lean_is_exclusive(x_4543)) { + lean_ctor_release(x_4543, 0); + lean_ctor_release(x_4543, 1); + x_4550 = x_4543; +} else { + lean_dec_ref(x_4543); + x_4550 = lean_box(0); +} +x_4551 = lean_ctor_get(x_4548, 1); +lean_inc(x_4551); +x_4552 = lean_nat_dec_eq(x_4537, x_4551); +lean_dec(x_4537); +if (x_4552 == 0) +{ +lean_object* x_4553; +lean_dec(x_4551); +if (lean_is_scalar(x_4550)) { + x_4553 = lean_alloc_ctor(1, 2, 0); +} else { + x_4553 = x_4550; +} +lean_ctor_set(x_4553, 0, x_4548); +lean_ctor_set(x_4553, 1, x_4549); +return x_4553; +} +else +{ +lean_object* x_4554; lean_object* x_4555; lean_object* x_4556; lean_object* x_4557; +lean_dec(x_4550); +lean_dec(x_4549); +x_4554 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_4555 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4555, 0, x_4554); +x_4556 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_4557 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4556, x_4555, x_4548); +if (lean_obj_tag(x_4557) == 0) +{ +lean_object* x_4558; lean_object* x_4559; lean_object* x_4560; lean_object* x_4561; +lean_dec(x_4551); +x_4558 = lean_ctor_get(x_4557, 0); +lean_inc(x_4558); +x_4559 = lean_ctor_get(x_4557, 1); +lean_inc(x_4559); +if (lean_is_exclusive(x_4557)) { + lean_ctor_release(x_4557, 0); + lean_ctor_release(x_4557, 1); + x_4560 = x_4557; +} else { + lean_dec_ref(x_4557); + x_4560 = lean_box(0); +} +if (lean_is_scalar(x_4560)) { + x_4561 = lean_alloc_ctor(0, 2, 0); +} else { + x_4561 = x_4560; +} +lean_ctor_set(x_4561, 0, x_4558); +lean_ctor_set(x_4561, 1, x_4559); +return x_4561; +} +else +{ +lean_object* x_4562; lean_object* x_4563; lean_object* x_4564; lean_object* x_4565; uint8_t x_4566; +x_4562 = lean_ctor_get(x_4557, 0); +lean_inc(x_4562); +x_4563 = lean_ctor_get(x_4557, 1); +lean_inc(x_4563); +if (lean_is_exclusive(x_4557)) { + lean_ctor_release(x_4557, 0); + lean_ctor_release(x_4557, 1); + x_4564 = x_4557; +} else { + lean_dec_ref(x_4557); + x_4564 = lean_box(0); +} +x_4565 = lean_ctor_get(x_4562, 1); +lean_inc(x_4565); +x_4566 = lean_nat_dec_eq(x_4551, x_4565); +lean_dec(x_4551); +if (x_4566 == 0) +{ +lean_object* x_4567; +lean_dec(x_4565); +if (lean_is_scalar(x_4564)) { + x_4567 = lean_alloc_ctor(1, 2, 0); +} else { + x_4567 = x_4564; +} +lean_ctor_set(x_4567, 0, x_4562); +lean_ctor_set(x_4567, 1, x_4563); +return x_4567; +} +else +{ +lean_object* x_4568; lean_object* x_4569; lean_object* x_4570; +lean_dec(x_4564); +lean_dec(x_4563); +x_4568 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_4569 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4569, 0, x_4568); +x_4570 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4556, x_4569, x_4562); +if (lean_obj_tag(x_4570) == 0) +{ +lean_object* x_4571; lean_object* x_4572; lean_object* x_4573; lean_object* x_4574; +lean_dec(x_4565); +x_4571 = lean_ctor_get(x_4570, 0); +lean_inc(x_4571); +x_4572 = lean_ctor_get(x_4570, 1); +lean_inc(x_4572); +if (lean_is_exclusive(x_4570)) { + lean_ctor_release(x_4570, 0); + lean_ctor_release(x_4570, 1); + x_4573 = x_4570; +} else { + lean_dec_ref(x_4570); + x_4573 = lean_box(0); +} +if (lean_is_scalar(x_4573)) { + x_4574 = lean_alloc_ctor(0, 2, 0); +} else { + x_4574 = x_4573; +} +lean_ctor_set(x_4574, 0, x_4571); +lean_ctor_set(x_4574, 1, x_4572); +return x_4574; +} +else +{ +lean_object* x_4575; lean_object* x_4576; lean_object* x_4577; lean_object* x_4578; uint8_t x_4579; +x_4575 = lean_ctor_get(x_4570, 0); +lean_inc(x_4575); +x_4576 = lean_ctor_get(x_4570, 1); +lean_inc(x_4576); +if (lean_is_exclusive(x_4570)) { + lean_ctor_release(x_4570, 0); + lean_ctor_release(x_4570, 1); + x_4577 = x_4570; +} else { + lean_dec_ref(x_4570); + x_4577 = lean_box(0); +} +x_4578 = lean_ctor_get(x_4575, 1); +lean_inc(x_4578); +x_4579 = lean_nat_dec_eq(x_4565, x_4578); +lean_dec(x_4565); +if (x_4579 == 0) +{ +lean_object* x_4580; +lean_dec(x_4578); +if (lean_is_scalar(x_4577)) { + x_4580 = lean_alloc_ctor(1, 2, 0); +} else { + x_4580 = x_4577; +} +lean_ctor_set(x_4580, 0, x_4575); +lean_ctor_set(x_4580, 1, x_4576); +return x_4580; +} +else +{ +lean_object* x_4581; lean_object* x_4582; lean_object* x_4583; lean_object* x_4584; +lean_dec(x_4577); +lean_dec(x_4576); +x_4581 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_4582 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4582, 0, x_4581); +x_4583 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_4584 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4583, x_4582, x_4575); +if (lean_obj_tag(x_4584) == 0) +{ +lean_object* x_4585; lean_object* x_4586; lean_object* x_4587; lean_object* x_4588; +lean_dec(x_4578); +x_4585 = lean_ctor_get(x_4584, 0); +lean_inc(x_4585); +x_4586 = lean_ctor_get(x_4584, 1); +lean_inc(x_4586); +if (lean_is_exclusive(x_4584)) { + lean_ctor_release(x_4584, 0); + lean_ctor_release(x_4584, 1); + x_4587 = x_4584; +} else { + lean_dec_ref(x_4584); + x_4587 = lean_box(0); +} +if (lean_is_scalar(x_4587)) { + x_4588 = lean_alloc_ctor(0, 2, 0); +} else { + x_4588 = x_4587; +} +lean_ctor_set(x_4588, 0, x_4585); +lean_ctor_set(x_4588, 1, x_4586); +return x_4588; +} +else +{ +lean_object* x_4589; lean_object* x_4590; lean_object* x_4591; lean_object* x_4592; uint8_t x_4593; +x_4589 = lean_ctor_get(x_4584, 0); +lean_inc(x_4589); +x_4590 = lean_ctor_get(x_4584, 1); +lean_inc(x_4590); +if (lean_is_exclusive(x_4584)) { + lean_ctor_release(x_4584, 0); + lean_ctor_release(x_4584, 1); + x_4591 = x_4584; +} else { + lean_dec_ref(x_4584); + x_4591 = lean_box(0); +} +x_4592 = lean_ctor_get(x_4589, 1); +lean_inc(x_4592); +x_4593 = lean_nat_dec_eq(x_4578, x_4592); +lean_dec(x_4578); +if (x_4593 == 0) +{ +lean_object* x_4594; +lean_dec(x_4592); +if (lean_is_scalar(x_4591)) { + x_4594 = lean_alloc_ctor(1, 2, 0); +} else { + x_4594 = x_4591; +} +lean_ctor_set(x_4594, 0, x_4589); +lean_ctor_set(x_4594, 1, x_4590); +return x_4594; +} +else +{ +lean_object* x_4595; lean_object* x_4596; lean_object* x_4597; lean_object* x_4598; +lean_dec(x_4591); +lean_dec(x_4590); +x_4595 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_4596 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4596, 0, x_4595); +x_4597 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_4598 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4597, x_4596, x_4589); +if (lean_obj_tag(x_4598) == 0) +{ +lean_object* x_4599; lean_object* x_4600; lean_object* x_4601; lean_object* x_4602; +lean_dec(x_4592); +x_4599 = lean_ctor_get(x_4598, 0); +lean_inc(x_4599); +x_4600 = lean_ctor_get(x_4598, 1); +lean_inc(x_4600); +if (lean_is_exclusive(x_4598)) { + lean_ctor_release(x_4598, 0); + lean_ctor_release(x_4598, 1); + x_4601 = x_4598; +} else { + lean_dec_ref(x_4598); + x_4601 = lean_box(0); +} +if (lean_is_scalar(x_4601)) { + x_4602 = lean_alloc_ctor(0, 2, 0); +} else { + x_4602 = x_4601; +} +lean_ctor_set(x_4602, 0, x_4599); +lean_ctor_set(x_4602, 1, x_4600); +return x_4602; +} +else +{ +lean_object* x_4603; lean_object* x_4604; lean_object* x_4605; lean_object* x_4606; uint8_t x_4607; +x_4603 = lean_ctor_get(x_4598, 0); +lean_inc(x_4603); +x_4604 = lean_ctor_get(x_4598, 1); +lean_inc(x_4604); +if (lean_is_exclusive(x_4598)) { + lean_ctor_release(x_4598, 0); + lean_ctor_release(x_4598, 1); + x_4605 = x_4598; +} else { + lean_dec_ref(x_4598); + x_4605 = lean_box(0); +} +x_4606 = lean_ctor_get(x_4603, 1); +lean_inc(x_4606); +x_4607 = lean_nat_dec_eq(x_4592, x_4606); +lean_dec(x_4592); +if (x_4607 == 0) +{ +lean_object* x_4608; +lean_dec(x_4606); +if (lean_is_scalar(x_4605)) { + x_4608 = lean_alloc_ctor(1, 2, 0); +} else { + x_4608 = x_4605; +} +lean_ctor_set(x_4608, 0, x_4603); +lean_ctor_set(x_4608, 1, x_4604); +return x_4608; +} +else +{ +lean_object* x_4609; lean_object* x_4610; lean_object* x_4611; lean_object* x_4612; +lean_dec(x_4605); +lean_dec(x_4604); +x_4609 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_4610 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4610, 0, x_4609); +x_4611 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_4612 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4611, x_4610, x_4603); +if (lean_obj_tag(x_4612) == 0) +{ +lean_object* x_4613; lean_object* x_4614; lean_object* x_4615; lean_object* x_4616; +lean_dec(x_4606); +x_4613 = lean_ctor_get(x_4612, 0); +lean_inc(x_4613); +x_4614 = lean_ctor_get(x_4612, 1); +lean_inc(x_4614); +if (lean_is_exclusive(x_4612)) { + lean_ctor_release(x_4612, 0); + lean_ctor_release(x_4612, 1); + x_4615 = x_4612; +} else { + lean_dec_ref(x_4612); + x_4615 = lean_box(0); +} +if (lean_is_scalar(x_4615)) { + x_4616 = lean_alloc_ctor(0, 2, 0); +} else { + x_4616 = x_4615; +} +lean_ctor_set(x_4616, 0, x_4613); +lean_ctor_set(x_4616, 1, x_4614); +return x_4616; +} +else +{ +lean_object* x_4617; lean_object* x_4618; lean_object* x_4619; lean_object* x_4620; uint8_t x_4621; +x_4617 = lean_ctor_get(x_4612, 0); +lean_inc(x_4617); +x_4618 = lean_ctor_get(x_4612, 1); +lean_inc(x_4618); +if (lean_is_exclusive(x_4612)) { + lean_ctor_release(x_4612, 0); + lean_ctor_release(x_4612, 1); + x_4619 = x_4612; +} else { + lean_dec_ref(x_4612); + x_4619 = lean_box(0); +} +x_4620 = lean_ctor_get(x_4617, 1); +lean_inc(x_4620); +x_4621 = lean_nat_dec_eq(x_4606, x_4620); +lean_dec(x_4606); +if (x_4621 == 0) +{ +lean_object* x_4622; +lean_dec(x_4620); +if (lean_is_scalar(x_4619)) { + x_4622 = lean_alloc_ctor(1, 2, 0); +} else { + x_4622 = x_4619; +} +lean_ctor_set(x_4622, 0, x_4617); +lean_ctor_set(x_4622, 1, x_4618); +return x_4622; +} +else +{ +lean_object* x_4623; lean_object* x_4624; lean_object* x_4625; lean_object* x_4626; +lean_dec(x_4619); +lean_dec(x_4618); +x_4623 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_4624 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4624, 0, x_4623); +x_4625 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_4626 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4625, x_4624, x_4617); +if (lean_obj_tag(x_4626) == 0) +{ +lean_object* x_4627; lean_object* x_4628; lean_object* x_4629; lean_object* x_4630; +lean_dec(x_4620); +x_4627 = lean_ctor_get(x_4626, 0); +lean_inc(x_4627); +x_4628 = lean_ctor_get(x_4626, 1); +lean_inc(x_4628); +if (lean_is_exclusive(x_4626)) { + lean_ctor_release(x_4626, 0); + lean_ctor_release(x_4626, 1); + x_4629 = x_4626; +} else { + lean_dec_ref(x_4626); + x_4629 = lean_box(0); +} +if (lean_is_scalar(x_4629)) { + x_4630 = lean_alloc_ctor(0, 2, 0); +} else { + x_4630 = x_4629; +} +lean_ctor_set(x_4630, 0, x_4627); +lean_ctor_set(x_4630, 1, x_4628); +return x_4630; +} +else +{ +lean_object* x_4631; lean_object* x_4632; lean_object* x_4633; lean_object* x_4634; uint8_t x_4635; +x_4631 = lean_ctor_get(x_4626, 0); +lean_inc(x_4631); +x_4632 = lean_ctor_get(x_4626, 1); +lean_inc(x_4632); +if (lean_is_exclusive(x_4626)) { + lean_ctor_release(x_4626, 0); + lean_ctor_release(x_4626, 1); + x_4633 = x_4626; +} else { + lean_dec_ref(x_4626); + x_4633 = lean_box(0); +} +x_4634 = lean_ctor_get(x_4631, 1); +lean_inc(x_4634); +x_4635 = lean_nat_dec_eq(x_4620, x_4634); +lean_dec(x_4620); +if (x_4635 == 0) +{ +lean_object* x_4636; +lean_dec(x_4634); +if (lean_is_scalar(x_4633)) { + x_4636 = lean_alloc_ctor(1, 2, 0); +} else { + x_4636 = x_4633; +} +lean_ctor_set(x_4636, 0, x_4631); +lean_ctor_set(x_4636, 1, x_4632); +return x_4636; +} +else +{ +lean_object* x_4637; lean_object* x_4638; lean_object* x_4639; lean_object* x_4640; +lean_dec(x_4633); +lean_dec(x_4632); +x_4637 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_4638 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4638, 0, x_4637); +x_4639 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_4640 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4639, x_4638, x_4631); +if (lean_obj_tag(x_4640) == 0) +{ +lean_object* x_4641; lean_object* x_4642; lean_object* x_4643; lean_object* x_4644; +lean_dec(x_4634); +x_4641 = lean_ctor_get(x_4640, 0); +lean_inc(x_4641); +x_4642 = lean_ctor_get(x_4640, 1); +lean_inc(x_4642); +if (lean_is_exclusive(x_4640)) { + lean_ctor_release(x_4640, 0); + lean_ctor_release(x_4640, 1); + x_4643 = x_4640; +} else { + lean_dec_ref(x_4640); + x_4643 = lean_box(0); +} +if (lean_is_scalar(x_4643)) { + x_4644 = lean_alloc_ctor(0, 2, 0); +} else { + x_4644 = x_4643; +} +lean_ctor_set(x_4644, 0, x_4641); +lean_ctor_set(x_4644, 1, x_4642); +return x_4644; +} +else +{ +lean_object* x_4645; lean_object* x_4646; lean_object* x_4647; lean_object* x_4648; uint8_t x_4649; +x_4645 = lean_ctor_get(x_4640, 0); +lean_inc(x_4645); +x_4646 = lean_ctor_get(x_4640, 1); +lean_inc(x_4646); +if (lean_is_exclusive(x_4640)) { + lean_ctor_release(x_4640, 0); + lean_ctor_release(x_4640, 1); + x_4647 = x_4640; +} else { + lean_dec_ref(x_4640); + x_4647 = lean_box(0); +} +x_4648 = lean_ctor_get(x_4645, 1); +lean_inc(x_4648); +x_4649 = lean_nat_dec_eq(x_4634, x_4648); +lean_dec(x_4634); +if (x_4649 == 0) +{ +lean_object* x_4650; +lean_dec(x_4648); +if (lean_is_scalar(x_4647)) { + x_4650 = lean_alloc_ctor(1, 2, 0); +} else { + x_4650 = x_4647; +} +lean_ctor_set(x_4650, 0, x_4645); +lean_ctor_set(x_4650, 1, x_4646); +return x_4650; +} +else +{ +lean_object* x_4651; lean_object* x_4652; lean_object* x_4653; lean_object* x_4654; +lean_dec(x_4647); +lean_dec(x_4646); +x_4651 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_4652 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4652, 0, x_4651); +x_4653 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_4654 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4653, x_4652, x_4645); +if (lean_obj_tag(x_4654) == 0) +{ +lean_object* x_4655; lean_object* x_4656; lean_object* x_4657; lean_object* x_4658; +lean_dec(x_4648); +x_4655 = lean_ctor_get(x_4654, 0); +lean_inc(x_4655); +x_4656 = lean_ctor_get(x_4654, 1); +lean_inc(x_4656); +if (lean_is_exclusive(x_4654)) { + lean_ctor_release(x_4654, 0); + lean_ctor_release(x_4654, 1); + x_4657 = x_4654; +} else { + lean_dec_ref(x_4654); + x_4657 = lean_box(0); +} +if (lean_is_scalar(x_4657)) { + x_4658 = lean_alloc_ctor(0, 2, 0); +} else { + x_4658 = x_4657; +} +lean_ctor_set(x_4658, 0, x_4655); +lean_ctor_set(x_4658, 1, x_4656); +return x_4658; +} +else +{ +lean_object* x_4659; lean_object* x_4660; lean_object* x_4661; lean_object* x_4662; uint8_t x_4663; +x_4659 = lean_ctor_get(x_4654, 0); +lean_inc(x_4659); +x_4660 = lean_ctor_get(x_4654, 1); +lean_inc(x_4660); +if (lean_is_exclusive(x_4654)) { + lean_ctor_release(x_4654, 0); + lean_ctor_release(x_4654, 1); + x_4661 = x_4654; +} else { + lean_dec_ref(x_4654); + x_4661 = lean_box(0); +} +x_4662 = lean_ctor_get(x_4659, 1); +lean_inc(x_4662); +x_4663 = lean_nat_dec_eq(x_4648, x_4662); +lean_dec(x_4648); +if (x_4663 == 0) +{ +lean_object* x_4664; +lean_dec(x_4662); +if (lean_is_scalar(x_4661)) { + x_4664 = lean_alloc_ctor(1, 2, 0); +} else { + x_4664 = x_4661; +} +lean_ctor_set(x_4664, 0, x_4659); +lean_ctor_set(x_4664, 1, x_4660); +return x_4664; +} +else +{ +lean_object* x_4665; lean_object* x_4666; lean_object* x_4667; lean_object* x_4668; +lean_dec(x_4661); +lean_dec(x_4660); +x_4665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_4666 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4666, 0, x_4665); +x_4667 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_4668 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4667, x_4666, x_4659); +if (lean_obj_tag(x_4668) == 0) +{ +lean_object* x_4669; lean_object* x_4670; lean_object* x_4671; lean_object* x_4672; +lean_dec(x_4662); +x_4669 = lean_ctor_get(x_4668, 0); +lean_inc(x_4669); +x_4670 = lean_ctor_get(x_4668, 1); +lean_inc(x_4670); +if (lean_is_exclusive(x_4668)) { + lean_ctor_release(x_4668, 0); + lean_ctor_release(x_4668, 1); + x_4671 = x_4668; +} else { + lean_dec_ref(x_4668); + x_4671 = lean_box(0); +} +if (lean_is_scalar(x_4671)) { + x_4672 = lean_alloc_ctor(0, 2, 0); +} else { + x_4672 = x_4671; +} +lean_ctor_set(x_4672, 0, x_4669); +lean_ctor_set(x_4672, 1, x_4670); +return x_4672; +} +else +{ +lean_object* x_4673; lean_object* x_4674; lean_object* x_4675; lean_object* x_4676; uint8_t x_4677; +x_4673 = lean_ctor_get(x_4668, 0); +lean_inc(x_4673); +x_4674 = lean_ctor_get(x_4668, 1); +lean_inc(x_4674); +if (lean_is_exclusive(x_4668)) { + lean_ctor_release(x_4668, 0); + lean_ctor_release(x_4668, 1); + x_4675 = x_4668; +} else { + lean_dec_ref(x_4668); + x_4675 = lean_box(0); +} +x_4676 = lean_ctor_get(x_4673, 1); +lean_inc(x_4676); +x_4677 = lean_nat_dec_eq(x_4662, x_4676); +lean_dec(x_4662); +if (x_4677 == 0) +{ +lean_object* x_4678; +lean_dec(x_4676); +if (lean_is_scalar(x_4675)) { + x_4678 = lean_alloc_ctor(1, 2, 0); +} else { + x_4678 = x_4675; +} +lean_ctor_set(x_4678, 0, x_4673); +lean_ctor_set(x_4678, 1, x_4674); +return x_4678; +} +else +{ +lean_object* x_4679; lean_object* x_4680; lean_object* x_4681; lean_object* x_4682; +lean_dec(x_4675); +lean_dec(x_4674); +x_4679 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_4680 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4680, 0, x_4679); +x_4681 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_4682 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4681, x_4680, x_4673); +if (lean_obj_tag(x_4682) == 0) +{ +lean_object* x_4683; lean_object* x_4684; lean_object* x_4685; lean_object* x_4686; +lean_dec(x_4676); +x_4683 = lean_ctor_get(x_4682, 0); +lean_inc(x_4683); +x_4684 = lean_ctor_get(x_4682, 1); +lean_inc(x_4684); +if (lean_is_exclusive(x_4682)) { + lean_ctor_release(x_4682, 0); + lean_ctor_release(x_4682, 1); + x_4685 = x_4682; +} else { + lean_dec_ref(x_4682); + x_4685 = lean_box(0); +} +if (lean_is_scalar(x_4685)) { + x_4686 = lean_alloc_ctor(0, 2, 0); +} else { + x_4686 = x_4685; +} +lean_ctor_set(x_4686, 0, x_4683); +lean_ctor_set(x_4686, 1, x_4684); +return x_4686; +} +else +{ +lean_object* x_4687; lean_object* x_4688; lean_object* x_4689; lean_object* x_4690; uint8_t x_4691; +x_4687 = lean_ctor_get(x_4682, 0); +lean_inc(x_4687); +x_4688 = lean_ctor_get(x_4682, 1); +lean_inc(x_4688); +if (lean_is_exclusive(x_4682)) { + lean_ctor_release(x_4682, 0); + lean_ctor_release(x_4682, 1); + x_4689 = x_4682; +} else { + lean_dec_ref(x_4682); + x_4689 = lean_box(0); +} +x_4690 = lean_ctor_get(x_4687, 1); +lean_inc(x_4690); +x_4691 = lean_nat_dec_eq(x_4676, x_4690); +lean_dec(x_4676); +if (x_4691 == 0) +{ +lean_object* x_4692; +lean_dec(x_4690); +if (lean_is_scalar(x_4689)) { + x_4692 = lean_alloc_ctor(1, 2, 0); +} else { + x_4692 = x_4689; +} +lean_ctor_set(x_4692, 0, x_4687); +lean_ctor_set(x_4692, 1, x_4688); +return x_4692; +} +else +{ +lean_object* x_4693; lean_object* x_4694; lean_object* x_4695; lean_object* x_4696; +lean_dec(x_4689); +lean_dec(x_4688); +x_4693 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_4694 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4694, 0, x_4693); +x_4695 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_4696 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4695, x_4694, x_4687); +if (lean_obj_tag(x_4696) == 0) +{ +lean_object* x_4697; lean_object* x_4698; lean_object* x_4699; lean_object* x_4700; +lean_dec(x_4690); +x_4697 = lean_ctor_get(x_4696, 0); +lean_inc(x_4697); +x_4698 = lean_ctor_get(x_4696, 1); +lean_inc(x_4698); +if (lean_is_exclusive(x_4696)) { + lean_ctor_release(x_4696, 0); + lean_ctor_release(x_4696, 1); + x_4699 = x_4696; +} else { + lean_dec_ref(x_4696); + x_4699 = lean_box(0); +} +if (lean_is_scalar(x_4699)) { + x_4700 = lean_alloc_ctor(0, 2, 0); +} else { + x_4700 = x_4699; +} +lean_ctor_set(x_4700, 0, x_4697); +lean_ctor_set(x_4700, 1, x_4698); +return x_4700; +} +else +{ +lean_object* x_4701; lean_object* x_4702; lean_object* x_4703; lean_object* x_4704; uint8_t x_4705; +x_4701 = lean_ctor_get(x_4696, 0); +lean_inc(x_4701); +x_4702 = lean_ctor_get(x_4696, 1); +lean_inc(x_4702); +if (lean_is_exclusive(x_4696)) { + lean_ctor_release(x_4696, 0); + lean_ctor_release(x_4696, 1); + x_4703 = x_4696; +} else { + lean_dec_ref(x_4696); + x_4703 = lean_box(0); +} +x_4704 = lean_ctor_get(x_4701, 1); +lean_inc(x_4704); +x_4705 = lean_nat_dec_eq(x_4690, x_4704); +lean_dec(x_4690); +if (x_4705 == 0) +{ +lean_object* x_4706; +lean_dec(x_4704); +if (lean_is_scalar(x_4703)) { + x_4706 = lean_alloc_ctor(1, 2, 0); +} else { + x_4706 = x_4703; +} +lean_ctor_set(x_4706, 0, x_4701); +lean_ctor_set(x_4706, 1, x_4702); +return x_4706; +} +else +{ +lean_object* x_4707; lean_object* x_4708; lean_object* x_4709; lean_object* x_4710; +lean_dec(x_4703); +lean_dec(x_4702); +x_4707 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_4708 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4708, 0, x_4707); +x_4709 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_4710 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4709, x_4708, x_4701); +if (lean_obj_tag(x_4710) == 0) +{ +lean_object* x_4711; lean_object* x_4712; lean_object* x_4713; lean_object* x_4714; +lean_dec(x_4704); +x_4711 = lean_ctor_get(x_4710, 0); +lean_inc(x_4711); +x_4712 = lean_ctor_get(x_4710, 1); +lean_inc(x_4712); +if (lean_is_exclusive(x_4710)) { + lean_ctor_release(x_4710, 0); + lean_ctor_release(x_4710, 1); + x_4713 = x_4710; +} else { + lean_dec_ref(x_4710); + x_4713 = lean_box(0); +} +if (lean_is_scalar(x_4713)) { + x_4714 = lean_alloc_ctor(0, 2, 0); +} else { + x_4714 = x_4713; +} +lean_ctor_set(x_4714, 0, x_4711); +lean_ctor_set(x_4714, 1, x_4712); +return x_4714; +} +else +{ +lean_object* x_4715; lean_object* x_4716; lean_object* x_4717; lean_object* x_4718; uint8_t x_4719; +x_4715 = lean_ctor_get(x_4710, 0); +lean_inc(x_4715); +x_4716 = lean_ctor_get(x_4710, 1); +lean_inc(x_4716); +if (lean_is_exclusive(x_4710)) { + lean_ctor_release(x_4710, 0); + lean_ctor_release(x_4710, 1); + x_4717 = x_4710; +} else { + lean_dec_ref(x_4710); + x_4717 = lean_box(0); +} +x_4718 = lean_ctor_get(x_4715, 1); +lean_inc(x_4718); +x_4719 = lean_nat_dec_eq(x_4704, x_4718); +lean_dec(x_4704); +if (x_4719 == 0) +{ +lean_object* x_4720; +lean_dec(x_4718); +if (lean_is_scalar(x_4717)) { + x_4720 = lean_alloc_ctor(1, 2, 0); +} else { + x_4720 = x_4717; +} +lean_ctor_set(x_4720, 0, x_4715); +lean_ctor_set(x_4720, 1, x_4716); +return x_4720; +} +else +{ +lean_object* x_4721; lean_object* x_4722; lean_object* x_4723; lean_object* x_4724; +lean_dec(x_4717); +lean_dec(x_4716); +x_4721 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_4722 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4722, 0, x_4721); +x_4723 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_4724 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4723, x_4722, x_4715); +if (lean_obj_tag(x_4724) == 0) +{ +lean_object* x_4725; lean_object* x_4726; lean_object* x_4727; lean_object* x_4728; +lean_dec(x_4718); +x_4725 = lean_ctor_get(x_4724, 0); +lean_inc(x_4725); +x_4726 = lean_ctor_get(x_4724, 1); +lean_inc(x_4726); +if (lean_is_exclusive(x_4724)) { + lean_ctor_release(x_4724, 0); + lean_ctor_release(x_4724, 1); + x_4727 = x_4724; +} else { + lean_dec_ref(x_4724); + x_4727 = lean_box(0); +} +if (lean_is_scalar(x_4727)) { + x_4728 = lean_alloc_ctor(0, 2, 0); +} else { + x_4728 = x_4727; +} +lean_ctor_set(x_4728, 0, x_4725); +lean_ctor_set(x_4728, 1, x_4726); +return x_4728; +} +else +{ +lean_object* x_4729; lean_object* x_4730; lean_object* x_4731; lean_object* x_4732; uint8_t x_4733; +x_4729 = lean_ctor_get(x_4724, 0); +lean_inc(x_4729); +x_4730 = lean_ctor_get(x_4724, 1); +lean_inc(x_4730); +if (lean_is_exclusive(x_4724)) { + lean_ctor_release(x_4724, 0); + lean_ctor_release(x_4724, 1); + x_4731 = x_4724; +} else { + lean_dec_ref(x_4724); + x_4731 = lean_box(0); +} +x_4732 = lean_ctor_get(x_4729, 1); +lean_inc(x_4732); +x_4733 = lean_nat_dec_eq(x_4718, x_4732); +lean_dec(x_4718); +if (x_4733 == 0) +{ +lean_object* x_4734; +lean_dec(x_4732); +if (lean_is_scalar(x_4731)) { + x_4734 = lean_alloc_ctor(1, 2, 0); +} else { + x_4734 = x_4731; +} +lean_ctor_set(x_4734, 0, x_4729); +lean_ctor_set(x_4734, 1, x_4730); +return x_4734; +} +else +{ +lean_object* x_4735; lean_object* x_4736; lean_object* x_4737; lean_object* x_4738; +lean_dec(x_4731); +lean_dec(x_4730); +x_4735 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_4736 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4736, 0, x_4735); +x_4737 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_4738 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4737, x_4736, x_4729); +if (lean_obj_tag(x_4738) == 0) +{ +lean_object* x_4739; lean_object* x_4740; lean_object* x_4741; lean_object* x_4742; +lean_dec(x_4732); +x_4739 = lean_ctor_get(x_4738, 0); +lean_inc(x_4739); +x_4740 = lean_ctor_get(x_4738, 1); +lean_inc(x_4740); +if (lean_is_exclusive(x_4738)) { + lean_ctor_release(x_4738, 0); + lean_ctor_release(x_4738, 1); + x_4741 = x_4738; +} else { + lean_dec_ref(x_4738); + x_4741 = lean_box(0); +} +if (lean_is_scalar(x_4741)) { + x_4742 = lean_alloc_ctor(0, 2, 0); +} else { + x_4742 = x_4741; +} +lean_ctor_set(x_4742, 0, x_4739); +lean_ctor_set(x_4742, 1, x_4740); +return x_4742; +} +else +{ +lean_object* x_4743; lean_object* x_4744; lean_object* x_4745; lean_object* x_4746; uint8_t x_4747; +x_4743 = lean_ctor_get(x_4738, 0); +lean_inc(x_4743); +x_4744 = lean_ctor_get(x_4738, 1); +lean_inc(x_4744); +if (lean_is_exclusive(x_4738)) { + lean_ctor_release(x_4738, 0); + lean_ctor_release(x_4738, 1); + x_4745 = x_4738; +} else { + lean_dec_ref(x_4738); + x_4745 = lean_box(0); +} +x_4746 = lean_ctor_get(x_4743, 1); +lean_inc(x_4746); +x_4747 = lean_nat_dec_eq(x_4732, x_4746); +lean_dec(x_4732); +if (x_4747 == 0) +{ +lean_object* x_4748; +lean_dec(x_4746); +if (lean_is_scalar(x_4745)) { + x_4748 = lean_alloc_ctor(1, 2, 0); +} else { + x_4748 = x_4745; +} +lean_ctor_set(x_4748, 0, x_4743); +lean_ctor_set(x_4748, 1, x_4744); +return x_4748; +} +else +{ +lean_object* x_4749; lean_object* x_4750; lean_object* x_4751; lean_object* x_4752; +lean_dec(x_4745); +lean_dec(x_4744); +x_4749 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_4750 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4750, 0, x_4749); +x_4751 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_4752 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4751, x_4750, x_4743); +if (lean_obj_tag(x_4752) == 0) +{ +lean_object* x_4753; lean_object* x_4754; lean_object* x_4755; lean_object* x_4756; +lean_dec(x_4746); +x_4753 = lean_ctor_get(x_4752, 0); +lean_inc(x_4753); +x_4754 = lean_ctor_get(x_4752, 1); +lean_inc(x_4754); +if (lean_is_exclusive(x_4752)) { + lean_ctor_release(x_4752, 0); + lean_ctor_release(x_4752, 1); + x_4755 = x_4752; +} else { + lean_dec_ref(x_4752); + x_4755 = lean_box(0); +} +if (lean_is_scalar(x_4755)) { + x_4756 = lean_alloc_ctor(0, 2, 0); +} else { + x_4756 = x_4755; +} +lean_ctor_set(x_4756, 0, x_4753); +lean_ctor_set(x_4756, 1, x_4754); +return x_4756; +} +else +{ +lean_object* x_4757; lean_object* x_4758; lean_object* x_4759; lean_object* x_4760; uint8_t x_4761; +x_4757 = lean_ctor_get(x_4752, 0); +lean_inc(x_4757); +x_4758 = lean_ctor_get(x_4752, 1); +lean_inc(x_4758); +if (lean_is_exclusive(x_4752)) { + lean_ctor_release(x_4752, 0); + lean_ctor_release(x_4752, 1); + x_4759 = x_4752; +} else { + lean_dec_ref(x_4752); + x_4759 = lean_box(0); +} +x_4760 = lean_ctor_get(x_4757, 1); +lean_inc(x_4760); +x_4761 = lean_nat_dec_eq(x_4746, x_4760); +lean_dec(x_4746); +if (x_4761 == 0) +{ +lean_object* x_4762; +lean_dec(x_4760); +if (lean_is_scalar(x_4759)) { + x_4762 = lean_alloc_ctor(1, 2, 0); +} else { + x_4762 = x_4759; +} +lean_ctor_set(x_4762, 0, x_4757); +lean_ctor_set(x_4762, 1, x_4758); +return x_4762; +} +else +{ +lean_object* x_4763; lean_object* x_4764; lean_object* x_4765; lean_object* x_4766; +lean_dec(x_4759); +lean_dec(x_4758); +x_4763 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_4764 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4764, 0, x_4763); +x_4765 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_4766 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4765, x_4764, x_4757); +if (lean_obj_tag(x_4766) == 0) +{ +lean_object* x_4767; lean_object* x_4768; lean_object* x_4769; lean_object* x_4770; +lean_dec(x_4760); +x_4767 = lean_ctor_get(x_4766, 0); +lean_inc(x_4767); +x_4768 = lean_ctor_get(x_4766, 1); +lean_inc(x_4768); +if (lean_is_exclusive(x_4766)) { + lean_ctor_release(x_4766, 0); + lean_ctor_release(x_4766, 1); + x_4769 = x_4766; +} else { + lean_dec_ref(x_4766); + x_4769 = lean_box(0); +} +if (lean_is_scalar(x_4769)) { + x_4770 = lean_alloc_ctor(0, 2, 0); +} else { + x_4770 = x_4769; +} +lean_ctor_set(x_4770, 0, x_4767); +lean_ctor_set(x_4770, 1, x_4768); +return x_4770; +} +else +{ +lean_object* x_4771; lean_object* x_4772; lean_object* x_4773; lean_object* x_4774; uint8_t x_4775; +x_4771 = lean_ctor_get(x_4766, 0); +lean_inc(x_4771); +x_4772 = lean_ctor_get(x_4766, 1); +lean_inc(x_4772); +if (lean_is_exclusive(x_4766)) { + lean_ctor_release(x_4766, 0); + lean_ctor_release(x_4766, 1); + x_4773 = x_4766; +} else { + lean_dec_ref(x_4766); + x_4773 = lean_box(0); +} +x_4774 = lean_ctor_get(x_4771, 1); +lean_inc(x_4774); +x_4775 = lean_nat_dec_eq(x_4760, x_4774); +lean_dec(x_4760); +if (x_4775 == 0) +{ +lean_object* x_4776; +lean_dec(x_4774); +if (lean_is_scalar(x_4773)) { + x_4776 = lean_alloc_ctor(1, 2, 0); +} else { + x_4776 = x_4773; +} +lean_ctor_set(x_4776, 0, x_4771); +lean_ctor_set(x_4776, 1, x_4772); +return x_4776; +} +else +{ +lean_object* x_4777; lean_object* x_4778; lean_object* x_4779; lean_object* x_4780; +lean_dec(x_4773); +lean_dec(x_4772); +x_4777 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_4778 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4778, 0, x_4777); +x_4779 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_4780 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4779, x_4778, x_4771); +if (lean_obj_tag(x_4780) == 0) +{ +lean_object* x_4781; lean_object* x_4782; lean_object* x_4783; lean_object* x_4784; +lean_dec(x_4774); +x_4781 = lean_ctor_get(x_4780, 0); +lean_inc(x_4781); +x_4782 = lean_ctor_get(x_4780, 1); +lean_inc(x_4782); +if (lean_is_exclusive(x_4780)) { + lean_ctor_release(x_4780, 0); + lean_ctor_release(x_4780, 1); + x_4783 = x_4780; +} else { + lean_dec_ref(x_4780); + x_4783 = lean_box(0); +} +if (lean_is_scalar(x_4783)) { + x_4784 = lean_alloc_ctor(0, 2, 0); +} else { + x_4784 = x_4783; +} +lean_ctor_set(x_4784, 0, x_4781); +lean_ctor_set(x_4784, 1, x_4782); +return x_4784; +} +else +{ +lean_object* x_4785; lean_object* x_4786; lean_object* x_4787; lean_object* x_4788; uint8_t x_4789; +x_4785 = lean_ctor_get(x_4780, 0); +lean_inc(x_4785); +x_4786 = lean_ctor_get(x_4780, 1); +lean_inc(x_4786); +if (lean_is_exclusive(x_4780)) { + lean_ctor_release(x_4780, 0); + lean_ctor_release(x_4780, 1); + x_4787 = x_4780; +} else { + lean_dec_ref(x_4780); + x_4787 = lean_box(0); +} +x_4788 = lean_ctor_get(x_4785, 1); +lean_inc(x_4788); +x_4789 = lean_nat_dec_eq(x_4774, x_4788); +lean_dec(x_4774); +if (x_4789 == 0) +{ +lean_object* x_4790; +lean_dec(x_4788); +if (lean_is_scalar(x_4787)) { + x_4790 = lean_alloc_ctor(1, 2, 0); +} else { + x_4790 = x_4787; +} +lean_ctor_set(x_4790, 0, x_4785); +lean_ctor_set(x_4790, 1, x_4786); +return x_4790; +} +else +{ +lean_object* x_4791; lean_object* x_4792; lean_object* x_4793; lean_object* x_4794; +lean_dec(x_4787); +lean_dec(x_4786); +x_4791 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_4792 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4792, 0, x_4791); +x_4793 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_4794 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4793, x_4792, x_4785); +if (lean_obj_tag(x_4794) == 0) +{ +lean_object* x_4795; lean_object* x_4796; lean_object* x_4797; lean_object* x_4798; +lean_dec(x_4788); +x_4795 = lean_ctor_get(x_4794, 0); +lean_inc(x_4795); +x_4796 = lean_ctor_get(x_4794, 1); +lean_inc(x_4796); +if (lean_is_exclusive(x_4794)) { + lean_ctor_release(x_4794, 0); + lean_ctor_release(x_4794, 1); + x_4797 = x_4794; +} else { + lean_dec_ref(x_4794); + x_4797 = lean_box(0); +} +if (lean_is_scalar(x_4797)) { + x_4798 = lean_alloc_ctor(0, 2, 0); +} else { + x_4798 = x_4797; +} +lean_ctor_set(x_4798, 0, x_4795); +lean_ctor_set(x_4798, 1, x_4796); +return x_4798; +} +else +{ +lean_object* x_4799; lean_object* x_4800; lean_object* x_4801; lean_object* x_4802; uint8_t x_4803; +x_4799 = lean_ctor_get(x_4794, 0); +lean_inc(x_4799); +x_4800 = lean_ctor_get(x_4794, 1); +lean_inc(x_4800); +if (lean_is_exclusive(x_4794)) { + lean_ctor_release(x_4794, 0); + lean_ctor_release(x_4794, 1); + x_4801 = x_4794; +} else { + lean_dec_ref(x_4794); + x_4801 = lean_box(0); +} +x_4802 = lean_ctor_get(x_4799, 1); +lean_inc(x_4802); +x_4803 = lean_nat_dec_eq(x_4788, x_4802); +lean_dec(x_4788); +if (x_4803 == 0) +{ +lean_object* x_4804; +lean_dec(x_4802); +if (lean_is_scalar(x_4801)) { + x_4804 = lean_alloc_ctor(1, 2, 0); +} else { + x_4804 = x_4801; +} +lean_ctor_set(x_4804, 0, x_4799); +lean_ctor_set(x_4804, 1, x_4800); +return x_4804; +} +else +{ +lean_object* x_4805; lean_object* x_4806; lean_object* x_4807; lean_object* x_4808; +lean_dec(x_4801); +lean_dec(x_4800); +x_4805 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_4806 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4806, 0, x_4805); +x_4807 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_4808 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4807, x_4806, x_4799); +if (lean_obj_tag(x_4808) == 0) +{ +lean_object* x_4809; lean_object* x_4810; lean_object* x_4811; lean_object* x_4812; +lean_dec(x_4802); +x_4809 = lean_ctor_get(x_4808, 0); +lean_inc(x_4809); +x_4810 = lean_ctor_get(x_4808, 1); +lean_inc(x_4810); +if (lean_is_exclusive(x_4808)) { + lean_ctor_release(x_4808, 0); + lean_ctor_release(x_4808, 1); + x_4811 = x_4808; +} else { + lean_dec_ref(x_4808); + x_4811 = lean_box(0); +} +if (lean_is_scalar(x_4811)) { + x_4812 = lean_alloc_ctor(0, 2, 0); +} else { + x_4812 = x_4811; +} +lean_ctor_set(x_4812, 0, x_4809); +lean_ctor_set(x_4812, 1, x_4810); +return x_4812; +} +else +{ +lean_object* x_4813; lean_object* x_4814; lean_object* x_4815; lean_object* x_4816; uint8_t x_4817; +x_4813 = lean_ctor_get(x_4808, 0); +lean_inc(x_4813); +x_4814 = lean_ctor_get(x_4808, 1); +lean_inc(x_4814); +if (lean_is_exclusive(x_4808)) { + lean_ctor_release(x_4808, 0); + lean_ctor_release(x_4808, 1); + x_4815 = x_4808; +} else { + lean_dec_ref(x_4808); + x_4815 = lean_box(0); +} +x_4816 = lean_ctor_get(x_4813, 1); +lean_inc(x_4816); +x_4817 = lean_nat_dec_eq(x_4802, x_4816); +lean_dec(x_4816); +lean_dec(x_4802); +if (x_4817 == 0) +{ +lean_object* x_4818; +if (lean_is_scalar(x_4815)) { + x_4818 = lean_alloc_ctor(1, 2, 0); +} else { + x_4818 = x_4815; +} +lean_ctor_set(x_4818, 0, x_4813); +lean_ctor_set(x_4818, 1, x_4814); +return x_4818; +} +else +{ +lean_object* x_4819; lean_object* x_4820; lean_object* x_4821; lean_object* x_4822; +lean_dec(x_4815); +lean_dec(x_4814); +x_4819 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_4820 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4820, 0, x_4819); +x_4821 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_4822 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4821, x_4820, x_4813); +return x_4822; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_4823; lean_object* x_4824; lean_object* x_4825; uint8_t x_4826; +x_4823 = lean_ctor_get(x_70, 0); +x_4824 = lean_ctor_get(x_70, 1); +lean_inc(x_4824); +lean_inc(x_4823); +lean_dec(x_70); +x_4825 = lean_ctor_get(x_4823, 1); +lean_inc(x_4825); +x_4826 = lean_nat_dec_eq(x_66, x_4825); +lean_dec(x_66); +if (x_4826 == 0) +{ +lean_object* x_4827; +lean_dec(x_4825); +x_4827 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4827, 0, x_4823); +lean_ctor_set(x_4827, 1, x_4824); +return x_4827; +} +else +{ +lean_object* x_4828; lean_object* x_4829; lean_object* x_4830; lean_object* x_4831; +lean_dec(x_4824); +x_4828 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_4829 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4829, 0, x_4828); +x_4830 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_4831 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4830, x_4829, x_4823); +if (lean_obj_tag(x_4831) == 0) +{ +lean_object* x_4832; lean_object* x_4833; lean_object* x_4834; lean_object* x_4835; +lean_dec(x_4825); +x_4832 = lean_ctor_get(x_4831, 0); +lean_inc(x_4832); +x_4833 = lean_ctor_get(x_4831, 1); +lean_inc(x_4833); +if (lean_is_exclusive(x_4831)) { + lean_ctor_release(x_4831, 0); + lean_ctor_release(x_4831, 1); + x_4834 = x_4831; +} else { + lean_dec_ref(x_4831); + x_4834 = lean_box(0); +} +if (lean_is_scalar(x_4834)) { + x_4835 = lean_alloc_ctor(0, 2, 0); +} else { + x_4835 = x_4834; +} +lean_ctor_set(x_4835, 0, x_4832); +lean_ctor_set(x_4835, 1, x_4833); +return x_4835; +} +else +{ +lean_object* x_4836; lean_object* x_4837; lean_object* x_4838; lean_object* x_4839; uint8_t x_4840; +x_4836 = lean_ctor_get(x_4831, 0); +lean_inc(x_4836); +x_4837 = lean_ctor_get(x_4831, 1); +lean_inc(x_4837); +if (lean_is_exclusive(x_4831)) { + lean_ctor_release(x_4831, 0); + lean_ctor_release(x_4831, 1); + x_4838 = x_4831; +} else { + lean_dec_ref(x_4831); + x_4838 = lean_box(0); +} +x_4839 = lean_ctor_get(x_4836, 1); +lean_inc(x_4839); +x_4840 = lean_nat_dec_eq(x_4825, x_4839); +lean_dec(x_4825); +if (x_4840 == 0) +{ +lean_object* x_4841; +lean_dec(x_4839); +if (lean_is_scalar(x_4838)) { + x_4841 = lean_alloc_ctor(1, 2, 0); +} else { + x_4841 = x_4838; +} +lean_ctor_set(x_4841, 0, x_4836); +lean_ctor_set(x_4841, 1, x_4837); +return x_4841; +} +else +{ +lean_object* x_4842; lean_object* x_4843; lean_object* x_4844; lean_object* x_4845; +lean_dec(x_4838); +lean_dec(x_4837); +x_4842 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_4843 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4843, 0, x_4842); +x_4844 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_4845 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4844, x_4843, x_4836); +if (lean_obj_tag(x_4845) == 0) +{ +lean_object* x_4846; lean_object* x_4847; lean_object* x_4848; lean_object* x_4849; +lean_dec(x_4839); +x_4846 = lean_ctor_get(x_4845, 0); +lean_inc(x_4846); +x_4847 = lean_ctor_get(x_4845, 1); +lean_inc(x_4847); +if (lean_is_exclusive(x_4845)) { + lean_ctor_release(x_4845, 0); + lean_ctor_release(x_4845, 1); + x_4848 = x_4845; +} else { + lean_dec_ref(x_4845); + x_4848 = lean_box(0); +} +if (lean_is_scalar(x_4848)) { + x_4849 = lean_alloc_ctor(0, 2, 0); +} else { + x_4849 = x_4848; +} +lean_ctor_set(x_4849, 0, x_4846); +lean_ctor_set(x_4849, 1, x_4847); +return x_4849; +} +else +{ +lean_object* x_4850; lean_object* x_4851; lean_object* x_4852; lean_object* x_4853; uint8_t x_4854; +x_4850 = lean_ctor_get(x_4845, 0); +lean_inc(x_4850); +x_4851 = lean_ctor_get(x_4845, 1); +lean_inc(x_4851); +if (lean_is_exclusive(x_4845)) { + lean_ctor_release(x_4845, 0); + lean_ctor_release(x_4845, 1); + x_4852 = x_4845; +} else { + lean_dec_ref(x_4845); + x_4852 = lean_box(0); +} +x_4853 = lean_ctor_get(x_4850, 1); +lean_inc(x_4853); +x_4854 = lean_nat_dec_eq(x_4839, x_4853); +lean_dec(x_4839); +if (x_4854 == 0) +{ +lean_object* x_4855; +lean_dec(x_4853); +if (lean_is_scalar(x_4852)) { + x_4855 = lean_alloc_ctor(1, 2, 0); +} else { + x_4855 = x_4852; +} +lean_ctor_set(x_4855, 0, x_4850); +lean_ctor_set(x_4855, 1, x_4851); +return x_4855; +} +else +{ +lean_object* x_4856; lean_object* x_4857; lean_object* x_4858; +lean_dec(x_4852); +lean_dec(x_4851); +x_4856 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_4857 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4857, 0, x_4856); +x_4858 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4844, x_4857, x_4850); +if (lean_obj_tag(x_4858) == 0) +{ +lean_object* x_4859; lean_object* x_4860; lean_object* x_4861; lean_object* x_4862; +lean_dec(x_4853); +x_4859 = lean_ctor_get(x_4858, 0); +lean_inc(x_4859); +x_4860 = lean_ctor_get(x_4858, 1); +lean_inc(x_4860); +if (lean_is_exclusive(x_4858)) { + lean_ctor_release(x_4858, 0); + lean_ctor_release(x_4858, 1); + x_4861 = x_4858; +} else { + lean_dec_ref(x_4858); + x_4861 = lean_box(0); +} +if (lean_is_scalar(x_4861)) { + x_4862 = lean_alloc_ctor(0, 2, 0); +} else { + x_4862 = x_4861; +} +lean_ctor_set(x_4862, 0, x_4859); +lean_ctor_set(x_4862, 1, x_4860); +return x_4862; +} +else +{ +lean_object* x_4863; lean_object* x_4864; lean_object* x_4865; lean_object* x_4866; uint8_t x_4867; +x_4863 = lean_ctor_get(x_4858, 0); +lean_inc(x_4863); +x_4864 = lean_ctor_get(x_4858, 1); +lean_inc(x_4864); +if (lean_is_exclusive(x_4858)) { + lean_ctor_release(x_4858, 0); + lean_ctor_release(x_4858, 1); + x_4865 = x_4858; +} else { + lean_dec_ref(x_4858); + x_4865 = lean_box(0); +} +x_4866 = lean_ctor_get(x_4863, 1); +lean_inc(x_4866); +x_4867 = lean_nat_dec_eq(x_4853, x_4866); +lean_dec(x_4853); +if (x_4867 == 0) +{ +lean_object* x_4868; +lean_dec(x_4866); +if (lean_is_scalar(x_4865)) { + x_4868 = lean_alloc_ctor(1, 2, 0); +} else { + x_4868 = x_4865; +} +lean_ctor_set(x_4868, 0, x_4863); +lean_ctor_set(x_4868, 1, x_4864); +return x_4868; +} +else +{ +lean_object* x_4869; lean_object* x_4870; lean_object* x_4871; lean_object* x_4872; +lean_dec(x_4865); +lean_dec(x_4864); +x_4869 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_4870 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4870, 0, x_4869); +x_4871 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_4872 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4871, x_4870, x_4863); +if (lean_obj_tag(x_4872) == 0) +{ +lean_object* x_4873; lean_object* x_4874; lean_object* x_4875; lean_object* x_4876; +lean_dec(x_4866); +x_4873 = lean_ctor_get(x_4872, 0); +lean_inc(x_4873); +x_4874 = lean_ctor_get(x_4872, 1); +lean_inc(x_4874); +if (lean_is_exclusive(x_4872)) { + lean_ctor_release(x_4872, 0); + lean_ctor_release(x_4872, 1); + x_4875 = x_4872; +} else { + lean_dec_ref(x_4872); + x_4875 = lean_box(0); +} +if (lean_is_scalar(x_4875)) { + x_4876 = lean_alloc_ctor(0, 2, 0); +} else { + x_4876 = x_4875; +} +lean_ctor_set(x_4876, 0, x_4873); +lean_ctor_set(x_4876, 1, x_4874); +return x_4876; +} +else +{ +lean_object* x_4877; lean_object* x_4878; lean_object* x_4879; lean_object* x_4880; uint8_t x_4881; +x_4877 = lean_ctor_get(x_4872, 0); +lean_inc(x_4877); +x_4878 = lean_ctor_get(x_4872, 1); +lean_inc(x_4878); +if (lean_is_exclusive(x_4872)) { + lean_ctor_release(x_4872, 0); + lean_ctor_release(x_4872, 1); + x_4879 = x_4872; +} else { + lean_dec_ref(x_4872); + x_4879 = lean_box(0); +} +x_4880 = lean_ctor_get(x_4877, 1); +lean_inc(x_4880); +x_4881 = lean_nat_dec_eq(x_4866, x_4880); +lean_dec(x_4866); +if (x_4881 == 0) +{ +lean_object* x_4882; +lean_dec(x_4880); +if (lean_is_scalar(x_4879)) { + x_4882 = lean_alloc_ctor(1, 2, 0); +} else { + x_4882 = x_4879; +} +lean_ctor_set(x_4882, 0, x_4877); +lean_ctor_set(x_4882, 1, x_4878); +return x_4882; +} +else +{ +lean_object* x_4883; lean_object* x_4884; lean_object* x_4885; lean_object* x_4886; +lean_dec(x_4879); +lean_dec(x_4878); +x_4883 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_4884 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4884, 0, x_4883); +x_4885 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_4886 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4885, x_4884, x_4877); +if (lean_obj_tag(x_4886) == 0) +{ +lean_object* x_4887; lean_object* x_4888; lean_object* x_4889; lean_object* x_4890; +lean_dec(x_4880); +x_4887 = lean_ctor_get(x_4886, 0); +lean_inc(x_4887); +x_4888 = lean_ctor_get(x_4886, 1); +lean_inc(x_4888); +if (lean_is_exclusive(x_4886)) { + lean_ctor_release(x_4886, 0); + lean_ctor_release(x_4886, 1); + x_4889 = x_4886; +} else { + lean_dec_ref(x_4886); + x_4889 = lean_box(0); +} +if (lean_is_scalar(x_4889)) { + x_4890 = lean_alloc_ctor(0, 2, 0); +} else { + x_4890 = x_4889; +} +lean_ctor_set(x_4890, 0, x_4887); +lean_ctor_set(x_4890, 1, x_4888); +return x_4890; +} +else +{ +lean_object* x_4891; lean_object* x_4892; lean_object* x_4893; lean_object* x_4894; uint8_t x_4895; +x_4891 = lean_ctor_get(x_4886, 0); +lean_inc(x_4891); +x_4892 = lean_ctor_get(x_4886, 1); +lean_inc(x_4892); +if (lean_is_exclusive(x_4886)) { + lean_ctor_release(x_4886, 0); + lean_ctor_release(x_4886, 1); + x_4893 = x_4886; +} else { + lean_dec_ref(x_4886); + x_4893 = lean_box(0); +} +x_4894 = lean_ctor_get(x_4891, 1); +lean_inc(x_4894); +x_4895 = lean_nat_dec_eq(x_4880, x_4894); +lean_dec(x_4880); +if (x_4895 == 0) +{ +lean_object* x_4896; +lean_dec(x_4894); +if (lean_is_scalar(x_4893)) { + x_4896 = lean_alloc_ctor(1, 2, 0); +} else { + x_4896 = x_4893; +} +lean_ctor_set(x_4896, 0, x_4891); +lean_ctor_set(x_4896, 1, x_4892); +return x_4896; +} +else +{ +lean_object* x_4897; lean_object* x_4898; lean_object* x_4899; lean_object* x_4900; +lean_dec(x_4893); +lean_dec(x_4892); +x_4897 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_4898 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4898, 0, x_4897); +x_4899 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_4900 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4899, x_4898, x_4891); +if (lean_obj_tag(x_4900) == 0) +{ +lean_object* x_4901; lean_object* x_4902; lean_object* x_4903; lean_object* x_4904; +lean_dec(x_4894); +x_4901 = lean_ctor_get(x_4900, 0); +lean_inc(x_4901); +x_4902 = lean_ctor_get(x_4900, 1); +lean_inc(x_4902); +if (lean_is_exclusive(x_4900)) { + lean_ctor_release(x_4900, 0); + lean_ctor_release(x_4900, 1); + x_4903 = x_4900; +} else { + lean_dec_ref(x_4900); + x_4903 = lean_box(0); +} +if (lean_is_scalar(x_4903)) { + x_4904 = lean_alloc_ctor(0, 2, 0); +} else { + x_4904 = x_4903; +} +lean_ctor_set(x_4904, 0, x_4901); +lean_ctor_set(x_4904, 1, x_4902); +return x_4904; +} +else +{ +lean_object* x_4905; lean_object* x_4906; lean_object* x_4907; lean_object* x_4908; uint8_t x_4909; +x_4905 = lean_ctor_get(x_4900, 0); +lean_inc(x_4905); +x_4906 = lean_ctor_get(x_4900, 1); +lean_inc(x_4906); +if (lean_is_exclusive(x_4900)) { + lean_ctor_release(x_4900, 0); + lean_ctor_release(x_4900, 1); + x_4907 = x_4900; +} else { + lean_dec_ref(x_4900); + x_4907 = lean_box(0); +} +x_4908 = lean_ctor_get(x_4905, 1); +lean_inc(x_4908); +x_4909 = lean_nat_dec_eq(x_4894, x_4908); +lean_dec(x_4894); +if (x_4909 == 0) +{ +lean_object* x_4910; +lean_dec(x_4908); +if (lean_is_scalar(x_4907)) { + x_4910 = lean_alloc_ctor(1, 2, 0); +} else { + x_4910 = x_4907; +} +lean_ctor_set(x_4910, 0, x_4905); +lean_ctor_set(x_4910, 1, x_4906); +return x_4910; +} +else +{ +lean_object* x_4911; lean_object* x_4912; lean_object* x_4913; lean_object* x_4914; +lean_dec(x_4907); +lean_dec(x_4906); +x_4911 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_4912 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4912, 0, x_4911); +x_4913 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_4914 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4913, x_4912, x_4905); +if (lean_obj_tag(x_4914) == 0) +{ +lean_object* x_4915; lean_object* x_4916; lean_object* x_4917; lean_object* x_4918; +lean_dec(x_4908); +x_4915 = lean_ctor_get(x_4914, 0); +lean_inc(x_4915); +x_4916 = lean_ctor_get(x_4914, 1); +lean_inc(x_4916); +if (lean_is_exclusive(x_4914)) { + lean_ctor_release(x_4914, 0); + lean_ctor_release(x_4914, 1); + x_4917 = x_4914; +} else { + lean_dec_ref(x_4914); + x_4917 = lean_box(0); +} +if (lean_is_scalar(x_4917)) { + x_4918 = lean_alloc_ctor(0, 2, 0); +} else { + x_4918 = x_4917; +} +lean_ctor_set(x_4918, 0, x_4915); +lean_ctor_set(x_4918, 1, x_4916); +return x_4918; +} +else +{ +lean_object* x_4919; lean_object* x_4920; lean_object* x_4921; lean_object* x_4922; uint8_t x_4923; +x_4919 = lean_ctor_get(x_4914, 0); +lean_inc(x_4919); +x_4920 = lean_ctor_get(x_4914, 1); +lean_inc(x_4920); +if (lean_is_exclusive(x_4914)) { + lean_ctor_release(x_4914, 0); + lean_ctor_release(x_4914, 1); + x_4921 = x_4914; +} else { + lean_dec_ref(x_4914); + x_4921 = lean_box(0); +} +x_4922 = lean_ctor_get(x_4919, 1); +lean_inc(x_4922); +x_4923 = lean_nat_dec_eq(x_4908, x_4922); +lean_dec(x_4908); +if (x_4923 == 0) +{ +lean_object* x_4924; +lean_dec(x_4922); +if (lean_is_scalar(x_4921)) { + x_4924 = lean_alloc_ctor(1, 2, 0); +} else { + x_4924 = x_4921; +} +lean_ctor_set(x_4924, 0, x_4919); +lean_ctor_set(x_4924, 1, x_4920); +return x_4924; +} +else +{ +lean_object* x_4925; lean_object* x_4926; lean_object* x_4927; +lean_dec(x_4921); +lean_dec(x_4920); +x_4925 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_4926 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4926, 0, x_4925); +x_4927 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4913, x_4926, x_4919); +if (lean_obj_tag(x_4927) == 0) +{ +lean_object* x_4928; lean_object* x_4929; lean_object* x_4930; lean_object* x_4931; +lean_dec(x_4922); +x_4928 = lean_ctor_get(x_4927, 0); +lean_inc(x_4928); +x_4929 = lean_ctor_get(x_4927, 1); +lean_inc(x_4929); +if (lean_is_exclusive(x_4927)) { + lean_ctor_release(x_4927, 0); + lean_ctor_release(x_4927, 1); + x_4930 = x_4927; +} else { + lean_dec_ref(x_4927); + x_4930 = lean_box(0); +} +if (lean_is_scalar(x_4930)) { + x_4931 = lean_alloc_ctor(0, 2, 0); +} else { + x_4931 = x_4930; +} +lean_ctor_set(x_4931, 0, x_4928); +lean_ctor_set(x_4931, 1, x_4929); +return x_4931; +} +else +{ +lean_object* x_4932; lean_object* x_4933; lean_object* x_4934; lean_object* x_4935; uint8_t x_4936; +x_4932 = lean_ctor_get(x_4927, 0); +lean_inc(x_4932); +x_4933 = lean_ctor_get(x_4927, 1); +lean_inc(x_4933); +if (lean_is_exclusive(x_4927)) { + lean_ctor_release(x_4927, 0); + lean_ctor_release(x_4927, 1); + x_4934 = x_4927; +} else { + lean_dec_ref(x_4927); + x_4934 = lean_box(0); +} +x_4935 = lean_ctor_get(x_4932, 1); +lean_inc(x_4935); +x_4936 = lean_nat_dec_eq(x_4922, x_4935); +lean_dec(x_4922); +if (x_4936 == 0) +{ +lean_object* x_4937; +lean_dec(x_4935); +if (lean_is_scalar(x_4934)) { + x_4937 = lean_alloc_ctor(1, 2, 0); +} else { + x_4937 = x_4934; +} +lean_ctor_set(x_4937, 0, x_4932); +lean_ctor_set(x_4937, 1, x_4933); +return x_4937; +} +else +{ +lean_object* x_4938; lean_object* x_4939; lean_object* x_4940; lean_object* x_4941; +lean_dec(x_4934); +lean_dec(x_4933); +x_4938 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_4939 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4939, 0, x_4938); +x_4940 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_4941 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4940, x_4939, x_4932); +if (lean_obj_tag(x_4941) == 0) +{ +lean_object* x_4942; lean_object* x_4943; lean_object* x_4944; lean_object* x_4945; +lean_dec(x_4935); +x_4942 = lean_ctor_get(x_4941, 0); +lean_inc(x_4942); +x_4943 = lean_ctor_get(x_4941, 1); +lean_inc(x_4943); +if (lean_is_exclusive(x_4941)) { + lean_ctor_release(x_4941, 0); + lean_ctor_release(x_4941, 1); + x_4944 = x_4941; +} else { + lean_dec_ref(x_4941); + x_4944 = lean_box(0); +} +if (lean_is_scalar(x_4944)) { + x_4945 = lean_alloc_ctor(0, 2, 0); +} else { + x_4945 = x_4944; +} +lean_ctor_set(x_4945, 0, x_4942); +lean_ctor_set(x_4945, 1, x_4943); +return x_4945; +} +else +{ +lean_object* x_4946; lean_object* x_4947; lean_object* x_4948; lean_object* x_4949; uint8_t x_4950; +x_4946 = lean_ctor_get(x_4941, 0); +lean_inc(x_4946); +x_4947 = lean_ctor_get(x_4941, 1); +lean_inc(x_4947); +if (lean_is_exclusive(x_4941)) { + lean_ctor_release(x_4941, 0); + lean_ctor_release(x_4941, 1); + x_4948 = x_4941; +} else { + lean_dec_ref(x_4941); + x_4948 = lean_box(0); +} +x_4949 = lean_ctor_get(x_4946, 1); +lean_inc(x_4949); +x_4950 = lean_nat_dec_eq(x_4935, x_4949); +lean_dec(x_4935); +if (x_4950 == 0) +{ +lean_object* x_4951; +lean_dec(x_4949); +if (lean_is_scalar(x_4948)) { + x_4951 = lean_alloc_ctor(1, 2, 0); +} else { + x_4951 = x_4948; +} +lean_ctor_set(x_4951, 0, x_4946); +lean_ctor_set(x_4951, 1, x_4947); +return x_4951; +} +else +{ +lean_object* x_4952; lean_object* x_4953; lean_object* x_4954; lean_object* x_4955; +lean_dec(x_4948); +lean_dec(x_4947); +x_4952 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_4953 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4953, 0, x_4952); +x_4954 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_4955 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4954, x_4953, x_4946); +if (lean_obj_tag(x_4955) == 0) +{ +lean_object* x_4956; lean_object* x_4957; lean_object* x_4958; lean_object* x_4959; +lean_dec(x_4949); +x_4956 = lean_ctor_get(x_4955, 0); +lean_inc(x_4956); +x_4957 = lean_ctor_get(x_4955, 1); +lean_inc(x_4957); +if (lean_is_exclusive(x_4955)) { + lean_ctor_release(x_4955, 0); + lean_ctor_release(x_4955, 1); + x_4958 = x_4955; +} else { + lean_dec_ref(x_4955); + x_4958 = lean_box(0); +} +if (lean_is_scalar(x_4958)) { + x_4959 = lean_alloc_ctor(0, 2, 0); +} else { + x_4959 = x_4958; +} +lean_ctor_set(x_4959, 0, x_4956); +lean_ctor_set(x_4959, 1, x_4957); +return x_4959; +} +else +{ +lean_object* x_4960; lean_object* x_4961; lean_object* x_4962; lean_object* x_4963; uint8_t x_4964; +x_4960 = lean_ctor_get(x_4955, 0); +lean_inc(x_4960); +x_4961 = lean_ctor_get(x_4955, 1); +lean_inc(x_4961); +if (lean_is_exclusive(x_4955)) { + lean_ctor_release(x_4955, 0); + lean_ctor_release(x_4955, 1); + x_4962 = x_4955; +} else { + lean_dec_ref(x_4955); + x_4962 = lean_box(0); +} +x_4963 = lean_ctor_get(x_4960, 1); +lean_inc(x_4963); +x_4964 = lean_nat_dec_eq(x_4949, x_4963); +lean_dec(x_4949); +if (x_4964 == 0) +{ +lean_object* x_4965; +lean_dec(x_4963); +if (lean_is_scalar(x_4962)) { + x_4965 = lean_alloc_ctor(1, 2, 0); +} else { + x_4965 = x_4962; +} +lean_ctor_set(x_4965, 0, x_4960); +lean_ctor_set(x_4965, 1, x_4961); +return x_4965; +} +else +{ +lean_object* x_4966; lean_object* x_4967; lean_object* x_4968; lean_object* x_4969; +lean_dec(x_4962); +lean_dec(x_4961); +x_4966 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_4967 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4967, 0, x_4966); +x_4968 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_4969 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4968, x_4967, x_4960); +if (lean_obj_tag(x_4969) == 0) +{ +lean_object* x_4970; lean_object* x_4971; lean_object* x_4972; lean_object* x_4973; +lean_dec(x_4963); +x_4970 = lean_ctor_get(x_4969, 0); +lean_inc(x_4970); +x_4971 = lean_ctor_get(x_4969, 1); +lean_inc(x_4971); +if (lean_is_exclusive(x_4969)) { + lean_ctor_release(x_4969, 0); + lean_ctor_release(x_4969, 1); + x_4972 = x_4969; +} else { + lean_dec_ref(x_4969); + x_4972 = lean_box(0); +} +if (lean_is_scalar(x_4972)) { + x_4973 = lean_alloc_ctor(0, 2, 0); +} else { + x_4973 = x_4972; +} +lean_ctor_set(x_4973, 0, x_4970); +lean_ctor_set(x_4973, 1, x_4971); +return x_4973; +} +else +{ +lean_object* x_4974; lean_object* x_4975; lean_object* x_4976; lean_object* x_4977; uint8_t x_4978; +x_4974 = lean_ctor_get(x_4969, 0); +lean_inc(x_4974); +x_4975 = lean_ctor_get(x_4969, 1); +lean_inc(x_4975); +if (lean_is_exclusive(x_4969)) { + lean_ctor_release(x_4969, 0); + lean_ctor_release(x_4969, 1); + x_4976 = x_4969; +} else { + lean_dec_ref(x_4969); + x_4976 = lean_box(0); +} +x_4977 = lean_ctor_get(x_4974, 1); +lean_inc(x_4977); +x_4978 = lean_nat_dec_eq(x_4963, x_4977); +lean_dec(x_4963); +if (x_4978 == 0) +{ +lean_object* x_4979; +lean_dec(x_4977); +if (lean_is_scalar(x_4976)) { + x_4979 = lean_alloc_ctor(1, 2, 0); +} else { + x_4979 = x_4976; +} +lean_ctor_set(x_4979, 0, x_4974); +lean_ctor_set(x_4979, 1, x_4975); +return x_4979; +} +else +{ +lean_object* x_4980; lean_object* x_4981; lean_object* x_4982; lean_object* x_4983; +lean_dec(x_4976); +lean_dec(x_4975); +x_4980 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_4981 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4981, 0, x_4980); +x_4982 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_4983 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4982, x_4981, x_4974); +if (lean_obj_tag(x_4983) == 0) +{ +lean_object* x_4984; lean_object* x_4985; lean_object* x_4986; lean_object* x_4987; +lean_dec(x_4977); +x_4984 = lean_ctor_get(x_4983, 0); +lean_inc(x_4984); +x_4985 = lean_ctor_get(x_4983, 1); +lean_inc(x_4985); +if (lean_is_exclusive(x_4983)) { + lean_ctor_release(x_4983, 0); + lean_ctor_release(x_4983, 1); + x_4986 = x_4983; +} else { + lean_dec_ref(x_4983); + x_4986 = lean_box(0); +} +if (lean_is_scalar(x_4986)) { + x_4987 = lean_alloc_ctor(0, 2, 0); +} else { + x_4987 = x_4986; +} +lean_ctor_set(x_4987, 0, x_4984); +lean_ctor_set(x_4987, 1, x_4985); +return x_4987; +} +else +{ +lean_object* x_4988; lean_object* x_4989; lean_object* x_4990; lean_object* x_4991; uint8_t x_4992; +x_4988 = lean_ctor_get(x_4983, 0); +lean_inc(x_4988); +x_4989 = lean_ctor_get(x_4983, 1); +lean_inc(x_4989); +if (lean_is_exclusive(x_4983)) { + lean_ctor_release(x_4983, 0); + lean_ctor_release(x_4983, 1); + x_4990 = x_4983; +} else { + lean_dec_ref(x_4983); + x_4990 = lean_box(0); +} +x_4991 = lean_ctor_get(x_4988, 1); +lean_inc(x_4991); +x_4992 = lean_nat_dec_eq(x_4977, x_4991); +lean_dec(x_4977); +if (x_4992 == 0) +{ +lean_object* x_4993; +lean_dec(x_4991); +if (lean_is_scalar(x_4990)) { + x_4993 = lean_alloc_ctor(1, 2, 0); +} else { + x_4993 = x_4990; +} +lean_ctor_set(x_4993, 0, x_4988); +lean_ctor_set(x_4993, 1, x_4989); +return x_4993; +} +else +{ +lean_object* x_4994; lean_object* x_4995; lean_object* x_4996; lean_object* x_4997; +lean_dec(x_4990); +lean_dec(x_4989); +x_4994 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_4995 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_4995, 0, x_4994); +x_4996 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_4997 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_4996, x_4995, x_4988); +if (lean_obj_tag(x_4997) == 0) +{ +lean_object* x_4998; lean_object* x_4999; lean_object* x_5000; lean_object* x_5001; +lean_dec(x_4991); +x_4998 = lean_ctor_get(x_4997, 0); +lean_inc(x_4998); +x_4999 = lean_ctor_get(x_4997, 1); +lean_inc(x_4999); +if (lean_is_exclusive(x_4997)) { + lean_ctor_release(x_4997, 0); + lean_ctor_release(x_4997, 1); + x_5000 = x_4997; +} else { + lean_dec_ref(x_4997); + x_5000 = lean_box(0); +} +if (lean_is_scalar(x_5000)) { + x_5001 = lean_alloc_ctor(0, 2, 0); +} else { + x_5001 = x_5000; +} +lean_ctor_set(x_5001, 0, x_4998); +lean_ctor_set(x_5001, 1, x_4999); +return x_5001; +} +else +{ +lean_object* x_5002; lean_object* x_5003; lean_object* x_5004; lean_object* x_5005; uint8_t x_5006; +x_5002 = lean_ctor_get(x_4997, 0); +lean_inc(x_5002); +x_5003 = lean_ctor_get(x_4997, 1); +lean_inc(x_5003); +if (lean_is_exclusive(x_4997)) { + lean_ctor_release(x_4997, 0); + lean_ctor_release(x_4997, 1); + x_5004 = x_4997; +} else { + lean_dec_ref(x_4997); + x_5004 = lean_box(0); +} +x_5005 = lean_ctor_get(x_5002, 1); +lean_inc(x_5005); +x_5006 = lean_nat_dec_eq(x_4991, x_5005); +lean_dec(x_4991); +if (x_5006 == 0) +{ +lean_object* x_5007; +lean_dec(x_5005); +if (lean_is_scalar(x_5004)) { + x_5007 = lean_alloc_ctor(1, 2, 0); +} else { + x_5007 = x_5004; +} +lean_ctor_set(x_5007, 0, x_5002); +lean_ctor_set(x_5007, 1, x_5003); +return x_5007; +} +else +{ +lean_object* x_5008; lean_object* x_5009; lean_object* x_5010; lean_object* x_5011; +lean_dec(x_5004); +lean_dec(x_5003); +x_5008 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_5009 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5009, 0, x_5008); +x_5010 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_5011 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5010, x_5009, x_5002); +if (lean_obj_tag(x_5011) == 0) +{ +lean_object* x_5012; lean_object* x_5013; lean_object* x_5014; lean_object* x_5015; +lean_dec(x_5005); +x_5012 = lean_ctor_get(x_5011, 0); +lean_inc(x_5012); +x_5013 = lean_ctor_get(x_5011, 1); +lean_inc(x_5013); +if (lean_is_exclusive(x_5011)) { + lean_ctor_release(x_5011, 0); + lean_ctor_release(x_5011, 1); + x_5014 = x_5011; +} else { + lean_dec_ref(x_5011); + x_5014 = lean_box(0); +} +if (lean_is_scalar(x_5014)) { + x_5015 = lean_alloc_ctor(0, 2, 0); +} else { + x_5015 = x_5014; +} +lean_ctor_set(x_5015, 0, x_5012); +lean_ctor_set(x_5015, 1, x_5013); +return x_5015; +} +else +{ +lean_object* x_5016; lean_object* x_5017; lean_object* x_5018; lean_object* x_5019; uint8_t x_5020; +x_5016 = lean_ctor_get(x_5011, 0); +lean_inc(x_5016); +x_5017 = lean_ctor_get(x_5011, 1); +lean_inc(x_5017); +if (lean_is_exclusive(x_5011)) { + lean_ctor_release(x_5011, 0); + lean_ctor_release(x_5011, 1); + x_5018 = x_5011; +} else { + lean_dec_ref(x_5011); + x_5018 = lean_box(0); +} +x_5019 = lean_ctor_get(x_5016, 1); +lean_inc(x_5019); +x_5020 = lean_nat_dec_eq(x_5005, x_5019); +lean_dec(x_5005); +if (x_5020 == 0) +{ +lean_object* x_5021; +lean_dec(x_5019); +if (lean_is_scalar(x_5018)) { + x_5021 = lean_alloc_ctor(1, 2, 0); +} else { + x_5021 = x_5018; +} +lean_ctor_set(x_5021, 0, x_5016); +lean_ctor_set(x_5021, 1, x_5017); +return x_5021; +} +else +{ +lean_object* x_5022; lean_object* x_5023; lean_object* x_5024; lean_object* x_5025; +lean_dec(x_5018); +lean_dec(x_5017); +x_5022 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_5023 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5023, 0, x_5022); +x_5024 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_5025 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5024, x_5023, x_5016); +if (lean_obj_tag(x_5025) == 0) +{ +lean_object* x_5026; lean_object* x_5027; lean_object* x_5028; lean_object* x_5029; +lean_dec(x_5019); +x_5026 = lean_ctor_get(x_5025, 0); +lean_inc(x_5026); +x_5027 = lean_ctor_get(x_5025, 1); +lean_inc(x_5027); +if (lean_is_exclusive(x_5025)) { + lean_ctor_release(x_5025, 0); + lean_ctor_release(x_5025, 1); + x_5028 = x_5025; +} else { + lean_dec_ref(x_5025); + x_5028 = lean_box(0); +} +if (lean_is_scalar(x_5028)) { + x_5029 = lean_alloc_ctor(0, 2, 0); +} else { + x_5029 = x_5028; +} +lean_ctor_set(x_5029, 0, x_5026); +lean_ctor_set(x_5029, 1, x_5027); +return x_5029; +} +else +{ +lean_object* x_5030; lean_object* x_5031; lean_object* x_5032; lean_object* x_5033; uint8_t x_5034; +x_5030 = lean_ctor_get(x_5025, 0); +lean_inc(x_5030); +x_5031 = lean_ctor_get(x_5025, 1); +lean_inc(x_5031); +if (lean_is_exclusive(x_5025)) { + lean_ctor_release(x_5025, 0); + lean_ctor_release(x_5025, 1); + x_5032 = x_5025; +} else { + lean_dec_ref(x_5025); + x_5032 = lean_box(0); +} +x_5033 = lean_ctor_get(x_5030, 1); +lean_inc(x_5033); +x_5034 = lean_nat_dec_eq(x_5019, x_5033); +lean_dec(x_5019); +if (x_5034 == 0) +{ +lean_object* x_5035; +lean_dec(x_5033); +if (lean_is_scalar(x_5032)) { + x_5035 = lean_alloc_ctor(1, 2, 0); +} else { + x_5035 = x_5032; +} +lean_ctor_set(x_5035, 0, x_5030); +lean_ctor_set(x_5035, 1, x_5031); +return x_5035; +} +else +{ +lean_object* x_5036; lean_object* x_5037; lean_object* x_5038; lean_object* x_5039; +lean_dec(x_5032); +lean_dec(x_5031); +x_5036 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_5037 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5037, 0, x_5036); +x_5038 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_5039 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5038, x_5037, x_5030); +if (lean_obj_tag(x_5039) == 0) +{ +lean_object* x_5040; lean_object* x_5041; lean_object* x_5042; lean_object* x_5043; +lean_dec(x_5033); +x_5040 = lean_ctor_get(x_5039, 0); +lean_inc(x_5040); +x_5041 = lean_ctor_get(x_5039, 1); +lean_inc(x_5041); +if (lean_is_exclusive(x_5039)) { + lean_ctor_release(x_5039, 0); + lean_ctor_release(x_5039, 1); + x_5042 = x_5039; +} else { + lean_dec_ref(x_5039); + x_5042 = lean_box(0); +} +if (lean_is_scalar(x_5042)) { + x_5043 = lean_alloc_ctor(0, 2, 0); +} else { + x_5043 = x_5042; +} +lean_ctor_set(x_5043, 0, x_5040); +lean_ctor_set(x_5043, 1, x_5041); +return x_5043; +} +else +{ +lean_object* x_5044; lean_object* x_5045; lean_object* x_5046; lean_object* x_5047; uint8_t x_5048; +x_5044 = lean_ctor_get(x_5039, 0); +lean_inc(x_5044); +x_5045 = lean_ctor_get(x_5039, 1); +lean_inc(x_5045); +if (lean_is_exclusive(x_5039)) { + lean_ctor_release(x_5039, 0); + lean_ctor_release(x_5039, 1); + x_5046 = x_5039; +} else { + lean_dec_ref(x_5039); + x_5046 = lean_box(0); +} +x_5047 = lean_ctor_get(x_5044, 1); +lean_inc(x_5047); +x_5048 = lean_nat_dec_eq(x_5033, x_5047); +lean_dec(x_5033); +if (x_5048 == 0) +{ +lean_object* x_5049; +lean_dec(x_5047); +if (lean_is_scalar(x_5046)) { + x_5049 = lean_alloc_ctor(1, 2, 0); +} else { + x_5049 = x_5046; +} +lean_ctor_set(x_5049, 0, x_5044); +lean_ctor_set(x_5049, 1, x_5045); +return x_5049; +} +else +{ +lean_object* x_5050; lean_object* x_5051; lean_object* x_5052; lean_object* x_5053; +lean_dec(x_5046); +lean_dec(x_5045); +x_5050 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_5051 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5051, 0, x_5050); +x_5052 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_5053 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5052, x_5051, x_5044); +if (lean_obj_tag(x_5053) == 0) +{ +lean_object* x_5054; lean_object* x_5055; lean_object* x_5056; lean_object* x_5057; +lean_dec(x_5047); +x_5054 = lean_ctor_get(x_5053, 0); +lean_inc(x_5054); +x_5055 = lean_ctor_get(x_5053, 1); +lean_inc(x_5055); +if (lean_is_exclusive(x_5053)) { + lean_ctor_release(x_5053, 0); + lean_ctor_release(x_5053, 1); + x_5056 = x_5053; +} else { + lean_dec_ref(x_5053); + x_5056 = lean_box(0); +} +if (lean_is_scalar(x_5056)) { + x_5057 = lean_alloc_ctor(0, 2, 0); +} else { + x_5057 = x_5056; +} +lean_ctor_set(x_5057, 0, x_5054); +lean_ctor_set(x_5057, 1, x_5055); +return x_5057; +} +else +{ +lean_object* x_5058; lean_object* x_5059; lean_object* x_5060; lean_object* x_5061; uint8_t x_5062; +x_5058 = lean_ctor_get(x_5053, 0); +lean_inc(x_5058); +x_5059 = lean_ctor_get(x_5053, 1); +lean_inc(x_5059); +if (lean_is_exclusive(x_5053)) { + lean_ctor_release(x_5053, 0); + lean_ctor_release(x_5053, 1); + x_5060 = x_5053; +} else { + lean_dec_ref(x_5053); + x_5060 = lean_box(0); +} +x_5061 = lean_ctor_get(x_5058, 1); +lean_inc(x_5061); +x_5062 = lean_nat_dec_eq(x_5047, x_5061); +lean_dec(x_5047); +if (x_5062 == 0) +{ +lean_object* x_5063; +lean_dec(x_5061); +if (lean_is_scalar(x_5060)) { + x_5063 = lean_alloc_ctor(1, 2, 0); +} else { + x_5063 = x_5060; +} +lean_ctor_set(x_5063, 0, x_5058); +lean_ctor_set(x_5063, 1, x_5059); +return x_5063; +} +else +{ +lean_object* x_5064; lean_object* x_5065; lean_object* x_5066; lean_object* x_5067; +lean_dec(x_5060); +lean_dec(x_5059); +x_5064 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_5065 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5065, 0, x_5064); +x_5066 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_5067 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5066, x_5065, x_5058); +if (lean_obj_tag(x_5067) == 0) +{ +lean_object* x_5068; lean_object* x_5069; lean_object* x_5070; lean_object* x_5071; +lean_dec(x_5061); +x_5068 = lean_ctor_get(x_5067, 0); +lean_inc(x_5068); +x_5069 = lean_ctor_get(x_5067, 1); +lean_inc(x_5069); +if (lean_is_exclusive(x_5067)) { + lean_ctor_release(x_5067, 0); + lean_ctor_release(x_5067, 1); + x_5070 = x_5067; +} else { + lean_dec_ref(x_5067); + x_5070 = lean_box(0); +} +if (lean_is_scalar(x_5070)) { + x_5071 = lean_alloc_ctor(0, 2, 0); +} else { + x_5071 = x_5070; +} +lean_ctor_set(x_5071, 0, x_5068); +lean_ctor_set(x_5071, 1, x_5069); +return x_5071; +} +else +{ +lean_object* x_5072; lean_object* x_5073; lean_object* x_5074; lean_object* x_5075; uint8_t x_5076; +x_5072 = lean_ctor_get(x_5067, 0); +lean_inc(x_5072); +x_5073 = lean_ctor_get(x_5067, 1); +lean_inc(x_5073); +if (lean_is_exclusive(x_5067)) { + lean_ctor_release(x_5067, 0); + lean_ctor_release(x_5067, 1); + x_5074 = x_5067; +} else { + lean_dec_ref(x_5067); + x_5074 = lean_box(0); +} +x_5075 = lean_ctor_get(x_5072, 1); +lean_inc(x_5075); +x_5076 = lean_nat_dec_eq(x_5061, x_5075); +lean_dec(x_5061); +if (x_5076 == 0) +{ +lean_object* x_5077; +lean_dec(x_5075); +if (lean_is_scalar(x_5074)) { + x_5077 = lean_alloc_ctor(1, 2, 0); +} else { + x_5077 = x_5074; +} +lean_ctor_set(x_5077, 0, x_5072); +lean_ctor_set(x_5077, 1, x_5073); +return x_5077; +} +else +{ +lean_object* x_5078; lean_object* x_5079; lean_object* x_5080; lean_object* x_5081; +lean_dec(x_5074); +lean_dec(x_5073); +x_5078 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_5079 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5079, 0, x_5078); +x_5080 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_5081 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5080, x_5079, x_5072); +if (lean_obj_tag(x_5081) == 0) +{ +lean_object* x_5082; lean_object* x_5083; lean_object* x_5084; lean_object* x_5085; +lean_dec(x_5075); +x_5082 = lean_ctor_get(x_5081, 0); +lean_inc(x_5082); +x_5083 = lean_ctor_get(x_5081, 1); +lean_inc(x_5083); +if (lean_is_exclusive(x_5081)) { + lean_ctor_release(x_5081, 0); + lean_ctor_release(x_5081, 1); + x_5084 = x_5081; +} else { + lean_dec_ref(x_5081); + x_5084 = lean_box(0); +} +if (lean_is_scalar(x_5084)) { + x_5085 = lean_alloc_ctor(0, 2, 0); +} else { + x_5085 = x_5084; +} +lean_ctor_set(x_5085, 0, x_5082); +lean_ctor_set(x_5085, 1, x_5083); +return x_5085; +} +else +{ +lean_object* x_5086; lean_object* x_5087; lean_object* x_5088; lean_object* x_5089; uint8_t x_5090; +x_5086 = lean_ctor_get(x_5081, 0); +lean_inc(x_5086); +x_5087 = lean_ctor_get(x_5081, 1); +lean_inc(x_5087); +if (lean_is_exclusive(x_5081)) { + lean_ctor_release(x_5081, 0); + lean_ctor_release(x_5081, 1); + x_5088 = x_5081; +} else { + lean_dec_ref(x_5081); + x_5088 = lean_box(0); +} +x_5089 = lean_ctor_get(x_5086, 1); +lean_inc(x_5089); +x_5090 = lean_nat_dec_eq(x_5075, x_5089); +lean_dec(x_5075); +if (x_5090 == 0) +{ +lean_object* x_5091; +lean_dec(x_5089); +if (lean_is_scalar(x_5088)) { + x_5091 = lean_alloc_ctor(1, 2, 0); +} else { + x_5091 = x_5088; +} +lean_ctor_set(x_5091, 0, x_5086); +lean_ctor_set(x_5091, 1, x_5087); +return x_5091; +} +else +{ +lean_object* x_5092; lean_object* x_5093; lean_object* x_5094; lean_object* x_5095; +lean_dec(x_5088); +lean_dec(x_5087); +x_5092 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_5093 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5093, 0, x_5092); +x_5094 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_5095 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5094, x_5093, x_5086); +if (lean_obj_tag(x_5095) == 0) +{ +lean_object* x_5096; lean_object* x_5097; lean_object* x_5098; lean_object* x_5099; +lean_dec(x_5089); +x_5096 = lean_ctor_get(x_5095, 0); +lean_inc(x_5096); +x_5097 = lean_ctor_get(x_5095, 1); +lean_inc(x_5097); +if (lean_is_exclusive(x_5095)) { + lean_ctor_release(x_5095, 0); + lean_ctor_release(x_5095, 1); + x_5098 = x_5095; +} else { + lean_dec_ref(x_5095); + x_5098 = lean_box(0); +} +if (lean_is_scalar(x_5098)) { + x_5099 = lean_alloc_ctor(0, 2, 0); +} else { + x_5099 = x_5098; +} +lean_ctor_set(x_5099, 0, x_5096); +lean_ctor_set(x_5099, 1, x_5097); +return x_5099; +} +else +{ +lean_object* x_5100; lean_object* x_5101; lean_object* x_5102; lean_object* x_5103; uint8_t x_5104; +x_5100 = lean_ctor_get(x_5095, 0); +lean_inc(x_5100); +x_5101 = lean_ctor_get(x_5095, 1); +lean_inc(x_5101); +if (lean_is_exclusive(x_5095)) { + lean_ctor_release(x_5095, 0); + lean_ctor_release(x_5095, 1); + x_5102 = x_5095; +} else { + lean_dec_ref(x_5095); + x_5102 = lean_box(0); +} +x_5103 = lean_ctor_get(x_5100, 1); +lean_inc(x_5103); +x_5104 = lean_nat_dec_eq(x_5089, x_5103); +lean_dec(x_5089); +if (x_5104 == 0) +{ +lean_object* x_5105; +lean_dec(x_5103); +if (lean_is_scalar(x_5102)) { + x_5105 = lean_alloc_ctor(1, 2, 0); +} else { + x_5105 = x_5102; +} +lean_ctor_set(x_5105, 0, x_5100); +lean_ctor_set(x_5105, 1, x_5101); +return x_5105; +} +else +{ +lean_object* x_5106; lean_object* x_5107; lean_object* x_5108; lean_object* x_5109; +lean_dec(x_5102); +lean_dec(x_5101); +x_5106 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_5107 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5107, 0, x_5106); +x_5108 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_5109 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5108, x_5107, x_5100); +if (lean_obj_tag(x_5109) == 0) +{ +lean_object* x_5110; lean_object* x_5111; lean_object* x_5112; lean_object* x_5113; +lean_dec(x_5103); +x_5110 = lean_ctor_get(x_5109, 0); +lean_inc(x_5110); +x_5111 = lean_ctor_get(x_5109, 1); +lean_inc(x_5111); +if (lean_is_exclusive(x_5109)) { + lean_ctor_release(x_5109, 0); + lean_ctor_release(x_5109, 1); + x_5112 = x_5109; +} else { + lean_dec_ref(x_5109); + x_5112 = lean_box(0); +} +if (lean_is_scalar(x_5112)) { + x_5113 = lean_alloc_ctor(0, 2, 0); +} else { + x_5113 = x_5112; +} +lean_ctor_set(x_5113, 0, x_5110); +lean_ctor_set(x_5113, 1, x_5111); +return x_5113; +} +else +{ +lean_object* x_5114; lean_object* x_5115; lean_object* x_5116; lean_object* x_5117; uint8_t x_5118; +x_5114 = lean_ctor_get(x_5109, 0); +lean_inc(x_5114); +x_5115 = lean_ctor_get(x_5109, 1); +lean_inc(x_5115); +if (lean_is_exclusive(x_5109)) { + lean_ctor_release(x_5109, 0); + lean_ctor_release(x_5109, 1); + x_5116 = x_5109; +} else { + lean_dec_ref(x_5109); + x_5116 = lean_box(0); +} +x_5117 = lean_ctor_get(x_5114, 1); +lean_inc(x_5117); +x_5118 = lean_nat_dec_eq(x_5103, x_5117); +lean_dec(x_5103); +if (x_5118 == 0) +{ +lean_object* x_5119; +lean_dec(x_5117); +if (lean_is_scalar(x_5116)) { + x_5119 = lean_alloc_ctor(1, 2, 0); +} else { + x_5119 = x_5116; +} +lean_ctor_set(x_5119, 0, x_5114); +lean_ctor_set(x_5119, 1, x_5115); +return x_5119; +} +else +{ +lean_object* x_5120; lean_object* x_5121; lean_object* x_5122; lean_object* x_5123; +lean_dec(x_5116); +lean_dec(x_5115); +x_5120 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_5121 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5121, 0, x_5120); +x_5122 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_5123 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5122, x_5121, x_5114); +if (lean_obj_tag(x_5123) == 0) +{ +lean_object* x_5124; lean_object* x_5125; lean_object* x_5126; lean_object* x_5127; +lean_dec(x_5117); +x_5124 = lean_ctor_get(x_5123, 0); +lean_inc(x_5124); +x_5125 = lean_ctor_get(x_5123, 1); +lean_inc(x_5125); +if (lean_is_exclusive(x_5123)) { + lean_ctor_release(x_5123, 0); + lean_ctor_release(x_5123, 1); + x_5126 = x_5123; +} else { + lean_dec_ref(x_5123); + x_5126 = lean_box(0); +} +if (lean_is_scalar(x_5126)) { + x_5127 = lean_alloc_ctor(0, 2, 0); +} else { + x_5127 = x_5126; +} +lean_ctor_set(x_5127, 0, x_5124); +lean_ctor_set(x_5127, 1, x_5125); +return x_5127; +} +else +{ +lean_object* x_5128; lean_object* x_5129; lean_object* x_5130; lean_object* x_5131; uint8_t x_5132; +x_5128 = lean_ctor_get(x_5123, 0); +lean_inc(x_5128); +x_5129 = lean_ctor_get(x_5123, 1); +lean_inc(x_5129); +if (lean_is_exclusive(x_5123)) { + lean_ctor_release(x_5123, 0); + lean_ctor_release(x_5123, 1); + x_5130 = x_5123; +} else { + lean_dec_ref(x_5123); + x_5130 = lean_box(0); +} +x_5131 = lean_ctor_get(x_5128, 1); +lean_inc(x_5131); +x_5132 = lean_nat_dec_eq(x_5117, x_5131); +lean_dec(x_5117); +if (x_5132 == 0) +{ +lean_object* x_5133; +lean_dec(x_5131); +if (lean_is_scalar(x_5130)) { + x_5133 = lean_alloc_ctor(1, 2, 0); +} else { + x_5133 = x_5130; +} +lean_ctor_set(x_5133, 0, x_5128); +lean_ctor_set(x_5133, 1, x_5129); +return x_5133; +} +else +{ +lean_object* x_5134; lean_object* x_5135; lean_object* x_5136; lean_object* x_5137; +lean_dec(x_5130); +lean_dec(x_5129); +x_5134 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_5135 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5135, 0, x_5134); +x_5136 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_5137 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5136, x_5135, x_5128); +if (lean_obj_tag(x_5137) == 0) +{ +lean_object* x_5138; lean_object* x_5139; lean_object* x_5140; lean_object* x_5141; +lean_dec(x_5131); +x_5138 = lean_ctor_get(x_5137, 0); +lean_inc(x_5138); +x_5139 = lean_ctor_get(x_5137, 1); +lean_inc(x_5139); +if (lean_is_exclusive(x_5137)) { + lean_ctor_release(x_5137, 0); + lean_ctor_release(x_5137, 1); + x_5140 = x_5137; +} else { + lean_dec_ref(x_5137); + x_5140 = lean_box(0); +} +if (lean_is_scalar(x_5140)) { + x_5141 = lean_alloc_ctor(0, 2, 0); +} else { + x_5141 = x_5140; +} +lean_ctor_set(x_5141, 0, x_5138); +lean_ctor_set(x_5141, 1, x_5139); +return x_5141; +} +else +{ +lean_object* x_5142; lean_object* x_5143; lean_object* x_5144; lean_object* x_5145; uint8_t x_5146; +x_5142 = lean_ctor_get(x_5137, 0); +lean_inc(x_5142); +x_5143 = lean_ctor_get(x_5137, 1); +lean_inc(x_5143); +if (lean_is_exclusive(x_5137)) { + lean_ctor_release(x_5137, 0); + lean_ctor_release(x_5137, 1); + x_5144 = x_5137; +} else { + lean_dec_ref(x_5137); + x_5144 = lean_box(0); +} +x_5145 = lean_ctor_get(x_5142, 1); +lean_inc(x_5145); +x_5146 = lean_nat_dec_eq(x_5131, x_5145); +lean_dec(x_5131); +if (x_5146 == 0) +{ +lean_object* x_5147; +lean_dec(x_5145); +if (lean_is_scalar(x_5144)) { + x_5147 = lean_alloc_ctor(1, 2, 0); +} else { + x_5147 = x_5144; +} +lean_ctor_set(x_5147, 0, x_5142); +lean_ctor_set(x_5147, 1, x_5143); +return x_5147; +} +else +{ +lean_object* x_5148; lean_object* x_5149; lean_object* x_5150; lean_object* x_5151; +lean_dec(x_5144); +lean_dec(x_5143); +x_5148 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_5149 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5149, 0, x_5148); +x_5150 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_5151 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5150, x_5149, x_5142); +if (lean_obj_tag(x_5151) == 0) +{ +lean_object* x_5152; lean_object* x_5153; lean_object* x_5154; lean_object* x_5155; +lean_dec(x_5145); +x_5152 = lean_ctor_get(x_5151, 0); +lean_inc(x_5152); +x_5153 = lean_ctor_get(x_5151, 1); +lean_inc(x_5153); +if (lean_is_exclusive(x_5151)) { + lean_ctor_release(x_5151, 0); + lean_ctor_release(x_5151, 1); + x_5154 = x_5151; +} else { + lean_dec_ref(x_5151); + x_5154 = lean_box(0); +} +if (lean_is_scalar(x_5154)) { + x_5155 = lean_alloc_ctor(0, 2, 0); +} else { + x_5155 = x_5154; +} +lean_ctor_set(x_5155, 0, x_5152); +lean_ctor_set(x_5155, 1, x_5153); +return x_5155; +} +else +{ +lean_object* x_5156; lean_object* x_5157; lean_object* x_5158; lean_object* x_5159; uint8_t x_5160; +x_5156 = lean_ctor_get(x_5151, 0); +lean_inc(x_5156); +x_5157 = lean_ctor_get(x_5151, 1); +lean_inc(x_5157); +if (lean_is_exclusive(x_5151)) { + lean_ctor_release(x_5151, 0); + lean_ctor_release(x_5151, 1); + x_5158 = x_5151; +} else { + lean_dec_ref(x_5151); + x_5158 = lean_box(0); +} +x_5159 = lean_ctor_get(x_5156, 1); +lean_inc(x_5159); +x_5160 = lean_nat_dec_eq(x_5145, x_5159); +lean_dec(x_5145); +if (x_5160 == 0) +{ +lean_object* x_5161; +lean_dec(x_5159); +if (lean_is_scalar(x_5158)) { + x_5161 = lean_alloc_ctor(1, 2, 0); +} else { + x_5161 = x_5158; +} +lean_ctor_set(x_5161, 0, x_5156); +lean_ctor_set(x_5161, 1, x_5157); +return x_5161; +} +else +{ +lean_object* x_5162; lean_object* x_5163; lean_object* x_5164; lean_object* x_5165; +lean_dec(x_5158); +lean_dec(x_5157); +x_5162 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_5163 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5163, 0, x_5162); +x_5164 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_5165 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5164, x_5163, x_5156); +if (lean_obj_tag(x_5165) == 0) +{ +lean_object* x_5166; lean_object* x_5167; lean_object* x_5168; lean_object* x_5169; +lean_dec(x_5159); +x_5166 = lean_ctor_get(x_5165, 0); +lean_inc(x_5166); +x_5167 = lean_ctor_get(x_5165, 1); +lean_inc(x_5167); +if (lean_is_exclusive(x_5165)) { + lean_ctor_release(x_5165, 0); + lean_ctor_release(x_5165, 1); + x_5168 = x_5165; +} else { + lean_dec_ref(x_5165); + x_5168 = lean_box(0); +} +if (lean_is_scalar(x_5168)) { + x_5169 = lean_alloc_ctor(0, 2, 0); +} else { + x_5169 = x_5168; +} +lean_ctor_set(x_5169, 0, x_5166); +lean_ctor_set(x_5169, 1, x_5167); +return x_5169; +} +else +{ +lean_object* x_5170; lean_object* x_5171; lean_object* x_5172; lean_object* x_5173; uint8_t x_5174; +x_5170 = lean_ctor_get(x_5165, 0); +lean_inc(x_5170); +x_5171 = lean_ctor_get(x_5165, 1); +lean_inc(x_5171); +if (lean_is_exclusive(x_5165)) { + lean_ctor_release(x_5165, 0); + lean_ctor_release(x_5165, 1); + x_5172 = x_5165; +} else { + lean_dec_ref(x_5165); + x_5172 = lean_box(0); +} +x_5173 = lean_ctor_get(x_5170, 1); +lean_inc(x_5173); +x_5174 = lean_nat_dec_eq(x_5159, x_5173); +lean_dec(x_5173); +lean_dec(x_5159); +if (x_5174 == 0) +{ +lean_object* x_5175; +if (lean_is_scalar(x_5172)) { + x_5175 = lean_alloc_ctor(1, 2, 0); +} else { + x_5175 = x_5172; +} +lean_ctor_set(x_5175, 0, x_5170); +lean_ctor_set(x_5175, 1, x_5171); +return x_5175; +} +else +{ +lean_object* x_5176; lean_object* x_5177; lean_object* x_5178; lean_object* x_5179; +lean_dec(x_5172); +lean_dec(x_5171); +x_5176 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_5177 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5177, 0, x_5176); +x_5178 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_5179 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5178, x_5177, x_5170); +return x_5179; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_5180; lean_object* x_5181; lean_object* x_5182; uint8_t x_5183; +x_5180 = lean_ctor_get(x_58, 0); +x_5181 = lean_ctor_get(x_58, 1); +lean_inc(x_5181); +lean_inc(x_5180); +lean_dec(x_58); +x_5182 = lean_ctor_get(x_5180, 1); +lean_inc(x_5182); +x_5183 = lean_nat_dec_eq(x_53, x_5182); +lean_dec(x_53); +if (x_5183 == 0) +{ +lean_object* x_5184; +lean_dec(x_5182); +x_5184 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5184, 0, x_5180); +lean_ctor_set(x_5184, 1, x_5181); +return x_5184; +} +else +{ +lean_object* x_5185; lean_object* x_5186; lean_object* x_5187; +lean_dec(x_5181); +x_5185 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_5186 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5186, 0, x_5185); +x_5187 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_57, x_5186, x_5180); +if (lean_obj_tag(x_5187) == 0) +{ +lean_object* x_5188; lean_object* x_5189; lean_object* x_5190; lean_object* x_5191; +lean_dec(x_5182); +x_5188 = lean_ctor_get(x_5187, 0); +lean_inc(x_5188); +x_5189 = lean_ctor_get(x_5187, 1); +lean_inc(x_5189); +if (lean_is_exclusive(x_5187)) { + lean_ctor_release(x_5187, 0); + lean_ctor_release(x_5187, 1); + x_5190 = x_5187; +} else { + lean_dec_ref(x_5187); + x_5190 = lean_box(0); +} +if (lean_is_scalar(x_5190)) { + x_5191 = lean_alloc_ctor(0, 2, 0); +} else { + x_5191 = x_5190; +} +lean_ctor_set(x_5191, 0, x_5188); +lean_ctor_set(x_5191, 1, x_5189); +return x_5191; +} +else +{ +lean_object* x_5192; lean_object* x_5193; lean_object* x_5194; lean_object* x_5195; uint8_t x_5196; +x_5192 = lean_ctor_get(x_5187, 0); +lean_inc(x_5192); +x_5193 = lean_ctor_get(x_5187, 1); +lean_inc(x_5193); +if (lean_is_exclusive(x_5187)) { + lean_ctor_release(x_5187, 0); + lean_ctor_release(x_5187, 1); + x_5194 = x_5187; +} else { + lean_dec_ref(x_5187); + x_5194 = lean_box(0); +} +x_5195 = lean_ctor_get(x_5192, 1); +lean_inc(x_5195); +x_5196 = lean_nat_dec_eq(x_5182, x_5195); +lean_dec(x_5182); +if (x_5196 == 0) +{ +lean_object* x_5197; +lean_dec(x_5195); +if (lean_is_scalar(x_5194)) { + x_5197 = lean_alloc_ctor(1, 2, 0); +} else { + x_5197 = x_5194; +} +lean_ctor_set(x_5197, 0, x_5192); +lean_ctor_set(x_5197, 1, x_5193); +return x_5197; +} +else +{ +lean_object* x_5198; lean_object* x_5199; lean_object* x_5200; lean_object* x_5201; +lean_dec(x_5194); +lean_dec(x_5193); +x_5198 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_5199 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5199, 0, x_5198); +x_5200 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_5201 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5200, x_5199, x_5192); +if (lean_obj_tag(x_5201) == 0) +{ +lean_object* x_5202; lean_object* x_5203; lean_object* x_5204; lean_object* x_5205; +lean_dec(x_5195); +x_5202 = lean_ctor_get(x_5201, 0); +lean_inc(x_5202); +x_5203 = lean_ctor_get(x_5201, 1); +lean_inc(x_5203); +if (lean_is_exclusive(x_5201)) { + lean_ctor_release(x_5201, 0); + lean_ctor_release(x_5201, 1); + x_5204 = x_5201; +} else { + lean_dec_ref(x_5201); + x_5204 = lean_box(0); +} +if (lean_is_scalar(x_5204)) { + x_5205 = lean_alloc_ctor(0, 2, 0); +} else { + x_5205 = x_5204; +} +lean_ctor_set(x_5205, 0, x_5202); +lean_ctor_set(x_5205, 1, x_5203); +return x_5205; +} +else +{ +lean_object* x_5206; lean_object* x_5207; lean_object* x_5208; lean_object* x_5209; uint8_t x_5210; +x_5206 = lean_ctor_get(x_5201, 0); +lean_inc(x_5206); +x_5207 = lean_ctor_get(x_5201, 1); +lean_inc(x_5207); +if (lean_is_exclusive(x_5201)) { + lean_ctor_release(x_5201, 0); + lean_ctor_release(x_5201, 1); + x_5208 = x_5201; +} else { + lean_dec_ref(x_5201); + x_5208 = lean_box(0); +} +x_5209 = lean_ctor_get(x_5206, 1); +lean_inc(x_5209); +x_5210 = lean_nat_dec_eq(x_5195, x_5209); +lean_dec(x_5195); +if (x_5210 == 0) +{ +lean_object* x_5211; +lean_dec(x_5209); +if (lean_is_scalar(x_5208)) { + x_5211 = lean_alloc_ctor(1, 2, 0); +} else { + x_5211 = x_5208; +} +lean_ctor_set(x_5211, 0, x_5206); +lean_ctor_set(x_5211, 1, x_5207); +return x_5211; +} +else +{ +lean_object* x_5212; lean_object* x_5213; lean_object* x_5214; lean_object* x_5215; +lean_dec(x_5208); +lean_dec(x_5207); +x_5212 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_5213 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5213, 0, x_5212); +x_5214 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_5215 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5214, x_5213, x_5206); +if (lean_obj_tag(x_5215) == 0) +{ +lean_object* x_5216; lean_object* x_5217; lean_object* x_5218; lean_object* x_5219; +lean_dec(x_5209); +x_5216 = lean_ctor_get(x_5215, 0); +lean_inc(x_5216); +x_5217 = lean_ctor_get(x_5215, 1); +lean_inc(x_5217); +if (lean_is_exclusive(x_5215)) { + lean_ctor_release(x_5215, 0); + lean_ctor_release(x_5215, 1); + x_5218 = x_5215; +} else { + lean_dec_ref(x_5215); + x_5218 = lean_box(0); +} +if (lean_is_scalar(x_5218)) { + x_5219 = lean_alloc_ctor(0, 2, 0); +} else { + x_5219 = x_5218; +} +lean_ctor_set(x_5219, 0, x_5216); +lean_ctor_set(x_5219, 1, x_5217); +return x_5219; +} +else +{ +lean_object* x_5220; lean_object* x_5221; lean_object* x_5222; lean_object* x_5223; uint8_t x_5224; +x_5220 = lean_ctor_get(x_5215, 0); +lean_inc(x_5220); +x_5221 = lean_ctor_get(x_5215, 1); +lean_inc(x_5221); +if (lean_is_exclusive(x_5215)) { + lean_ctor_release(x_5215, 0); + lean_ctor_release(x_5215, 1); + x_5222 = x_5215; +} else { + lean_dec_ref(x_5215); + x_5222 = lean_box(0); +} +x_5223 = lean_ctor_get(x_5220, 1); +lean_inc(x_5223); +x_5224 = lean_nat_dec_eq(x_5209, x_5223); +lean_dec(x_5209); +if (x_5224 == 0) +{ +lean_object* x_5225; +lean_dec(x_5223); +if (lean_is_scalar(x_5222)) { + x_5225 = lean_alloc_ctor(1, 2, 0); +} else { + x_5225 = x_5222; +} +lean_ctor_set(x_5225, 0, x_5220); +lean_ctor_set(x_5225, 1, x_5221); +return x_5225; +} +else +{ +lean_object* x_5226; lean_object* x_5227; lean_object* x_5228; +lean_dec(x_5222); +lean_dec(x_5221); +x_5226 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_5227 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5227, 0, x_5226); +x_5228 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5214, x_5227, x_5220); +if (lean_obj_tag(x_5228) == 0) +{ +lean_object* x_5229; lean_object* x_5230; lean_object* x_5231; lean_object* x_5232; +lean_dec(x_5223); +x_5229 = lean_ctor_get(x_5228, 0); +lean_inc(x_5229); +x_5230 = lean_ctor_get(x_5228, 1); +lean_inc(x_5230); +if (lean_is_exclusive(x_5228)) { + lean_ctor_release(x_5228, 0); + lean_ctor_release(x_5228, 1); + x_5231 = x_5228; +} else { + lean_dec_ref(x_5228); + x_5231 = lean_box(0); +} +if (lean_is_scalar(x_5231)) { + x_5232 = lean_alloc_ctor(0, 2, 0); +} else { + x_5232 = x_5231; +} +lean_ctor_set(x_5232, 0, x_5229); +lean_ctor_set(x_5232, 1, x_5230); +return x_5232; +} +else +{ +lean_object* x_5233; lean_object* x_5234; lean_object* x_5235; lean_object* x_5236; uint8_t x_5237; +x_5233 = lean_ctor_get(x_5228, 0); +lean_inc(x_5233); +x_5234 = lean_ctor_get(x_5228, 1); +lean_inc(x_5234); +if (lean_is_exclusive(x_5228)) { + lean_ctor_release(x_5228, 0); + lean_ctor_release(x_5228, 1); + x_5235 = x_5228; +} else { + lean_dec_ref(x_5228); + x_5235 = lean_box(0); +} +x_5236 = lean_ctor_get(x_5233, 1); +lean_inc(x_5236); +x_5237 = lean_nat_dec_eq(x_5223, x_5236); +lean_dec(x_5223); +if (x_5237 == 0) +{ +lean_object* x_5238; +lean_dec(x_5236); +if (lean_is_scalar(x_5235)) { + x_5238 = lean_alloc_ctor(1, 2, 0); +} else { + x_5238 = x_5235; +} +lean_ctor_set(x_5238, 0, x_5233); +lean_ctor_set(x_5238, 1, x_5234); +return x_5238; +} +else +{ +lean_object* x_5239; lean_object* x_5240; lean_object* x_5241; lean_object* x_5242; +lean_dec(x_5235); +lean_dec(x_5234); +x_5239 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_5240 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5240, 0, x_5239); +x_5241 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_5242 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5241, x_5240, x_5233); +if (lean_obj_tag(x_5242) == 0) +{ +lean_object* x_5243; lean_object* x_5244; lean_object* x_5245; lean_object* x_5246; +lean_dec(x_5236); +x_5243 = lean_ctor_get(x_5242, 0); +lean_inc(x_5243); +x_5244 = lean_ctor_get(x_5242, 1); +lean_inc(x_5244); +if (lean_is_exclusive(x_5242)) { + lean_ctor_release(x_5242, 0); + lean_ctor_release(x_5242, 1); + x_5245 = x_5242; +} else { + lean_dec_ref(x_5242); + x_5245 = lean_box(0); +} +if (lean_is_scalar(x_5245)) { + x_5246 = lean_alloc_ctor(0, 2, 0); +} else { + x_5246 = x_5245; +} +lean_ctor_set(x_5246, 0, x_5243); +lean_ctor_set(x_5246, 1, x_5244); +return x_5246; +} +else +{ +lean_object* x_5247; lean_object* x_5248; lean_object* x_5249; lean_object* x_5250; uint8_t x_5251; +x_5247 = lean_ctor_get(x_5242, 0); +lean_inc(x_5247); +x_5248 = lean_ctor_get(x_5242, 1); +lean_inc(x_5248); +if (lean_is_exclusive(x_5242)) { + lean_ctor_release(x_5242, 0); + lean_ctor_release(x_5242, 1); + x_5249 = x_5242; +} else { + lean_dec_ref(x_5242); + x_5249 = lean_box(0); +} +x_5250 = lean_ctor_get(x_5247, 1); +lean_inc(x_5250); +x_5251 = lean_nat_dec_eq(x_5236, x_5250); +lean_dec(x_5236); +if (x_5251 == 0) +{ +lean_object* x_5252; +lean_dec(x_5250); +if (lean_is_scalar(x_5249)) { + x_5252 = lean_alloc_ctor(1, 2, 0); +} else { + x_5252 = x_5249; +} +lean_ctor_set(x_5252, 0, x_5247); +lean_ctor_set(x_5252, 1, x_5248); +return x_5252; +} +else +{ +lean_object* x_5253; lean_object* x_5254; lean_object* x_5255; lean_object* x_5256; +lean_dec(x_5249); +lean_dec(x_5248); +x_5253 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_5254 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5254, 0, x_5253); +x_5255 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_5256 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5255, x_5254, x_5247); +if (lean_obj_tag(x_5256) == 0) +{ +lean_object* x_5257; lean_object* x_5258; lean_object* x_5259; lean_object* x_5260; +lean_dec(x_5250); +x_5257 = lean_ctor_get(x_5256, 0); +lean_inc(x_5257); +x_5258 = lean_ctor_get(x_5256, 1); +lean_inc(x_5258); +if (lean_is_exclusive(x_5256)) { + lean_ctor_release(x_5256, 0); + lean_ctor_release(x_5256, 1); + x_5259 = x_5256; +} else { + lean_dec_ref(x_5256); + x_5259 = lean_box(0); +} +if (lean_is_scalar(x_5259)) { + x_5260 = lean_alloc_ctor(0, 2, 0); +} else { + x_5260 = x_5259; +} +lean_ctor_set(x_5260, 0, x_5257); +lean_ctor_set(x_5260, 1, x_5258); +return x_5260; +} +else +{ +lean_object* x_5261; lean_object* x_5262; lean_object* x_5263; lean_object* x_5264; uint8_t x_5265; +x_5261 = lean_ctor_get(x_5256, 0); +lean_inc(x_5261); +x_5262 = lean_ctor_get(x_5256, 1); +lean_inc(x_5262); +if (lean_is_exclusive(x_5256)) { + lean_ctor_release(x_5256, 0); + lean_ctor_release(x_5256, 1); + x_5263 = x_5256; +} else { + lean_dec_ref(x_5256); + x_5263 = lean_box(0); +} +x_5264 = lean_ctor_get(x_5261, 1); +lean_inc(x_5264); +x_5265 = lean_nat_dec_eq(x_5250, x_5264); +lean_dec(x_5250); +if (x_5265 == 0) +{ +lean_object* x_5266; +lean_dec(x_5264); +if (lean_is_scalar(x_5263)) { + x_5266 = lean_alloc_ctor(1, 2, 0); +} else { + x_5266 = x_5263; +} +lean_ctor_set(x_5266, 0, x_5261); +lean_ctor_set(x_5266, 1, x_5262); +return x_5266; +} +else +{ +lean_object* x_5267; lean_object* x_5268; lean_object* x_5269; lean_object* x_5270; +lean_dec(x_5263); +lean_dec(x_5262); +x_5267 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_5268 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5268, 0, x_5267); +x_5269 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_5270 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5269, x_5268, x_5261); +if (lean_obj_tag(x_5270) == 0) +{ +lean_object* x_5271; lean_object* x_5272; lean_object* x_5273; lean_object* x_5274; +lean_dec(x_5264); +x_5271 = lean_ctor_get(x_5270, 0); +lean_inc(x_5271); +x_5272 = lean_ctor_get(x_5270, 1); +lean_inc(x_5272); +if (lean_is_exclusive(x_5270)) { + lean_ctor_release(x_5270, 0); + lean_ctor_release(x_5270, 1); + x_5273 = x_5270; +} else { + lean_dec_ref(x_5270); + x_5273 = lean_box(0); +} +if (lean_is_scalar(x_5273)) { + x_5274 = lean_alloc_ctor(0, 2, 0); +} else { + x_5274 = x_5273; +} +lean_ctor_set(x_5274, 0, x_5271); +lean_ctor_set(x_5274, 1, x_5272); +return x_5274; +} +else +{ +lean_object* x_5275; lean_object* x_5276; lean_object* x_5277; lean_object* x_5278; uint8_t x_5279; +x_5275 = lean_ctor_get(x_5270, 0); +lean_inc(x_5275); +x_5276 = lean_ctor_get(x_5270, 1); +lean_inc(x_5276); +if (lean_is_exclusive(x_5270)) { + lean_ctor_release(x_5270, 0); + lean_ctor_release(x_5270, 1); + x_5277 = x_5270; +} else { + lean_dec_ref(x_5270); + x_5277 = lean_box(0); +} +x_5278 = lean_ctor_get(x_5275, 1); +lean_inc(x_5278); +x_5279 = lean_nat_dec_eq(x_5264, x_5278); +lean_dec(x_5264); +if (x_5279 == 0) +{ +lean_object* x_5280; +lean_dec(x_5278); +if (lean_is_scalar(x_5277)) { + x_5280 = lean_alloc_ctor(1, 2, 0); +} else { + x_5280 = x_5277; +} +lean_ctor_set(x_5280, 0, x_5275); +lean_ctor_set(x_5280, 1, x_5276); +return x_5280; +} +else +{ +lean_object* x_5281; lean_object* x_5282; lean_object* x_5283; lean_object* x_5284; +lean_dec(x_5277); +lean_dec(x_5276); +x_5281 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_5282 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5282, 0, x_5281); +x_5283 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_5284 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5283, x_5282, x_5275); +if (lean_obj_tag(x_5284) == 0) +{ +lean_object* x_5285; lean_object* x_5286; lean_object* x_5287; lean_object* x_5288; +lean_dec(x_5278); +x_5285 = lean_ctor_get(x_5284, 0); +lean_inc(x_5285); +x_5286 = lean_ctor_get(x_5284, 1); +lean_inc(x_5286); +if (lean_is_exclusive(x_5284)) { + lean_ctor_release(x_5284, 0); + lean_ctor_release(x_5284, 1); + x_5287 = x_5284; +} else { + lean_dec_ref(x_5284); + x_5287 = lean_box(0); +} +if (lean_is_scalar(x_5287)) { + x_5288 = lean_alloc_ctor(0, 2, 0); +} else { + x_5288 = x_5287; +} +lean_ctor_set(x_5288, 0, x_5285); +lean_ctor_set(x_5288, 1, x_5286); +return x_5288; +} +else +{ +lean_object* x_5289; lean_object* x_5290; lean_object* x_5291; lean_object* x_5292; uint8_t x_5293; +x_5289 = lean_ctor_get(x_5284, 0); +lean_inc(x_5289); +x_5290 = lean_ctor_get(x_5284, 1); +lean_inc(x_5290); +if (lean_is_exclusive(x_5284)) { + lean_ctor_release(x_5284, 0); + lean_ctor_release(x_5284, 1); + x_5291 = x_5284; +} else { + lean_dec_ref(x_5284); + x_5291 = lean_box(0); +} +x_5292 = lean_ctor_get(x_5289, 1); +lean_inc(x_5292); +x_5293 = lean_nat_dec_eq(x_5278, x_5292); +lean_dec(x_5278); +if (x_5293 == 0) +{ +lean_object* x_5294; +lean_dec(x_5292); +if (lean_is_scalar(x_5291)) { + x_5294 = lean_alloc_ctor(1, 2, 0); +} else { + x_5294 = x_5291; +} +lean_ctor_set(x_5294, 0, x_5289); +lean_ctor_set(x_5294, 1, x_5290); +return x_5294; +} +else +{ +lean_object* x_5295; lean_object* x_5296; lean_object* x_5297; +lean_dec(x_5291); +lean_dec(x_5290); +x_5295 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_5296 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5296, 0, x_5295); +x_5297 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5283, x_5296, x_5289); +if (lean_obj_tag(x_5297) == 0) +{ +lean_object* x_5298; lean_object* x_5299; lean_object* x_5300; lean_object* x_5301; +lean_dec(x_5292); +x_5298 = lean_ctor_get(x_5297, 0); +lean_inc(x_5298); +x_5299 = lean_ctor_get(x_5297, 1); +lean_inc(x_5299); +if (lean_is_exclusive(x_5297)) { + lean_ctor_release(x_5297, 0); + lean_ctor_release(x_5297, 1); + x_5300 = x_5297; +} else { + lean_dec_ref(x_5297); + x_5300 = lean_box(0); +} +if (lean_is_scalar(x_5300)) { + x_5301 = lean_alloc_ctor(0, 2, 0); +} else { + x_5301 = x_5300; +} +lean_ctor_set(x_5301, 0, x_5298); +lean_ctor_set(x_5301, 1, x_5299); +return x_5301; +} +else +{ +lean_object* x_5302; lean_object* x_5303; lean_object* x_5304; lean_object* x_5305; uint8_t x_5306; +x_5302 = lean_ctor_get(x_5297, 0); +lean_inc(x_5302); +x_5303 = lean_ctor_get(x_5297, 1); +lean_inc(x_5303); +if (lean_is_exclusive(x_5297)) { + lean_ctor_release(x_5297, 0); + lean_ctor_release(x_5297, 1); + x_5304 = x_5297; +} else { + lean_dec_ref(x_5297); + x_5304 = lean_box(0); +} +x_5305 = lean_ctor_get(x_5302, 1); +lean_inc(x_5305); +x_5306 = lean_nat_dec_eq(x_5292, x_5305); +lean_dec(x_5292); +if (x_5306 == 0) +{ +lean_object* x_5307; +lean_dec(x_5305); +if (lean_is_scalar(x_5304)) { + x_5307 = lean_alloc_ctor(1, 2, 0); +} else { + x_5307 = x_5304; +} +lean_ctor_set(x_5307, 0, x_5302); +lean_ctor_set(x_5307, 1, x_5303); +return x_5307; +} +else +{ +lean_object* x_5308; lean_object* x_5309; lean_object* x_5310; lean_object* x_5311; +lean_dec(x_5304); +lean_dec(x_5303); +x_5308 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_5309 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5309, 0, x_5308); +x_5310 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_5311 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5310, x_5309, x_5302); +if (lean_obj_tag(x_5311) == 0) +{ +lean_object* x_5312; lean_object* x_5313; lean_object* x_5314; lean_object* x_5315; +lean_dec(x_5305); +x_5312 = lean_ctor_get(x_5311, 0); +lean_inc(x_5312); +x_5313 = lean_ctor_get(x_5311, 1); +lean_inc(x_5313); +if (lean_is_exclusive(x_5311)) { + lean_ctor_release(x_5311, 0); + lean_ctor_release(x_5311, 1); + x_5314 = x_5311; +} else { + lean_dec_ref(x_5311); + x_5314 = lean_box(0); +} +if (lean_is_scalar(x_5314)) { + x_5315 = lean_alloc_ctor(0, 2, 0); +} else { + x_5315 = x_5314; +} +lean_ctor_set(x_5315, 0, x_5312); +lean_ctor_set(x_5315, 1, x_5313); +return x_5315; +} +else +{ +lean_object* x_5316; lean_object* x_5317; lean_object* x_5318; lean_object* x_5319; uint8_t x_5320; +x_5316 = lean_ctor_get(x_5311, 0); +lean_inc(x_5316); +x_5317 = lean_ctor_get(x_5311, 1); +lean_inc(x_5317); +if (lean_is_exclusive(x_5311)) { + lean_ctor_release(x_5311, 0); + lean_ctor_release(x_5311, 1); + x_5318 = x_5311; +} else { + lean_dec_ref(x_5311); + x_5318 = lean_box(0); +} +x_5319 = lean_ctor_get(x_5316, 1); +lean_inc(x_5319); +x_5320 = lean_nat_dec_eq(x_5305, x_5319); +lean_dec(x_5305); +if (x_5320 == 0) +{ +lean_object* x_5321; +lean_dec(x_5319); +if (lean_is_scalar(x_5318)) { + x_5321 = lean_alloc_ctor(1, 2, 0); +} else { + x_5321 = x_5318; +} +lean_ctor_set(x_5321, 0, x_5316); +lean_ctor_set(x_5321, 1, x_5317); +return x_5321; +} +else +{ +lean_object* x_5322; lean_object* x_5323; lean_object* x_5324; lean_object* x_5325; +lean_dec(x_5318); +lean_dec(x_5317); +x_5322 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_5323 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5323, 0, x_5322); +x_5324 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_5325 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5324, x_5323, x_5316); +if (lean_obj_tag(x_5325) == 0) +{ +lean_object* x_5326; lean_object* x_5327; lean_object* x_5328; lean_object* x_5329; +lean_dec(x_5319); +x_5326 = lean_ctor_get(x_5325, 0); +lean_inc(x_5326); +x_5327 = lean_ctor_get(x_5325, 1); +lean_inc(x_5327); +if (lean_is_exclusive(x_5325)) { + lean_ctor_release(x_5325, 0); + lean_ctor_release(x_5325, 1); + x_5328 = x_5325; +} else { + lean_dec_ref(x_5325); + x_5328 = lean_box(0); +} +if (lean_is_scalar(x_5328)) { + x_5329 = lean_alloc_ctor(0, 2, 0); +} else { + x_5329 = x_5328; +} +lean_ctor_set(x_5329, 0, x_5326); +lean_ctor_set(x_5329, 1, x_5327); +return x_5329; +} +else +{ +lean_object* x_5330; lean_object* x_5331; lean_object* x_5332; lean_object* x_5333; uint8_t x_5334; +x_5330 = lean_ctor_get(x_5325, 0); +lean_inc(x_5330); +x_5331 = lean_ctor_get(x_5325, 1); +lean_inc(x_5331); +if (lean_is_exclusive(x_5325)) { + lean_ctor_release(x_5325, 0); + lean_ctor_release(x_5325, 1); + x_5332 = x_5325; +} else { + lean_dec_ref(x_5325); + x_5332 = lean_box(0); +} +x_5333 = lean_ctor_get(x_5330, 1); +lean_inc(x_5333); +x_5334 = lean_nat_dec_eq(x_5319, x_5333); +lean_dec(x_5319); +if (x_5334 == 0) +{ +lean_object* x_5335; +lean_dec(x_5333); +if (lean_is_scalar(x_5332)) { + x_5335 = lean_alloc_ctor(1, 2, 0); +} else { + x_5335 = x_5332; +} +lean_ctor_set(x_5335, 0, x_5330); +lean_ctor_set(x_5335, 1, x_5331); +return x_5335; +} +else +{ +lean_object* x_5336; lean_object* x_5337; lean_object* x_5338; lean_object* x_5339; +lean_dec(x_5332); +lean_dec(x_5331); +x_5336 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_5337 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5337, 0, x_5336); +x_5338 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_5339 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5338, x_5337, x_5330); +if (lean_obj_tag(x_5339) == 0) +{ +lean_object* x_5340; lean_object* x_5341; lean_object* x_5342; lean_object* x_5343; +lean_dec(x_5333); +x_5340 = lean_ctor_get(x_5339, 0); +lean_inc(x_5340); +x_5341 = lean_ctor_get(x_5339, 1); +lean_inc(x_5341); +if (lean_is_exclusive(x_5339)) { + lean_ctor_release(x_5339, 0); + lean_ctor_release(x_5339, 1); + x_5342 = x_5339; +} else { + lean_dec_ref(x_5339); + x_5342 = lean_box(0); +} +if (lean_is_scalar(x_5342)) { + x_5343 = lean_alloc_ctor(0, 2, 0); +} else { + x_5343 = x_5342; +} +lean_ctor_set(x_5343, 0, x_5340); +lean_ctor_set(x_5343, 1, x_5341); +return x_5343; +} +else +{ +lean_object* x_5344; lean_object* x_5345; lean_object* x_5346; lean_object* x_5347; uint8_t x_5348; +x_5344 = lean_ctor_get(x_5339, 0); +lean_inc(x_5344); +x_5345 = lean_ctor_get(x_5339, 1); +lean_inc(x_5345); +if (lean_is_exclusive(x_5339)) { + lean_ctor_release(x_5339, 0); + lean_ctor_release(x_5339, 1); + x_5346 = x_5339; +} else { + lean_dec_ref(x_5339); + x_5346 = lean_box(0); +} +x_5347 = lean_ctor_get(x_5344, 1); +lean_inc(x_5347); +x_5348 = lean_nat_dec_eq(x_5333, x_5347); +lean_dec(x_5333); +if (x_5348 == 0) +{ +lean_object* x_5349; +lean_dec(x_5347); +if (lean_is_scalar(x_5346)) { + x_5349 = lean_alloc_ctor(1, 2, 0); +} else { + x_5349 = x_5346; +} +lean_ctor_set(x_5349, 0, x_5344); +lean_ctor_set(x_5349, 1, x_5345); +return x_5349; +} +else +{ +lean_object* x_5350; lean_object* x_5351; lean_object* x_5352; lean_object* x_5353; +lean_dec(x_5346); +lean_dec(x_5345); +x_5350 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_5351 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5351, 0, x_5350); +x_5352 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_5353 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5352, x_5351, x_5344); +if (lean_obj_tag(x_5353) == 0) +{ +lean_object* x_5354; lean_object* x_5355; lean_object* x_5356; lean_object* x_5357; +lean_dec(x_5347); +x_5354 = lean_ctor_get(x_5353, 0); +lean_inc(x_5354); +x_5355 = lean_ctor_get(x_5353, 1); +lean_inc(x_5355); +if (lean_is_exclusive(x_5353)) { + lean_ctor_release(x_5353, 0); + lean_ctor_release(x_5353, 1); + x_5356 = x_5353; +} else { + lean_dec_ref(x_5353); + x_5356 = lean_box(0); +} +if (lean_is_scalar(x_5356)) { + x_5357 = lean_alloc_ctor(0, 2, 0); +} else { + x_5357 = x_5356; +} +lean_ctor_set(x_5357, 0, x_5354); +lean_ctor_set(x_5357, 1, x_5355); +return x_5357; +} +else +{ +lean_object* x_5358; lean_object* x_5359; lean_object* x_5360; lean_object* x_5361; uint8_t x_5362; +x_5358 = lean_ctor_get(x_5353, 0); +lean_inc(x_5358); +x_5359 = lean_ctor_get(x_5353, 1); +lean_inc(x_5359); +if (lean_is_exclusive(x_5353)) { + lean_ctor_release(x_5353, 0); + lean_ctor_release(x_5353, 1); + x_5360 = x_5353; +} else { + lean_dec_ref(x_5353); + x_5360 = lean_box(0); +} +x_5361 = lean_ctor_get(x_5358, 1); +lean_inc(x_5361); +x_5362 = lean_nat_dec_eq(x_5347, x_5361); +lean_dec(x_5347); +if (x_5362 == 0) +{ +lean_object* x_5363; +lean_dec(x_5361); +if (lean_is_scalar(x_5360)) { + x_5363 = lean_alloc_ctor(1, 2, 0); +} else { + x_5363 = x_5360; +} +lean_ctor_set(x_5363, 0, x_5358); +lean_ctor_set(x_5363, 1, x_5359); +return x_5363; +} +else +{ +lean_object* x_5364; lean_object* x_5365; lean_object* x_5366; lean_object* x_5367; +lean_dec(x_5360); +lean_dec(x_5359); +x_5364 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_5365 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5365, 0, x_5364); +x_5366 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_5367 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5366, x_5365, x_5358); +if (lean_obj_tag(x_5367) == 0) +{ +lean_object* x_5368; lean_object* x_5369; lean_object* x_5370; lean_object* x_5371; +lean_dec(x_5361); +x_5368 = lean_ctor_get(x_5367, 0); +lean_inc(x_5368); +x_5369 = lean_ctor_get(x_5367, 1); +lean_inc(x_5369); +if (lean_is_exclusive(x_5367)) { + lean_ctor_release(x_5367, 0); + lean_ctor_release(x_5367, 1); + x_5370 = x_5367; +} else { + lean_dec_ref(x_5367); + x_5370 = lean_box(0); +} +if (lean_is_scalar(x_5370)) { + x_5371 = lean_alloc_ctor(0, 2, 0); +} else { + x_5371 = x_5370; +} +lean_ctor_set(x_5371, 0, x_5368); +lean_ctor_set(x_5371, 1, x_5369); +return x_5371; +} +else +{ +lean_object* x_5372; lean_object* x_5373; lean_object* x_5374; lean_object* x_5375; uint8_t x_5376; +x_5372 = lean_ctor_get(x_5367, 0); +lean_inc(x_5372); +x_5373 = lean_ctor_get(x_5367, 1); +lean_inc(x_5373); +if (lean_is_exclusive(x_5367)) { + lean_ctor_release(x_5367, 0); + lean_ctor_release(x_5367, 1); + x_5374 = x_5367; +} else { + lean_dec_ref(x_5367); + x_5374 = lean_box(0); +} +x_5375 = lean_ctor_get(x_5372, 1); +lean_inc(x_5375); +x_5376 = lean_nat_dec_eq(x_5361, x_5375); +lean_dec(x_5361); +if (x_5376 == 0) +{ +lean_object* x_5377; +lean_dec(x_5375); +if (lean_is_scalar(x_5374)) { + x_5377 = lean_alloc_ctor(1, 2, 0); +} else { + x_5377 = x_5374; +} +lean_ctor_set(x_5377, 0, x_5372); +lean_ctor_set(x_5377, 1, x_5373); +return x_5377; +} +else +{ +lean_object* x_5378; lean_object* x_5379; lean_object* x_5380; lean_object* x_5381; +lean_dec(x_5374); +lean_dec(x_5373); +x_5378 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_5379 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5379, 0, x_5378); +x_5380 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_5381 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5380, x_5379, x_5372); +if (lean_obj_tag(x_5381) == 0) +{ +lean_object* x_5382; lean_object* x_5383; lean_object* x_5384; lean_object* x_5385; +lean_dec(x_5375); +x_5382 = lean_ctor_get(x_5381, 0); +lean_inc(x_5382); +x_5383 = lean_ctor_get(x_5381, 1); +lean_inc(x_5383); +if (lean_is_exclusive(x_5381)) { + lean_ctor_release(x_5381, 0); + lean_ctor_release(x_5381, 1); + x_5384 = x_5381; +} else { + lean_dec_ref(x_5381); + x_5384 = lean_box(0); +} +if (lean_is_scalar(x_5384)) { + x_5385 = lean_alloc_ctor(0, 2, 0); +} else { + x_5385 = x_5384; +} +lean_ctor_set(x_5385, 0, x_5382); +lean_ctor_set(x_5385, 1, x_5383); +return x_5385; +} +else +{ +lean_object* x_5386; lean_object* x_5387; lean_object* x_5388; lean_object* x_5389; uint8_t x_5390; +x_5386 = lean_ctor_get(x_5381, 0); +lean_inc(x_5386); +x_5387 = lean_ctor_get(x_5381, 1); +lean_inc(x_5387); +if (lean_is_exclusive(x_5381)) { + lean_ctor_release(x_5381, 0); + lean_ctor_release(x_5381, 1); + x_5388 = x_5381; +} else { + lean_dec_ref(x_5381); + x_5388 = lean_box(0); +} +x_5389 = lean_ctor_get(x_5386, 1); +lean_inc(x_5389); +x_5390 = lean_nat_dec_eq(x_5375, x_5389); +lean_dec(x_5375); +if (x_5390 == 0) +{ +lean_object* x_5391; +lean_dec(x_5389); +if (lean_is_scalar(x_5388)) { + x_5391 = lean_alloc_ctor(1, 2, 0); +} else { + x_5391 = x_5388; +} +lean_ctor_set(x_5391, 0, x_5386); +lean_ctor_set(x_5391, 1, x_5387); +return x_5391; +} +else +{ +lean_object* x_5392; lean_object* x_5393; lean_object* x_5394; lean_object* x_5395; +lean_dec(x_5388); +lean_dec(x_5387); +x_5392 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_5393 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5393, 0, x_5392); +x_5394 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_5395 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5394, x_5393, x_5386); +if (lean_obj_tag(x_5395) == 0) +{ +lean_object* x_5396; lean_object* x_5397; lean_object* x_5398; lean_object* x_5399; +lean_dec(x_5389); +x_5396 = lean_ctor_get(x_5395, 0); +lean_inc(x_5396); +x_5397 = lean_ctor_get(x_5395, 1); +lean_inc(x_5397); +if (lean_is_exclusive(x_5395)) { + lean_ctor_release(x_5395, 0); + lean_ctor_release(x_5395, 1); + x_5398 = x_5395; +} else { + lean_dec_ref(x_5395); + x_5398 = lean_box(0); +} +if (lean_is_scalar(x_5398)) { + x_5399 = lean_alloc_ctor(0, 2, 0); +} else { + x_5399 = x_5398; +} +lean_ctor_set(x_5399, 0, x_5396); +lean_ctor_set(x_5399, 1, x_5397); +return x_5399; +} +else +{ +lean_object* x_5400; lean_object* x_5401; lean_object* x_5402; lean_object* x_5403; uint8_t x_5404; +x_5400 = lean_ctor_get(x_5395, 0); +lean_inc(x_5400); +x_5401 = lean_ctor_get(x_5395, 1); +lean_inc(x_5401); +if (lean_is_exclusive(x_5395)) { + lean_ctor_release(x_5395, 0); + lean_ctor_release(x_5395, 1); + x_5402 = x_5395; +} else { + lean_dec_ref(x_5395); + x_5402 = lean_box(0); +} +x_5403 = lean_ctor_get(x_5400, 1); +lean_inc(x_5403); +x_5404 = lean_nat_dec_eq(x_5389, x_5403); +lean_dec(x_5389); +if (x_5404 == 0) +{ +lean_object* x_5405; +lean_dec(x_5403); +if (lean_is_scalar(x_5402)) { + x_5405 = lean_alloc_ctor(1, 2, 0); +} else { + x_5405 = x_5402; +} +lean_ctor_set(x_5405, 0, x_5400); +lean_ctor_set(x_5405, 1, x_5401); +return x_5405; +} +else +{ +lean_object* x_5406; lean_object* x_5407; lean_object* x_5408; lean_object* x_5409; +lean_dec(x_5402); +lean_dec(x_5401); +x_5406 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_5407 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5407, 0, x_5406); +x_5408 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_5409 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5408, x_5407, x_5400); +if (lean_obj_tag(x_5409) == 0) +{ +lean_object* x_5410; lean_object* x_5411; lean_object* x_5412; lean_object* x_5413; +lean_dec(x_5403); +x_5410 = lean_ctor_get(x_5409, 0); +lean_inc(x_5410); +x_5411 = lean_ctor_get(x_5409, 1); +lean_inc(x_5411); +if (lean_is_exclusive(x_5409)) { + lean_ctor_release(x_5409, 0); + lean_ctor_release(x_5409, 1); + x_5412 = x_5409; +} else { + lean_dec_ref(x_5409); + x_5412 = lean_box(0); +} +if (lean_is_scalar(x_5412)) { + x_5413 = lean_alloc_ctor(0, 2, 0); +} else { + x_5413 = x_5412; +} +lean_ctor_set(x_5413, 0, x_5410); +lean_ctor_set(x_5413, 1, x_5411); +return x_5413; +} +else +{ +lean_object* x_5414; lean_object* x_5415; lean_object* x_5416; lean_object* x_5417; uint8_t x_5418; +x_5414 = lean_ctor_get(x_5409, 0); +lean_inc(x_5414); +x_5415 = lean_ctor_get(x_5409, 1); +lean_inc(x_5415); +if (lean_is_exclusive(x_5409)) { + lean_ctor_release(x_5409, 0); + lean_ctor_release(x_5409, 1); + x_5416 = x_5409; +} else { + lean_dec_ref(x_5409); + x_5416 = lean_box(0); +} +x_5417 = lean_ctor_get(x_5414, 1); +lean_inc(x_5417); +x_5418 = lean_nat_dec_eq(x_5403, x_5417); +lean_dec(x_5403); +if (x_5418 == 0) +{ +lean_object* x_5419; +lean_dec(x_5417); +if (lean_is_scalar(x_5416)) { + x_5419 = lean_alloc_ctor(1, 2, 0); +} else { + x_5419 = x_5416; +} +lean_ctor_set(x_5419, 0, x_5414); +lean_ctor_set(x_5419, 1, x_5415); +return x_5419; +} +else +{ +lean_object* x_5420; lean_object* x_5421; lean_object* x_5422; lean_object* x_5423; +lean_dec(x_5416); +lean_dec(x_5415); +x_5420 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_5421 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5421, 0, x_5420); +x_5422 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_5423 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5422, x_5421, x_5414); +if (lean_obj_tag(x_5423) == 0) +{ +lean_object* x_5424; lean_object* x_5425; lean_object* x_5426; lean_object* x_5427; +lean_dec(x_5417); +x_5424 = lean_ctor_get(x_5423, 0); +lean_inc(x_5424); +x_5425 = lean_ctor_get(x_5423, 1); +lean_inc(x_5425); +if (lean_is_exclusive(x_5423)) { + lean_ctor_release(x_5423, 0); + lean_ctor_release(x_5423, 1); + x_5426 = x_5423; +} else { + lean_dec_ref(x_5423); + x_5426 = lean_box(0); +} +if (lean_is_scalar(x_5426)) { + x_5427 = lean_alloc_ctor(0, 2, 0); +} else { + x_5427 = x_5426; +} +lean_ctor_set(x_5427, 0, x_5424); +lean_ctor_set(x_5427, 1, x_5425); +return x_5427; +} +else +{ +lean_object* x_5428; lean_object* x_5429; lean_object* x_5430; lean_object* x_5431; uint8_t x_5432; +x_5428 = lean_ctor_get(x_5423, 0); +lean_inc(x_5428); +x_5429 = lean_ctor_get(x_5423, 1); +lean_inc(x_5429); +if (lean_is_exclusive(x_5423)) { + lean_ctor_release(x_5423, 0); + lean_ctor_release(x_5423, 1); + x_5430 = x_5423; +} else { + lean_dec_ref(x_5423); + x_5430 = lean_box(0); +} +x_5431 = lean_ctor_get(x_5428, 1); +lean_inc(x_5431); +x_5432 = lean_nat_dec_eq(x_5417, x_5431); +lean_dec(x_5417); +if (x_5432 == 0) +{ +lean_object* x_5433; +lean_dec(x_5431); +if (lean_is_scalar(x_5430)) { + x_5433 = lean_alloc_ctor(1, 2, 0); +} else { + x_5433 = x_5430; +} +lean_ctor_set(x_5433, 0, x_5428); +lean_ctor_set(x_5433, 1, x_5429); +return x_5433; +} +else +{ +lean_object* x_5434; lean_object* x_5435; lean_object* x_5436; lean_object* x_5437; +lean_dec(x_5430); +lean_dec(x_5429); +x_5434 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_5435 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5435, 0, x_5434); +x_5436 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_5437 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5436, x_5435, x_5428); +if (lean_obj_tag(x_5437) == 0) +{ +lean_object* x_5438; lean_object* x_5439; lean_object* x_5440; lean_object* x_5441; +lean_dec(x_5431); +x_5438 = lean_ctor_get(x_5437, 0); +lean_inc(x_5438); +x_5439 = lean_ctor_get(x_5437, 1); +lean_inc(x_5439); +if (lean_is_exclusive(x_5437)) { + lean_ctor_release(x_5437, 0); + lean_ctor_release(x_5437, 1); + x_5440 = x_5437; +} else { + lean_dec_ref(x_5437); + x_5440 = lean_box(0); +} +if (lean_is_scalar(x_5440)) { + x_5441 = lean_alloc_ctor(0, 2, 0); +} else { + x_5441 = x_5440; +} +lean_ctor_set(x_5441, 0, x_5438); +lean_ctor_set(x_5441, 1, x_5439); +return x_5441; +} +else +{ +lean_object* x_5442; lean_object* x_5443; lean_object* x_5444; lean_object* x_5445; uint8_t x_5446; +x_5442 = lean_ctor_get(x_5437, 0); +lean_inc(x_5442); +x_5443 = lean_ctor_get(x_5437, 1); +lean_inc(x_5443); +if (lean_is_exclusive(x_5437)) { + lean_ctor_release(x_5437, 0); + lean_ctor_release(x_5437, 1); + x_5444 = x_5437; +} else { + lean_dec_ref(x_5437); + x_5444 = lean_box(0); +} +x_5445 = lean_ctor_get(x_5442, 1); +lean_inc(x_5445); +x_5446 = lean_nat_dec_eq(x_5431, x_5445); +lean_dec(x_5431); +if (x_5446 == 0) +{ +lean_object* x_5447; +lean_dec(x_5445); +if (lean_is_scalar(x_5444)) { + x_5447 = lean_alloc_ctor(1, 2, 0); +} else { + x_5447 = x_5444; +} +lean_ctor_set(x_5447, 0, x_5442); +lean_ctor_set(x_5447, 1, x_5443); +return x_5447; +} +else +{ +lean_object* x_5448; lean_object* x_5449; lean_object* x_5450; lean_object* x_5451; +lean_dec(x_5444); +lean_dec(x_5443); +x_5448 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_5449 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5449, 0, x_5448); +x_5450 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_5451 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5450, x_5449, x_5442); +if (lean_obj_tag(x_5451) == 0) +{ +lean_object* x_5452; lean_object* x_5453; lean_object* x_5454; lean_object* x_5455; +lean_dec(x_5445); +x_5452 = lean_ctor_get(x_5451, 0); +lean_inc(x_5452); +x_5453 = lean_ctor_get(x_5451, 1); +lean_inc(x_5453); +if (lean_is_exclusive(x_5451)) { + lean_ctor_release(x_5451, 0); + lean_ctor_release(x_5451, 1); + x_5454 = x_5451; +} else { + lean_dec_ref(x_5451); + x_5454 = lean_box(0); +} +if (lean_is_scalar(x_5454)) { + x_5455 = lean_alloc_ctor(0, 2, 0); +} else { + x_5455 = x_5454; +} +lean_ctor_set(x_5455, 0, x_5452); +lean_ctor_set(x_5455, 1, x_5453); +return x_5455; +} +else +{ +lean_object* x_5456; lean_object* x_5457; lean_object* x_5458; lean_object* x_5459; uint8_t x_5460; +x_5456 = lean_ctor_get(x_5451, 0); +lean_inc(x_5456); +x_5457 = lean_ctor_get(x_5451, 1); +lean_inc(x_5457); +if (lean_is_exclusive(x_5451)) { + lean_ctor_release(x_5451, 0); + lean_ctor_release(x_5451, 1); + x_5458 = x_5451; +} else { + lean_dec_ref(x_5451); + x_5458 = lean_box(0); +} +x_5459 = lean_ctor_get(x_5456, 1); +lean_inc(x_5459); +x_5460 = lean_nat_dec_eq(x_5445, x_5459); +lean_dec(x_5445); +if (x_5460 == 0) +{ +lean_object* x_5461; +lean_dec(x_5459); +if (lean_is_scalar(x_5458)) { + x_5461 = lean_alloc_ctor(1, 2, 0); +} else { + x_5461 = x_5458; +} +lean_ctor_set(x_5461, 0, x_5456); +lean_ctor_set(x_5461, 1, x_5457); +return x_5461; +} +else +{ +lean_object* x_5462; lean_object* x_5463; lean_object* x_5464; lean_object* x_5465; +lean_dec(x_5458); +lean_dec(x_5457); +x_5462 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_5463 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5463, 0, x_5462); +x_5464 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_5465 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5464, x_5463, x_5456); +if (lean_obj_tag(x_5465) == 0) +{ +lean_object* x_5466; lean_object* x_5467; lean_object* x_5468; lean_object* x_5469; +lean_dec(x_5459); +x_5466 = lean_ctor_get(x_5465, 0); +lean_inc(x_5466); +x_5467 = lean_ctor_get(x_5465, 1); +lean_inc(x_5467); +if (lean_is_exclusive(x_5465)) { + lean_ctor_release(x_5465, 0); + lean_ctor_release(x_5465, 1); + x_5468 = x_5465; +} else { + lean_dec_ref(x_5465); + x_5468 = lean_box(0); +} +if (lean_is_scalar(x_5468)) { + x_5469 = lean_alloc_ctor(0, 2, 0); +} else { + x_5469 = x_5468; +} +lean_ctor_set(x_5469, 0, x_5466); +lean_ctor_set(x_5469, 1, x_5467); +return x_5469; +} +else +{ +lean_object* x_5470; lean_object* x_5471; lean_object* x_5472; lean_object* x_5473; uint8_t x_5474; +x_5470 = lean_ctor_get(x_5465, 0); +lean_inc(x_5470); +x_5471 = lean_ctor_get(x_5465, 1); +lean_inc(x_5471); +if (lean_is_exclusive(x_5465)) { + lean_ctor_release(x_5465, 0); + lean_ctor_release(x_5465, 1); + x_5472 = x_5465; +} else { + lean_dec_ref(x_5465); + x_5472 = lean_box(0); +} +x_5473 = lean_ctor_get(x_5470, 1); +lean_inc(x_5473); +x_5474 = lean_nat_dec_eq(x_5459, x_5473); +lean_dec(x_5459); +if (x_5474 == 0) +{ +lean_object* x_5475; +lean_dec(x_5473); +if (lean_is_scalar(x_5472)) { + x_5475 = lean_alloc_ctor(1, 2, 0); +} else { + x_5475 = x_5472; +} +lean_ctor_set(x_5475, 0, x_5470); +lean_ctor_set(x_5475, 1, x_5471); +return x_5475; +} +else +{ +lean_object* x_5476; lean_object* x_5477; lean_object* x_5478; lean_object* x_5479; +lean_dec(x_5472); +lean_dec(x_5471); +x_5476 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_5477 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5477, 0, x_5476); +x_5478 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_5479 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5478, x_5477, x_5470); +if (lean_obj_tag(x_5479) == 0) +{ +lean_object* x_5480; lean_object* x_5481; lean_object* x_5482; lean_object* x_5483; +lean_dec(x_5473); +x_5480 = lean_ctor_get(x_5479, 0); +lean_inc(x_5480); +x_5481 = lean_ctor_get(x_5479, 1); +lean_inc(x_5481); +if (lean_is_exclusive(x_5479)) { + lean_ctor_release(x_5479, 0); + lean_ctor_release(x_5479, 1); + x_5482 = x_5479; +} else { + lean_dec_ref(x_5479); + x_5482 = lean_box(0); +} +if (lean_is_scalar(x_5482)) { + x_5483 = lean_alloc_ctor(0, 2, 0); +} else { + x_5483 = x_5482; +} +lean_ctor_set(x_5483, 0, x_5480); +lean_ctor_set(x_5483, 1, x_5481); +return x_5483; +} +else +{ +lean_object* x_5484; lean_object* x_5485; lean_object* x_5486; lean_object* x_5487; uint8_t x_5488; +x_5484 = lean_ctor_get(x_5479, 0); +lean_inc(x_5484); +x_5485 = lean_ctor_get(x_5479, 1); +lean_inc(x_5485); +if (lean_is_exclusive(x_5479)) { + lean_ctor_release(x_5479, 0); + lean_ctor_release(x_5479, 1); + x_5486 = x_5479; +} else { + lean_dec_ref(x_5479); + x_5486 = lean_box(0); +} +x_5487 = lean_ctor_get(x_5484, 1); +lean_inc(x_5487); +x_5488 = lean_nat_dec_eq(x_5473, x_5487); +lean_dec(x_5473); +if (x_5488 == 0) +{ +lean_object* x_5489; +lean_dec(x_5487); +if (lean_is_scalar(x_5486)) { + x_5489 = lean_alloc_ctor(1, 2, 0); +} else { + x_5489 = x_5486; +} +lean_ctor_set(x_5489, 0, x_5484); +lean_ctor_set(x_5489, 1, x_5485); +return x_5489; +} +else +{ +lean_object* x_5490; lean_object* x_5491; lean_object* x_5492; lean_object* x_5493; +lean_dec(x_5486); +lean_dec(x_5485); +x_5490 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_5491 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5491, 0, x_5490); +x_5492 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_5493 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5492, x_5491, x_5484); +if (lean_obj_tag(x_5493) == 0) +{ +lean_object* x_5494; lean_object* x_5495; lean_object* x_5496; lean_object* x_5497; +lean_dec(x_5487); +x_5494 = lean_ctor_get(x_5493, 0); +lean_inc(x_5494); +x_5495 = lean_ctor_get(x_5493, 1); +lean_inc(x_5495); +if (lean_is_exclusive(x_5493)) { + lean_ctor_release(x_5493, 0); + lean_ctor_release(x_5493, 1); + x_5496 = x_5493; +} else { + lean_dec_ref(x_5493); + x_5496 = lean_box(0); +} +if (lean_is_scalar(x_5496)) { + x_5497 = lean_alloc_ctor(0, 2, 0); +} else { + x_5497 = x_5496; +} +lean_ctor_set(x_5497, 0, x_5494); +lean_ctor_set(x_5497, 1, x_5495); +return x_5497; +} +else +{ +lean_object* x_5498; lean_object* x_5499; lean_object* x_5500; lean_object* x_5501; uint8_t x_5502; +x_5498 = lean_ctor_get(x_5493, 0); +lean_inc(x_5498); +x_5499 = lean_ctor_get(x_5493, 1); +lean_inc(x_5499); +if (lean_is_exclusive(x_5493)) { + lean_ctor_release(x_5493, 0); + lean_ctor_release(x_5493, 1); + x_5500 = x_5493; +} else { + lean_dec_ref(x_5493); + x_5500 = lean_box(0); +} +x_5501 = lean_ctor_get(x_5498, 1); +lean_inc(x_5501); +x_5502 = lean_nat_dec_eq(x_5487, x_5501); +lean_dec(x_5487); +if (x_5502 == 0) +{ +lean_object* x_5503; +lean_dec(x_5501); +if (lean_is_scalar(x_5500)) { + x_5503 = lean_alloc_ctor(1, 2, 0); +} else { + x_5503 = x_5500; +} +lean_ctor_set(x_5503, 0, x_5498); +lean_ctor_set(x_5503, 1, x_5499); +return x_5503; +} +else +{ +lean_object* x_5504; lean_object* x_5505; lean_object* x_5506; lean_object* x_5507; +lean_dec(x_5500); +lean_dec(x_5499); +x_5504 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_5505 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5505, 0, x_5504); +x_5506 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_5507 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5506, x_5505, x_5498); +if (lean_obj_tag(x_5507) == 0) +{ +lean_object* x_5508; lean_object* x_5509; lean_object* x_5510; lean_object* x_5511; +lean_dec(x_5501); +x_5508 = lean_ctor_get(x_5507, 0); +lean_inc(x_5508); +x_5509 = lean_ctor_get(x_5507, 1); +lean_inc(x_5509); +if (lean_is_exclusive(x_5507)) { + lean_ctor_release(x_5507, 0); + lean_ctor_release(x_5507, 1); + x_5510 = x_5507; +} else { + lean_dec_ref(x_5507); + x_5510 = lean_box(0); +} +if (lean_is_scalar(x_5510)) { + x_5511 = lean_alloc_ctor(0, 2, 0); +} else { + x_5511 = x_5510; +} +lean_ctor_set(x_5511, 0, x_5508); +lean_ctor_set(x_5511, 1, x_5509); +return x_5511; +} +else +{ +lean_object* x_5512; lean_object* x_5513; lean_object* x_5514; lean_object* x_5515; uint8_t x_5516; +x_5512 = lean_ctor_get(x_5507, 0); +lean_inc(x_5512); +x_5513 = lean_ctor_get(x_5507, 1); +lean_inc(x_5513); +if (lean_is_exclusive(x_5507)) { + lean_ctor_release(x_5507, 0); + lean_ctor_release(x_5507, 1); + x_5514 = x_5507; +} else { + lean_dec_ref(x_5507); + x_5514 = lean_box(0); +} +x_5515 = lean_ctor_get(x_5512, 1); +lean_inc(x_5515); +x_5516 = lean_nat_dec_eq(x_5501, x_5515); +lean_dec(x_5501); +if (x_5516 == 0) +{ +lean_object* x_5517; +lean_dec(x_5515); +if (lean_is_scalar(x_5514)) { + x_5517 = lean_alloc_ctor(1, 2, 0); +} else { + x_5517 = x_5514; +} +lean_ctor_set(x_5517, 0, x_5512); +lean_ctor_set(x_5517, 1, x_5513); +return x_5517; +} +else +{ +lean_object* x_5518; lean_object* x_5519; lean_object* x_5520; lean_object* x_5521; +lean_dec(x_5514); +lean_dec(x_5513); +x_5518 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_5519 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5519, 0, x_5518); +x_5520 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_5521 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5520, x_5519, x_5512); +if (lean_obj_tag(x_5521) == 0) +{ +lean_object* x_5522; lean_object* x_5523; lean_object* x_5524; lean_object* x_5525; +lean_dec(x_5515); +x_5522 = lean_ctor_get(x_5521, 0); +lean_inc(x_5522); +x_5523 = lean_ctor_get(x_5521, 1); +lean_inc(x_5523); +if (lean_is_exclusive(x_5521)) { + lean_ctor_release(x_5521, 0); + lean_ctor_release(x_5521, 1); + x_5524 = x_5521; +} else { + lean_dec_ref(x_5521); + x_5524 = lean_box(0); +} +if (lean_is_scalar(x_5524)) { + x_5525 = lean_alloc_ctor(0, 2, 0); +} else { + x_5525 = x_5524; +} +lean_ctor_set(x_5525, 0, x_5522); +lean_ctor_set(x_5525, 1, x_5523); +return x_5525; +} +else +{ +lean_object* x_5526; lean_object* x_5527; lean_object* x_5528; lean_object* x_5529; uint8_t x_5530; +x_5526 = lean_ctor_get(x_5521, 0); +lean_inc(x_5526); +x_5527 = lean_ctor_get(x_5521, 1); +lean_inc(x_5527); +if (lean_is_exclusive(x_5521)) { + lean_ctor_release(x_5521, 0); + lean_ctor_release(x_5521, 1); + x_5528 = x_5521; +} else { + lean_dec_ref(x_5521); + x_5528 = lean_box(0); +} +x_5529 = lean_ctor_get(x_5526, 1); +lean_inc(x_5529); +x_5530 = lean_nat_dec_eq(x_5515, x_5529); +lean_dec(x_5515); +if (x_5530 == 0) +{ +lean_object* x_5531; +lean_dec(x_5529); +if (lean_is_scalar(x_5528)) { + x_5531 = lean_alloc_ctor(1, 2, 0); +} else { + x_5531 = x_5528; +} +lean_ctor_set(x_5531, 0, x_5526); +lean_ctor_set(x_5531, 1, x_5527); +return x_5531; +} +else +{ +lean_object* x_5532; lean_object* x_5533; lean_object* x_5534; lean_object* x_5535; +lean_dec(x_5528); +lean_dec(x_5527); +x_5532 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_5533 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5533, 0, x_5532); +x_5534 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_5535 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5534, x_5533, x_5526); +if (lean_obj_tag(x_5535) == 0) +{ +lean_object* x_5536; lean_object* x_5537; lean_object* x_5538; lean_object* x_5539; +lean_dec(x_5529); +x_5536 = lean_ctor_get(x_5535, 0); +lean_inc(x_5536); +x_5537 = lean_ctor_get(x_5535, 1); +lean_inc(x_5537); +if (lean_is_exclusive(x_5535)) { + lean_ctor_release(x_5535, 0); + lean_ctor_release(x_5535, 1); + x_5538 = x_5535; +} else { + lean_dec_ref(x_5535); + x_5538 = lean_box(0); +} +if (lean_is_scalar(x_5538)) { + x_5539 = lean_alloc_ctor(0, 2, 0); +} else { + x_5539 = x_5538; +} +lean_ctor_set(x_5539, 0, x_5536); +lean_ctor_set(x_5539, 1, x_5537); +return x_5539; +} +else +{ +lean_object* x_5540; lean_object* x_5541; lean_object* x_5542; lean_object* x_5543; uint8_t x_5544; +x_5540 = lean_ctor_get(x_5535, 0); +lean_inc(x_5540); +x_5541 = lean_ctor_get(x_5535, 1); +lean_inc(x_5541); +if (lean_is_exclusive(x_5535)) { + lean_ctor_release(x_5535, 0); + lean_ctor_release(x_5535, 1); + x_5542 = x_5535; +} else { + lean_dec_ref(x_5535); + x_5542 = lean_box(0); +} +x_5543 = lean_ctor_get(x_5540, 1); +lean_inc(x_5543); +x_5544 = lean_nat_dec_eq(x_5529, x_5543); +lean_dec(x_5543); +lean_dec(x_5529); +if (x_5544 == 0) +{ +lean_object* x_5545; +if (lean_is_scalar(x_5542)) { + x_5545 = lean_alloc_ctor(1, 2, 0); +} else { + x_5545 = x_5542; +} +lean_ctor_set(x_5545, 0, x_5540); +lean_ctor_set(x_5545, 1, x_5541); +return x_5545; +} +else +{ +lean_object* x_5546; lean_object* x_5547; lean_object* x_5548; lean_object* x_5549; +lean_dec(x_5542); +lean_dec(x_5541); +x_5546 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_5547 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5547, 0, x_5546); +x_5548 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_5549 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5548, x_5547, x_5540); +return x_5549; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_5550; lean_object* x_5551; lean_object* x_5552; uint8_t x_5553; +x_5550 = lean_ctor_get(x_45, 0); +x_5551 = lean_ctor_get(x_45, 1); +lean_inc(x_5551); +lean_inc(x_5550); +lean_dec(x_45); +x_5552 = lean_ctor_get(x_5550, 1); +lean_inc(x_5552); +x_5553 = lean_nat_dec_eq(x_40, x_5552); +lean_dec(x_40); +if (x_5553 == 0) +{ +lean_object* x_5554; +lean_dec(x_5552); +x_5554 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5554, 0, x_5550); +lean_ctor_set(x_5554, 1, x_5551); +return x_5554; +} +else +{ +lean_object* x_5555; lean_object* x_5556; lean_object* x_5557; lean_object* x_5558; +lean_dec(x_5551); +x_5555 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +x_5556 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5556, 0, x_5555); +x_5557 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +x_5558 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5557, x_5556, x_5550); +if (lean_obj_tag(x_5558) == 0) +{ +lean_object* x_5559; lean_object* x_5560; lean_object* x_5561; lean_object* x_5562; +lean_dec(x_5552); +x_5559 = lean_ctor_get(x_5558, 0); +lean_inc(x_5559); +x_5560 = lean_ctor_get(x_5558, 1); +lean_inc(x_5560); +if (lean_is_exclusive(x_5558)) { + lean_ctor_release(x_5558, 0); + lean_ctor_release(x_5558, 1); + x_5561 = x_5558; +} else { + lean_dec_ref(x_5558); + x_5561 = lean_box(0); +} +if (lean_is_scalar(x_5561)) { + x_5562 = lean_alloc_ctor(0, 2, 0); +} else { + x_5562 = x_5561; +} +lean_ctor_set(x_5562, 0, x_5559); +lean_ctor_set(x_5562, 1, x_5560); +return x_5562; +} +else +{ +lean_object* x_5563; lean_object* x_5564; lean_object* x_5565; lean_object* x_5566; uint8_t x_5567; +x_5563 = lean_ctor_get(x_5558, 0); +lean_inc(x_5563); +x_5564 = lean_ctor_get(x_5558, 1); +lean_inc(x_5564); +if (lean_is_exclusive(x_5558)) { + lean_ctor_release(x_5558, 0); + lean_ctor_release(x_5558, 1); + x_5565 = x_5558; +} else { + lean_dec_ref(x_5558); + x_5565 = lean_box(0); +} +x_5566 = lean_ctor_get(x_5563, 1); +lean_inc(x_5566); +x_5567 = lean_nat_dec_eq(x_5552, x_5566); +lean_dec(x_5552); +if (x_5567 == 0) +{ +lean_object* x_5568; +lean_dec(x_5566); +if (lean_is_scalar(x_5565)) { + x_5568 = lean_alloc_ctor(1, 2, 0); +} else { + x_5568 = x_5565; +} +lean_ctor_set(x_5568, 0, x_5563); +lean_ctor_set(x_5568, 1, x_5564); +return x_5568; +} +else +{ +lean_object* x_5569; lean_object* x_5570; lean_object* x_5571; +lean_dec(x_5565); +lean_dec(x_5564); +x_5569 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_5570 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5570, 0, x_5569); +x_5571 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5557, x_5570, x_5563); +if (lean_obj_tag(x_5571) == 0) +{ +lean_object* x_5572; lean_object* x_5573; lean_object* x_5574; lean_object* x_5575; +lean_dec(x_5566); +x_5572 = lean_ctor_get(x_5571, 0); +lean_inc(x_5572); +x_5573 = lean_ctor_get(x_5571, 1); +lean_inc(x_5573); +if (lean_is_exclusive(x_5571)) { + lean_ctor_release(x_5571, 0); + lean_ctor_release(x_5571, 1); + x_5574 = x_5571; +} else { + lean_dec_ref(x_5571); + x_5574 = lean_box(0); +} +if (lean_is_scalar(x_5574)) { + x_5575 = lean_alloc_ctor(0, 2, 0); +} else { + x_5575 = x_5574; +} +lean_ctor_set(x_5575, 0, x_5572); +lean_ctor_set(x_5575, 1, x_5573); +return x_5575; +} +else +{ +lean_object* x_5576; lean_object* x_5577; lean_object* x_5578; lean_object* x_5579; uint8_t x_5580; +x_5576 = lean_ctor_get(x_5571, 0); +lean_inc(x_5576); +x_5577 = lean_ctor_get(x_5571, 1); +lean_inc(x_5577); +if (lean_is_exclusive(x_5571)) { + lean_ctor_release(x_5571, 0); + lean_ctor_release(x_5571, 1); + x_5578 = x_5571; +} else { + lean_dec_ref(x_5571); + x_5578 = lean_box(0); +} +x_5579 = lean_ctor_get(x_5576, 1); +lean_inc(x_5579); +x_5580 = lean_nat_dec_eq(x_5566, x_5579); +lean_dec(x_5566); +if (x_5580 == 0) +{ +lean_object* x_5581; +lean_dec(x_5579); +if (lean_is_scalar(x_5578)) { + x_5581 = lean_alloc_ctor(1, 2, 0); +} else { + x_5581 = x_5578; +} +lean_ctor_set(x_5581, 0, x_5576); +lean_ctor_set(x_5581, 1, x_5577); +return x_5581; +} +else +{ +lean_object* x_5582; lean_object* x_5583; lean_object* x_5584; lean_object* x_5585; +lean_dec(x_5578); +lean_dec(x_5577); +x_5582 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_5583 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5583, 0, x_5582); +x_5584 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_5585 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5584, x_5583, x_5576); +if (lean_obj_tag(x_5585) == 0) +{ +lean_object* x_5586; lean_object* x_5587; lean_object* x_5588; lean_object* x_5589; +lean_dec(x_5579); +x_5586 = lean_ctor_get(x_5585, 0); +lean_inc(x_5586); +x_5587 = lean_ctor_get(x_5585, 1); +lean_inc(x_5587); +if (lean_is_exclusive(x_5585)) { + lean_ctor_release(x_5585, 0); + lean_ctor_release(x_5585, 1); + x_5588 = x_5585; +} else { + lean_dec_ref(x_5585); + x_5588 = lean_box(0); +} +if (lean_is_scalar(x_5588)) { + x_5589 = lean_alloc_ctor(0, 2, 0); +} else { + x_5589 = x_5588; +} +lean_ctor_set(x_5589, 0, x_5586); +lean_ctor_set(x_5589, 1, x_5587); +return x_5589; +} +else +{ +lean_object* x_5590; lean_object* x_5591; lean_object* x_5592; lean_object* x_5593; uint8_t x_5594; +x_5590 = lean_ctor_get(x_5585, 0); +lean_inc(x_5590); +x_5591 = lean_ctor_get(x_5585, 1); +lean_inc(x_5591); +if (lean_is_exclusive(x_5585)) { + lean_ctor_release(x_5585, 0); + lean_ctor_release(x_5585, 1); + x_5592 = x_5585; +} else { + lean_dec_ref(x_5585); + x_5592 = lean_box(0); +} +x_5593 = lean_ctor_get(x_5590, 1); +lean_inc(x_5593); +x_5594 = lean_nat_dec_eq(x_5579, x_5593); +lean_dec(x_5579); +if (x_5594 == 0) +{ +lean_object* x_5595; +lean_dec(x_5593); +if (lean_is_scalar(x_5592)) { + x_5595 = lean_alloc_ctor(1, 2, 0); +} else { + x_5595 = x_5592; +} +lean_ctor_set(x_5595, 0, x_5590); +lean_ctor_set(x_5595, 1, x_5591); +return x_5595; +} +else +{ +lean_object* x_5596; lean_object* x_5597; lean_object* x_5598; lean_object* x_5599; +lean_dec(x_5592); +lean_dec(x_5591); +x_5596 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_5597 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5597, 0, x_5596); +x_5598 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_5599 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5598, x_5597, x_5590); +if (lean_obj_tag(x_5599) == 0) +{ +lean_object* x_5600; lean_object* x_5601; lean_object* x_5602; lean_object* x_5603; +lean_dec(x_5593); +x_5600 = lean_ctor_get(x_5599, 0); +lean_inc(x_5600); +x_5601 = lean_ctor_get(x_5599, 1); +lean_inc(x_5601); +if (lean_is_exclusive(x_5599)) { + lean_ctor_release(x_5599, 0); + lean_ctor_release(x_5599, 1); + x_5602 = x_5599; +} else { + lean_dec_ref(x_5599); + x_5602 = lean_box(0); +} +if (lean_is_scalar(x_5602)) { + x_5603 = lean_alloc_ctor(0, 2, 0); +} else { + x_5603 = x_5602; +} +lean_ctor_set(x_5603, 0, x_5600); +lean_ctor_set(x_5603, 1, x_5601); +return x_5603; +} +else +{ +lean_object* x_5604; lean_object* x_5605; lean_object* x_5606; lean_object* x_5607; uint8_t x_5608; +x_5604 = lean_ctor_get(x_5599, 0); +lean_inc(x_5604); +x_5605 = lean_ctor_get(x_5599, 1); +lean_inc(x_5605); +if (lean_is_exclusive(x_5599)) { + lean_ctor_release(x_5599, 0); + lean_ctor_release(x_5599, 1); + x_5606 = x_5599; +} else { + lean_dec_ref(x_5599); + x_5606 = lean_box(0); +} +x_5607 = lean_ctor_get(x_5604, 1); +lean_inc(x_5607); +x_5608 = lean_nat_dec_eq(x_5593, x_5607); +lean_dec(x_5593); +if (x_5608 == 0) +{ +lean_object* x_5609; +lean_dec(x_5607); +if (lean_is_scalar(x_5606)) { + x_5609 = lean_alloc_ctor(1, 2, 0); +} else { + x_5609 = x_5606; +} +lean_ctor_set(x_5609, 0, x_5604); +lean_ctor_set(x_5609, 1, x_5605); +return x_5609; +} +else +{ +lean_object* x_5610; lean_object* x_5611; lean_object* x_5612; +lean_dec(x_5606); +lean_dec(x_5605); +x_5610 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_5611 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5611, 0, x_5610); +x_5612 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5598, x_5611, x_5604); +if (lean_obj_tag(x_5612) == 0) +{ +lean_object* x_5613; lean_object* x_5614; lean_object* x_5615; lean_object* x_5616; +lean_dec(x_5607); +x_5613 = lean_ctor_get(x_5612, 0); +lean_inc(x_5613); +x_5614 = lean_ctor_get(x_5612, 1); +lean_inc(x_5614); +if (lean_is_exclusive(x_5612)) { + lean_ctor_release(x_5612, 0); + lean_ctor_release(x_5612, 1); + x_5615 = x_5612; +} else { + lean_dec_ref(x_5612); + x_5615 = lean_box(0); +} +if (lean_is_scalar(x_5615)) { + x_5616 = lean_alloc_ctor(0, 2, 0); +} else { + x_5616 = x_5615; +} +lean_ctor_set(x_5616, 0, x_5613); +lean_ctor_set(x_5616, 1, x_5614); +return x_5616; +} +else +{ +lean_object* x_5617; lean_object* x_5618; lean_object* x_5619; lean_object* x_5620; uint8_t x_5621; +x_5617 = lean_ctor_get(x_5612, 0); +lean_inc(x_5617); +x_5618 = lean_ctor_get(x_5612, 1); +lean_inc(x_5618); +if (lean_is_exclusive(x_5612)) { + lean_ctor_release(x_5612, 0); + lean_ctor_release(x_5612, 1); + x_5619 = x_5612; +} else { + lean_dec_ref(x_5612); + x_5619 = lean_box(0); +} +x_5620 = lean_ctor_get(x_5617, 1); +lean_inc(x_5620); +x_5621 = lean_nat_dec_eq(x_5607, x_5620); +lean_dec(x_5607); +if (x_5621 == 0) +{ +lean_object* x_5622; +lean_dec(x_5620); +if (lean_is_scalar(x_5619)) { + x_5622 = lean_alloc_ctor(1, 2, 0); +} else { + x_5622 = x_5619; +} +lean_ctor_set(x_5622, 0, x_5617); +lean_ctor_set(x_5622, 1, x_5618); +return x_5622; +} +else +{ +lean_object* x_5623; lean_object* x_5624; lean_object* x_5625; lean_object* x_5626; +lean_dec(x_5619); +lean_dec(x_5618); +x_5623 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_5624 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5624, 0, x_5623); +x_5625 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_5626 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5625, x_5624, x_5617); +if (lean_obj_tag(x_5626) == 0) +{ +lean_object* x_5627; lean_object* x_5628; lean_object* x_5629; lean_object* x_5630; +lean_dec(x_5620); +x_5627 = lean_ctor_get(x_5626, 0); +lean_inc(x_5627); +x_5628 = lean_ctor_get(x_5626, 1); +lean_inc(x_5628); +if (lean_is_exclusive(x_5626)) { + lean_ctor_release(x_5626, 0); + lean_ctor_release(x_5626, 1); + x_5629 = x_5626; +} else { + lean_dec_ref(x_5626); + x_5629 = lean_box(0); +} +if (lean_is_scalar(x_5629)) { + x_5630 = lean_alloc_ctor(0, 2, 0); +} else { + x_5630 = x_5629; +} +lean_ctor_set(x_5630, 0, x_5627); +lean_ctor_set(x_5630, 1, x_5628); +return x_5630; +} +else +{ +lean_object* x_5631; lean_object* x_5632; lean_object* x_5633; lean_object* x_5634; uint8_t x_5635; +x_5631 = lean_ctor_get(x_5626, 0); +lean_inc(x_5631); +x_5632 = lean_ctor_get(x_5626, 1); +lean_inc(x_5632); +if (lean_is_exclusive(x_5626)) { + lean_ctor_release(x_5626, 0); + lean_ctor_release(x_5626, 1); + x_5633 = x_5626; +} else { + lean_dec_ref(x_5626); + x_5633 = lean_box(0); +} +x_5634 = lean_ctor_get(x_5631, 1); +lean_inc(x_5634); +x_5635 = lean_nat_dec_eq(x_5620, x_5634); +lean_dec(x_5620); +if (x_5635 == 0) +{ +lean_object* x_5636; +lean_dec(x_5634); +if (lean_is_scalar(x_5633)) { + x_5636 = lean_alloc_ctor(1, 2, 0); +} else { + x_5636 = x_5633; +} +lean_ctor_set(x_5636, 0, x_5631); +lean_ctor_set(x_5636, 1, x_5632); +return x_5636; +} +else +{ +lean_object* x_5637; lean_object* x_5638; lean_object* x_5639; lean_object* x_5640; +lean_dec(x_5633); +lean_dec(x_5632); +x_5637 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_5638 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5638, 0, x_5637); +x_5639 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_5640 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5639, x_5638, x_5631); +if (lean_obj_tag(x_5640) == 0) +{ +lean_object* x_5641; lean_object* x_5642; lean_object* x_5643; lean_object* x_5644; +lean_dec(x_5634); +x_5641 = lean_ctor_get(x_5640, 0); +lean_inc(x_5641); +x_5642 = lean_ctor_get(x_5640, 1); +lean_inc(x_5642); +if (lean_is_exclusive(x_5640)) { + lean_ctor_release(x_5640, 0); + lean_ctor_release(x_5640, 1); + x_5643 = x_5640; +} else { + lean_dec_ref(x_5640); + x_5643 = lean_box(0); +} +if (lean_is_scalar(x_5643)) { + x_5644 = lean_alloc_ctor(0, 2, 0); +} else { + x_5644 = x_5643; +} +lean_ctor_set(x_5644, 0, x_5641); +lean_ctor_set(x_5644, 1, x_5642); +return x_5644; +} +else +{ +lean_object* x_5645; lean_object* x_5646; lean_object* x_5647; lean_object* x_5648; uint8_t x_5649; +x_5645 = lean_ctor_get(x_5640, 0); +lean_inc(x_5645); +x_5646 = lean_ctor_get(x_5640, 1); +lean_inc(x_5646); +if (lean_is_exclusive(x_5640)) { + lean_ctor_release(x_5640, 0); + lean_ctor_release(x_5640, 1); + x_5647 = x_5640; +} else { + lean_dec_ref(x_5640); + x_5647 = lean_box(0); +} +x_5648 = lean_ctor_get(x_5645, 1); +lean_inc(x_5648); +x_5649 = lean_nat_dec_eq(x_5634, x_5648); +lean_dec(x_5634); +if (x_5649 == 0) +{ +lean_object* x_5650; +lean_dec(x_5648); +if (lean_is_scalar(x_5647)) { + x_5650 = lean_alloc_ctor(1, 2, 0); +} else { + x_5650 = x_5647; +} +lean_ctor_set(x_5650, 0, x_5645); +lean_ctor_set(x_5650, 1, x_5646); +return x_5650; +} +else +{ +lean_object* x_5651; lean_object* x_5652; lean_object* x_5653; lean_object* x_5654; +lean_dec(x_5647); +lean_dec(x_5646); +x_5651 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_5652 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5652, 0, x_5651); +x_5653 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_5654 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5653, x_5652, x_5645); +if (lean_obj_tag(x_5654) == 0) +{ +lean_object* x_5655; lean_object* x_5656; lean_object* x_5657; lean_object* x_5658; +lean_dec(x_5648); +x_5655 = lean_ctor_get(x_5654, 0); +lean_inc(x_5655); +x_5656 = lean_ctor_get(x_5654, 1); +lean_inc(x_5656); +if (lean_is_exclusive(x_5654)) { + lean_ctor_release(x_5654, 0); + lean_ctor_release(x_5654, 1); + x_5657 = x_5654; +} else { + lean_dec_ref(x_5654); + x_5657 = lean_box(0); +} +if (lean_is_scalar(x_5657)) { + x_5658 = lean_alloc_ctor(0, 2, 0); +} else { + x_5658 = x_5657; +} +lean_ctor_set(x_5658, 0, x_5655); +lean_ctor_set(x_5658, 1, x_5656); +return x_5658; +} +else +{ +lean_object* x_5659; lean_object* x_5660; lean_object* x_5661; lean_object* x_5662; uint8_t x_5663; +x_5659 = lean_ctor_get(x_5654, 0); +lean_inc(x_5659); +x_5660 = lean_ctor_get(x_5654, 1); +lean_inc(x_5660); +if (lean_is_exclusive(x_5654)) { + lean_ctor_release(x_5654, 0); + lean_ctor_release(x_5654, 1); + x_5661 = x_5654; +} else { + lean_dec_ref(x_5654); + x_5661 = lean_box(0); +} +x_5662 = lean_ctor_get(x_5659, 1); +lean_inc(x_5662); +x_5663 = lean_nat_dec_eq(x_5648, x_5662); +lean_dec(x_5648); +if (x_5663 == 0) +{ +lean_object* x_5664; +lean_dec(x_5662); +if (lean_is_scalar(x_5661)) { + x_5664 = lean_alloc_ctor(1, 2, 0); +} else { + x_5664 = x_5661; +} +lean_ctor_set(x_5664, 0, x_5659); +lean_ctor_set(x_5664, 1, x_5660); +return x_5664; +} +else +{ +lean_object* x_5665; lean_object* x_5666; lean_object* x_5667; lean_object* x_5668; +lean_dec(x_5661); +lean_dec(x_5660); +x_5665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_5666 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5666, 0, x_5665); +x_5667 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_5668 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5667, x_5666, x_5659); +if (lean_obj_tag(x_5668) == 0) +{ +lean_object* x_5669; lean_object* x_5670; lean_object* x_5671; lean_object* x_5672; +lean_dec(x_5662); +x_5669 = lean_ctor_get(x_5668, 0); +lean_inc(x_5669); +x_5670 = lean_ctor_get(x_5668, 1); +lean_inc(x_5670); +if (lean_is_exclusive(x_5668)) { + lean_ctor_release(x_5668, 0); + lean_ctor_release(x_5668, 1); + x_5671 = x_5668; +} else { + lean_dec_ref(x_5668); + x_5671 = lean_box(0); +} +if (lean_is_scalar(x_5671)) { + x_5672 = lean_alloc_ctor(0, 2, 0); +} else { + x_5672 = x_5671; +} +lean_ctor_set(x_5672, 0, x_5669); +lean_ctor_set(x_5672, 1, x_5670); +return x_5672; +} +else +{ +lean_object* x_5673; lean_object* x_5674; lean_object* x_5675; lean_object* x_5676; uint8_t x_5677; +x_5673 = lean_ctor_get(x_5668, 0); +lean_inc(x_5673); +x_5674 = lean_ctor_get(x_5668, 1); +lean_inc(x_5674); +if (lean_is_exclusive(x_5668)) { + lean_ctor_release(x_5668, 0); + lean_ctor_release(x_5668, 1); + x_5675 = x_5668; +} else { + lean_dec_ref(x_5668); + x_5675 = lean_box(0); +} +x_5676 = lean_ctor_get(x_5673, 1); +lean_inc(x_5676); +x_5677 = lean_nat_dec_eq(x_5662, x_5676); +lean_dec(x_5662); +if (x_5677 == 0) +{ +lean_object* x_5678; +lean_dec(x_5676); +if (lean_is_scalar(x_5675)) { + x_5678 = lean_alloc_ctor(1, 2, 0); +} else { + x_5678 = x_5675; +} +lean_ctor_set(x_5678, 0, x_5673); +lean_ctor_set(x_5678, 1, x_5674); +return x_5678; +} +else +{ +lean_object* x_5679; lean_object* x_5680; lean_object* x_5681; +lean_dec(x_5675); +lean_dec(x_5674); +x_5679 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_5680 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5680, 0, x_5679); +x_5681 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5667, x_5680, x_5673); +if (lean_obj_tag(x_5681) == 0) +{ +lean_object* x_5682; lean_object* x_5683; lean_object* x_5684; lean_object* x_5685; +lean_dec(x_5676); +x_5682 = lean_ctor_get(x_5681, 0); +lean_inc(x_5682); +x_5683 = lean_ctor_get(x_5681, 1); +lean_inc(x_5683); +if (lean_is_exclusive(x_5681)) { + lean_ctor_release(x_5681, 0); + lean_ctor_release(x_5681, 1); + x_5684 = x_5681; +} else { + lean_dec_ref(x_5681); + x_5684 = lean_box(0); +} +if (lean_is_scalar(x_5684)) { + x_5685 = lean_alloc_ctor(0, 2, 0); +} else { + x_5685 = x_5684; +} +lean_ctor_set(x_5685, 0, x_5682); +lean_ctor_set(x_5685, 1, x_5683); +return x_5685; +} +else +{ +lean_object* x_5686; lean_object* x_5687; lean_object* x_5688; lean_object* x_5689; uint8_t x_5690; +x_5686 = lean_ctor_get(x_5681, 0); +lean_inc(x_5686); +x_5687 = lean_ctor_get(x_5681, 1); +lean_inc(x_5687); +if (lean_is_exclusive(x_5681)) { + lean_ctor_release(x_5681, 0); + lean_ctor_release(x_5681, 1); + x_5688 = x_5681; +} else { + lean_dec_ref(x_5681); + x_5688 = lean_box(0); +} +x_5689 = lean_ctor_get(x_5686, 1); +lean_inc(x_5689); +x_5690 = lean_nat_dec_eq(x_5676, x_5689); +lean_dec(x_5676); +if (x_5690 == 0) +{ +lean_object* x_5691; +lean_dec(x_5689); +if (lean_is_scalar(x_5688)) { + x_5691 = lean_alloc_ctor(1, 2, 0); +} else { + x_5691 = x_5688; +} +lean_ctor_set(x_5691, 0, x_5686); +lean_ctor_set(x_5691, 1, x_5687); +return x_5691; +} +else +{ +lean_object* x_5692; lean_object* x_5693; lean_object* x_5694; lean_object* x_5695; +lean_dec(x_5688); +lean_dec(x_5687); +x_5692 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_5693 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5693, 0, x_5692); +x_5694 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_5695 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5694, x_5693, x_5686); +if (lean_obj_tag(x_5695) == 0) +{ +lean_object* x_5696; lean_object* x_5697; lean_object* x_5698; lean_object* x_5699; +lean_dec(x_5689); +x_5696 = lean_ctor_get(x_5695, 0); +lean_inc(x_5696); +x_5697 = lean_ctor_get(x_5695, 1); +lean_inc(x_5697); +if (lean_is_exclusive(x_5695)) { + lean_ctor_release(x_5695, 0); + lean_ctor_release(x_5695, 1); + x_5698 = x_5695; +} else { + lean_dec_ref(x_5695); + x_5698 = lean_box(0); +} +if (lean_is_scalar(x_5698)) { + x_5699 = lean_alloc_ctor(0, 2, 0); +} else { + x_5699 = x_5698; +} +lean_ctor_set(x_5699, 0, x_5696); +lean_ctor_set(x_5699, 1, x_5697); +return x_5699; +} +else +{ +lean_object* x_5700; lean_object* x_5701; lean_object* x_5702; lean_object* x_5703; uint8_t x_5704; +x_5700 = lean_ctor_get(x_5695, 0); +lean_inc(x_5700); +x_5701 = lean_ctor_get(x_5695, 1); +lean_inc(x_5701); +if (lean_is_exclusive(x_5695)) { + lean_ctor_release(x_5695, 0); + lean_ctor_release(x_5695, 1); + x_5702 = x_5695; +} else { + lean_dec_ref(x_5695); + x_5702 = lean_box(0); +} +x_5703 = lean_ctor_get(x_5700, 1); +lean_inc(x_5703); +x_5704 = lean_nat_dec_eq(x_5689, x_5703); +lean_dec(x_5689); +if (x_5704 == 0) +{ +lean_object* x_5705; +lean_dec(x_5703); +if (lean_is_scalar(x_5702)) { + x_5705 = lean_alloc_ctor(1, 2, 0); +} else { + x_5705 = x_5702; +} +lean_ctor_set(x_5705, 0, x_5700); +lean_ctor_set(x_5705, 1, x_5701); +return x_5705; +} +else +{ +lean_object* x_5706; lean_object* x_5707; lean_object* x_5708; lean_object* x_5709; +lean_dec(x_5702); +lean_dec(x_5701); +x_5706 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_5707 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5707, 0, x_5706); +x_5708 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_5709 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5708, x_5707, x_5700); +if (lean_obj_tag(x_5709) == 0) +{ +lean_object* x_5710; lean_object* x_5711; lean_object* x_5712; lean_object* x_5713; +lean_dec(x_5703); +x_5710 = lean_ctor_get(x_5709, 0); +lean_inc(x_5710); +x_5711 = lean_ctor_get(x_5709, 1); +lean_inc(x_5711); +if (lean_is_exclusive(x_5709)) { + lean_ctor_release(x_5709, 0); + lean_ctor_release(x_5709, 1); + x_5712 = x_5709; +} else { + lean_dec_ref(x_5709); + x_5712 = lean_box(0); +} +if (lean_is_scalar(x_5712)) { + x_5713 = lean_alloc_ctor(0, 2, 0); +} else { + x_5713 = x_5712; +} +lean_ctor_set(x_5713, 0, x_5710); +lean_ctor_set(x_5713, 1, x_5711); +return x_5713; +} +else +{ +lean_object* x_5714; lean_object* x_5715; lean_object* x_5716; lean_object* x_5717; uint8_t x_5718; +x_5714 = lean_ctor_get(x_5709, 0); +lean_inc(x_5714); +x_5715 = lean_ctor_get(x_5709, 1); +lean_inc(x_5715); +if (lean_is_exclusive(x_5709)) { + lean_ctor_release(x_5709, 0); + lean_ctor_release(x_5709, 1); + x_5716 = x_5709; +} else { + lean_dec_ref(x_5709); + x_5716 = lean_box(0); +} +x_5717 = lean_ctor_get(x_5714, 1); +lean_inc(x_5717); +x_5718 = lean_nat_dec_eq(x_5703, x_5717); +lean_dec(x_5703); +if (x_5718 == 0) +{ +lean_object* x_5719; +lean_dec(x_5717); +if (lean_is_scalar(x_5716)) { + x_5719 = lean_alloc_ctor(1, 2, 0); +} else { + x_5719 = x_5716; +} +lean_ctor_set(x_5719, 0, x_5714); +lean_ctor_set(x_5719, 1, x_5715); +return x_5719; +} +else +{ +lean_object* x_5720; lean_object* x_5721; lean_object* x_5722; lean_object* x_5723; +lean_dec(x_5716); +lean_dec(x_5715); +x_5720 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_5721 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5721, 0, x_5720); +x_5722 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_5723 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5722, x_5721, x_5714); +if (lean_obj_tag(x_5723) == 0) +{ +lean_object* x_5724; lean_object* x_5725; lean_object* x_5726; lean_object* x_5727; +lean_dec(x_5717); +x_5724 = lean_ctor_get(x_5723, 0); +lean_inc(x_5724); +x_5725 = lean_ctor_get(x_5723, 1); +lean_inc(x_5725); +if (lean_is_exclusive(x_5723)) { + lean_ctor_release(x_5723, 0); + lean_ctor_release(x_5723, 1); + x_5726 = x_5723; +} else { + lean_dec_ref(x_5723); + x_5726 = lean_box(0); +} +if (lean_is_scalar(x_5726)) { + x_5727 = lean_alloc_ctor(0, 2, 0); +} else { + x_5727 = x_5726; +} +lean_ctor_set(x_5727, 0, x_5724); +lean_ctor_set(x_5727, 1, x_5725); +return x_5727; +} +else +{ +lean_object* x_5728; lean_object* x_5729; lean_object* x_5730; lean_object* x_5731; uint8_t x_5732; +x_5728 = lean_ctor_get(x_5723, 0); +lean_inc(x_5728); +x_5729 = lean_ctor_get(x_5723, 1); +lean_inc(x_5729); +if (lean_is_exclusive(x_5723)) { + lean_ctor_release(x_5723, 0); + lean_ctor_release(x_5723, 1); + x_5730 = x_5723; +} else { + lean_dec_ref(x_5723); + x_5730 = lean_box(0); +} +x_5731 = lean_ctor_get(x_5728, 1); +lean_inc(x_5731); +x_5732 = lean_nat_dec_eq(x_5717, x_5731); +lean_dec(x_5717); +if (x_5732 == 0) +{ +lean_object* x_5733; +lean_dec(x_5731); +if (lean_is_scalar(x_5730)) { + x_5733 = lean_alloc_ctor(1, 2, 0); +} else { + x_5733 = x_5730; +} +lean_ctor_set(x_5733, 0, x_5728); +lean_ctor_set(x_5733, 1, x_5729); +return x_5733; +} +else +{ +lean_object* x_5734; lean_object* x_5735; lean_object* x_5736; lean_object* x_5737; +lean_dec(x_5730); +lean_dec(x_5729); +x_5734 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_5735 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5735, 0, x_5734); +x_5736 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_5737 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5736, x_5735, x_5728); +if (lean_obj_tag(x_5737) == 0) +{ +lean_object* x_5738; lean_object* x_5739; lean_object* x_5740; lean_object* x_5741; +lean_dec(x_5731); +x_5738 = lean_ctor_get(x_5737, 0); +lean_inc(x_5738); +x_5739 = lean_ctor_get(x_5737, 1); +lean_inc(x_5739); +if (lean_is_exclusive(x_5737)) { + lean_ctor_release(x_5737, 0); + lean_ctor_release(x_5737, 1); + x_5740 = x_5737; +} else { + lean_dec_ref(x_5737); + x_5740 = lean_box(0); +} +if (lean_is_scalar(x_5740)) { + x_5741 = lean_alloc_ctor(0, 2, 0); +} else { + x_5741 = x_5740; +} +lean_ctor_set(x_5741, 0, x_5738); +lean_ctor_set(x_5741, 1, x_5739); +return x_5741; +} +else +{ +lean_object* x_5742; lean_object* x_5743; lean_object* x_5744; lean_object* x_5745; uint8_t x_5746; +x_5742 = lean_ctor_get(x_5737, 0); +lean_inc(x_5742); +x_5743 = lean_ctor_get(x_5737, 1); +lean_inc(x_5743); +if (lean_is_exclusive(x_5737)) { + lean_ctor_release(x_5737, 0); + lean_ctor_release(x_5737, 1); + x_5744 = x_5737; +} else { + lean_dec_ref(x_5737); + x_5744 = lean_box(0); +} +x_5745 = lean_ctor_get(x_5742, 1); +lean_inc(x_5745); +x_5746 = lean_nat_dec_eq(x_5731, x_5745); +lean_dec(x_5731); +if (x_5746 == 0) +{ +lean_object* x_5747; +lean_dec(x_5745); +if (lean_is_scalar(x_5744)) { + x_5747 = lean_alloc_ctor(1, 2, 0); +} else { + x_5747 = x_5744; +} +lean_ctor_set(x_5747, 0, x_5742); +lean_ctor_set(x_5747, 1, x_5743); +return x_5747; +} +else +{ +lean_object* x_5748; lean_object* x_5749; lean_object* x_5750; lean_object* x_5751; +lean_dec(x_5744); +lean_dec(x_5743); +x_5748 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_5749 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5749, 0, x_5748); +x_5750 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_5751 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5750, x_5749, x_5742); +if (lean_obj_tag(x_5751) == 0) +{ +lean_object* x_5752; lean_object* x_5753; lean_object* x_5754; lean_object* x_5755; +lean_dec(x_5745); +x_5752 = lean_ctor_get(x_5751, 0); +lean_inc(x_5752); +x_5753 = lean_ctor_get(x_5751, 1); +lean_inc(x_5753); +if (lean_is_exclusive(x_5751)) { + lean_ctor_release(x_5751, 0); + lean_ctor_release(x_5751, 1); + x_5754 = x_5751; +} else { + lean_dec_ref(x_5751); + x_5754 = lean_box(0); +} +if (lean_is_scalar(x_5754)) { + x_5755 = lean_alloc_ctor(0, 2, 0); +} else { + x_5755 = x_5754; +} +lean_ctor_set(x_5755, 0, x_5752); +lean_ctor_set(x_5755, 1, x_5753); +return x_5755; +} +else +{ +lean_object* x_5756; lean_object* x_5757; lean_object* x_5758; lean_object* x_5759; uint8_t x_5760; +x_5756 = lean_ctor_get(x_5751, 0); +lean_inc(x_5756); +x_5757 = lean_ctor_get(x_5751, 1); +lean_inc(x_5757); +if (lean_is_exclusive(x_5751)) { + lean_ctor_release(x_5751, 0); + lean_ctor_release(x_5751, 1); + x_5758 = x_5751; +} else { + lean_dec_ref(x_5751); + x_5758 = lean_box(0); +} +x_5759 = lean_ctor_get(x_5756, 1); +lean_inc(x_5759); +x_5760 = lean_nat_dec_eq(x_5745, x_5759); +lean_dec(x_5745); +if (x_5760 == 0) +{ +lean_object* x_5761; +lean_dec(x_5759); +if (lean_is_scalar(x_5758)) { + x_5761 = lean_alloc_ctor(1, 2, 0); +} else { + x_5761 = x_5758; +} +lean_ctor_set(x_5761, 0, x_5756); +lean_ctor_set(x_5761, 1, x_5757); +return x_5761; +} +else +{ +lean_object* x_5762; lean_object* x_5763; lean_object* x_5764; lean_object* x_5765; +lean_dec(x_5758); +lean_dec(x_5757); +x_5762 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_5763 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5763, 0, x_5762); +x_5764 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_5765 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5764, x_5763, x_5756); +if (lean_obj_tag(x_5765) == 0) +{ +lean_object* x_5766; lean_object* x_5767; lean_object* x_5768; lean_object* x_5769; +lean_dec(x_5759); +x_5766 = lean_ctor_get(x_5765, 0); +lean_inc(x_5766); +x_5767 = lean_ctor_get(x_5765, 1); +lean_inc(x_5767); +if (lean_is_exclusive(x_5765)) { + lean_ctor_release(x_5765, 0); + lean_ctor_release(x_5765, 1); + x_5768 = x_5765; +} else { + lean_dec_ref(x_5765); + x_5768 = lean_box(0); +} +if (lean_is_scalar(x_5768)) { + x_5769 = lean_alloc_ctor(0, 2, 0); +} else { + x_5769 = x_5768; +} +lean_ctor_set(x_5769, 0, x_5766); +lean_ctor_set(x_5769, 1, x_5767); +return x_5769; +} +else +{ +lean_object* x_5770; lean_object* x_5771; lean_object* x_5772; lean_object* x_5773; uint8_t x_5774; +x_5770 = lean_ctor_get(x_5765, 0); +lean_inc(x_5770); +x_5771 = lean_ctor_get(x_5765, 1); +lean_inc(x_5771); +if (lean_is_exclusive(x_5765)) { + lean_ctor_release(x_5765, 0); + lean_ctor_release(x_5765, 1); + x_5772 = x_5765; +} else { + lean_dec_ref(x_5765); + x_5772 = lean_box(0); +} +x_5773 = lean_ctor_get(x_5770, 1); +lean_inc(x_5773); +x_5774 = lean_nat_dec_eq(x_5759, x_5773); +lean_dec(x_5759); +if (x_5774 == 0) +{ +lean_object* x_5775; +lean_dec(x_5773); +if (lean_is_scalar(x_5772)) { + x_5775 = lean_alloc_ctor(1, 2, 0); +} else { + x_5775 = x_5772; +} +lean_ctor_set(x_5775, 0, x_5770); +lean_ctor_set(x_5775, 1, x_5771); +return x_5775; +} +else +{ +lean_object* x_5776; lean_object* x_5777; lean_object* x_5778; lean_object* x_5779; +lean_dec(x_5772); +lean_dec(x_5771); +x_5776 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_5777 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5777, 0, x_5776); +x_5778 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_5779 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5778, x_5777, x_5770); +if (lean_obj_tag(x_5779) == 0) +{ +lean_object* x_5780; lean_object* x_5781; lean_object* x_5782; lean_object* x_5783; +lean_dec(x_5773); +x_5780 = lean_ctor_get(x_5779, 0); +lean_inc(x_5780); +x_5781 = lean_ctor_get(x_5779, 1); +lean_inc(x_5781); +if (lean_is_exclusive(x_5779)) { + lean_ctor_release(x_5779, 0); + lean_ctor_release(x_5779, 1); + x_5782 = x_5779; +} else { + lean_dec_ref(x_5779); + x_5782 = lean_box(0); +} +if (lean_is_scalar(x_5782)) { + x_5783 = lean_alloc_ctor(0, 2, 0); +} else { + x_5783 = x_5782; +} +lean_ctor_set(x_5783, 0, x_5780); +lean_ctor_set(x_5783, 1, x_5781); +return x_5783; +} +else +{ +lean_object* x_5784; lean_object* x_5785; lean_object* x_5786; lean_object* x_5787; uint8_t x_5788; +x_5784 = lean_ctor_get(x_5779, 0); +lean_inc(x_5784); +x_5785 = lean_ctor_get(x_5779, 1); +lean_inc(x_5785); +if (lean_is_exclusive(x_5779)) { + lean_ctor_release(x_5779, 0); + lean_ctor_release(x_5779, 1); + x_5786 = x_5779; +} else { + lean_dec_ref(x_5779); + x_5786 = lean_box(0); +} +x_5787 = lean_ctor_get(x_5784, 1); +lean_inc(x_5787); +x_5788 = lean_nat_dec_eq(x_5773, x_5787); +lean_dec(x_5773); +if (x_5788 == 0) +{ +lean_object* x_5789; +lean_dec(x_5787); +if (lean_is_scalar(x_5786)) { + x_5789 = lean_alloc_ctor(1, 2, 0); +} else { + x_5789 = x_5786; +} +lean_ctor_set(x_5789, 0, x_5784); +lean_ctor_set(x_5789, 1, x_5785); +return x_5789; +} +else +{ +lean_object* x_5790; lean_object* x_5791; lean_object* x_5792; lean_object* x_5793; +lean_dec(x_5786); +lean_dec(x_5785); +x_5790 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_5791 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5791, 0, x_5790); +x_5792 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_5793 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5792, x_5791, x_5784); +if (lean_obj_tag(x_5793) == 0) +{ +lean_object* x_5794; lean_object* x_5795; lean_object* x_5796; lean_object* x_5797; +lean_dec(x_5787); +x_5794 = lean_ctor_get(x_5793, 0); +lean_inc(x_5794); +x_5795 = lean_ctor_get(x_5793, 1); +lean_inc(x_5795); +if (lean_is_exclusive(x_5793)) { + lean_ctor_release(x_5793, 0); + lean_ctor_release(x_5793, 1); + x_5796 = x_5793; +} else { + lean_dec_ref(x_5793); + x_5796 = lean_box(0); +} +if (lean_is_scalar(x_5796)) { + x_5797 = lean_alloc_ctor(0, 2, 0); +} else { + x_5797 = x_5796; +} +lean_ctor_set(x_5797, 0, x_5794); +lean_ctor_set(x_5797, 1, x_5795); +return x_5797; +} +else +{ +lean_object* x_5798; lean_object* x_5799; lean_object* x_5800; lean_object* x_5801; uint8_t x_5802; +x_5798 = lean_ctor_get(x_5793, 0); +lean_inc(x_5798); +x_5799 = lean_ctor_get(x_5793, 1); +lean_inc(x_5799); +if (lean_is_exclusive(x_5793)) { + lean_ctor_release(x_5793, 0); + lean_ctor_release(x_5793, 1); + x_5800 = x_5793; +} else { + lean_dec_ref(x_5793); + x_5800 = lean_box(0); +} +x_5801 = lean_ctor_get(x_5798, 1); +lean_inc(x_5801); +x_5802 = lean_nat_dec_eq(x_5787, x_5801); +lean_dec(x_5787); +if (x_5802 == 0) +{ +lean_object* x_5803; +lean_dec(x_5801); +if (lean_is_scalar(x_5800)) { + x_5803 = lean_alloc_ctor(1, 2, 0); +} else { + x_5803 = x_5800; +} +lean_ctor_set(x_5803, 0, x_5798); +lean_ctor_set(x_5803, 1, x_5799); +return x_5803; +} +else +{ +lean_object* x_5804; lean_object* x_5805; lean_object* x_5806; lean_object* x_5807; +lean_dec(x_5800); +lean_dec(x_5799); +x_5804 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_5805 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5805, 0, x_5804); +x_5806 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_5807 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5806, x_5805, x_5798); +if (lean_obj_tag(x_5807) == 0) +{ +lean_object* x_5808; lean_object* x_5809; lean_object* x_5810; lean_object* x_5811; +lean_dec(x_5801); +x_5808 = lean_ctor_get(x_5807, 0); +lean_inc(x_5808); +x_5809 = lean_ctor_get(x_5807, 1); +lean_inc(x_5809); +if (lean_is_exclusive(x_5807)) { + lean_ctor_release(x_5807, 0); + lean_ctor_release(x_5807, 1); + x_5810 = x_5807; +} else { + lean_dec_ref(x_5807); + x_5810 = lean_box(0); +} +if (lean_is_scalar(x_5810)) { + x_5811 = lean_alloc_ctor(0, 2, 0); +} else { + x_5811 = x_5810; +} +lean_ctor_set(x_5811, 0, x_5808); +lean_ctor_set(x_5811, 1, x_5809); +return x_5811; +} +else +{ +lean_object* x_5812; lean_object* x_5813; lean_object* x_5814; lean_object* x_5815; uint8_t x_5816; +x_5812 = lean_ctor_get(x_5807, 0); +lean_inc(x_5812); +x_5813 = lean_ctor_get(x_5807, 1); +lean_inc(x_5813); +if (lean_is_exclusive(x_5807)) { + lean_ctor_release(x_5807, 0); + lean_ctor_release(x_5807, 1); + x_5814 = x_5807; +} else { + lean_dec_ref(x_5807); + x_5814 = lean_box(0); +} +x_5815 = lean_ctor_get(x_5812, 1); +lean_inc(x_5815); +x_5816 = lean_nat_dec_eq(x_5801, x_5815); +lean_dec(x_5801); +if (x_5816 == 0) +{ +lean_object* x_5817; +lean_dec(x_5815); +if (lean_is_scalar(x_5814)) { + x_5817 = lean_alloc_ctor(1, 2, 0); +} else { + x_5817 = x_5814; +} +lean_ctor_set(x_5817, 0, x_5812); +lean_ctor_set(x_5817, 1, x_5813); +return x_5817; +} +else +{ +lean_object* x_5818; lean_object* x_5819; lean_object* x_5820; lean_object* x_5821; +lean_dec(x_5814); +lean_dec(x_5813); +x_5818 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_5819 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5819, 0, x_5818); +x_5820 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_5821 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5820, x_5819, x_5812); +if (lean_obj_tag(x_5821) == 0) +{ +lean_object* x_5822; lean_object* x_5823; lean_object* x_5824; lean_object* x_5825; +lean_dec(x_5815); +x_5822 = lean_ctor_get(x_5821, 0); +lean_inc(x_5822); +x_5823 = lean_ctor_get(x_5821, 1); +lean_inc(x_5823); +if (lean_is_exclusive(x_5821)) { + lean_ctor_release(x_5821, 0); + lean_ctor_release(x_5821, 1); + x_5824 = x_5821; +} else { + lean_dec_ref(x_5821); + x_5824 = lean_box(0); +} +if (lean_is_scalar(x_5824)) { + x_5825 = lean_alloc_ctor(0, 2, 0); +} else { + x_5825 = x_5824; +} +lean_ctor_set(x_5825, 0, x_5822); +lean_ctor_set(x_5825, 1, x_5823); +return x_5825; +} +else +{ +lean_object* x_5826; lean_object* x_5827; lean_object* x_5828; lean_object* x_5829; uint8_t x_5830; +x_5826 = lean_ctor_get(x_5821, 0); +lean_inc(x_5826); +x_5827 = lean_ctor_get(x_5821, 1); +lean_inc(x_5827); +if (lean_is_exclusive(x_5821)) { + lean_ctor_release(x_5821, 0); + lean_ctor_release(x_5821, 1); + x_5828 = x_5821; +} else { + lean_dec_ref(x_5821); + x_5828 = lean_box(0); +} +x_5829 = lean_ctor_get(x_5826, 1); +lean_inc(x_5829); +x_5830 = lean_nat_dec_eq(x_5815, x_5829); +lean_dec(x_5815); +if (x_5830 == 0) +{ +lean_object* x_5831; +lean_dec(x_5829); +if (lean_is_scalar(x_5828)) { + x_5831 = lean_alloc_ctor(1, 2, 0); +} else { + x_5831 = x_5828; +} +lean_ctor_set(x_5831, 0, x_5826); +lean_ctor_set(x_5831, 1, x_5827); +return x_5831; +} +else +{ +lean_object* x_5832; lean_object* x_5833; lean_object* x_5834; lean_object* x_5835; +lean_dec(x_5828); +lean_dec(x_5827); +x_5832 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_5833 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5833, 0, x_5832); +x_5834 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_5835 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5834, x_5833, x_5826); +if (lean_obj_tag(x_5835) == 0) +{ +lean_object* x_5836; lean_object* x_5837; lean_object* x_5838; lean_object* x_5839; +lean_dec(x_5829); +x_5836 = lean_ctor_get(x_5835, 0); +lean_inc(x_5836); +x_5837 = lean_ctor_get(x_5835, 1); +lean_inc(x_5837); +if (lean_is_exclusive(x_5835)) { + lean_ctor_release(x_5835, 0); + lean_ctor_release(x_5835, 1); + x_5838 = x_5835; +} else { + lean_dec_ref(x_5835); + x_5838 = lean_box(0); +} +if (lean_is_scalar(x_5838)) { + x_5839 = lean_alloc_ctor(0, 2, 0); +} else { + x_5839 = x_5838; +} +lean_ctor_set(x_5839, 0, x_5836); +lean_ctor_set(x_5839, 1, x_5837); +return x_5839; +} +else +{ +lean_object* x_5840; lean_object* x_5841; lean_object* x_5842; lean_object* x_5843; uint8_t x_5844; +x_5840 = lean_ctor_get(x_5835, 0); +lean_inc(x_5840); +x_5841 = lean_ctor_get(x_5835, 1); +lean_inc(x_5841); +if (lean_is_exclusive(x_5835)) { + lean_ctor_release(x_5835, 0); + lean_ctor_release(x_5835, 1); + x_5842 = x_5835; +} else { + lean_dec_ref(x_5835); + x_5842 = lean_box(0); +} +x_5843 = lean_ctor_get(x_5840, 1); +lean_inc(x_5843); +x_5844 = lean_nat_dec_eq(x_5829, x_5843); +lean_dec(x_5829); +if (x_5844 == 0) +{ +lean_object* x_5845; +lean_dec(x_5843); +if (lean_is_scalar(x_5842)) { + x_5845 = lean_alloc_ctor(1, 2, 0); +} else { + x_5845 = x_5842; +} +lean_ctor_set(x_5845, 0, x_5840); +lean_ctor_set(x_5845, 1, x_5841); +return x_5845; +} +else +{ +lean_object* x_5846; lean_object* x_5847; lean_object* x_5848; lean_object* x_5849; +lean_dec(x_5842); +lean_dec(x_5841); +x_5846 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_5847 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5847, 0, x_5846); +x_5848 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_5849 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5848, x_5847, x_5840); +if (lean_obj_tag(x_5849) == 0) +{ +lean_object* x_5850; lean_object* x_5851; lean_object* x_5852; lean_object* x_5853; +lean_dec(x_5843); +x_5850 = lean_ctor_get(x_5849, 0); +lean_inc(x_5850); +x_5851 = lean_ctor_get(x_5849, 1); +lean_inc(x_5851); +if (lean_is_exclusive(x_5849)) { + lean_ctor_release(x_5849, 0); + lean_ctor_release(x_5849, 1); + x_5852 = x_5849; +} else { + lean_dec_ref(x_5849); + x_5852 = lean_box(0); +} +if (lean_is_scalar(x_5852)) { + x_5853 = lean_alloc_ctor(0, 2, 0); +} else { + x_5853 = x_5852; +} +lean_ctor_set(x_5853, 0, x_5850); +lean_ctor_set(x_5853, 1, x_5851); +return x_5853; +} +else +{ +lean_object* x_5854; lean_object* x_5855; lean_object* x_5856; lean_object* x_5857; uint8_t x_5858; +x_5854 = lean_ctor_get(x_5849, 0); +lean_inc(x_5854); +x_5855 = lean_ctor_get(x_5849, 1); +lean_inc(x_5855); +if (lean_is_exclusive(x_5849)) { + lean_ctor_release(x_5849, 0); + lean_ctor_release(x_5849, 1); + x_5856 = x_5849; +} else { + lean_dec_ref(x_5849); + x_5856 = lean_box(0); +} +x_5857 = lean_ctor_get(x_5854, 1); +lean_inc(x_5857); +x_5858 = lean_nat_dec_eq(x_5843, x_5857); +lean_dec(x_5843); +if (x_5858 == 0) +{ +lean_object* x_5859; +lean_dec(x_5857); +if (lean_is_scalar(x_5856)) { + x_5859 = lean_alloc_ctor(1, 2, 0); +} else { + x_5859 = x_5856; +} +lean_ctor_set(x_5859, 0, x_5854); +lean_ctor_set(x_5859, 1, x_5855); +return x_5859; +} +else +{ +lean_object* x_5860; lean_object* x_5861; lean_object* x_5862; lean_object* x_5863; +lean_dec(x_5856); +lean_dec(x_5855); +x_5860 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_5861 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5861, 0, x_5860); +x_5862 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_5863 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5862, x_5861, x_5854); +if (lean_obj_tag(x_5863) == 0) +{ +lean_object* x_5864; lean_object* x_5865; lean_object* x_5866; lean_object* x_5867; +lean_dec(x_5857); +x_5864 = lean_ctor_get(x_5863, 0); +lean_inc(x_5864); +x_5865 = lean_ctor_get(x_5863, 1); +lean_inc(x_5865); +if (lean_is_exclusive(x_5863)) { + lean_ctor_release(x_5863, 0); + lean_ctor_release(x_5863, 1); + x_5866 = x_5863; +} else { + lean_dec_ref(x_5863); + x_5866 = lean_box(0); +} +if (lean_is_scalar(x_5866)) { + x_5867 = lean_alloc_ctor(0, 2, 0); +} else { + x_5867 = x_5866; +} +lean_ctor_set(x_5867, 0, x_5864); +lean_ctor_set(x_5867, 1, x_5865); +return x_5867; +} +else +{ +lean_object* x_5868; lean_object* x_5869; lean_object* x_5870; lean_object* x_5871; uint8_t x_5872; +x_5868 = lean_ctor_get(x_5863, 0); +lean_inc(x_5868); +x_5869 = lean_ctor_get(x_5863, 1); +lean_inc(x_5869); +if (lean_is_exclusive(x_5863)) { + lean_ctor_release(x_5863, 0); + lean_ctor_release(x_5863, 1); + x_5870 = x_5863; +} else { + lean_dec_ref(x_5863); + x_5870 = lean_box(0); +} +x_5871 = lean_ctor_get(x_5868, 1); +lean_inc(x_5871); +x_5872 = lean_nat_dec_eq(x_5857, x_5871); +lean_dec(x_5857); +if (x_5872 == 0) +{ +lean_object* x_5873; +lean_dec(x_5871); +if (lean_is_scalar(x_5870)) { + x_5873 = lean_alloc_ctor(1, 2, 0); +} else { + x_5873 = x_5870; +} +lean_ctor_set(x_5873, 0, x_5868); +lean_ctor_set(x_5873, 1, x_5869); +return x_5873; +} +else +{ +lean_object* x_5874; lean_object* x_5875; lean_object* x_5876; lean_object* x_5877; +lean_dec(x_5870); +lean_dec(x_5869); +x_5874 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_5875 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5875, 0, x_5874); +x_5876 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_5877 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5876, x_5875, x_5868); +if (lean_obj_tag(x_5877) == 0) +{ +lean_object* x_5878; lean_object* x_5879; lean_object* x_5880; lean_object* x_5881; +lean_dec(x_5871); +x_5878 = lean_ctor_get(x_5877, 0); +lean_inc(x_5878); +x_5879 = lean_ctor_get(x_5877, 1); +lean_inc(x_5879); +if (lean_is_exclusive(x_5877)) { + lean_ctor_release(x_5877, 0); + lean_ctor_release(x_5877, 1); + x_5880 = x_5877; +} else { + lean_dec_ref(x_5877); + x_5880 = lean_box(0); +} +if (lean_is_scalar(x_5880)) { + x_5881 = lean_alloc_ctor(0, 2, 0); +} else { + x_5881 = x_5880; +} +lean_ctor_set(x_5881, 0, x_5878); +lean_ctor_set(x_5881, 1, x_5879); +return x_5881; +} +else +{ +lean_object* x_5882; lean_object* x_5883; lean_object* x_5884; lean_object* x_5885; uint8_t x_5886; +x_5882 = lean_ctor_get(x_5877, 0); +lean_inc(x_5882); +x_5883 = lean_ctor_get(x_5877, 1); +lean_inc(x_5883); +if (lean_is_exclusive(x_5877)) { + lean_ctor_release(x_5877, 0); + lean_ctor_release(x_5877, 1); + x_5884 = x_5877; +} else { + lean_dec_ref(x_5877); + x_5884 = lean_box(0); +} +x_5885 = lean_ctor_get(x_5882, 1); +lean_inc(x_5885); +x_5886 = lean_nat_dec_eq(x_5871, x_5885); +lean_dec(x_5871); +if (x_5886 == 0) +{ +lean_object* x_5887; +lean_dec(x_5885); +if (lean_is_scalar(x_5884)) { + x_5887 = lean_alloc_ctor(1, 2, 0); +} else { + x_5887 = x_5884; +} +lean_ctor_set(x_5887, 0, x_5882); +lean_ctor_set(x_5887, 1, x_5883); +return x_5887; +} +else +{ +lean_object* x_5888; lean_object* x_5889; lean_object* x_5890; lean_object* x_5891; +lean_dec(x_5884); +lean_dec(x_5883); +x_5888 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_5889 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5889, 0, x_5888); +x_5890 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_5891 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5890, x_5889, x_5882); +if (lean_obj_tag(x_5891) == 0) +{ +lean_object* x_5892; lean_object* x_5893; lean_object* x_5894; lean_object* x_5895; +lean_dec(x_5885); +x_5892 = lean_ctor_get(x_5891, 0); +lean_inc(x_5892); +x_5893 = lean_ctor_get(x_5891, 1); +lean_inc(x_5893); +if (lean_is_exclusive(x_5891)) { + lean_ctor_release(x_5891, 0); + lean_ctor_release(x_5891, 1); + x_5894 = x_5891; +} else { + lean_dec_ref(x_5891); + x_5894 = lean_box(0); +} +if (lean_is_scalar(x_5894)) { + x_5895 = lean_alloc_ctor(0, 2, 0); +} else { + x_5895 = x_5894; +} +lean_ctor_set(x_5895, 0, x_5892); +lean_ctor_set(x_5895, 1, x_5893); +return x_5895; +} +else +{ +lean_object* x_5896; lean_object* x_5897; lean_object* x_5898; lean_object* x_5899; uint8_t x_5900; +x_5896 = lean_ctor_get(x_5891, 0); +lean_inc(x_5896); +x_5897 = lean_ctor_get(x_5891, 1); +lean_inc(x_5897); +if (lean_is_exclusive(x_5891)) { + lean_ctor_release(x_5891, 0); + lean_ctor_release(x_5891, 1); + x_5898 = x_5891; +} else { + lean_dec_ref(x_5891); + x_5898 = lean_box(0); +} +x_5899 = lean_ctor_get(x_5896, 1); +lean_inc(x_5899); +x_5900 = lean_nat_dec_eq(x_5885, x_5899); +lean_dec(x_5885); +if (x_5900 == 0) +{ +lean_object* x_5901; +lean_dec(x_5899); +if (lean_is_scalar(x_5898)) { + x_5901 = lean_alloc_ctor(1, 2, 0); +} else { + x_5901 = x_5898; +} +lean_ctor_set(x_5901, 0, x_5896); +lean_ctor_set(x_5901, 1, x_5897); +return x_5901; +} +else +{ +lean_object* x_5902; lean_object* x_5903; lean_object* x_5904; lean_object* x_5905; +lean_dec(x_5898); +lean_dec(x_5897); +x_5902 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_5903 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5903, 0, x_5902); +x_5904 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_5905 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5904, x_5903, x_5896); +if (lean_obj_tag(x_5905) == 0) +{ +lean_object* x_5906; lean_object* x_5907; lean_object* x_5908; lean_object* x_5909; +lean_dec(x_5899); +x_5906 = lean_ctor_get(x_5905, 0); +lean_inc(x_5906); +x_5907 = lean_ctor_get(x_5905, 1); +lean_inc(x_5907); +if (lean_is_exclusive(x_5905)) { + lean_ctor_release(x_5905, 0); + lean_ctor_release(x_5905, 1); + x_5908 = x_5905; +} else { + lean_dec_ref(x_5905); + x_5908 = lean_box(0); +} +if (lean_is_scalar(x_5908)) { + x_5909 = lean_alloc_ctor(0, 2, 0); +} else { + x_5909 = x_5908; +} +lean_ctor_set(x_5909, 0, x_5906); +lean_ctor_set(x_5909, 1, x_5907); +return x_5909; +} +else +{ +lean_object* x_5910; lean_object* x_5911; lean_object* x_5912; lean_object* x_5913; uint8_t x_5914; +x_5910 = lean_ctor_get(x_5905, 0); +lean_inc(x_5910); +x_5911 = lean_ctor_get(x_5905, 1); +lean_inc(x_5911); +if (lean_is_exclusive(x_5905)) { + lean_ctor_release(x_5905, 0); + lean_ctor_release(x_5905, 1); + x_5912 = x_5905; +} else { + lean_dec_ref(x_5905); + x_5912 = lean_box(0); +} +x_5913 = lean_ctor_get(x_5910, 1); +lean_inc(x_5913); +x_5914 = lean_nat_dec_eq(x_5899, x_5913); +lean_dec(x_5899); +if (x_5914 == 0) +{ +lean_object* x_5915; +lean_dec(x_5913); +if (lean_is_scalar(x_5912)) { + x_5915 = lean_alloc_ctor(1, 2, 0); +} else { + x_5915 = x_5912; +} +lean_ctor_set(x_5915, 0, x_5910); +lean_ctor_set(x_5915, 1, x_5911); +return x_5915; +} +else +{ +lean_object* x_5916; lean_object* x_5917; lean_object* x_5918; lean_object* x_5919; +lean_dec(x_5912); +lean_dec(x_5911); +x_5916 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_5917 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5917, 0, x_5916); +x_5918 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_5919 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5918, x_5917, x_5910); +if (lean_obj_tag(x_5919) == 0) +{ +lean_object* x_5920; lean_object* x_5921; lean_object* x_5922; lean_object* x_5923; +lean_dec(x_5913); +x_5920 = lean_ctor_get(x_5919, 0); +lean_inc(x_5920); +x_5921 = lean_ctor_get(x_5919, 1); +lean_inc(x_5921); +if (lean_is_exclusive(x_5919)) { + lean_ctor_release(x_5919, 0); + lean_ctor_release(x_5919, 1); + x_5922 = x_5919; +} else { + lean_dec_ref(x_5919); + x_5922 = lean_box(0); +} +if (lean_is_scalar(x_5922)) { + x_5923 = lean_alloc_ctor(0, 2, 0); +} else { + x_5923 = x_5922; +} +lean_ctor_set(x_5923, 0, x_5920); +lean_ctor_set(x_5923, 1, x_5921); +return x_5923; +} +else +{ +lean_object* x_5924; lean_object* x_5925; lean_object* x_5926; lean_object* x_5927; uint8_t x_5928; +x_5924 = lean_ctor_get(x_5919, 0); +lean_inc(x_5924); +x_5925 = lean_ctor_get(x_5919, 1); +lean_inc(x_5925); +if (lean_is_exclusive(x_5919)) { + lean_ctor_release(x_5919, 0); + lean_ctor_release(x_5919, 1); + x_5926 = x_5919; +} else { + lean_dec_ref(x_5919); + x_5926 = lean_box(0); +} +x_5927 = lean_ctor_get(x_5924, 1); +lean_inc(x_5927); +x_5928 = lean_nat_dec_eq(x_5913, x_5927); +lean_dec(x_5927); +lean_dec(x_5913); +if (x_5928 == 0) +{ +lean_object* x_5929; +if (lean_is_scalar(x_5926)) { + x_5929 = lean_alloc_ctor(1, 2, 0); +} else { + x_5929 = x_5926; +} +lean_ctor_set(x_5929, 0, x_5924); +lean_ctor_set(x_5929, 1, x_5925); +return x_5929; +} +else +{ +lean_object* x_5930; lean_object* x_5931; lean_object* x_5932; lean_object* x_5933; +lean_dec(x_5926); +lean_dec(x_5925); +x_5930 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_5931 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5931, 0, x_5930); +x_5932 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_5933 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5932, x_5931, x_5924); +return x_5933; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_5934; lean_object* x_5935; lean_object* x_5936; uint8_t x_5937; +x_5934 = lean_ctor_get(x_32, 0); +x_5935 = lean_ctor_get(x_32, 1); +lean_inc(x_5935); +lean_inc(x_5934); +lean_dec(x_32); +x_5936 = lean_ctor_get(x_5934, 1); +lean_inc(x_5936); +x_5937 = lean_nat_dec_eq(x_27, x_5936); +lean_dec(x_27); +if (x_5937 == 0) +{ +lean_object* x_5938; +lean_dec(x_5936); +x_5938 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5938, 0, x_5934); +lean_ctor_set(x_5938, 1, x_5935); +return x_5938; +} +else +{ +lean_object* x_5939; lean_object* x_5940; lean_object* x_5941; lean_object* x_5942; +lean_dec(x_5935); +x_5939 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12; +x_5940 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5940, 0, x_5939); +x_5941 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11; +x_5942 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5941, x_5940, x_5934); +if (lean_obj_tag(x_5942) == 0) +{ +lean_object* x_5943; lean_object* x_5944; lean_object* x_5945; lean_object* x_5946; +lean_dec(x_5936); +x_5943 = lean_ctor_get(x_5942, 0); +lean_inc(x_5943); +x_5944 = lean_ctor_get(x_5942, 1); +lean_inc(x_5944); +if (lean_is_exclusive(x_5942)) { + lean_ctor_release(x_5942, 0); + lean_ctor_release(x_5942, 1); + x_5945 = x_5942; +} else { + lean_dec_ref(x_5942); + x_5945 = lean_box(0); +} +if (lean_is_scalar(x_5945)) { + x_5946 = lean_alloc_ctor(0, 2, 0); +} else { + x_5946 = x_5945; +} +lean_ctor_set(x_5946, 0, x_5943); +lean_ctor_set(x_5946, 1, x_5944); +return x_5946; +} +else +{ +lean_object* x_5947; lean_object* x_5948; lean_object* x_5949; lean_object* x_5950; uint8_t x_5951; +x_5947 = lean_ctor_get(x_5942, 0); +lean_inc(x_5947); +x_5948 = lean_ctor_get(x_5942, 1); +lean_inc(x_5948); +if (lean_is_exclusive(x_5942)) { + lean_ctor_release(x_5942, 0); + lean_ctor_release(x_5942, 1); + x_5949 = x_5942; +} else { + lean_dec_ref(x_5942); + x_5949 = lean_box(0); +} +x_5950 = lean_ctor_get(x_5947, 1); +lean_inc(x_5950); +x_5951 = lean_nat_dec_eq(x_5936, x_5950); +lean_dec(x_5936); +if (x_5951 == 0) +{ +lean_object* x_5952; +lean_dec(x_5950); +if (lean_is_scalar(x_5949)) { + x_5952 = lean_alloc_ctor(1, 2, 0); +} else { + x_5952 = x_5949; +} +lean_ctor_set(x_5952, 0, x_5947); +lean_ctor_set(x_5952, 1, x_5948); +return x_5952; +} +else +{ +lean_object* x_5953; lean_object* x_5954; lean_object* x_5955; lean_object* x_5956; +lean_dec(x_5949); +lean_dec(x_5948); +x_5953 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +x_5954 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5954, 0, x_5953); +x_5955 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +x_5956 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5955, x_5954, x_5947); +if (lean_obj_tag(x_5956) == 0) +{ +lean_object* x_5957; lean_object* x_5958; lean_object* x_5959; lean_object* x_5960; +lean_dec(x_5950); +x_5957 = lean_ctor_get(x_5956, 0); +lean_inc(x_5957); +x_5958 = lean_ctor_get(x_5956, 1); +lean_inc(x_5958); +if (lean_is_exclusive(x_5956)) { + lean_ctor_release(x_5956, 0); + lean_ctor_release(x_5956, 1); + x_5959 = x_5956; +} else { + lean_dec_ref(x_5956); + x_5959 = lean_box(0); +} +if (lean_is_scalar(x_5959)) { + x_5960 = lean_alloc_ctor(0, 2, 0); +} else { + x_5960 = x_5959; +} +lean_ctor_set(x_5960, 0, x_5957); +lean_ctor_set(x_5960, 1, x_5958); +return x_5960; +} +else +{ +lean_object* x_5961; lean_object* x_5962; lean_object* x_5963; lean_object* x_5964; uint8_t x_5965; +x_5961 = lean_ctor_get(x_5956, 0); +lean_inc(x_5961); +x_5962 = lean_ctor_get(x_5956, 1); +lean_inc(x_5962); +if (lean_is_exclusive(x_5956)) { + lean_ctor_release(x_5956, 0); + lean_ctor_release(x_5956, 1); + x_5963 = x_5956; +} else { + lean_dec_ref(x_5956); + x_5963 = lean_box(0); +} +x_5964 = lean_ctor_get(x_5961, 1); +lean_inc(x_5964); +x_5965 = lean_nat_dec_eq(x_5950, x_5964); +lean_dec(x_5950); +if (x_5965 == 0) +{ +lean_object* x_5966; +lean_dec(x_5964); +if (lean_is_scalar(x_5963)) { + x_5966 = lean_alloc_ctor(1, 2, 0); +} else { + x_5966 = x_5963; +} +lean_ctor_set(x_5966, 0, x_5961); +lean_ctor_set(x_5966, 1, x_5962); +return x_5966; +} +else +{ +lean_object* x_5967; lean_object* x_5968; lean_object* x_5969; +lean_dec(x_5963); +lean_dec(x_5962); +x_5967 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_5968 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5968, 0, x_5967); +x_5969 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5955, x_5968, x_5961); +if (lean_obj_tag(x_5969) == 0) +{ +lean_object* x_5970; lean_object* x_5971; lean_object* x_5972; lean_object* x_5973; +lean_dec(x_5964); +x_5970 = lean_ctor_get(x_5969, 0); +lean_inc(x_5970); +x_5971 = lean_ctor_get(x_5969, 1); +lean_inc(x_5971); +if (lean_is_exclusive(x_5969)) { + lean_ctor_release(x_5969, 0); + lean_ctor_release(x_5969, 1); + x_5972 = x_5969; +} else { + lean_dec_ref(x_5969); + x_5972 = lean_box(0); +} +if (lean_is_scalar(x_5972)) { + x_5973 = lean_alloc_ctor(0, 2, 0); +} else { + x_5973 = x_5972; +} +lean_ctor_set(x_5973, 0, x_5970); +lean_ctor_set(x_5973, 1, x_5971); +return x_5973; +} +else +{ +lean_object* x_5974; lean_object* x_5975; lean_object* x_5976; lean_object* x_5977; uint8_t x_5978; +x_5974 = lean_ctor_get(x_5969, 0); +lean_inc(x_5974); +x_5975 = lean_ctor_get(x_5969, 1); +lean_inc(x_5975); +if (lean_is_exclusive(x_5969)) { + lean_ctor_release(x_5969, 0); + lean_ctor_release(x_5969, 1); + x_5976 = x_5969; +} else { + lean_dec_ref(x_5969); + x_5976 = lean_box(0); +} +x_5977 = lean_ctor_get(x_5974, 1); +lean_inc(x_5977); +x_5978 = lean_nat_dec_eq(x_5964, x_5977); +lean_dec(x_5964); +if (x_5978 == 0) +{ +lean_object* x_5979; +lean_dec(x_5977); +if (lean_is_scalar(x_5976)) { + x_5979 = lean_alloc_ctor(1, 2, 0); +} else { + x_5979 = x_5976; +} +lean_ctor_set(x_5979, 0, x_5974); +lean_ctor_set(x_5979, 1, x_5975); +return x_5979; +} +else +{ +lean_object* x_5980; lean_object* x_5981; lean_object* x_5982; lean_object* x_5983; +lean_dec(x_5976); +lean_dec(x_5975); +x_5980 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_5981 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5981, 0, x_5980); +x_5982 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_5983 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5982, x_5981, x_5974); +if (lean_obj_tag(x_5983) == 0) +{ +lean_object* x_5984; lean_object* x_5985; lean_object* x_5986; lean_object* x_5987; +lean_dec(x_5977); +x_5984 = lean_ctor_get(x_5983, 0); +lean_inc(x_5984); +x_5985 = lean_ctor_get(x_5983, 1); +lean_inc(x_5985); +if (lean_is_exclusive(x_5983)) { + lean_ctor_release(x_5983, 0); + lean_ctor_release(x_5983, 1); + x_5986 = x_5983; +} else { + lean_dec_ref(x_5983); + x_5986 = lean_box(0); +} +if (lean_is_scalar(x_5986)) { + x_5987 = lean_alloc_ctor(0, 2, 0); +} else { + x_5987 = x_5986; +} +lean_ctor_set(x_5987, 0, x_5984); +lean_ctor_set(x_5987, 1, x_5985); +return x_5987; +} +else +{ +lean_object* x_5988; lean_object* x_5989; lean_object* x_5990; lean_object* x_5991; uint8_t x_5992; +x_5988 = lean_ctor_get(x_5983, 0); +lean_inc(x_5988); +x_5989 = lean_ctor_get(x_5983, 1); +lean_inc(x_5989); +if (lean_is_exclusive(x_5983)) { + lean_ctor_release(x_5983, 0); + lean_ctor_release(x_5983, 1); + x_5990 = x_5983; +} else { + lean_dec_ref(x_5983); + x_5990 = lean_box(0); +} +x_5991 = lean_ctor_get(x_5988, 1); +lean_inc(x_5991); +x_5992 = lean_nat_dec_eq(x_5977, x_5991); +lean_dec(x_5977); +if (x_5992 == 0) +{ +lean_object* x_5993; +lean_dec(x_5991); +if (lean_is_scalar(x_5990)) { + x_5993 = lean_alloc_ctor(1, 2, 0); +} else { + x_5993 = x_5990; +} +lean_ctor_set(x_5993, 0, x_5988); +lean_ctor_set(x_5993, 1, x_5989); +return x_5993; +} +else +{ +lean_object* x_5994; lean_object* x_5995; lean_object* x_5996; lean_object* x_5997; +lean_dec(x_5990); +lean_dec(x_5989); +x_5994 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_5995 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_5995, 0, x_5994); +x_5996 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_5997 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5996, x_5995, x_5988); +if (lean_obj_tag(x_5997) == 0) +{ +lean_object* x_5998; lean_object* x_5999; lean_object* x_6000; lean_object* x_6001; +lean_dec(x_5991); +x_5998 = lean_ctor_get(x_5997, 0); +lean_inc(x_5998); +x_5999 = lean_ctor_get(x_5997, 1); +lean_inc(x_5999); +if (lean_is_exclusive(x_5997)) { + lean_ctor_release(x_5997, 0); + lean_ctor_release(x_5997, 1); + x_6000 = x_5997; +} else { + lean_dec_ref(x_5997); + x_6000 = lean_box(0); +} +if (lean_is_scalar(x_6000)) { + x_6001 = lean_alloc_ctor(0, 2, 0); +} else { + x_6001 = x_6000; +} +lean_ctor_set(x_6001, 0, x_5998); +lean_ctor_set(x_6001, 1, x_5999); +return x_6001; +} +else +{ +lean_object* x_6002; lean_object* x_6003; lean_object* x_6004; lean_object* x_6005; uint8_t x_6006; +x_6002 = lean_ctor_get(x_5997, 0); +lean_inc(x_6002); +x_6003 = lean_ctor_get(x_5997, 1); +lean_inc(x_6003); +if (lean_is_exclusive(x_5997)) { + lean_ctor_release(x_5997, 0); + lean_ctor_release(x_5997, 1); + x_6004 = x_5997; +} else { + lean_dec_ref(x_5997); + x_6004 = lean_box(0); +} +x_6005 = lean_ctor_get(x_6002, 1); +lean_inc(x_6005); +x_6006 = lean_nat_dec_eq(x_5991, x_6005); +lean_dec(x_5991); +if (x_6006 == 0) +{ +lean_object* x_6007; +lean_dec(x_6005); +if (lean_is_scalar(x_6004)) { + x_6007 = lean_alloc_ctor(1, 2, 0); +} else { + x_6007 = x_6004; +} +lean_ctor_set(x_6007, 0, x_6002); +lean_ctor_set(x_6007, 1, x_6003); +return x_6007; +} +else +{ +lean_object* x_6008; lean_object* x_6009; lean_object* x_6010; +lean_dec(x_6004); +lean_dec(x_6003); +x_6008 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_6009 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6009, 0, x_6008); +x_6010 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_5996, x_6009, x_6002); +if (lean_obj_tag(x_6010) == 0) +{ +lean_object* x_6011; lean_object* x_6012; lean_object* x_6013; lean_object* x_6014; +lean_dec(x_6005); +x_6011 = lean_ctor_get(x_6010, 0); +lean_inc(x_6011); +x_6012 = lean_ctor_get(x_6010, 1); +lean_inc(x_6012); +if (lean_is_exclusive(x_6010)) { + lean_ctor_release(x_6010, 0); + lean_ctor_release(x_6010, 1); + x_6013 = x_6010; +} else { + lean_dec_ref(x_6010); + x_6013 = lean_box(0); +} +if (lean_is_scalar(x_6013)) { + x_6014 = lean_alloc_ctor(0, 2, 0); +} else { + x_6014 = x_6013; +} +lean_ctor_set(x_6014, 0, x_6011); +lean_ctor_set(x_6014, 1, x_6012); +return x_6014; +} +else +{ +lean_object* x_6015; lean_object* x_6016; lean_object* x_6017; lean_object* x_6018; uint8_t x_6019; +x_6015 = lean_ctor_get(x_6010, 0); +lean_inc(x_6015); +x_6016 = lean_ctor_get(x_6010, 1); +lean_inc(x_6016); +if (lean_is_exclusive(x_6010)) { + lean_ctor_release(x_6010, 0); + lean_ctor_release(x_6010, 1); + x_6017 = x_6010; +} else { + lean_dec_ref(x_6010); + x_6017 = lean_box(0); +} +x_6018 = lean_ctor_get(x_6015, 1); +lean_inc(x_6018); +x_6019 = lean_nat_dec_eq(x_6005, x_6018); +lean_dec(x_6005); +if (x_6019 == 0) +{ +lean_object* x_6020; +lean_dec(x_6018); +if (lean_is_scalar(x_6017)) { + x_6020 = lean_alloc_ctor(1, 2, 0); +} else { + x_6020 = x_6017; +} +lean_ctor_set(x_6020, 0, x_6015); +lean_ctor_set(x_6020, 1, x_6016); +return x_6020; +} +else +{ +lean_object* x_6021; lean_object* x_6022; lean_object* x_6023; lean_object* x_6024; +lean_dec(x_6017); +lean_dec(x_6016); +x_6021 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_6022 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6022, 0, x_6021); +x_6023 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_6024 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6023, x_6022, x_6015); +if (lean_obj_tag(x_6024) == 0) +{ +lean_object* x_6025; lean_object* x_6026; lean_object* x_6027; lean_object* x_6028; +lean_dec(x_6018); +x_6025 = lean_ctor_get(x_6024, 0); +lean_inc(x_6025); +x_6026 = lean_ctor_get(x_6024, 1); +lean_inc(x_6026); +if (lean_is_exclusive(x_6024)) { + lean_ctor_release(x_6024, 0); + lean_ctor_release(x_6024, 1); + x_6027 = x_6024; +} else { + lean_dec_ref(x_6024); + x_6027 = lean_box(0); +} +if (lean_is_scalar(x_6027)) { + x_6028 = lean_alloc_ctor(0, 2, 0); +} else { + x_6028 = x_6027; +} +lean_ctor_set(x_6028, 0, x_6025); +lean_ctor_set(x_6028, 1, x_6026); +return x_6028; +} +else +{ +lean_object* x_6029; lean_object* x_6030; lean_object* x_6031; lean_object* x_6032; uint8_t x_6033; +x_6029 = lean_ctor_get(x_6024, 0); +lean_inc(x_6029); +x_6030 = lean_ctor_get(x_6024, 1); +lean_inc(x_6030); +if (lean_is_exclusive(x_6024)) { + lean_ctor_release(x_6024, 0); + lean_ctor_release(x_6024, 1); + x_6031 = x_6024; +} else { + lean_dec_ref(x_6024); + x_6031 = lean_box(0); +} +x_6032 = lean_ctor_get(x_6029, 1); +lean_inc(x_6032); +x_6033 = lean_nat_dec_eq(x_6018, x_6032); +lean_dec(x_6018); +if (x_6033 == 0) +{ +lean_object* x_6034; +lean_dec(x_6032); +if (lean_is_scalar(x_6031)) { + x_6034 = lean_alloc_ctor(1, 2, 0); +} else { + x_6034 = x_6031; +} +lean_ctor_set(x_6034, 0, x_6029); +lean_ctor_set(x_6034, 1, x_6030); +return x_6034; +} +else +{ +lean_object* x_6035; lean_object* x_6036; lean_object* x_6037; lean_object* x_6038; +lean_dec(x_6031); +lean_dec(x_6030); +x_6035 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_6036 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6036, 0, x_6035); +x_6037 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_6038 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6037, x_6036, x_6029); +if (lean_obj_tag(x_6038) == 0) +{ +lean_object* x_6039; lean_object* x_6040; lean_object* x_6041; lean_object* x_6042; +lean_dec(x_6032); +x_6039 = lean_ctor_get(x_6038, 0); +lean_inc(x_6039); +x_6040 = lean_ctor_get(x_6038, 1); +lean_inc(x_6040); +if (lean_is_exclusive(x_6038)) { + lean_ctor_release(x_6038, 0); + lean_ctor_release(x_6038, 1); + x_6041 = x_6038; +} else { + lean_dec_ref(x_6038); + x_6041 = lean_box(0); +} +if (lean_is_scalar(x_6041)) { + x_6042 = lean_alloc_ctor(0, 2, 0); +} else { + x_6042 = x_6041; +} +lean_ctor_set(x_6042, 0, x_6039); +lean_ctor_set(x_6042, 1, x_6040); +return x_6042; +} +else +{ +lean_object* x_6043; lean_object* x_6044; lean_object* x_6045; lean_object* x_6046; uint8_t x_6047; +x_6043 = lean_ctor_get(x_6038, 0); +lean_inc(x_6043); +x_6044 = lean_ctor_get(x_6038, 1); +lean_inc(x_6044); +if (lean_is_exclusive(x_6038)) { + lean_ctor_release(x_6038, 0); + lean_ctor_release(x_6038, 1); + x_6045 = x_6038; +} else { + lean_dec_ref(x_6038); + x_6045 = lean_box(0); +} +x_6046 = lean_ctor_get(x_6043, 1); +lean_inc(x_6046); +x_6047 = lean_nat_dec_eq(x_6032, x_6046); +lean_dec(x_6032); +if (x_6047 == 0) +{ +lean_object* x_6048; +lean_dec(x_6046); +if (lean_is_scalar(x_6045)) { + x_6048 = lean_alloc_ctor(1, 2, 0); +} else { + x_6048 = x_6045; +} +lean_ctor_set(x_6048, 0, x_6043); +lean_ctor_set(x_6048, 1, x_6044); +return x_6048; +} +else +{ +lean_object* x_6049; lean_object* x_6050; lean_object* x_6051; lean_object* x_6052; +lean_dec(x_6045); +lean_dec(x_6044); +x_6049 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_6050 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6050, 0, x_6049); +x_6051 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_6052 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6051, x_6050, x_6043); +if (lean_obj_tag(x_6052) == 0) +{ +lean_object* x_6053; lean_object* x_6054; lean_object* x_6055; lean_object* x_6056; +lean_dec(x_6046); +x_6053 = lean_ctor_get(x_6052, 0); +lean_inc(x_6053); +x_6054 = lean_ctor_get(x_6052, 1); +lean_inc(x_6054); +if (lean_is_exclusive(x_6052)) { + lean_ctor_release(x_6052, 0); + lean_ctor_release(x_6052, 1); + x_6055 = x_6052; +} else { + lean_dec_ref(x_6052); + x_6055 = lean_box(0); +} +if (lean_is_scalar(x_6055)) { + x_6056 = lean_alloc_ctor(0, 2, 0); +} else { + x_6056 = x_6055; +} +lean_ctor_set(x_6056, 0, x_6053); +lean_ctor_set(x_6056, 1, x_6054); +return x_6056; +} +else +{ +lean_object* x_6057; lean_object* x_6058; lean_object* x_6059; lean_object* x_6060; uint8_t x_6061; +x_6057 = lean_ctor_get(x_6052, 0); +lean_inc(x_6057); +x_6058 = lean_ctor_get(x_6052, 1); +lean_inc(x_6058); +if (lean_is_exclusive(x_6052)) { + lean_ctor_release(x_6052, 0); + lean_ctor_release(x_6052, 1); + x_6059 = x_6052; +} else { + lean_dec_ref(x_6052); + x_6059 = lean_box(0); +} +x_6060 = lean_ctor_get(x_6057, 1); +lean_inc(x_6060); +x_6061 = lean_nat_dec_eq(x_6046, x_6060); +lean_dec(x_6046); +if (x_6061 == 0) +{ +lean_object* x_6062; +lean_dec(x_6060); +if (lean_is_scalar(x_6059)) { + x_6062 = lean_alloc_ctor(1, 2, 0); +} else { + x_6062 = x_6059; +} +lean_ctor_set(x_6062, 0, x_6057); +lean_ctor_set(x_6062, 1, x_6058); +return x_6062; +} +else +{ +lean_object* x_6063; lean_object* x_6064; lean_object* x_6065; lean_object* x_6066; +lean_dec(x_6059); +lean_dec(x_6058); +x_6063 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_6064 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6064, 0, x_6063); +x_6065 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_6066 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6065, x_6064, x_6057); +if (lean_obj_tag(x_6066) == 0) +{ +lean_object* x_6067; lean_object* x_6068; lean_object* x_6069; lean_object* x_6070; +lean_dec(x_6060); +x_6067 = lean_ctor_get(x_6066, 0); +lean_inc(x_6067); +x_6068 = lean_ctor_get(x_6066, 1); +lean_inc(x_6068); +if (lean_is_exclusive(x_6066)) { + lean_ctor_release(x_6066, 0); + lean_ctor_release(x_6066, 1); + x_6069 = x_6066; +} else { + lean_dec_ref(x_6066); + x_6069 = lean_box(0); +} +if (lean_is_scalar(x_6069)) { + x_6070 = lean_alloc_ctor(0, 2, 0); +} else { + x_6070 = x_6069; +} +lean_ctor_set(x_6070, 0, x_6067); +lean_ctor_set(x_6070, 1, x_6068); +return x_6070; +} +else +{ +lean_object* x_6071; lean_object* x_6072; lean_object* x_6073; lean_object* x_6074; uint8_t x_6075; +x_6071 = lean_ctor_get(x_6066, 0); +lean_inc(x_6071); +x_6072 = lean_ctor_get(x_6066, 1); +lean_inc(x_6072); +if (lean_is_exclusive(x_6066)) { + lean_ctor_release(x_6066, 0); + lean_ctor_release(x_6066, 1); + x_6073 = x_6066; +} else { + lean_dec_ref(x_6066); + x_6073 = lean_box(0); +} +x_6074 = lean_ctor_get(x_6071, 1); +lean_inc(x_6074); +x_6075 = lean_nat_dec_eq(x_6060, x_6074); +lean_dec(x_6060); +if (x_6075 == 0) +{ +lean_object* x_6076; +lean_dec(x_6074); +if (lean_is_scalar(x_6073)) { + x_6076 = lean_alloc_ctor(1, 2, 0); +} else { + x_6076 = x_6073; +} +lean_ctor_set(x_6076, 0, x_6071); +lean_ctor_set(x_6076, 1, x_6072); +return x_6076; +} +else +{ +lean_object* x_6077; lean_object* x_6078; lean_object* x_6079; +lean_dec(x_6073); +lean_dec(x_6072); +x_6077 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_6078 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6078, 0, x_6077); +x_6079 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6065, x_6078, x_6071); +if (lean_obj_tag(x_6079) == 0) +{ +lean_object* x_6080; lean_object* x_6081; lean_object* x_6082; lean_object* x_6083; +lean_dec(x_6074); +x_6080 = lean_ctor_get(x_6079, 0); +lean_inc(x_6080); +x_6081 = lean_ctor_get(x_6079, 1); +lean_inc(x_6081); +if (lean_is_exclusive(x_6079)) { + lean_ctor_release(x_6079, 0); + lean_ctor_release(x_6079, 1); + x_6082 = x_6079; +} else { + lean_dec_ref(x_6079); + x_6082 = lean_box(0); +} +if (lean_is_scalar(x_6082)) { + x_6083 = lean_alloc_ctor(0, 2, 0); +} else { + x_6083 = x_6082; +} +lean_ctor_set(x_6083, 0, x_6080); +lean_ctor_set(x_6083, 1, x_6081); +return x_6083; +} +else +{ +lean_object* x_6084; lean_object* x_6085; lean_object* x_6086; lean_object* x_6087; uint8_t x_6088; +x_6084 = lean_ctor_get(x_6079, 0); +lean_inc(x_6084); +x_6085 = lean_ctor_get(x_6079, 1); +lean_inc(x_6085); +if (lean_is_exclusive(x_6079)) { + lean_ctor_release(x_6079, 0); + lean_ctor_release(x_6079, 1); + x_6086 = x_6079; +} else { + lean_dec_ref(x_6079); + x_6086 = lean_box(0); +} +x_6087 = lean_ctor_get(x_6084, 1); +lean_inc(x_6087); +x_6088 = lean_nat_dec_eq(x_6074, x_6087); +lean_dec(x_6074); +if (x_6088 == 0) +{ +lean_object* x_6089; +lean_dec(x_6087); +if (lean_is_scalar(x_6086)) { + x_6089 = lean_alloc_ctor(1, 2, 0); +} else { + x_6089 = x_6086; +} +lean_ctor_set(x_6089, 0, x_6084); +lean_ctor_set(x_6089, 1, x_6085); +return x_6089; +} +else +{ +lean_object* x_6090; lean_object* x_6091; lean_object* x_6092; lean_object* x_6093; +lean_dec(x_6086); +lean_dec(x_6085); +x_6090 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_6091 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6091, 0, x_6090); +x_6092 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_6093 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6092, x_6091, x_6084); +if (lean_obj_tag(x_6093) == 0) +{ +lean_object* x_6094; lean_object* x_6095; lean_object* x_6096; lean_object* x_6097; +lean_dec(x_6087); +x_6094 = lean_ctor_get(x_6093, 0); +lean_inc(x_6094); +x_6095 = lean_ctor_get(x_6093, 1); +lean_inc(x_6095); +if (lean_is_exclusive(x_6093)) { + lean_ctor_release(x_6093, 0); + lean_ctor_release(x_6093, 1); + x_6096 = x_6093; +} else { + lean_dec_ref(x_6093); + x_6096 = lean_box(0); +} +if (lean_is_scalar(x_6096)) { + x_6097 = lean_alloc_ctor(0, 2, 0); +} else { + x_6097 = x_6096; +} +lean_ctor_set(x_6097, 0, x_6094); +lean_ctor_set(x_6097, 1, x_6095); +return x_6097; +} +else +{ +lean_object* x_6098; lean_object* x_6099; lean_object* x_6100; lean_object* x_6101; uint8_t x_6102; +x_6098 = lean_ctor_get(x_6093, 0); +lean_inc(x_6098); +x_6099 = lean_ctor_get(x_6093, 1); +lean_inc(x_6099); +if (lean_is_exclusive(x_6093)) { + lean_ctor_release(x_6093, 0); + lean_ctor_release(x_6093, 1); + x_6100 = x_6093; +} else { + lean_dec_ref(x_6093); + x_6100 = lean_box(0); +} +x_6101 = lean_ctor_get(x_6098, 1); +lean_inc(x_6101); +x_6102 = lean_nat_dec_eq(x_6087, x_6101); +lean_dec(x_6087); +if (x_6102 == 0) +{ +lean_object* x_6103; +lean_dec(x_6101); +if (lean_is_scalar(x_6100)) { + x_6103 = lean_alloc_ctor(1, 2, 0); +} else { + x_6103 = x_6100; +} +lean_ctor_set(x_6103, 0, x_6098); +lean_ctor_set(x_6103, 1, x_6099); +return x_6103; +} +else +{ +lean_object* x_6104; lean_object* x_6105; lean_object* x_6106; lean_object* x_6107; +lean_dec(x_6100); +lean_dec(x_6099); +x_6104 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_6105 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6105, 0, x_6104); +x_6106 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_6107 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6106, x_6105, x_6098); +if (lean_obj_tag(x_6107) == 0) +{ +lean_object* x_6108; lean_object* x_6109; lean_object* x_6110; lean_object* x_6111; +lean_dec(x_6101); +x_6108 = lean_ctor_get(x_6107, 0); +lean_inc(x_6108); +x_6109 = lean_ctor_get(x_6107, 1); +lean_inc(x_6109); +if (lean_is_exclusive(x_6107)) { + lean_ctor_release(x_6107, 0); + lean_ctor_release(x_6107, 1); + x_6110 = x_6107; +} else { + lean_dec_ref(x_6107); + x_6110 = lean_box(0); +} +if (lean_is_scalar(x_6110)) { + x_6111 = lean_alloc_ctor(0, 2, 0); +} else { + x_6111 = x_6110; +} +lean_ctor_set(x_6111, 0, x_6108); +lean_ctor_set(x_6111, 1, x_6109); +return x_6111; +} +else +{ +lean_object* x_6112; lean_object* x_6113; lean_object* x_6114; lean_object* x_6115; uint8_t x_6116; +x_6112 = lean_ctor_get(x_6107, 0); +lean_inc(x_6112); +x_6113 = lean_ctor_get(x_6107, 1); +lean_inc(x_6113); +if (lean_is_exclusive(x_6107)) { + lean_ctor_release(x_6107, 0); + lean_ctor_release(x_6107, 1); + x_6114 = x_6107; +} else { + lean_dec_ref(x_6107); + x_6114 = lean_box(0); +} +x_6115 = lean_ctor_get(x_6112, 1); +lean_inc(x_6115); +x_6116 = lean_nat_dec_eq(x_6101, x_6115); +lean_dec(x_6101); +if (x_6116 == 0) +{ +lean_object* x_6117; +lean_dec(x_6115); +if (lean_is_scalar(x_6114)) { + x_6117 = lean_alloc_ctor(1, 2, 0); +} else { + x_6117 = x_6114; +} +lean_ctor_set(x_6117, 0, x_6112); +lean_ctor_set(x_6117, 1, x_6113); +return x_6117; +} +else +{ +lean_object* x_6118; lean_object* x_6119; lean_object* x_6120; lean_object* x_6121; +lean_dec(x_6114); +lean_dec(x_6113); +x_6118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_6119 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6119, 0, x_6118); +x_6120 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_6121 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6120, x_6119, x_6112); +if (lean_obj_tag(x_6121) == 0) +{ +lean_object* x_6122; lean_object* x_6123; lean_object* x_6124; lean_object* x_6125; +lean_dec(x_6115); +x_6122 = lean_ctor_get(x_6121, 0); +lean_inc(x_6122); +x_6123 = lean_ctor_get(x_6121, 1); +lean_inc(x_6123); +if (lean_is_exclusive(x_6121)) { + lean_ctor_release(x_6121, 0); + lean_ctor_release(x_6121, 1); + x_6124 = x_6121; +} else { + lean_dec_ref(x_6121); + x_6124 = lean_box(0); +} +if (lean_is_scalar(x_6124)) { + x_6125 = lean_alloc_ctor(0, 2, 0); +} else { + x_6125 = x_6124; +} +lean_ctor_set(x_6125, 0, x_6122); +lean_ctor_set(x_6125, 1, x_6123); +return x_6125; +} +else +{ +lean_object* x_6126; lean_object* x_6127; lean_object* x_6128; lean_object* x_6129; uint8_t x_6130; +x_6126 = lean_ctor_get(x_6121, 0); +lean_inc(x_6126); +x_6127 = lean_ctor_get(x_6121, 1); +lean_inc(x_6127); +if (lean_is_exclusive(x_6121)) { + lean_ctor_release(x_6121, 0); + lean_ctor_release(x_6121, 1); + x_6128 = x_6121; +} else { + lean_dec_ref(x_6121); + x_6128 = lean_box(0); +} +x_6129 = lean_ctor_get(x_6126, 1); +lean_inc(x_6129); +x_6130 = lean_nat_dec_eq(x_6115, x_6129); +lean_dec(x_6115); +if (x_6130 == 0) +{ +lean_object* x_6131; +lean_dec(x_6129); +if (lean_is_scalar(x_6128)) { + x_6131 = lean_alloc_ctor(1, 2, 0); +} else { + x_6131 = x_6128; +} +lean_ctor_set(x_6131, 0, x_6126); +lean_ctor_set(x_6131, 1, x_6127); +return x_6131; +} +else +{ +lean_object* x_6132; lean_object* x_6133; lean_object* x_6134; lean_object* x_6135; +lean_dec(x_6128); +lean_dec(x_6127); +x_6132 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_6133 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6133, 0, x_6132); +x_6134 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_6135 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6134, x_6133, x_6126); +if (lean_obj_tag(x_6135) == 0) +{ +lean_object* x_6136; lean_object* x_6137; lean_object* x_6138; lean_object* x_6139; +lean_dec(x_6129); +x_6136 = lean_ctor_get(x_6135, 0); +lean_inc(x_6136); +x_6137 = lean_ctor_get(x_6135, 1); +lean_inc(x_6137); +if (lean_is_exclusive(x_6135)) { + lean_ctor_release(x_6135, 0); + lean_ctor_release(x_6135, 1); + x_6138 = x_6135; +} else { + lean_dec_ref(x_6135); + x_6138 = lean_box(0); +} +if (lean_is_scalar(x_6138)) { + x_6139 = lean_alloc_ctor(0, 2, 0); +} else { + x_6139 = x_6138; +} +lean_ctor_set(x_6139, 0, x_6136); +lean_ctor_set(x_6139, 1, x_6137); +return x_6139; +} +else +{ +lean_object* x_6140; lean_object* x_6141; lean_object* x_6142; lean_object* x_6143; uint8_t x_6144; +x_6140 = lean_ctor_get(x_6135, 0); +lean_inc(x_6140); +x_6141 = lean_ctor_get(x_6135, 1); +lean_inc(x_6141); +if (lean_is_exclusive(x_6135)) { + lean_ctor_release(x_6135, 0); + lean_ctor_release(x_6135, 1); + x_6142 = x_6135; +} else { + lean_dec_ref(x_6135); + x_6142 = lean_box(0); +} +x_6143 = lean_ctor_get(x_6140, 1); +lean_inc(x_6143); +x_6144 = lean_nat_dec_eq(x_6129, x_6143); +lean_dec(x_6129); +if (x_6144 == 0) +{ +lean_object* x_6145; +lean_dec(x_6143); +if (lean_is_scalar(x_6142)) { + x_6145 = lean_alloc_ctor(1, 2, 0); +} else { + x_6145 = x_6142; +} +lean_ctor_set(x_6145, 0, x_6140); +lean_ctor_set(x_6145, 1, x_6141); +return x_6145; +} +else +{ +lean_object* x_6146; lean_object* x_6147; lean_object* x_6148; lean_object* x_6149; +lean_dec(x_6142); +lean_dec(x_6141); +x_6146 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_6147 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6147, 0, x_6146); +x_6148 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_6149 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6148, x_6147, x_6140); +if (lean_obj_tag(x_6149) == 0) +{ +lean_object* x_6150; lean_object* x_6151; lean_object* x_6152; lean_object* x_6153; +lean_dec(x_6143); +x_6150 = lean_ctor_get(x_6149, 0); +lean_inc(x_6150); +x_6151 = lean_ctor_get(x_6149, 1); +lean_inc(x_6151); +if (lean_is_exclusive(x_6149)) { + lean_ctor_release(x_6149, 0); + lean_ctor_release(x_6149, 1); + x_6152 = x_6149; +} else { + lean_dec_ref(x_6149); + x_6152 = lean_box(0); +} +if (lean_is_scalar(x_6152)) { + x_6153 = lean_alloc_ctor(0, 2, 0); +} else { + x_6153 = x_6152; +} +lean_ctor_set(x_6153, 0, x_6150); +lean_ctor_set(x_6153, 1, x_6151); +return x_6153; +} +else +{ +lean_object* x_6154; lean_object* x_6155; lean_object* x_6156; lean_object* x_6157; uint8_t x_6158; +x_6154 = lean_ctor_get(x_6149, 0); +lean_inc(x_6154); +x_6155 = lean_ctor_get(x_6149, 1); +lean_inc(x_6155); +if (lean_is_exclusive(x_6149)) { + lean_ctor_release(x_6149, 0); + lean_ctor_release(x_6149, 1); + x_6156 = x_6149; +} else { + lean_dec_ref(x_6149); + x_6156 = lean_box(0); +} +x_6157 = lean_ctor_get(x_6154, 1); +lean_inc(x_6157); +x_6158 = lean_nat_dec_eq(x_6143, x_6157); +lean_dec(x_6143); +if (x_6158 == 0) +{ +lean_object* x_6159; +lean_dec(x_6157); +if (lean_is_scalar(x_6156)) { + x_6159 = lean_alloc_ctor(1, 2, 0); +} else { + x_6159 = x_6156; +} +lean_ctor_set(x_6159, 0, x_6154); +lean_ctor_set(x_6159, 1, x_6155); +return x_6159; +} +else +{ +lean_object* x_6160; lean_object* x_6161; lean_object* x_6162; lean_object* x_6163; +lean_dec(x_6156); +lean_dec(x_6155); +x_6160 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_6161 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6161, 0, x_6160); +x_6162 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_6163 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6162, x_6161, x_6154); +if (lean_obj_tag(x_6163) == 0) +{ +lean_object* x_6164; lean_object* x_6165; lean_object* x_6166; lean_object* x_6167; +lean_dec(x_6157); +x_6164 = lean_ctor_get(x_6163, 0); +lean_inc(x_6164); +x_6165 = lean_ctor_get(x_6163, 1); +lean_inc(x_6165); +if (lean_is_exclusive(x_6163)) { + lean_ctor_release(x_6163, 0); + lean_ctor_release(x_6163, 1); + x_6166 = x_6163; +} else { + lean_dec_ref(x_6163); + x_6166 = lean_box(0); +} +if (lean_is_scalar(x_6166)) { + x_6167 = lean_alloc_ctor(0, 2, 0); +} else { + x_6167 = x_6166; +} +lean_ctor_set(x_6167, 0, x_6164); +lean_ctor_set(x_6167, 1, x_6165); +return x_6167; +} +else +{ +lean_object* x_6168; lean_object* x_6169; lean_object* x_6170; lean_object* x_6171; uint8_t x_6172; +x_6168 = lean_ctor_get(x_6163, 0); +lean_inc(x_6168); +x_6169 = lean_ctor_get(x_6163, 1); +lean_inc(x_6169); +if (lean_is_exclusive(x_6163)) { + lean_ctor_release(x_6163, 0); + lean_ctor_release(x_6163, 1); + x_6170 = x_6163; +} else { + lean_dec_ref(x_6163); + x_6170 = lean_box(0); +} +x_6171 = lean_ctor_get(x_6168, 1); +lean_inc(x_6171); +x_6172 = lean_nat_dec_eq(x_6157, x_6171); +lean_dec(x_6157); +if (x_6172 == 0) +{ +lean_object* x_6173; +lean_dec(x_6171); +if (lean_is_scalar(x_6170)) { + x_6173 = lean_alloc_ctor(1, 2, 0); +} else { + x_6173 = x_6170; +} +lean_ctor_set(x_6173, 0, x_6168); +lean_ctor_set(x_6173, 1, x_6169); +return x_6173; +} +else +{ +lean_object* x_6174; lean_object* x_6175; lean_object* x_6176; lean_object* x_6177; +lean_dec(x_6170); +lean_dec(x_6169); +x_6174 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_6175 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6175, 0, x_6174); +x_6176 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_6177 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6176, x_6175, x_6168); +if (lean_obj_tag(x_6177) == 0) +{ +lean_object* x_6178; lean_object* x_6179; lean_object* x_6180; lean_object* x_6181; +lean_dec(x_6171); +x_6178 = lean_ctor_get(x_6177, 0); +lean_inc(x_6178); +x_6179 = lean_ctor_get(x_6177, 1); +lean_inc(x_6179); +if (lean_is_exclusive(x_6177)) { + lean_ctor_release(x_6177, 0); + lean_ctor_release(x_6177, 1); + x_6180 = x_6177; +} else { + lean_dec_ref(x_6177); + x_6180 = lean_box(0); +} +if (lean_is_scalar(x_6180)) { + x_6181 = lean_alloc_ctor(0, 2, 0); +} else { + x_6181 = x_6180; +} +lean_ctor_set(x_6181, 0, x_6178); +lean_ctor_set(x_6181, 1, x_6179); +return x_6181; +} +else +{ +lean_object* x_6182; lean_object* x_6183; lean_object* x_6184; lean_object* x_6185; uint8_t x_6186; +x_6182 = lean_ctor_get(x_6177, 0); +lean_inc(x_6182); +x_6183 = lean_ctor_get(x_6177, 1); +lean_inc(x_6183); +if (lean_is_exclusive(x_6177)) { + lean_ctor_release(x_6177, 0); + lean_ctor_release(x_6177, 1); + x_6184 = x_6177; +} else { + lean_dec_ref(x_6177); + x_6184 = lean_box(0); +} +x_6185 = lean_ctor_get(x_6182, 1); +lean_inc(x_6185); +x_6186 = lean_nat_dec_eq(x_6171, x_6185); +lean_dec(x_6171); +if (x_6186 == 0) +{ +lean_object* x_6187; +lean_dec(x_6185); +if (lean_is_scalar(x_6184)) { + x_6187 = lean_alloc_ctor(1, 2, 0); +} else { + x_6187 = x_6184; +} +lean_ctor_set(x_6187, 0, x_6182); +lean_ctor_set(x_6187, 1, x_6183); +return x_6187; +} +else +{ +lean_object* x_6188; lean_object* x_6189; lean_object* x_6190; lean_object* x_6191; +lean_dec(x_6184); +lean_dec(x_6183); +x_6188 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_6189 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6189, 0, x_6188); +x_6190 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_6191 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6190, x_6189, x_6182); +if (lean_obj_tag(x_6191) == 0) +{ +lean_object* x_6192; lean_object* x_6193; lean_object* x_6194; lean_object* x_6195; +lean_dec(x_6185); +x_6192 = lean_ctor_get(x_6191, 0); +lean_inc(x_6192); +x_6193 = lean_ctor_get(x_6191, 1); +lean_inc(x_6193); +if (lean_is_exclusive(x_6191)) { + lean_ctor_release(x_6191, 0); + lean_ctor_release(x_6191, 1); + x_6194 = x_6191; +} else { + lean_dec_ref(x_6191); + x_6194 = lean_box(0); +} +if (lean_is_scalar(x_6194)) { + x_6195 = lean_alloc_ctor(0, 2, 0); +} else { + x_6195 = x_6194; +} +lean_ctor_set(x_6195, 0, x_6192); +lean_ctor_set(x_6195, 1, x_6193); +return x_6195; +} +else +{ +lean_object* x_6196; lean_object* x_6197; lean_object* x_6198; lean_object* x_6199; uint8_t x_6200; +x_6196 = lean_ctor_get(x_6191, 0); +lean_inc(x_6196); +x_6197 = lean_ctor_get(x_6191, 1); +lean_inc(x_6197); +if (lean_is_exclusive(x_6191)) { + lean_ctor_release(x_6191, 0); + lean_ctor_release(x_6191, 1); + x_6198 = x_6191; +} else { + lean_dec_ref(x_6191); + x_6198 = lean_box(0); +} +x_6199 = lean_ctor_get(x_6196, 1); +lean_inc(x_6199); +x_6200 = lean_nat_dec_eq(x_6185, x_6199); +lean_dec(x_6185); +if (x_6200 == 0) +{ +lean_object* x_6201; +lean_dec(x_6199); +if (lean_is_scalar(x_6198)) { + x_6201 = lean_alloc_ctor(1, 2, 0); +} else { + x_6201 = x_6198; +} +lean_ctor_set(x_6201, 0, x_6196); +lean_ctor_set(x_6201, 1, x_6197); +return x_6201; +} +else +{ +lean_object* x_6202; lean_object* x_6203; lean_object* x_6204; lean_object* x_6205; +lean_dec(x_6198); +lean_dec(x_6197); +x_6202 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_6203 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6203, 0, x_6202); +x_6204 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_6205 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6204, x_6203, x_6196); +if (lean_obj_tag(x_6205) == 0) +{ +lean_object* x_6206; lean_object* x_6207; lean_object* x_6208; lean_object* x_6209; +lean_dec(x_6199); +x_6206 = lean_ctor_get(x_6205, 0); +lean_inc(x_6206); +x_6207 = lean_ctor_get(x_6205, 1); +lean_inc(x_6207); +if (lean_is_exclusive(x_6205)) { + lean_ctor_release(x_6205, 0); + lean_ctor_release(x_6205, 1); + x_6208 = x_6205; +} else { + lean_dec_ref(x_6205); + x_6208 = lean_box(0); +} +if (lean_is_scalar(x_6208)) { + x_6209 = lean_alloc_ctor(0, 2, 0); +} else { + x_6209 = x_6208; +} +lean_ctor_set(x_6209, 0, x_6206); +lean_ctor_set(x_6209, 1, x_6207); +return x_6209; +} +else +{ +lean_object* x_6210; lean_object* x_6211; lean_object* x_6212; lean_object* x_6213; uint8_t x_6214; +x_6210 = lean_ctor_get(x_6205, 0); +lean_inc(x_6210); +x_6211 = lean_ctor_get(x_6205, 1); +lean_inc(x_6211); +if (lean_is_exclusive(x_6205)) { + lean_ctor_release(x_6205, 0); + lean_ctor_release(x_6205, 1); + x_6212 = x_6205; +} else { + lean_dec_ref(x_6205); + x_6212 = lean_box(0); +} +x_6213 = lean_ctor_get(x_6210, 1); +lean_inc(x_6213); +x_6214 = lean_nat_dec_eq(x_6199, x_6213); +lean_dec(x_6199); +if (x_6214 == 0) +{ +lean_object* x_6215; +lean_dec(x_6213); +if (lean_is_scalar(x_6212)) { + x_6215 = lean_alloc_ctor(1, 2, 0); +} else { + x_6215 = x_6212; +} +lean_ctor_set(x_6215, 0, x_6210); +lean_ctor_set(x_6215, 1, x_6211); +return x_6215; +} +else +{ +lean_object* x_6216; lean_object* x_6217; lean_object* x_6218; lean_object* x_6219; +lean_dec(x_6212); +lean_dec(x_6211); +x_6216 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_6217 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6217, 0, x_6216); +x_6218 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_6219 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6218, x_6217, x_6210); +if (lean_obj_tag(x_6219) == 0) +{ +lean_object* x_6220; lean_object* x_6221; lean_object* x_6222; lean_object* x_6223; +lean_dec(x_6213); +x_6220 = lean_ctor_get(x_6219, 0); +lean_inc(x_6220); +x_6221 = lean_ctor_get(x_6219, 1); +lean_inc(x_6221); +if (lean_is_exclusive(x_6219)) { + lean_ctor_release(x_6219, 0); + lean_ctor_release(x_6219, 1); + x_6222 = x_6219; +} else { + lean_dec_ref(x_6219); + x_6222 = lean_box(0); +} +if (lean_is_scalar(x_6222)) { + x_6223 = lean_alloc_ctor(0, 2, 0); +} else { + x_6223 = x_6222; +} +lean_ctor_set(x_6223, 0, x_6220); +lean_ctor_set(x_6223, 1, x_6221); +return x_6223; +} +else +{ +lean_object* x_6224; lean_object* x_6225; lean_object* x_6226; lean_object* x_6227; uint8_t x_6228; +x_6224 = lean_ctor_get(x_6219, 0); +lean_inc(x_6224); +x_6225 = lean_ctor_get(x_6219, 1); +lean_inc(x_6225); +if (lean_is_exclusive(x_6219)) { + lean_ctor_release(x_6219, 0); + lean_ctor_release(x_6219, 1); + x_6226 = x_6219; +} else { + lean_dec_ref(x_6219); + x_6226 = lean_box(0); +} +x_6227 = lean_ctor_get(x_6224, 1); +lean_inc(x_6227); +x_6228 = lean_nat_dec_eq(x_6213, x_6227); +lean_dec(x_6213); +if (x_6228 == 0) +{ +lean_object* x_6229; +lean_dec(x_6227); +if (lean_is_scalar(x_6226)) { + x_6229 = lean_alloc_ctor(1, 2, 0); +} else { + x_6229 = x_6226; +} +lean_ctor_set(x_6229, 0, x_6224); +lean_ctor_set(x_6229, 1, x_6225); +return x_6229; +} +else +{ +lean_object* x_6230; lean_object* x_6231; lean_object* x_6232; lean_object* x_6233; +lean_dec(x_6226); +lean_dec(x_6225); +x_6230 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_6231 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6231, 0, x_6230); +x_6232 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_6233 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6232, x_6231, x_6224); +if (lean_obj_tag(x_6233) == 0) +{ +lean_object* x_6234; lean_object* x_6235; lean_object* x_6236; lean_object* x_6237; +lean_dec(x_6227); +x_6234 = lean_ctor_get(x_6233, 0); +lean_inc(x_6234); +x_6235 = lean_ctor_get(x_6233, 1); +lean_inc(x_6235); +if (lean_is_exclusive(x_6233)) { + lean_ctor_release(x_6233, 0); + lean_ctor_release(x_6233, 1); + x_6236 = x_6233; +} else { + lean_dec_ref(x_6233); + x_6236 = lean_box(0); +} +if (lean_is_scalar(x_6236)) { + x_6237 = lean_alloc_ctor(0, 2, 0); +} else { + x_6237 = x_6236; +} +lean_ctor_set(x_6237, 0, x_6234); +lean_ctor_set(x_6237, 1, x_6235); +return x_6237; +} +else +{ +lean_object* x_6238; lean_object* x_6239; lean_object* x_6240; lean_object* x_6241; uint8_t x_6242; +x_6238 = lean_ctor_get(x_6233, 0); +lean_inc(x_6238); +x_6239 = lean_ctor_get(x_6233, 1); +lean_inc(x_6239); +if (lean_is_exclusive(x_6233)) { + lean_ctor_release(x_6233, 0); + lean_ctor_release(x_6233, 1); + x_6240 = x_6233; +} else { + lean_dec_ref(x_6233); + x_6240 = lean_box(0); +} +x_6241 = lean_ctor_get(x_6238, 1); +lean_inc(x_6241); +x_6242 = lean_nat_dec_eq(x_6227, x_6241); +lean_dec(x_6227); +if (x_6242 == 0) +{ +lean_object* x_6243; +lean_dec(x_6241); +if (lean_is_scalar(x_6240)) { + x_6243 = lean_alloc_ctor(1, 2, 0); +} else { + x_6243 = x_6240; +} +lean_ctor_set(x_6243, 0, x_6238); +lean_ctor_set(x_6243, 1, x_6239); +return x_6243; +} +else +{ +lean_object* x_6244; lean_object* x_6245; lean_object* x_6246; lean_object* x_6247; +lean_dec(x_6240); +lean_dec(x_6239); +x_6244 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_6245 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6245, 0, x_6244); +x_6246 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_6247 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6246, x_6245, x_6238); +if (lean_obj_tag(x_6247) == 0) +{ +lean_object* x_6248; lean_object* x_6249; lean_object* x_6250; lean_object* x_6251; +lean_dec(x_6241); +x_6248 = lean_ctor_get(x_6247, 0); +lean_inc(x_6248); +x_6249 = lean_ctor_get(x_6247, 1); +lean_inc(x_6249); +if (lean_is_exclusive(x_6247)) { + lean_ctor_release(x_6247, 0); + lean_ctor_release(x_6247, 1); + x_6250 = x_6247; +} else { + lean_dec_ref(x_6247); + x_6250 = lean_box(0); +} +if (lean_is_scalar(x_6250)) { + x_6251 = lean_alloc_ctor(0, 2, 0); +} else { + x_6251 = x_6250; +} +lean_ctor_set(x_6251, 0, x_6248); +lean_ctor_set(x_6251, 1, x_6249); +return x_6251; +} +else +{ +lean_object* x_6252; lean_object* x_6253; lean_object* x_6254; lean_object* x_6255; uint8_t x_6256; +x_6252 = lean_ctor_get(x_6247, 0); +lean_inc(x_6252); +x_6253 = lean_ctor_get(x_6247, 1); +lean_inc(x_6253); +if (lean_is_exclusive(x_6247)) { + lean_ctor_release(x_6247, 0); + lean_ctor_release(x_6247, 1); + x_6254 = x_6247; +} else { + lean_dec_ref(x_6247); + x_6254 = lean_box(0); +} +x_6255 = lean_ctor_get(x_6252, 1); +lean_inc(x_6255); +x_6256 = lean_nat_dec_eq(x_6241, x_6255); +lean_dec(x_6241); +if (x_6256 == 0) +{ +lean_object* x_6257; +lean_dec(x_6255); +if (lean_is_scalar(x_6254)) { + x_6257 = lean_alloc_ctor(1, 2, 0); +} else { + x_6257 = x_6254; +} +lean_ctor_set(x_6257, 0, x_6252); +lean_ctor_set(x_6257, 1, x_6253); +return x_6257; +} +else +{ +lean_object* x_6258; lean_object* x_6259; lean_object* x_6260; lean_object* x_6261; +lean_dec(x_6254); +lean_dec(x_6253); +x_6258 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_6259 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6259, 0, x_6258); +x_6260 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_6261 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6260, x_6259, x_6252); +if (lean_obj_tag(x_6261) == 0) +{ +lean_object* x_6262; lean_object* x_6263; lean_object* x_6264; lean_object* x_6265; +lean_dec(x_6255); +x_6262 = lean_ctor_get(x_6261, 0); +lean_inc(x_6262); +x_6263 = lean_ctor_get(x_6261, 1); +lean_inc(x_6263); +if (lean_is_exclusive(x_6261)) { + lean_ctor_release(x_6261, 0); + lean_ctor_release(x_6261, 1); + x_6264 = x_6261; +} else { + lean_dec_ref(x_6261); + x_6264 = lean_box(0); +} +if (lean_is_scalar(x_6264)) { + x_6265 = lean_alloc_ctor(0, 2, 0); +} else { + x_6265 = x_6264; +} +lean_ctor_set(x_6265, 0, x_6262); +lean_ctor_set(x_6265, 1, x_6263); +return x_6265; +} +else +{ +lean_object* x_6266; lean_object* x_6267; lean_object* x_6268; lean_object* x_6269; uint8_t x_6270; +x_6266 = lean_ctor_get(x_6261, 0); +lean_inc(x_6266); +x_6267 = lean_ctor_get(x_6261, 1); +lean_inc(x_6267); +if (lean_is_exclusive(x_6261)) { + lean_ctor_release(x_6261, 0); + lean_ctor_release(x_6261, 1); + x_6268 = x_6261; +} else { + lean_dec_ref(x_6261); + x_6268 = lean_box(0); +} +x_6269 = lean_ctor_get(x_6266, 1); +lean_inc(x_6269); +x_6270 = lean_nat_dec_eq(x_6255, x_6269); +lean_dec(x_6255); +if (x_6270 == 0) +{ +lean_object* x_6271; +lean_dec(x_6269); +if (lean_is_scalar(x_6268)) { + x_6271 = lean_alloc_ctor(1, 2, 0); +} else { + x_6271 = x_6268; +} +lean_ctor_set(x_6271, 0, x_6266); +lean_ctor_set(x_6271, 1, x_6267); +return x_6271; +} +else +{ +lean_object* x_6272; lean_object* x_6273; lean_object* x_6274; lean_object* x_6275; +lean_dec(x_6268); +lean_dec(x_6267); +x_6272 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_6273 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6273, 0, x_6272); +x_6274 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_6275 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6274, x_6273, x_6266); +if (lean_obj_tag(x_6275) == 0) +{ +lean_object* x_6276; lean_object* x_6277; lean_object* x_6278; lean_object* x_6279; +lean_dec(x_6269); +x_6276 = lean_ctor_get(x_6275, 0); +lean_inc(x_6276); +x_6277 = lean_ctor_get(x_6275, 1); +lean_inc(x_6277); +if (lean_is_exclusive(x_6275)) { + lean_ctor_release(x_6275, 0); + lean_ctor_release(x_6275, 1); + x_6278 = x_6275; +} else { + lean_dec_ref(x_6275); + x_6278 = lean_box(0); +} +if (lean_is_scalar(x_6278)) { + x_6279 = lean_alloc_ctor(0, 2, 0); +} else { + x_6279 = x_6278; +} +lean_ctor_set(x_6279, 0, x_6276); +lean_ctor_set(x_6279, 1, x_6277); +return x_6279; +} +else +{ +lean_object* x_6280; lean_object* x_6281; lean_object* x_6282; lean_object* x_6283; uint8_t x_6284; +x_6280 = lean_ctor_get(x_6275, 0); +lean_inc(x_6280); +x_6281 = lean_ctor_get(x_6275, 1); +lean_inc(x_6281); +if (lean_is_exclusive(x_6275)) { + lean_ctor_release(x_6275, 0); + lean_ctor_release(x_6275, 1); + x_6282 = x_6275; +} else { + lean_dec_ref(x_6275); + x_6282 = lean_box(0); +} +x_6283 = lean_ctor_get(x_6280, 1); +lean_inc(x_6283); +x_6284 = lean_nat_dec_eq(x_6269, x_6283); +lean_dec(x_6269); +if (x_6284 == 0) +{ +lean_object* x_6285; +lean_dec(x_6283); +if (lean_is_scalar(x_6282)) { + x_6285 = lean_alloc_ctor(1, 2, 0); +} else { + x_6285 = x_6282; +} +lean_ctor_set(x_6285, 0, x_6280); +lean_ctor_set(x_6285, 1, x_6281); +return x_6285; +} +else +{ +lean_object* x_6286; lean_object* x_6287; lean_object* x_6288; lean_object* x_6289; +lean_dec(x_6282); +lean_dec(x_6281); +x_6286 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_6287 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6287, 0, x_6286); +x_6288 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_6289 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6288, x_6287, x_6280); +if (lean_obj_tag(x_6289) == 0) +{ +lean_object* x_6290; lean_object* x_6291; lean_object* x_6292; lean_object* x_6293; +lean_dec(x_6283); +x_6290 = lean_ctor_get(x_6289, 0); +lean_inc(x_6290); +x_6291 = lean_ctor_get(x_6289, 1); +lean_inc(x_6291); +if (lean_is_exclusive(x_6289)) { + lean_ctor_release(x_6289, 0); + lean_ctor_release(x_6289, 1); + x_6292 = x_6289; +} else { + lean_dec_ref(x_6289); + x_6292 = lean_box(0); +} +if (lean_is_scalar(x_6292)) { + x_6293 = lean_alloc_ctor(0, 2, 0); +} else { + x_6293 = x_6292; +} +lean_ctor_set(x_6293, 0, x_6290); +lean_ctor_set(x_6293, 1, x_6291); +return x_6293; +} +else +{ +lean_object* x_6294; lean_object* x_6295; lean_object* x_6296; lean_object* x_6297; uint8_t x_6298; +x_6294 = lean_ctor_get(x_6289, 0); +lean_inc(x_6294); +x_6295 = lean_ctor_get(x_6289, 1); +lean_inc(x_6295); +if (lean_is_exclusive(x_6289)) { + lean_ctor_release(x_6289, 0); + lean_ctor_release(x_6289, 1); + x_6296 = x_6289; +} else { + lean_dec_ref(x_6289); + x_6296 = lean_box(0); +} +x_6297 = lean_ctor_get(x_6294, 1); +lean_inc(x_6297); +x_6298 = lean_nat_dec_eq(x_6283, x_6297); +lean_dec(x_6283); +if (x_6298 == 0) +{ +lean_object* x_6299; +lean_dec(x_6297); +if (lean_is_scalar(x_6296)) { + x_6299 = lean_alloc_ctor(1, 2, 0); +} else { + x_6299 = x_6296; +} +lean_ctor_set(x_6299, 0, x_6294); +lean_ctor_set(x_6299, 1, x_6295); +return x_6299; +} +else +{ +lean_object* x_6300; lean_object* x_6301; lean_object* x_6302; lean_object* x_6303; +lean_dec(x_6296); +lean_dec(x_6295); +x_6300 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_6301 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6301, 0, x_6300); +x_6302 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_6303 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6302, x_6301, x_6294); +if (lean_obj_tag(x_6303) == 0) +{ +lean_object* x_6304; lean_object* x_6305; lean_object* x_6306; lean_object* x_6307; +lean_dec(x_6297); +x_6304 = lean_ctor_get(x_6303, 0); +lean_inc(x_6304); +x_6305 = lean_ctor_get(x_6303, 1); +lean_inc(x_6305); +if (lean_is_exclusive(x_6303)) { + lean_ctor_release(x_6303, 0); + lean_ctor_release(x_6303, 1); + x_6306 = x_6303; +} else { + lean_dec_ref(x_6303); + x_6306 = lean_box(0); +} +if (lean_is_scalar(x_6306)) { + x_6307 = lean_alloc_ctor(0, 2, 0); +} else { + x_6307 = x_6306; +} +lean_ctor_set(x_6307, 0, x_6304); +lean_ctor_set(x_6307, 1, x_6305); +return x_6307; +} +else +{ +lean_object* x_6308; lean_object* x_6309; lean_object* x_6310; lean_object* x_6311; uint8_t x_6312; +x_6308 = lean_ctor_get(x_6303, 0); +lean_inc(x_6308); +x_6309 = lean_ctor_get(x_6303, 1); +lean_inc(x_6309); +if (lean_is_exclusive(x_6303)) { + lean_ctor_release(x_6303, 0); + lean_ctor_release(x_6303, 1); + x_6310 = x_6303; +} else { + lean_dec_ref(x_6303); + x_6310 = lean_box(0); +} +x_6311 = lean_ctor_get(x_6308, 1); +lean_inc(x_6311); +x_6312 = lean_nat_dec_eq(x_6297, x_6311); +lean_dec(x_6297); +if (x_6312 == 0) +{ +lean_object* x_6313; +lean_dec(x_6311); +if (lean_is_scalar(x_6310)) { + x_6313 = lean_alloc_ctor(1, 2, 0); +} else { + x_6313 = x_6310; +} +lean_ctor_set(x_6313, 0, x_6308); +lean_ctor_set(x_6313, 1, x_6309); +return x_6313; +} +else +{ +lean_object* x_6314; lean_object* x_6315; lean_object* x_6316; lean_object* x_6317; +lean_dec(x_6310); +lean_dec(x_6309); +x_6314 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_6315 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6315, 0, x_6314); +x_6316 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_6317 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6316, x_6315, x_6308); +if (lean_obj_tag(x_6317) == 0) +{ +lean_object* x_6318; lean_object* x_6319; lean_object* x_6320; lean_object* x_6321; +lean_dec(x_6311); +x_6318 = lean_ctor_get(x_6317, 0); +lean_inc(x_6318); +x_6319 = lean_ctor_get(x_6317, 1); +lean_inc(x_6319); +if (lean_is_exclusive(x_6317)) { + lean_ctor_release(x_6317, 0); + lean_ctor_release(x_6317, 1); + x_6320 = x_6317; +} else { + lean_dec_ref(x_6317); + x_6320 = lean_box(0); +} +if (lean_is_scalar(x_6320)) { + x_6321 = lean_alloc_ctor(0, 2, 0); +} else { + x_6321 = x_6320; +} +lean_ctor_set(x_6321, 0, x_6318); +lean_ctor_set(x_6321, 1, x_6319); +return x_6321; +} +else +{ +lean_object* x_6322; lean_object* x_6323; lean_object* x_6324; lean_object* x_6325; uint8_t x_6326; +x_6322 = lean_ctor_get(x_6317, 0); +lean_inc(x_6322); +x_6323 = lean_ctor_get(x_6317, 1); +lean_inc(x_6323); +if (lean_is_exclusive(x_6317)) { + lean_ctor_release(x_6317, 0); + lean_ctor_release(x_6317, 1); + x_6324 = x_6317; +} else { + lean_dec_ref(x_6317); + x_6324 = lean_box(0); +} +x_6325 = lean_ctor_get(x_6322, 1); +lean_inc(x_6325); +x_6326 = lean_nat_dec_eq(x_6311, x_6325); +lean_dec(x_6325); +lean_dec(x_6311); +if (x_6326 == 0) +{ +lean_object* x_6327; +if (lean_is_scalar(x_6324)) { + x_6327 = lean_alloc_ctor(1, 2, 0); +} else { + x_6327 = x_6324; +} +lean_ctor_set(x_6327, 0, x_6322); +lean_ctor_set(x_6327, 1, x_6323); +return x_6327; +} +else +{ +lean_object* x_6328; lean_object* x_6329; lean_object* x_6330; lean_object* x_6331; +lean_dec(x_6324); +lean_dec(x_6323); +x_6328 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_6329 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6329, 0, x_6328); +x_6330 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_6331 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6330, x_6329, x_6322); +return x_6331; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_6332; lean_object* x_6333; lean_object* x_6334; uint8_t x_6335; +x_6332 = lean_ctor_get(x_19, 0); +x_6333 = lean_ctor_get(x_19, 1); +lean_inc(x_6333); +lean_inc(x_6332); +lean_dec(x_19); +x_6334 = lean_ctor_get(x_6332, 1); +lean_inc(x_6334); +x_6335 = lean_nat_dec_eq(x_14, x_6334); +lean_dec(x_14); +if (x_6335 == 0) +{ +lean_object* x_6336; +lean_dec(x_6334); +x_6336 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6336, 0, x_6332); +lean_ctor_set(x_6336, 1, x_6333); +return x_6336; +} +else +{ +lean_object* x_6337; lean_object* x_6338; lean_object* x_6339; lean_object* x_6340; +lean_dec(x_6333); +x_6337 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9; +x_6338 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6338, 0, x_6337); +x_6339 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8; +x_6340 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6339, x_6338, x_6332); +if (lean_obj_tag(x_6340) == 0) +{ +lean_object* x_6341; lean_object* x_6342; lean_object* x_6343; lean_object* x_6344; +lean_dec(x_6334); +x_6341 = lean_ctor_get(x_6340, 0); +lean_inc(x_6341); +x_6342 = lean_ctor_get(x_6340, 1); +lean_inc(x_6342); +if (lean_is_exclusive(x_6340)) { + lean_ctor_release(x_6340, 0); + lean_ctor_release(x_6340, 1); + x_6343 = x_6340; +} else { + lean_dec_ref(x_6340); + x_6343 = lean_box(0); +} +if (lean_is_scalar(x_6343)) { + x_6344 = lean_alloc_ctor(0, 2, 0); +} else { + x_6344 = x_6343; +} +lean_ctor_set(x_6344, 0, x_6341); +lean_ctor_set(x_6344, 1, x_6342); +return x_6344; +} +else +{ +lean_object* x_6345; lean_object* x_6346; lean_object* x_6347; lean_object* x_6348; uint8_t x_6349; +x_6345 = lean_ctor_get(x_6340, 0); +lean_inc(x_6345); +x_6346 = lean_ctor_get(x_6340, 1); +lean_inc(x_6346); +if (lean_is_exclusive(x_6340)) { + lean_ctor_release(x_6340, 0); + lean_ctor_release(x_6340, 1); + x_6347 = x_6340; +} else { + lean_dec_ref(x_6340); + x_6347 = lean_box(0); +} +x_6348 = lean_ctor_get(x_6345, 1); +lean_inc(x_6348); +x_6349 = lean_nat_dec_eq(x_6334, x_6348); +lean_dec(x_6334); +if (x_6349 == 0) +{ +lean_object* x_6350; +lean_dec(x_6348); +if (lean_is_scalar(x_6347)) { + x_6350 = lean_alloc_ctor(1, 2, 0); +} else { + x_6350 = x_6347; +} +lean_ctor_set(x_6350, 0, x_6345); +lean_ctor_set(x_6350, 1, x_6346); +return x_6350; +} +else +{ +lean_object* x_6351; lean_object* x_6352; lean_object* x_6353; lean_object* x_6354; +lean_dec(x_6347); +lean_dec(x_6346); +x_6351 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12; +x_6352 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6352, 0, x_6351); +x_6353 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11; +x_6354 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6353, x_6352, x_6345); +if (lean_obj_tag(x_6354) == 0) +{ +lean_object* x_6355; lean_object* x_6356; lean_object* x_6357; lean_object* x_6358; +lean_dec(x_6348); +x_6355 = lean_ctor_get(x_6354, 0); +lean_inc(x_6355); +x_6356 = lean_ctor_get(x_6354, 1); +lean_inc(x_6356); +if (lean_is_exclusive(x_6354)) { + lean_ctor_release(x_6354, 0); + lean_ctor_release(x_6354, 1); + x_6357 = x_6354; +} else { + lean_dec_ref(x_6354); + x_6357 = lean_box(0); +} +if (lean_is_scalar(x_6357)) { + x_6358 = lean_alloc_ctor(0, 2, 0); +} else { + x_6358 = x_6357; +} +lean_ctor_set(x_6358, 0, x_6355); +lean_ctor_set(x_6358, 1, x_6356); +return x_6358; +} +else +{ +lean_object* x_6359; lean_object* x_6360; lean_object* x_6361; lean_object* x_6362; uint8_t x_6363; +x_6359 = lean_ctor_get(x_6354, 0); +lean_inc(x_6359); +x_6360 = lean_ctor_get(x_6354, 1); +lean_inc(x_6360); +if (lean_is_exclusive(x_6354)) { + lean_ctor_release(x_6354, 0); + lean_ctor_release(x_6354, 1); + x_6361 = x_6354; +} else { + lean_dec_ref(x_6354); + x_6361 = lean_box(0); +} +x_6362 = lean_ctor_get(x_6359, 1); +lean_inc(x_6362); +x_6363 = lean_nat_dec_eq(x_6348, x_6362); +lean_dec(x_6348); +if (x_6363 == 0) +{ +lean_object* x_6364; +lean_dec(x_6362); +if (lean_is_scalar(x_6361)) { + x_6364 = lean_alloc_ctor(1, 2, 0); +} else { + x_6364 = x_6361; +} +lean_ctor_set(x_6364, 0, x_6359); +lean_ctor_set(x_6364, 1, x_6360); +return x_6364; +} +else +{ +lean_object* x_6365; lean_object* x_6366; lean_object* x_6367; lean_object* x_6368; +lean_dec(x_6361); +lean_dec(x_6360); +x_6365 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +x_6366 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6366, 0, x_6365); +x_6367 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +x_6368 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6367, x_6366, x_6359); +if (lean_obj_tag(x_6368) == 0) +{ +lean_object* x_6369; lean_object* x_6370; lean_object* x_6371; lean_object* x_6372; +lean_dec(x_6362); +x_6369 = lean_ctor_get(x_6368, 0); +lean_inc(x_6369); +x_6370 = lean_ctor_get(x_6368, 1); +lean_inc(x_6370); +if (lean_is_exclusive(x_6368)) { + lean_ctor_release(x_6368, 0); + lean_ctor_release(x_6368, 1); + x_6371 = x_6368; +} else { + lean_dec_ref(x_6368); + x_6371 = lean_box(0); +} +if (lean_is_scalar(x_6371)) { + x_6372 = lean_alloc_ctor(0, 2, 0); +} else { + x_6372 = x_6371; +} +lean_ctor_set(x_6372, 0, x_6369); +lean_ctor_set(x_6372, 1, x_6370); +return x_6372; +} +else +{ +lean_object* x_6373; lean_object* x_6374; lean_object* x_6375; lean_object* x_6376; uint8_t x_6377; +x_6373 = lean_ctor_get(x_6368, 0); +lean_inc(x_6373); +x_6374 = lean_ctor_get(x_6368, 1); +lean_inc(x_6374); +if (lean_is_exclusive(x_6368)) { + lean_ctor_release(x_6368, 0); + lean_ctor_release(x_6368, 1); + x_6375 = x_6368; +} else { + lean_dec_ref(x_6368); + x_6375 = lean_box(0); +} +x_6376 = lean_ctor_get(x_6373, 1); +lean_inc(x_6376); +x_6377 = lean_nat_dec_eq(x_6362, x_6376); +lean_dec(x_6362); +if (x_6377 == 0) +{ +lean_object* x_6378; +lean_dec(x_6376); +if (lean_is_scalar(x_6375)) { + x_6378 = lean_alloc_ctor(1, 2, 0); +} else { + x_6378 = x_6375; +} +lean_ctor_set(x_6378, 0, x_6373); +lean_ctor_set(x_6378, 1, x_6374); +return x_6378; +} +else +{ +lean_object* x_6379; lean_object* x_6380; lean_object* x_6381; +lean_dec(x_6375); +lean_dec(x_6374); +x_6379 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_6380 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6380, 0, x_6379); +x_6381 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6367, x_6380, x_6373); +if (lean_obj_tag(x_6381) == 0) +{ +lean_object* x_6382; lean_object* x_6383; lean_object* x_6384; lean_object* x_6385; +lean_dec(x_6376); +x_6382 = lean_ctor_get(x_6381, 0); +lean_inc(x_6382); +x_6383 = lean_ctor_get(x_6381, 1); +lean_inc(x_6383); +if (lean_is_exclusive(x_6381)) { + lean_ctor_release(x_6381, 0); + lean_ctor_release(x_6381, 1); + x_6384 = x_6381; +} else { + lean_dec_ref(x_6381); + x_6384 = lean_box(0); +} +if (lean_is_scalar(x_6384)) { + x_6385 = lean_alloc_ctor(0, 2, 0); +} else { + x_6385 = x_6384; +} +lean_ctor_set(x_6385, 0, x_6382); +lean_ctor_set(x_6385, 1, x_6383); +return x_6385; +} +else +{ +lean_object* x_6386; lean_object* x_6387; lean_object* x_6388; lean_object* x_6389; uint8_t x_6390; +x_6386 = lean_ctor_get(x_6381, 0); +lean_inc(x_6386); +x_6387 = lean_ctor_get(x_6381, 1); +lean_inc(x_6387); +if (lean_is_exclusive(x_6381)) { + lean_ctor_release(x_6381, 0); + lean_ctor_release(x_6381, 1); + x_6388 = x_6381; +} else { + lean_dec_ref(x_6381); + x_6388 = lean_box(0); +} +x_6389 = lean_ctor_get(x_6386, 1); +lean_inc(x_6389); +x_6390 = lean_nat_dec_eq(x_6376, x_6389); +lean_dec(x_6376); +if (x_6390 == 0) +{ +lean_object* x_6391; +lean_dec(x_6389); +if (lean_is_scalar(x_6388)) { + x_6391 = lean_alloc_ctor(1, 2, 0); +} else { + x_6391 = x_6388; +} +lean_ctor_set(x_6391, 0, x_6386); +lean_ctor_set(x_6391, 1, x_6387); +return x_6391; +} +else +{ +lean_object* x_6392; lean_object* x_6393; lean_object* x_6394; lean_object* x_6395; +lean_dec(x_6388); +lean_dec(x_6387); +x_6392 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_6393 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6393, 0, x_6392); +x_6394 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_6395 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6394, x_6393, x_6386); +if (lean_obj_tag(x_6395) == 0) +{ +lean_object* x_6396; lean_object* x_6397; lean_object* x_6398; lean_object* x_6399; +lean_dec(x_6389); +x_6396 = lean_ctor_get(x_6395, 0); +lean_inc(x_6396); +x_6397 = lean_ctor_get(x_6395, 1); +lean_inc(x_6397); +if (lean_is_exclusive(x_6395)) { + lean_ctor_release(x_6395, 0); + lean_ctor_release(x_6395, 1); + x_6398 = x_6395; +} else { + lean_dec_ref(x_6395); + x_6398 = lean_box(0); +} +if (lean_is_scalar(x_6398)) { + x_6399 = lean_alloc_ctor(0, 2, 0); +} else { + x_6399 = x_6398; +} +lean_ctor_set(x_6399, 0, x_6396); +lean_ctor_set(x_6399, 1, x_6397); +return x_6399; +} +else +{ +lean_object* x_6400; lean_object* x_6401; lean_object* x_6402; lean_object* x_6403; uint8_t x_6404; +x_6400 = lean_ctor_get(x_6395, 0); +lean_inc(x_6400); +x_6401 = lean_ctor_get(x_6395, 1); +lean_inc(x_6401); +if (lean_is_exclusive(x_6395)) { + lean_ctor_release(x_6395, 0); + lean_ctor_release(x_6395, 1); + x_6402 = x_6395; +} else { + lean_dec_ref(x_6395); + x_6402 = lean_box(0); +} +x_6403 = lean_ctor_get(x_6400, 1); +lean_inc(x_6403); +x_6404 = lean_nat_dec_eq(x_6389, x_6403); +lean_dec(x_6389); +if (x_6404 == 0) +{ +lean_object* x_6405; +lean_dec(x_6403); +if (lean_is_scalar(x_6402)) { + x_6405 = lean_alloc_ctor(1, 2, 0); +} else { + x_6405 = x_6402; +} +lean_ctor_set(x_6405, 0, x_6400); +lean_ctor_set(x_6405, 1, x_6401); +return x_6405; +} +else +{ +lean_object* x_6406; lean_object* x_6407; lean_object* x_6408; lean_object* x_6409; +lean_dec(x_6402); +lean_dec(x_6401); +x_6406 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_6407 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6407, 0, x_6406); +x_6408 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_6409 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6408, x_6407, x_6400); +if (lean_obj_tag(x_6409) == 0) +{ +lean_object* x_6410; lean_object* x_6411; lean_object* x_6412; lean_object* x_6413; +lean_dec(x_6403); +x_6410 = lean_ctor_get(x_6409, 0); +lean_inc(x_6410); +x_6411 = lean_ctor_get(x_6409, 1); +lean_inc(x_6411); +if (lean_is_exclusive(x_6409)) { + lean_ctor_release(x_6409, 0); + lean_ctor_release(x_6409, 1); + x_6412 = x_6409; +} else { + lean_dec_ref(x_6409); + x_6412 = lean_box(0); +} +if (lean_is_scalar(x_6412)) { + x_6413 = lean_alloc_ctor(0, 2, 0); +} else { + x_6413 = x_6412; +} +lean_ctor_set(x_6413, 0, x_6410); +lean_ctor_set(x_6413, 1, x_6411); +return x_6413; +} +else +{ +lean_object* x_6414; lean_object* x_6415; lean_object* x_6416; lean_object* x_6417; uint8_t x_6418; +x_6414 = lean_ctor_get(x_6409, 0); +lean_inc(x_6414); +x_6415 = lean_ctor_get(x_6409, 1); +lean_inc(x_6415); +if (lean_is_exclusive(x_6409)) { + lean_ctor_release(x_6409, 0); + lean_ctor_release(x_6409, 1); + x_6416 = x_6409; +} else { + lean_dec_ref(x_6409); + x_6416 = lean_box(0); +} +x_6417 = lean_ctor_get(x_6414, 1); +lean_inc(x_6417); +x_6418 = lean_nat_dec_eq(x_6403, x_6417); +lean_dec(x_6403); +if (x_6418 == 0) +{ +lean_object* x_6419; +lean_dec(x_6417); +if (lean_is_scalar(x_6416)) { + x_6419 = lean_alloc_ctor(1, 2, 0); +} else { + x_6419 = x_6416; +} +lean_ctor_set(x_6419, 0, x_6414); +lean_ctor_set(x_6419, 1, x_6415); +return x_6419; +} +else +{ +lean_object* x_6420; lean_object* x_6421; lean_object* x_6422; +lean_dec(x_6416); +lean_dec(x_6415); +x_6420 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_6421 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6421, 0, x_6420); +x_6422 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6408, x_6421, x_6414); +if (lean_obj_tag(x_6422) == 0) +{ +lean_object* x_6423; lean_object* x_6424; lean_object* x_6425; lean_object* x_6426; +lean_dec(x_6417); +x_6423 = lean_ctor_get(x_6422, 0); +lean_inc(x_6423); +x_6424 = lean_ctor_get(x_6422, 1); +lean_inc(x_6424); +if (lean_is_exclusive(x_6422)) { + lean_ctor_release(x_6422, 0); + lean_ctor_release(x_6422, 1); + x_6425 = x_6422; +} else { + lean_dec_ref(x_6422); + x_6425 = lean_box(0); +} +if (lean_is_scalar(x_6425)) { + x_6426 = lean_alloc_ctor(0, 2, 0); +} else { + x_6426 = x_6425; +} +lean_ctor_set(x_6426, 0, x_6423); +lean_ctor_set(x_6426, 1, x_6424); +return x_6426; +} +else +{ +lean_object* x_6427; lean_object* x_6428; lean_object* x_6429; lean_object* x_6430; uint8_t x_6431; +x_6427 = lean_ctor_get(x_6422, 0); +lean_inc(x_6427); +x_6428 = lean_ctor_get(x_6422, 1); +lean_inc(x_6428); +if (lean_is_exclusive(x_6422)) { + lean_ctor_release(x_6422, 0); + lean_ctor_release(x_6422, 1); + x_6429 = x_6422; +} else { + lean_dec_ref(x_6422); + x_6429 = lean_box(0); +} +x_6430 = lean_ctor_get(x_6427, 1); +lean_inc(x_6430); +x_6431 = lean_nat_dec_eq(x_6417, x_6430); +lean_dec(x_6417); +if (x_6431 == 0) +{ +lean_object* x_6432; +lean_dec(x_6430); +if (lean_is_scalar(x_6429)) { + x_6432 = lean_alloc_ctor(1, 2, 0); +} else { + x_6432 = x_6429; +} +lean_ctor_set(x_6432, 0, x_6427); +lean_ctor_set(x_6432, 1, x_6428); +return x_6432; +} +else +{ +lean_object* x_6433; lean_object* x_6434; lean_object* x_6435; lean_object* x_6436; +lean_dec(x_6429); +lean_dec(x_6428); +x_6433 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_6434 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6434, 0, x_6433); +x_6435 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_6436 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6435, x_6434, x_6427); +if (lean_obj_tag(x_6436) == 0) +{ +lean_object* x_6437; lean_object* x_6438; lean_object* x_6439; lean_object* x_6440; +lean_dec(x_6430); +x_6437 = lean_ctor_get(x_6436, 0); +lean_inc(x_6437); +x_6438 = lean_ctor_get(x_6436, 1); +lean_inc(x_6438); +if (lean_is_exclusive(x_6436)) { + lean_ctor_release(x_6436, 0); + lean_ctor_release(x_6436, 1); + x_6439 = x_6436; +} else { + lean_dec_ref(x_6436); + x_6439 = lean_box(0); +} +if (lean_is_scalar(x_6439)) { + x_6440 = lean_alloc_ctor(0, 2, 0); +} else { + x_6440 = x_6439; +} +lean_ctor_set(x_6440, 0, x_6437); +lean_ctor_set(x_6440, 1, x_6438); +return x_6440; +} +else +{ +lean_object* x_6441; lean_object* x_6442; lean_object* x_6443; lean_object* x_6444; uint8_t x_6445; +x_6441 = lean_ctor_get(x_6436, 0); +lean_inc(x_6441); +x_6442 = lean_ctor_get(x_6436, 1); +lean_inc(x_6442); +if (lean_is_exclusive(x_6436)) { + lean_ctor_release(x_6436, 0); + lean_ctor_release(x_6436, 1); + x_6443 = x_6436; +} else { + lean_dec_ref(x_6436); + x_6443 = lean_box(0); +} +x_6444 = lean_ctor_get(x_6441, 1); +lean_inc(x_6444); +x_6445 = lean_nat_dec_eq(x_6430, x_6444); +lean_dec(x_6430); +if (x_6445 == 0) +{ +lean_object* x_6446; +lean_dec(x_6444); +if (lean_is_scalar(x_6443)) { + x_6446 = lean_alloc_ctor(1, 2, 0); +} else { + x_6446 = x_6443; +} +lean_ctor_set(x_6446, 0, x_6441); +lean_ctor_set(x_6446, 1, x_6442); +return x_6446; +} +else +{ +lean_object* x_6447; lean_object* x_6448; lean_object* x_6449; lean_object* x_6450; +lean_dec(x_6443); +lean_dec(x_6442); +x_6447 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_6448 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6448, 0, x_6447); +x_6449 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_6450 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6449, x_6448, x_6441); +if (lean_obj_tag(x_6450) == 0) +{ +lean_object* x_6451; lean_object* x_6452; lean_object* x_6453; lean_object* x_6454; +lean_dec(x_6444); +x_6451 = lean_ctor_get(x_6450, 0); +lean_inc(x_6451); +x_6452 = lean_ctor_get(x_6450, 1); +lean_inc(x_6452); +if (lean_is_exclusive(x_6450)) { + lean_ctor_release(x_6450, 0); + lean_ctor_release(x_6450, 1); + x_6453 = x_6450; +} else { + lean_dec_ref(x_6450); + x_6453 = lean_box(0); +} +if (lean_is_scalar(x_6453)) { + x_6454 = lean_alloc_ctor(0, 2, 0); +} else { + x_6454 = x_6453; +} +lean_ctor_set(x_6454, 0, x_6451); +lean_ctor_set(x_6454, 1, x_6452); +return x_6454; +} +else +{ +lean_object* x_6455; lean_object* x_6456; lean_object* x_6457; lean_object* x_6458; uint8_t x_6459; +x_6455 = lean_ctor_get(x_6450, 0); +lean_inc(x_6455); +x_6456 = lean_ctor_get(x_6450, 1); +lean_inc(x_6456); +if (lean_is_exclusive(x_6450)) { + lean_ctor_release(x_6450, 0); + lean_ctor_release(x_6450, 1); + x_6457 = x_6450; +} else { + lean_dec_ref(x_6450); + x_6457 = lean_box(0); +} +x_6458 = lean_ctor_get(x_6455, 1); +lean_inc(x_6458); +x_6459 = lean_nat_dec_eq(x_6444, x_6458); +lean_dec(x_6444); +if (x_6459 == 0) +{ +lean_object* x_6460; +lean_dec(x_6458); +if (lean_is_scalar(x_6457)) { + x_6460 = lean_alloc_ctor(1, 2, 0); +} else { + x_6460 = x_6457; +} +lean_ctor_set(x_6460, 0, x_6455); +lean_ctor_set(x_6460, 1, x_6456); +return x_6460; +} +else +{ +lean_object* x_6461; lean_object* x_6462; lean_object* x_6463; lean_object* x_6464; +lean_dec(x_6457); +lean_dec(x_6456); +x_6461 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_6462 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6462, 0, x_6461); +x_6463 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_6464 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6463, x_6462, x_6455); +if (lean_obj_tag(x_6464) == 0) +{ +lean_object* x_6465; lean_object* x_6466; lean_object* x_6467; lean_object* x_6468; +lean_dec(x_6458); +x_6465 = lean_ctor_get(x_6464, 0); +lean_inc(x_6465); +x_6466 = lean_ctor_get(x_6464, 1); +lean_inc(x_6466); +if (lean_is_exclusive(x_6464)) { + lean_ctor_release(x_6464, 0); + lean_ctor_release(x_6464, 1); + x_6467 = x_6464; +} else { + lean_dec_ref(x_6464); + x_6467 = lean_box(0); +} +if (lean_is_scalar(x_6467)) { + x_6468 = lean_alloc_ctor(0, 2, 0); +} else { + x_6468 = x_6467; +} +lean_ctor_set(x_6468, 0, x_6465); +lean_ctor_set(x_6468, 1, x_6466); +return x_6468; +} +else +{ +lean_object* x_6469; lean_object* x_6470; lean_object* x_6471; lean_object* x_6472; uint8_t x_6473; +x_6469 = lean_ctor_get(x_6464, 0); +lean_inc(x_6469); +x_6470 = lean_ctor_get(x_6464, 1); +lean_inc(x_6470); +if (lean_is_exclusive(x_6464)) { + lean_ctor_release(x_6464, 0); + lean_ctor_release(x_6464, 1); + x_6471 = x_6464; +} else { + lean_dec_ref(x_6464); + x_6471 = lean_box(0); +} +x_6472 = lean_ctor_get(x_6469, 1); +lean_inc(x_6472); +x_6473 = lean_nat_dec_eq(x_6458, x_6472); +lean_dec(x_6458); +if (x_6473 == 0) +{ +lean_object* x_6474; +lean_dec(x_6472); +if (lean_is_scalar(x_6471)) { + x_6474 = lean_alloc_ctor(1, 2, 0); +} else { + x_6474 = x_6471; +} +lean_ctor_set(x_6474, 0, x_6469); +lean_ctor_set(x_6474, 1, x_6470); +return x_6474; +} +else +{ +lean_object* x_6475; lean_object* x_6476; lean_object* x_6477; lean_object* x_6478; +lean_dec(x_6471); +lean_dec(x_6470); +x_6475 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_6476 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6476, 0, x_6475); +x_6477 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_6478 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6477, x_6476, x_6469); +if (lean_obj_tag(x_6478) == 0) +{ +lean_object* x_6479; lean_object* x_6480; lean_object* x_6481; lean_object* x_6482; +lean_dec(x_6472); +x_6479 = lean_ctor_get(x_6478, 0); +lean_inc(x_6479); +x_6480 = lean_ctor_get(x_6478, 1); +lean_inc(x_6480); +if (lean_is_exclusive(x_6478)) { + lean_ctor_release(x_6478, 0); + lean_ctor_release(x_6478, 1); + x_6481 = x_6478; +} else { + lean_dec_ref(x_6478); + x_6481 = lean_box(0); +} +if (lean_is_scalar(x_6481)) { + x_6482 = lean_alloc_ctor(0, 2, 0); +} else { + x_6482 = x_6481; +} +lean_ctor_set(x_6482, 0, x_6479); +lean_ctor_set(x_6482, 1, x_6480); +return x_6482; +} +else +{ +lean_object* x_6483; lean_object* x_6484; lean_object* x_6485; lean_object* x_6486; uint8_t x_6487; +x_6483 = lean_ctor_get(x_6478, 0); +lean_inc(x_6483); +x_6484 = lean_ctor_get(x_6478, 1); +lean_inc(x_6484); +if (lean_is_exclusive(x_6478)) { + lean_ctor_release(x_6478, 0); + lean_ctor_release(x_6478, 1); + x_6485 = x_6478; +} else { + lean_dec_ref(x_6478); + x_6485 = lean_box(0); +} +x_6486 = lean_ctor_get(x_6483, 1); +lean_inc(x_6486); +x_6487 = lean_nat_dec_eq(x_6472, x_6486); +lean_dec(x_6472); +if (x_6487 == 0) +{ +lean_object* x_6488; +lean_dec(x_6486); +if (lean_is_scalar(x_6485)) { + x_6488 = lean_alloc_ctor(1, 2, 0); +} else { + x_6488 = x_6485; +} +lean_ctor_set(x_6488, 0, x_6483); +lean_ctor_set(x_6488, 1, x_6484); +return x_6488; +} +else +{ +lean_object* x_6489; lean_object* x_6490; lean_object* x_6491; +lean_dec(x_6485); +lean_dec(x_6484); +x_6489 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_6490 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6490, 0, x_6489); +x_6491 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6477, x_6490, x_6483); +if (lean_obj_tag(x_6491) == 0) +{ +lean_object* x_6492; lean_object* x_6493; lean_object* x_6494; lean_object* x_6495; +lean_dec(x_6486); +x_6492 = lean_ctor_get(x_6491, 0); +lean_inc(x_6492); +x_6493 = lean_ctor_get(x_6491, 1); +lean_inc(x_6493); +if (lean_is_exclusive(x_6491)) { + lean_ctor_release(x_6491, 0); + lean_ctor_release(x_6491, 1); + x_6494 = x_6491; +} else { + lean_dec_ref(x_6491); + x_6494 = lean_box(0); +} +if (lean_is_scalar(x_6494)) { + x_6495 = lean_alloc_ctor(0, 2, 0); +} else { + x_6495 = x_6494; +} +lean_ctor_set(x_6495, 0, x_6492); +lean_ctor_set(x_6495, 1, x_6493); +return x_6495; +} +else +{ +lean_object* x_6496; lean_object* x_6497; lean_object* x_6498; lean_object* x_6499; uint8_t x_6500; +x_6496 = lean_ctor_get(x_6491, 0); +lean_inc(x_6496); +x_6497 = lean_ctor_get(x_6491, 1); +lean_inc(x_6497); +if (lean_is_exclusive(x_6491)) { + lean_ctor_release(x_6491, 0); + lean_ctor_release(x_6491, 1); + x_6498 = x_6491; +} else { + lean_dec_ref(x_6491); + x_6498 = lean_box(0); +} +x_6499 = lean_ctor_get(x_6496, 1); +lean_inc(x_6499); +x_6500 = lean_nat_dec_eq(x_6486, x_6499); +lean_dec(x_6486); +if (x_6500 == 0) +{ +lean_object* x_6501; +lean_dec(x_6499); +if (lean_is_scalar(x_6498)) { + x_6501 = lean_alloc_ctor(1, 2, 0); +} else { + x_6501 = x_6498; +} +lean_ctor_set(x_6501, 0, x_6496); +lean_ctor_set(x_6501, 1, x_6497); +return x_6501; +} +else +{ +lean_object* x_6502; lean_object* x_6503; lean_object* x_6504; lean_object* x_6505; +lean_dec(x_6498); +lean_dec(x_6497); +x_6502 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_6503 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6503, 0, x_6502); +x_6504 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_6505 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6504, x_6503, x_6496); +if (lean_obj_tag(x_6505) == 0) +{ +lean_object* x_6506; lean_object* x_6507; lean_object* x_6508; lean_object* x_6509; +lean_dec(x_6499); +x_6506 = lean_ctor_get(x_6505, 0); +lean_inc(x_6506); +x_6507 = lean_ctor_get(x_6505, 1); +lean_inc(x_6507); +if (lean_is_exclusive(x_6505)) { + lean_ctor_release(x_6505, 0); + lean_ctor_release(x_6505, 1); + x_6508 = x_6505; +} else { + lean_dec_ref(x_6505); + x_6508 = lean_box(0); +} +if (lean_is_scalar(x_6508)) { + x_6509 = lean_alloc_ctor(0, 2, 0); +} else { + x_6509 = x_6508; +} +lean_ctor_set(x_6509, 0, x_6506); +lean_ctor_set(x_6509, 1, x_6507); +return x_6509; +} +else +{ +lean_object* x_6510; lean_object* x_6511; lean_object* x_6512; lean_object* x_6513; uint8_t x_6514; +x_6510 = lean_ctor_get(x_6505, 0); +lean_inc(x_6510); +x_6511 = lean_ctor_get(x_6505, 1); +lean_inc(x_6511); +if (lean_is_exclusive(x_6505)) { + lean_ctor_release(x_6505, 0); + lean_ctor_release(x_6505, 1); + x_6512 = x_6505; +} else { + lean_dec_ref(x_6505); + x_6512 = lean_box(0); +} +x_6513 = lean_ctor_get(x_6510, 1); +lean_inc(x_6513); +x_6514 = lean_nat_dec_eq(x_6499, x_6513); +lean_dec(x_6499); +if (x_6514 == 0) +{ +lean_object* x_6515; +lean_dec(x_6513); +if (lean_is_scalar(x_6512)) { + x_6515 = lean_alloc_ctor(1, 2, 0); +} else { + x_6515 = x_6512; +} +lean_ctor_set(x_6515, 0, x_6510); +lean_ctor_set(x_6515, 1, x_6511); +return x_6515; +} +else +{ +lean_object* x_6516; lean_object* x_6517; lean_object* x_6518; lean_object* x_6519; +lean_dec(x_6512); +lean_dec(x_6511); +x_6516 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_6517 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6517, 0, x_6516); +x_6518 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_6519 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6518, x_6517, x_6510); +if (lean_obj_tag(x_6519) == 0) +{ +lean_object* x_6520; lean_object* x_6521; lean_object* x_6522; lean_object* x_6523; +lean_dec(x_6513); +x_6520 = lean_ctor_get(x_6519, 0); +lean_inc(x_6520); +x_6521 = lean_ctor_get(x_6519, 1); +lean_inc(x_6521); +if (lean_is_exclusive(x_6519)) { + lean_ctor_release(x_6519, 0); + lean_ctor_release(x_6519, 1); + x_6522 = x_6519; +} else { + lean_dec_ref(x_6519); + x_6522 = lean_box(0); +} +if (lean_is_scalar(x_6522)) { + x_6523 = lean_alloc_ctor(0, 2, 0); +} else { + x_6523 = x_6522; +} +lean_ctor_set(x_6523, 0, x_6520); +lean_ctor_set(x_6523, 1, x_6521); +return x_6523; +} +else +{ +lean_object* x_6524; lean_object* x_6525; lean_object* x_6526; lean_object* x_6527; uint8_t x_6528; +x_6524 = lean_ctor_get(x_6519, 0); +lean_inc(x_6524); +x_6525 = lean_ctor_get(x_6519, 1); +lean_inc(x_6525); +if (lean_is_exclusive(x_6519)) { + lean_ctor_release(x_6519, 0); + lean_ctor_release(x_6519, 1); + x_6526 = x_6519; +} else { + lean_dec_ref(x_6519); + x_6526 = lean_box(0); +} +x_6527 = lean_ctor_get(x_6524, 1); +lean_inc(x_6527); +x_6528 = lean_nat_dec_eq(x_6513, x_6527); +lean_dec(x_6513); +if (x_6528 == 0) +{ +lean_object* x_6529; +lean_dec(x_6527); +if (lean_is_scalar(x_6526)) { + x_6529 = lean_alloc_ctor(1, 2, 0); +} else { + x_6529 = x_6526; +} +lean_ctor_set(x_6529, 0, x_6524); +lean_ctor_set(x_6529, 1, x_6525); +return x_6529; +} +else +{ +lean_object* x_6530; lean_object* x_6531; lean_object* x_6532; lean_object* x_6533; +lean_dec(x_6526); +lean_dec(x_6525); +x_6530 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_6531 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6531, 0, x_6530); +x_6532 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_6533 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6532, x_6531, x_6524); +if (lean_obj_tag(x_6533) == 0) +{ +lean_object* x_6534; lean_object* x_6535; lean_object* x_6536; lean_object* x_6537; +lean_dec(x_6527); +x_6534 = lean_ctor_get(x_6533, 0); +lean_inc(x_6534); +x_6535 = lean_ctor_get(x_6533, 1); +lean_inc(x_6535); +if (lean_is_exclusive(x_6533)) { + lean_ctor_release(x_6533, 0); + lean_ctor_release(x_6533, 1); + x_6536 = x_6533; +} else { + lean_dec_ref(x_6533); + x_6536 = lean_box(0); +} +if (lean_is_scalar(x_6536)) { + x_6537 = lean_alloc_ctor(0, 2, 0); +} else { + x_6537 = x_6536; +} +lean_ctor_set(x_6537, 0, x_6534); +lean_ctor_set(x_6537, 1, x_6535); +return x_6537; +} +else +{ +lean_object* x_6538; lean_object* x_6539; lean_object* x_6540; lean_object* x_6541; uint8_t x_6542; +x_6538 = lean_ctor_get(x_6533, 0); +lean_inc(x_6538); +x_6539 = lean_ctor_get(x_6533, 1); +lean_inc(x_6539); +if (lean_is_exclusive(x_6533)) { + lean_ctor_release(x_6533, 0); + lean_ctor_release(x_6533, 1); + x_6540 = x_6533; +} else { + lean_dec_ref(x_6533); + x_6540 = lean_box(0); +} +x_6541 = lean_ctor_get(x_6538, 1); +lean_inc(x_6541); +x_6542 = lean_nat_dec_eq(x_6527, x_6541); +lean_dec(x_6527); +if (x_6542 == 0) +{ +lean_object* x_6543; +lean_dec(x_6541); +if (lean_is_scalar(x_6540)) { + x_6543 = lean_alloc_ctor(1, 2, 0); +} else { + x_6543 = x_6540; +} +lean_ctor_set(x_6543, 0, x_6538); +lean_ctor_set(x_6543, 1, x_6539); +return x_6543; +} +else +{ +lean_object* x_6544; lean_object* x_6545; lean_object* x_6546; lean_object* x_6547; +lean_dec(x_6540); +lean_dec(x_6539); +x_6544 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_6545 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6545, 0, x_6544); +x_6546 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_6547 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6546, x_6545, x_6538); +if (lean_obj_tag(x_6547) == 0) +{ +lean_object* x_6548; lean_object* x_6549; lean_object* x_6550; lean_object* x_6551; +lean_dec(x_6541); +x_6548 = lean_ctor_get(x_6547, 0); +lean_inc(x_6548); +x_6549 = lean_ctor_get(x_6547, 1); +lean_inc(x_6549); +if (lean_is_exclusive(x_6547)) { + lean_ctor_release(x_6547, 0); + lean_ctor_release(x_6547, 1); + x_6550 = x_6547; +} else { + lean_dec_ref(x_6547); + x_6550 = lean_box(0); +} +if (lean_is_scalar(x_6550)) { + x_6551 = lean_alloc_ctor(0, 2, 0); +} else { + x_6551 = x_6550; +} +lean_ctor_set(x_6551, 0, x_6548); +lean_ctor_set(x_6551, 1, x_6549); +return x_6551; +} +else +{ +lean_object* x_6552; lean_object* x_6553; lean_object* x_6554; lean_object* x_6555; uint8_t x_6556; +x_6552 = lean_ctor_get(x_6547, 0); +lean_inc(x_6552); +x_6553 = lean_ctor_get(x_6547, 1); +lean_inc(x_6553); +if (lean_is_exclusive(x_6547)) { + lean_ctor_release(x_6547, 0); + lean_ctor_release(x_6547, 1); + x_6554 = x_6547; +} else { + lean_dec_ref(x_6547); + x_6554 = lean_box(0); +} +x_6555 = lean_ctor_get(x_6552, 1); +lean_inc(x_6555); +x_6556 = lean_nat_dec_eq(x_6541, x_6555); +lean_dec(x_6541); +if (x_6556 == 0) +{ +lean_object* x_6557; +lean_dec(x_6555); +if (lean_is_scalar(x_6554)) { + x_6557 = lean_alloc_ctor(1, 2, 0); +} else { + x_6557 = x_6554; +} +lean_ctor_set(x_6557, 0, x_6552); +lean_ctor_set(x_6557, 1, x_6553); +return x_6557; +} +else +{ +lean_object* x_6558; lean_object* x_6559; lean_object* x_6560; lean_object* x_6561; +lean_dec(x_6554); +lean_dec(x_6553); +x_6558 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_6559 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6559, 0, x_6558); +x_6560 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_6561 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6560, x_6559, x_6552); +if (lean_obj_tag(x_6561) == 0) +{ +lean_object* x_6562; lean_object* x_6563; lean_object* x_6564; lean_object* x_6565; +lean_dec(x_6555); +x_6562 = lean_ctor_get(x_6561, 0); +lean_inc(x_6562); +x_6563 = lean_ctor_get(x_6561, 1); +lean_inc(x_6563); +if (lean_is_exclusive(x_6561)) { + lean_ctor_release(x_6561, 0); + lean_ctor_release(x_6561, 1); + x_6564 = x_6561; +} else { + lean_dec_ref(x_6561); + x_6564 = lean_box(0); +} +if (lean_is_scalar(x_6564)) { + x_6565 = lean_alloc_ctor(0, 2, 0); +} else { + x_6565 = x_6564; +} +lean_ctor_set(x_6565, 0, x_6562); +lean_ctor_set(x_6565, 1, x_6563); +return x_6565; +} +else +{ +lean_object* x_6566; lean_object* x_6567; lean_object* x_6568; lean_object* x_6569; uint8_t x_6570; +x_6566 = lean_ctor_get(x_6561, 0); +lean_inc(x_6566); +x_6567 = lean_ctor_get(x_6561, 1); +lean_inc(x_6567); +if (lean_is_exclusive(x_6561)) { + lean_ctor_release(x_6561, 0); + lean_ctor_release(x_6561, 1); + x_6568 = x_6561; +} else { + lean_dec_ref(x_6561); + x_6568 = lean_box(0); +} +x_6569 = lean_ctor_get(x_6566, 1); +lean_inc(x_6569); +x_6570 = lean_nat_dec_eq(x_6555, x_6569); +lean_dec(x_6555); +if (x_6570 == 0) +{ +lean_object* x_6571; +lean_dec(x_6569); +if (lean_is_scalar(x_6568)) { + x_6571 = lean_alloc_ctor(1, 2, 0); +} else { + x_6571 = x_6568; +} +lean_ctor_set(x_6571, 0, x_6566); +lean_ctor_set(x_6571, 1, x_6567); +return x_6571; +} +else +{ +lean_object* x_6572; lean_object* x_6573; lean_object* x_6574; lean_object* x_6575; +lean_dec(x_6568); +lean_dec(x_6567); +x_6572 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_6573 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6573, 0, x_6572); +x_6574 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_6575 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6574, x_6573, x_6566); +if (lean_obj_tag(x_6575) == 0) +{ +lean_object* x_6576; lean_object* x_6577; lean_object* x_6578; lean_object* x_6579; +lean_dec(x_6569); +x_6576 = lean_ctor_get(x_6575, 0); +lean_inc(x_6576); +x_6577 = lean_ctor_get(x_6575, 1); +lean_inc(x_6577); +if (lean_is_exclusive(x_6575)) { + lean_ctor_release(x_6575, 0); + lean_ctor_release(x_6575, 1); + x_6578 = x_6575; +} else { + lean_dec_ref(x_6575); + x_6578 = lean_box(0); +} +if (lean_is_scalar(x_6578)) { + x_6579 = lean_alloc_ctor(0, 2, 0); +} else { + x_6579 = x_6578; +} +lean_ctor_set(x_6579, 0, x_6576); +lean_ctor_set(x_6579, 1, x_6577); +return x_6579; +} +else +{ +lean_object* x_6580; lean_object* x_6581; lean_object* x_6582; lean_object* x_6583; uint8_t x_6584; +x_6580 = lean_ctor_get(x_6575, 0); +lean_inc(x_6580); +x_6581 = lean_ctor_get(x_6575, 1); +lean_inc(x_6581); +if (lean_is_exclusive(x_6575)) { + lean_ctor_release(x_6575, 0); + lean_ctor_release(x_6575, 1); + x_6582 = x_6575; +} else { + lean_dec_ref(x_6575); + x_6582 = lean_box(0); +} +x_6583 = lean_ctor_get(x_6580, 1); +lean_inc(x_6583); +x_6584 = lean_nat_dec_eq(x_6569, x_6583); +lean_dec(x_6569); +if (x_6584 == 0) +{ +lean_object* x_6585; +lean_dec(x_6583); +if (lean_is_scalar(x_6582)) { + x_6585 = lean_alloc_ctor(1, 2, 0); +} else { + x_6585 = x_6582; +} +lean_ctor_set(x_6585, 0, x_6580); +lean_ctor_set(x_6585, 1, x_6581); +return x_6585; +} +else +{ +lean_object* x_6586; lean_object* x_6587; lean_object* x_6588; lean_object* x_6589; +lean_dec(x_6582); +lean_dec(x_6581); +x_6586 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_6587 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6587, 0, x_6586); +x_6588 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_6589 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6588, x_6587, x_6580); +if (lean_obj_tag(x_6589) == 0) +{ +lean_object* x_6590; lean_object* x_6591; lean_object* x_6592; lean_object* x_6593; +lean_dec(x_6583); +x_6590 = lean_ctor_get(x_6589, 0); +lean_inc(x_6590); +x_6591 = lean_ctor_get(x_6589, 1); +lean_inc(x_6591); +if (lean_is_exclusive(x_6589)) { + lean_ctor_release(x_6589, 0); + lean_ctor_release(x_6589, 1); + x_6592 = x_6589; +} else { + lean_dec_ref(x_6589); + x_6592 = lean_box(0); +} +if (lean_is_scalar(x_6592)) { + x_6593 = lean_alloc_ctor(0, 2, 0); +} else { + x_6593 = x_6592; +} +lean_ctor_set(x_6593, 0, x_6590); +lean_ctor_set(x_6593, 1, x_6591); +return x_6593; +} +else +{ +lean_object* x_6594; lean_object* x_6595; lean_object* x_6596; lean_object* x_6597; uint8_t x_6598; +x_6594 = lean_ctor_get(x_6589, 0); +lean_inc(x_6594); +x_6595 = lean_ctor_get(x_6589, 1); +lean_inc(x_6595); +if (lean_is_exclusive(x_6589)) { + lean_ctor_release(x_6589, 0); + lean_ctor_release(x_6589, 1); + x_6596 = x_6589; +} else { + lean_dec_ref(x_6589); + x_6596 = lean_box(0); +} +x_6597 = lean_ctor_get(x_6594, 1); +lean_inc(x_6597); +x_6598 = lean_nat_dec_eq(x_6583, x_6597); +lean_dec(x_6583); +if (x_6598 == 0) +{ +lean_object* x_6599; +lean_dec(x_6597); +if (lean_is_scalar(x_6596)) { + x_6599 = lean_alloc_ctor(1, 2, 0); +} else { + x_6599 = x_6596; +} +lean_ctor_set(x_6599, 0, x_6594); +lean_ctor_set(x_6599, 1, x_6595); +return x_6599; +} +else +{ +lean_object* x_6600; lean_object* x_6601; lean_object* x_6602; lean_object* x_6603; +lean_dec(x_6596); +lean_dec(x_6595); +x_6600 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_6601 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6601, 0, x_6600); +x_6602 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_6603 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6602, x_6601, x_6594); +if (lean_obj_tag(x_6603) == 0) +{ +lean_object* x_6604; lean_object* x_6605; lean_object* x_6606; lean_object* x_6607; +lean_dec(x_6597); +x_6604 = lean_ctor_get(x_6603, 0); +lean_inc(x_6604); +x_6605 = lean_ctor_get(x_6603, 1); +lean_inc(x_6605); +if (lean_is_exclusive(x_6603)) { + lean_ctor_release(x_6603, 0); + lean_ctor_release(x_6603, 1); + x_6606 = x_6603; +} else { + lean_dec_ref(x_6603); + x_6606 = lean_box(0); +} +if (lean_is_scalar(x_6606)) { + x_6607 = lean_alloc_ctor(0, 2, 0); +} else { + x_6607 = x_6606; +} +lean_ctor_set(x_6607, 0, x_6604); +lean_ctor_set(x_6607, 1, x_6605); +return x_6607; +} +else +{ +lean_object* x_6608; lean_object* x_6609; lean_object* x_6610; lean_object* x_6611; uint8_t x_6612; +x_6608 = lean_ctor_get(x_6603, 0); +lean_inc(x_6608); +x_6609 = lean_ctor_get(x_6603, 1); +lean_inc(x_6609); +if (lean_is_exclusive(x_6603)) { + lean_ctor_release(x_6603, 0); + lean_ctor_release(x_6603, 1); + x_6610 = x_6603; +} else { + lean_dec_ref(x_6603); + x_6610 = lean_box(0); +} +x_6611 = lean_ctor_get(x_6608, 1); +lean_inc(x_6611); +x_6612 = lean_nat_dec_eq(x_6597, x_6611); +lean_dec(x_6597); +if (x_6612 == 0) +{ +lean_object* x_6613; +lean_dec(x_6611); +if (lean_is_scalar(x_6610)) { + x_6613 = lean_alloc_ctor(1, 2, 0); +} else { + x_6613 = x_6610; +} +lean_ctor_set(x_6613, 0, x_6608); +lean_ctor_set(x_6613, 1, x_6609); +return x_6613; +} +else +{ +lean_object* x_6614; lean_object* x_6615; lean_object* x_6616; lean_object* x_6617; +lean_dec(x_6610); +lean_dec(x_6609); +x_6614 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_6615 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6615, 0, x_6614); +x_6616 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_6617 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6616, x_6615, x_6608); +if (lean_obj_tag(x_6617) == 0) +{ +lean_object* x_6618; lean_object* x_6619; lean_object* x_6620; lean_object* x_6621; +lean_dec(x_6611); +x_6618 = lean_ctor_get(x_6617, 0); +lean_inc(x_6618); +x_6619 = lean_ctor_get(x_6617, 1); +lean_inc(x_6619); +if (lean_is_exclusive(x_6617)) { + lean_ctor_release(x_6617, 0); + lean_ctor_release(x_6617, 1); + x_6620 = x_6617; +} else { + lean_dec_ref(x_6617); + x_6620 = lean_box(0); +} +if (lean_is_scalar(x_6620)) { + x_6621 = lean_alloc_ctor(0, 2, 0); +} else { + x_6621 = x_6620; +} +lean_ctor_set(x_6621, 0, x_6618); +lean_ctor_set(x_6621, 1, x_6619); +return x_6621; +} +else +{ +lean_object* x_6622; lean_object* x_6623; lean_object* x_6624; lean_object* x_6625; uint8_t x_6626; +x_6622 = lean_ctor_get(x_6617, 0); +lean_inc(x_6622); +x_6623 = lean_ctor_get(x_6617, 1); +lean_inc(x_6623); +if (lean_is_exclusive(x_6617)) { + lean_ctor_release(x_6617, 0); + lean_ctor_release(x_6617, 1); + x_6624 = x_6617; +} else { + lean_dec_ref(x_6617); + x_6624 = lean_box(0); +} +x_6625 = lean_ctor_get(x_6622, 1); +lean_inc(x_6625); +x_6626 = lean_nat_dec_eq(x_6611, x_6625); +lean_dec(x_6611); +if (x_6626 == 0) +{ +lean_object* x_6627; +lean_dec(x_6625); +if (lean_is_scalar(x_6624)) { + x_6627 = lean_alloc_ctor(1, 2, 0); +} else { + x_6627 = x_6624; +} +lean_ctor_set(x_6627, 0, x_6622); +lean_ctor_set(x_6627, 1, x_6623); +return x_6627; +} +else +{ +lean_object* x_6628; lean_object* x_6629; lean_object* x_6630; lean_object* x_6631; +lean_dec(x_6624); +lean_dec(x_6623); +x_6628 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_6629 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6629, 0, x_6628); +x_6630 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_6631 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6630, x_6629, x_6622); +if (lean_obj_tag(x_6631) == 0) +{ +lean_object* x_6632; lean_object* x_6633; lean_object* x_6634; lean_object* x_6635; +lean_dec(x_6625); +x_6632 = lean_ctor_get(x_6631, 0); +lean_inc(x_6632); +x_6633 = lean_ctor_get(x_6631, 1); +lean_inc(x_6633); +if (lean_is_exclusive(x_6631)) { + lean_ctor_release(x_6631, 0); + lean_ctor_release(x_6631, 1); + x_6634 = x_6631; +} else { + lean_dec_ref(x_6631); + x_6634 = lean_box(0); +} +if (lean_is_scalar(x_6634)) { + x_6635 = lean_alloc_ctor(0, 2, 0); +} else { + x_6635 = x_6634; +} +lean_ctor_set(x_6635, 0, x_6632); +lean_ctor_set(x_6635, 1, x_6633); +return x_6635; +} +else +{ +lean_object* x_6636; lean_object* x_6637; lean_object* x_6638; lean_object* x_6639; uint8_t x_6640; +x_6636 = lean_ctor_get(x_6631, 0); +lean_inc(x_6636); +x_6637 = lean_ctor_get(x_6631, 1); +lean_inc(x_6637); +if (lean_is_exclusive(x_6631)) { + lean_ctor_release(x_6631, 0); + lean_ctor_release(x_6631, 1); + x_6638 = x_6631; +} else { + lean_dec_ref(x_6631); + x_6638 = lean_box(0); +} +x_6639 = lean_ctor_get(x_6636, 1); +lean_inc(x_6639); +x_6640 = lean_nat_dec_eq(x_6625, x_6639); +lean_dec(x_6625); +if (x_6640 == 0) +{ +lean_object* x_6641; +lean_dec(x_6639); +if (lean_is_scalar(x_6638)) { + x_6641 = lean_alloc_ctor(1, 2, 0); +} else { + x_6641 = x_6638; +} +lean_ctor_set(x_6641, 0, x_6636); +lean_ctor_set(x_6641, 1, x_6637); +return x_6641; +} +else +{ +lean_object* x_6642; lean_object* x_6643; lean_object* x_6644; lean_object* x_6645; +lean_dec(x_6638); +lean_dec(x_6637); +x_6642 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_6643 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6643, 0, x_6642); +x_6644 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_6645 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6644, x_6643, x_6636); +if (lean_obj_tag(x_6645) == 0) +{ +lean_object* x_6646; lean_object* x_6647; lean_object* x_6648; lean_object* x_6649; +lean_dec(x_6639); +x_6646 = lean_ctor_get(x_6645, 0); +lean_inc(x_6646); +x_6647 = lean_ctor_get(x_6645, 1); +lean_inc(x_6647); +if (lean_is_exclusive(x_6645)) { + lean_ctor_release(x_6645, 0); + lean_ctor_release(x_6645, 1); + x_6648 = x_6645; +} else { + lean_dec_ref(x_6645); + x_6648 = lean_box(0); +} +if (lean_is_scalar(x_6648)) { + x_6649 = lean_alloc_ctor(0, 2, 0); +} else { + x_6649 = x_6648; +} +lean_ctor_set(x_6649, 0, x_6646); +lean_ctor_set(x_6649, 1, x_6647); +return x_6649; +} +else +{ +lean_object* x_6650; lean_object* x_6651; lean_object* x_6652; lean_object* x_6653; uint8_t x_6654; +x_6650 = lean_ctor_get(x_6645, 0); +lean_inc(x_6650); +x_6651 = lean_ctor_get(x_6645, 1); +lean_inc(x_6651); +if (lean_is_exclusive(x_6645)) { + lean_ctor_release(x_6645, 0); + lean_ctor_release(x_6645, 1); + x_6652 = x_6645; +} else { + lean_dec_ref(x_6645); + x_6652 = lean_box(0); +} +x_6653 = lean_ctor_get(x_6650, 1); +lean_inc(x_6653); +x_6654 = lean_nat_dec_eq(x_6639, x_6653); +lean_dec(x_6639); +if (x_6654 == 0) +{ +lean_object* x_6655; +lean_dec(x_6653); +if (lean_is_scalar(x_6652)) { + x_6655 = lean_alloc_ctor(1, 2, 0); +} else { + x_6655 = x_6652; +} +lean_ctor_set(x_6655, 0, x_6650); +lean_ctor_set(x_6655, 1, x_6651); +return x_6655; +} +else +{ +lean_object* x_6656; lean_object* x_6657; lean_object* x_6658; lean_object* x_6659; +lean_dec(x_6652); +lean_dec(x_6651); +x_6656 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_6657 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6657, 0, x_6656); +x_6658 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_6659 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6658, x_6657, x_6650); +if (lean_obj_tag(x_6659) == 0) +{ +lean_object* x_6660; lean_object* x_6661; lean_object* x_6662; lean_object* x_6663; +lean_dec(x_6653); +x_6660 = lean_ctor_get(x_6659, 0); +lean_inc(x_6660); +x_6661 = lean_ctor_get(x_6659, 1); +lean_inc(x_6661); +if (lean_is_exclusive(x_6659)) { + lean_ctor_release(x_6659, 0); + lean_ctor_release(x_6659, 1); + x_6662 = x_6659; +} else { + lean_dec_ref(x_6659); + x_6662 = lean_box(0); +} +if (lean_is_scalar(x_6662)) { + x_6663 = lean_alloc_ctor(0, 2, 0); +} else { + x_6663 = x_6662; +} +lean_ctor_set(x_6663, 0, x_6660); +lean_ctor_set(x_6663, 1, x_6661); +return x_6663; +} +else +{ +lean_object* x_6664; lean_object* x_6665; lean_object* x_6666; lean_object* x_6667; uint8_t x_6668; +x_6664 = lean_ctor_get(x_6659, 0); +lean_inc(x_6664); +x_6665 = lean_ctor_get(x_6659, 1); +lean_inc(x_6665); +if (lean_is_exclusive(x_6659)) { + lean_ctor_release(x_6659, 0); + lean_ctor_release(x_6659, 1); + x_6666 = x_6659; +} else { + lean_dec_ref(x_6659); + x_6666 = lean_box(0); +} +x_6667 = lean_ctor_get(x_6664, 1); +lean_inc(x_6667); +x_6668 = lean_nat_dec_eq(x_6653, x_6667); +lean_dec(x_6653); +if (x_6668 == 0) +{ +lean_object* x_6669; +lean_dec(x_6667); +if (lean_is_scalar(x_6666)) { + x_6669 = lean_alloc_ctor(1, 2, 0); +} else { + x_6669 = x_6666; +} +lean_ctor_set(x_6669, 0, x_6664); +lean_ctor_set(x_6669, 1, x_6665); +return x_6669; +} +else +{ +lean_object* x_6670; lean_object* x_6671; lean_object* x_6672; lean_object* x_6673; +lean_dec(x_6666); +lean_dec(x_6665); +x_6670 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_6671 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6671, 0, x_6670); +x_6672 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_6673 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6672, x_6671, x_6664); +if (lean_obj_tag(x_6673) == 0) +{ +lean_object* x_6674; lean_object* x_6675; lean_object* x_6676; lean_object* x_6677; +lean_dec(x_6667); +x_6674 = lean_ctor_get(x_6673, 0); +lean_inc(x_6674); +x_6675 = lean_ctor_get(x_6673, 1); +lean_inc(x_6675); +if (lean_is_exclusive(x_6673)) { + lean_ctor_release(x_6673, 0); + lean_ctor_release(x_6673, 1); + x_6676 = x_6673; +} else { + lean_dec_ref(x_6673); + x_6676 = lean_box(0); +} +if (lean_is_scalar(x_6676)) { + x_6677 = lean_alloc_ctor(0, 2, 0); +} else { + x_6677 = x_6676; +} +lean_ctor_set(x_6677, 0, x_6674); +lean_ctor_set(x_6677, 1, x_6675); +return x_6677; +} +else +{ +lean_object* x_6678; lean_object* x_6679; lean_object* x_6680; lean_object* x_6681; uint8_t x_6682; +x_6678 = lean_ctor_get(x_6673, 0); +lean_inc(x_6678); +x_6679 = lean_ctor_get(x_6673, 1); +lean_inc(x_6679); +if (lean_is_exclusive(x_6673)) { + lean_ctor_release(x_6673, 0); + lean_ctor_release(x_6673, 1); + x_6680 = x_6673; +} else { + lean_dec_ref(x_6673); + x_6680 = lean_box(0); +} +x_6681 = lean_ctor_get(x_6678, 1); +lean_inc(x_6681); +x_6682 = lean_nat_dec_eq(x_6667, x_6681); +lean_dec(x_6667); +if (x_6682 == 0) +{ +lean_object* x_6683; +lean_dec(x_6681); +if (lean_is_scalar(x_6680)) { + x_6683 = lean_alloc_ctor(1, 2, 0); +} else { + x_6683 = x_6680; +} +lean_ctor_set(x_6683, 0, x_6678); +lean_ctor_set(x_6683, 1, x_6679); +return x_6683; +} +else +{ +lean_object* x_6684; lean_object* x_6685; lean_object* x_6686; lean_object* x_6687; +lean_dec(x_6680); +lean_dec(x_6679); +x_6684 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_6685 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6685, 0, x_6684); +x_6686 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_6687 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6686, x_6685, x_6678); +if (lean_obj_tag(x_6687) == 0) +{ +lean_object* x_6688; lean_object* x_6689; lean_object* x_6690; lean_object* x_6691; +lean_dec(x_6681); +x_6688 = lean_ctor_get(x_6687, 0); +lean_inc(x_6688); +x_6689 = lean_ctor_get(x_6687, 1); +lean_inc(x_6689); +if (lean_is_exclusive(x_6687)) { + lean_ctor_release(x_6687, 0); + lean_ctor_release(x_6687, 1); + x_6690 = x_6687; +} else { + lean_dec_ref(x_6687); + x_6690 = lean_box(0); +} +if (lean_is_scalar(x_6690)) { + x_6691 = lean_alloc_ctor(0, 2, 0); +} else { + x_6691 = x_6690; +} +lean_ctor_set(x_6691, 0, x_6688); +lean_ctor_set(x_6691, 1, x_6689); +return x_6691; +} +else +{ +lean_object* x_6692; lean_object* x_6693; lean_object* x_6694; lean_object* x_6695; uint8_t x_6696; +x_6692 = lean_ctor_get(x_6687, 0); +lean_inc(x_6692); +x_6693 = lean_ctor_get(x_6687, 1); +lean_inc(x_6693); +if (lean_is_exclusive(x_6687)) { + lean_ctor_release(x_6687, 0); + lean_ctor_release(x_6687, 1); + x_6694 = x_6687; +} else { + lean_dec_ref(x_6687); + x_6694 = lean_box(0); +} +x_6695 = lean_ctor_get(x_6692, 1); +lean_inc(x_6695); +x_6696 = lean_nat_dec_eq(x_6681, x_6695); +lean_dec(x_6681); +if (x_6696 == 0) +{ +lean_object* x_6697; +lean_dec(x_6695); +if (lean_is_scalar(x_6694)) { + x_6697 = lean_alloc_ctor(1, 2, 0); +} else { + x_6697 = x_6694; +} +lean_ctor_set(x_6697, 0, x_6692); +lean_ctor_set(x_6697, 1, x_6693); +return x_6697; +} +else +{ +lean_object* x_6698; lean_object* x_6699; lean_object* x_6700; lean_object* x_6701; +lean_dec(x_6694); +lean_dec(x_6693); +x_6698 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_6699 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6699, 0, x_6698); +x_6700 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_6701 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6700, x_6699, x_6692); +if (lean_obj_tag(x_6701) == 0) +{ +lean_object* x_6702; lean_object* x_6703; lean_object* x_6704; lean_object* x_6705; +lean_dec(x_6695); +x_6702 = lean_ctor_get(x_6701, 0); +lean_inc(x_6702); +x_6703 = lean_ctor_get(x_6701, 1); +lean_inc(x_6703); +if (lean_is_exclusive(x_6701)) { + lean_ctor_release(x_6701, 0); + lean_ctor_release(x_6701, 1); + x_6704 = x_6701; +} else { + lean_dec_ref(x_6701); + x_6704 = lean_box(0); +} +if (lean_is_scalar(x_6704)) { + x_6705 = lean_alloc_ctor(0, 2, 0); +} else { + x_6705 = x_6704; +} +lean_ctor_set(x_6705, 0, x_6702); +lean_ctor_set(x_6705, 1, x_6703); +return x_6705; +} +else +{ +lean_object* x_6706; lean_object* x_6707; lean_object* x_6708; lean_object* x_6709; uint8_t x_6710; +x_6706 = lean_ctor_get(x_6701, 0); +lean_inc(x_6706); +x_6707 = lean_ctor_get(x_6701, 1); +lean_inc(x_6707); +if (lean_is_exclusive(x_6701)) { + lean_ctor_release(x_6701, 0); + lean_ctor_release(x_6701, 1); + x_6708 = x_6701; +} else { + lean_dec_ref(x_6701); + x_6708 = lean_box(0); +} +x_6709 = lean_ctor_get(x_6706, 1); +lean_inc(x_6709); +x_6710 = lean_nat_dec_eq(x_6695, x_6709); +lean_dec(x_6695); +if (x_6710 == 0) +{ +lean_object* x_6711; +lean_dec(x_6709); +if (lean_is_scalar(x_6708)) { + x_6711 = lean_alloc_ctor(1, 2, 0); +} else { + x_6711 = x_6708; +} +lean_ctor_set(x_6711, 0, x_6706); +lean_ctor_set(x_6711, 1, x_6707); +return x_6711; +} +else +{ +lean_object* x_6712; lean_object* x_6713; lean_object* x_6714; lean_object* x_6715; +lean_dec(x_6708); +lean_dec(x_6707); +x_6712 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_6713 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6713, 0, x_6712); +x_6714 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_6715 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6714, x_6713, x_6706); +if (lean_obj_tag(x_6715) == 0) +{ +lean_object* x_6716; lean_object* x_6717; lean_object* x_6718; lean_object* x_6719; +lean_dec(x_6709); +x_6716 = lean_ctor_get(x_6715, 0); +lean_inc(x_6716); +x_6717 = lean_ctor_get(x_6715, 1); +lean_inc(x_6717); +if (lean_is_exclusive(x_6715)) { + lean_ctor_release(x_6715, 0); + lean_ctor_release(x_6715, 1); + x_6718 = x_6715; +} else { + lean_dec_ref(x_6715); + x_6718 = lean_box(0); +} +if (lean_is_scalar(x_6718)) { + x_6719 = lean_alloc_ctor(0, 2, 0); +} else { + x_6719 = x_6718; +} +lean_ctor_set(x_6719, 0, x_6716); +lean_ctor_set(x_6719, 1, x_6717); +return x_6719; +} +else +{ +lean_object* x_6720; lean_object* x_6721; lean_object* x_6722; lean_object* x_6723; uint8_t x_6724; +x_6720 = lean_ctor_get(x_6715, 0); +lean_inc(x_6720); +x_6721 = lean_ctor_get(x_6715, 1); +lean_inc(x_6721); +if (lean_is_exclusive(x_6715)) { + lean_ctor_release(x_6715, 0); + lean_ctor_release(x_6715, 1); + x_6722 = x_6715; +} else { + lean_dec_ref(x_6715); + x_6722 = lean_box(0); +} +x_6723 = lean_ctor_get(x_6720, 1); +lean_inc(x_6723); +x_6724 = lean_nat_dec_eq(x_6709, x_6723); +lean_dec(x_6709); +if (x_6724 == 0) +{ +lean_object* x_6725; +lean_dec(x_6723); +if (lean_is_scalar(x_6722)) { + x_6725 = lean_alloc_ctor(1, 2, 0); +} else { + x_6725 = x_6722; +} +lean_ctor_set(x_6725, 0, x_6720); +lean_ctor_set(x_6725, 1, x_6721); +return x_6725; +} +else +{ +lean_object* x_6726; lean_object* x_6727; lean_object* x_6728; lean_object* x_6729; +lean_dec(x_6722); +lean_dec(x_6721); +x_6726 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_6727 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6727, 0, x_6726); +x_6728 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_6729 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6728, x_6727, x_6720); +if (lean_obj_tag(x_6729) == 0) +{ +lean_object* x_6730; lean_object* x_6731; lean_object* x_6732; lean_object* x_6733; +lean_dec(x_6723); +x_6730 = lean_ctor_get(x_6729, 0); +lean_inc(x_6730); +x_6731 = lean_ctor_get(x_6729, 1); +lean_inc(x_6731); +if (lean_is_exclusive(x_6729)) { + lean_ctor_release(x_6729, 0); + lean_ctor_release(x_6729, 1); + x_6732 = x_6729; +} else { + lean_dec_ref(x_6729); + x_6732 = lean_box(0); +} +if (lean_is_scalar(x_6732)) { + x_6733 = lean_alloc_ctor(0, 2, 0); +} else { + x_6733 = x_6732; +} +lean_ctor_set(x_6733, 0, x_6730); +lean_ctor_set(x_6733, 1, x_6731); +return x_6733; +} +else +{ +lean_object* x_6734; lean_object* x_6735; lean_object* x_6736; lean_object* x_6737; uint8_t x_6738; +x_6734 = lean_ctor_get(x_6729, 0); +lean_inc(x_6734); +x_6735 = lean_ctor_get(x_6729, 1); +lean_inc(x_6735); +if (lean_is_exclusive(x_6729)) { + lean_ctor_release(x_6729, 0); + lean_ctor_release(x_6729, 1); + x_6736 = x_6729; +} else { + lean_dec_ref(x_6729); + x_6736 = lean_box(0); +} +x_6737 = lean_ctor_get(x_6734, 1); +lean_inc(x_6737); +x_6738 = lean_nat_dec_eq(x_6723, x_6737); +lean_dec(x_6737); +lean_dec(x_6723); +if (x_6738 == 0) +{ +lean_object* x_6739; +if (lean_is_scalar(x_6736)) { + x_6739 = lean_alloc_ctor(1, 2, 0); +} else { + x_6739 = x_6736; +} +lean_ctor_set(x_6739, 0, x_6734); +lean_ctor_set(x_6739, 1, x_6735); +return x_6739; +} +else +{ +lean_object* x_6740; lean_object* x_6741; lean_object* x_6742; lean_object* x_6743; +lean_dec(x_6736); +lean_dec(x_6735); +x_6740 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_6741 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6741, 0, x_6740); +x_6742 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_6743 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6742, x_6741, x_6734); +return x_6743; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_6744; lean_object* x_6745; lean_object* x_6746; lean_object* x_6747; uint8_t x_6748; +x_6744 = lean_ctor_get(x_5, 0); +x_6745 = lean_ctor_get(x_5, 1); +lean_inc(x_6745); +lean_inc(x_6744); +lean_dec(x_5); +x_6746 = lean_ctor_get(x_1, 1); +lean_inc(x_6746); +lean_dec(x_1); +x_6747 = lean_ctor_get(x_6744, 1); +lean_inc(x_6747); +x_6748 = lean_nat_dec_eq(x_6746, x_6747); +lean_dec(x_6746); +if (x_6748 == 0) +{ +lean_object* x_6749; +lean_dec(x_6747); +x_6749 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6749, 0, x_6744); +lean_ctor_set(x_6749, 1, x_6745); +return x_6749; +} +else +{ +lean_object* x_6750; lean_object* x_6751; lean_object* x_6752; lean_object* x_6753; +lean_dec(x_6745); +x_6750 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6; +x_6751 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6751, 0, x_6750); +x_6752 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5; +x_6753 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6752, x_6751, x_6744); +if (lean_obj_tag(x_6753) == 0) +{ +lean_object* x_6754; lean_object* x_6755; lean_object* x_6756; lean_object* x_6757; +lean_dec(x_6747); +x_6754 = lean_ctor_get(x_6753, 0); +lean_inc(x_6754); +x_6755 = lean_ctor_get(x_6753, 1); +lean_inc(x_6755); +if (lean_is_exclusive(x_6753)) { + lean_ctor_release(x_6753, 0); + lean_ctor_release(x_6753, 1); + x_6756 = x_6753; +} else { + lean_dec_ref(x_6753); + x_6756 = lean_box(0); +} +if (lean_is_scalar(x_6756)) { + x_6757 = lean_alloc_ctor(0, 2, 0); +} else { + x_6757 = x_6756; +} +lean_ctor_set(x_6757, 0, x_6754); +lean_ctor_set(x_6757, 1, x_6755); +return x_6757; +} +else +{ +lean_object* x_6758; lean_object* x_6759; lean_object* x_6760; lean_object* x_6761; uint8_t x_6762; +x_6758 = lean_ctor_get(x_6753, 0); +lean_inc(x_6758); +x_6759 = lean_ctor_get(x_6753, 1); +lean_inc(x_6759); +if (lean_is_exclusive(x_6753)) { + lean_ctor_release(x_6753, 0); + lean_ctor_release(x_6753, 1); + x_6760 = x_6753; +} else { + lean_dec_ref(x_6753); + x_6760 = lean_box(0); +} +x_6761 = lean_ctor_get(x_6758, 1); +lean_inc(x_6761); +x_6762 = lean_nat_dec_eq(x_6747, x_6761); +lean_dec(x_6747); +if (x_6762 == 0) +{ +lean_object* x_6763; +lean_dec(x_6761); +if (lean_is_scalar(x_6760)) { + x_6763 = lean_alloc_ctor(1, 2, 0); +} else { + x_6763 = x_6760; +} +lean_ctor_set(x_6763, 0, x_6758); +lean_ctor_set(x_6763, 1, x_6759); +return x_6763; +} +else +{ +lean_object* x_6764; lean_object* x_6765; lean_object* x_6766; lean_object* x_6767; +lean_dec(x_6760); +lean_dec(x_6759); +x_6764 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9; +x_6765 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6765, 0, x_6764); +x_6766 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8; +x_6767 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6766, x_6765, x_6758); +if (lean_obj_tag(x_6767) == 0) +{ +lean_object* x_6768; lean_object* x_6769; lean_object* x_6770; lean_object* x_6771; +lean_dec(x_6761); +x_6768 = lean_ctor_get(x_6767, 0); +lean_inc(x_6768); +x_6769 = lean_ctor_get(x_6767, 1); +lean_inc(x_6769); +if (lean_is_exclusive(x_6767)) { + lean_ctor_release(x_6767, 0); + lean_ctor_release(x_6767, 1); + x_6770 = x_6767; +} else { + lean_dec_ref(x_6767); + x_6770 = lean_box(0); +} +if (lean_is_scalar(x_6770)) { + x_6771 = lean_alloc_ctor(0, 2, 0); +} else { + x_6771 = x_6770; +} +lean_ctor_set(x_6771, 0, x_6768); +lean_ctor_set(x_6771, 1, x_6769); +return x_6771; +} +else +{ +lean_object* x_6772; lean_object* x_6773; lean_object* x_6774; lean_object* x_6775; uint8_t x_6776; +x_6772 = lean_ctor_get(x_6767, 0); +lean_inc(x_6772); +x_6773 = lean_ctor_get(x_6767, 1); +lean_inc(x_6773); +if (lean_is_exclusive(x_6767)) { + lean_ctor_release(x_6767, 0); + lean_ctor_release(x_6767, 1); + x_6774 = x_6767; +} else { + lean_dec_ref(x_6767); + x_6774 = lean_box(0); +} +x_6775 = lean_ctor_get(x_6772, 1); +lean_inc(x_6775); +x_6776 = lean_nat_dec_eq(x_6761, x_6775); +lean_dec(x_6761); +if (x_6776 == 0) +{ +lean_object* x_6777; +lean_dec(x_6775); +if (lean_is_scalar(x_6774)) { + x_6777 = lean_alloc_ctor(1, 2, 0); +} else { + x_6777 = x_6774; +} +lean_ctor_set(x_6777, 0, x_6772); +lean_ctor_set(x_6777, 1, x_6773); +return x_6777; +} +else +{ +lean_object* x_6778; lean_object* x_6779; lean_object* x_6780; lean_object* x_6781; +lean_dec(x_6774); +lean_dec(x_6773); +x_6778 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12; +x_6779 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6779, 0, x_6778); +x_6780 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11; +x_6781 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6780, x_6779, x_6772); +if (lean_obj_tag(x_6781) == 0) +{ +lean_object* x_6782; lean_object* x_6783; lean_object* x_6784; lean_object* x_6785; +lean_dec(x_6775); +x_6782 = lean_ctor_get(x_6781, 0); +lean_inc(x_6782); +x_6783 = lean_ctor_get(x_6781, 1); +lean_inc(x_6783); +if (lean_is_exclusive(x_6781)) { + lean_ctor_release(x_6781, 0); + lean_ctor_release(x_6781, 1); + x_6784 = x_6781; +} else { + lean_dec_ref(x_6781); + x_6784 = lean_box(0); +} +if (lean_is_scalar(x_6784)) { + x_6785 = lean_alloc_ctor(0, 2, 0); +} else { + x_6785 = x_6784; +} +lean_ctor_set(x_6785, 0, x_6782); +lean_ctor_set(x_6785, 1, x_6783); +return x_6785; +} +else +{ +lean_object* x_6786; lean_object* x_6787; lean_object* x_6788; lean_object* x_6789; uint8_t x_6790; +x_6786 = lean_ctor_get(x_6781, 0); +lean_inc(x_6786); +x_6787 = lean_ctor_get(x_6781, 1); +lean_inc(x_6787); +if (lean_is_exclusive(x_6781)) { + lean_ctor_release(x_6781, 0); + lean_ctor_release(x_6781, 1); + x_6788 = x_6781; +} else { + lean_dec_ref(x_6781); + x_6788 = lean_box(0); +} +x_6789 = lean_ctor_get(x_6786, 1); +lean_inc(x_6789); +x_6790 = lean_nat_dec_eq(x_6775, x_6789); +lean_dec(x_6775); +if (x_6790 == 0) +{ +lean_object* x_6791; +lean_dec(x_6789); +if (lean_is_scalar(x_6788)) { + x_6791 = lean_alloc_ctor(1, 2, 0); +} else { + x_6791 = x_6788; +} +lean_ctor_set(x_6791, 0, x_6786); +lean_ctor_set(x_6791, 1, x_6787); +return x_6791; +} +else +{ +lean_object* x_6792; lean_object* x_6793; lean_object* x_6794; lean_object* x_6795; +lean_dec(x_6788); +lean_dec(x_6787); +x_6792 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15; +x_6793 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6793, 0, x_6792); +x_6794 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14; +x_6795 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6794, x_6793, x_6786); +if (lean_obj_tag(x_6795) == 0) +{ +lean_object* x_6796; lean_object* x_6797; lean_object* x_6798; lean_object* x_6799; +lean_dec(x_6789); +x_6796 = lean_ctor_get(x_6795, 0); +lean_inc(x_6796); +x_6797 = lean_ctor_get(x_6795, 1); +lean_inc(x_6797); +if (lean_is_exclusive(x_6795)) { + lean_ctor_release(x_6795, 0); + lean_ctor_release(x_6795, 1); + x_6798 = x_6795; +} else { + lean_dec_ref(x_6795); + x_6798 = lean_box(0); +} +if (lean_is_scalar(x_6798)) { + x_6799 = lean_alloc_ctor(0, 2, 0); +} else { + x_6799 = x_6798; +} +lean_ctor_set(x_6799, 0, x_6796); +lean_ctor_set(x_6799, 1, x_6797); +return x_6799; +} +else +{ +lean_object* x_6800; lean_object* x_6801; lean_object* x_6802; lean_object* x_6803; uint8_t x_6804; +x_6800 = lean_ctor_get(x_6795, 0); +lean_inc(x_6800); +x_6801 = lean_ctor_get(x_6795, 1); +lean_inc(x_6801); +if (lean_is_exclusive(x_6795)) { + lean_ctor_release(x_6795, 0); + lean_ctor_release(x_6795, 1); + x_6802 = x_6795; +} else { + lean_dec_ref(x_6795); + x_6802 = lean_box(0); +} +x_6803 = lean_ctor_get(x_6800, 1); +lean_inc(x_6803); +x_6804 = lean_nat_dec_eq(x_6789, x_6803); +lean_dec(x_6789); +if (x_6804 == 0) +{ +lean_object* x_6805; +lean_dec(x_6803); +if (lean_is_scalar(x_6802)) { + x_6805 = lean_alloc_ctor(1, 2, 0); +} else { + x_6805 = x_6802; +} +lean_ctor_set(x_6805, 0, x_6800); +lean_ctor_set(x_6805, 1, x_6801); +return x_6805; +} +else +{ +lean_object* x_6806; lean_object* x_6807; lean_object* x_6808; +lean_dec(x_6802); +lean_dec(x_6801); +x_6806 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16; +x_6807 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6807, 0, x_6806); +x_6808 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6794, x_6807, x_6800); +if (lean_obj_tag(x_6808) == 0) +{ +lean_object* x_6809; lean_object* x_6810; lean_object* x_6811; lean_object* x_6812; +lean_dec(x_6803); +x_6809 = lean_ctor_get(x_6808, 0); +lean_inc(x_6809); +x_6810 = lean_ctor_get(x_6808, 1); +lean_inc(x_6810); +if (lean_is_exclusive(x_6808)) { + lean_ctor_release(x_6808, 0); + lean_ctor_release(x_6808, 1); + x_6811 = x_6808; +} else { + lean_dec_ref(x_6808); + x_6811 = lean_box(0); +} +if (lean_is_scalar(x_6811)) { + x_6812 = lean_alloc_ctor(0, 2, 0); +} else { + x_6812 = x_6811; +} +lean_ctor_set(x_6812, 0, x_6809); +lean_ctor_set(x_6812, 1, x_6810); +return x_6812; +} +else +{ +lean_object* x_6813; lean_object* x_6814; lean_object* x_6815; lean_object* x_6816; uint8_t x_6817; +x_6813 = lean_ctor_get(x_6808, 0); +lean_inc(x_6813); +x_6814 = lean_ctor_get(x_6808, 1); +lean_inc(x_6814); +if (lean_is_exclusive(x_6808)) { + lean_ctor_release(x_6808, 0); + lean_ctor_release(x_6808, 1); + x_6815 = x_6808; +} else { + lean_dec_ref(x_6808); + x_6815 = lean_box(0); +} +x_6816 = lean_ctor_get(x_6813, 1); +lean_inc(x_6816); +x_6817 = lean_nat_dec_eq(x_6803, x_6816); +lean_dec(x_6803); +if (x_6817 == 0) +{ +lean_object* x_6818; +lean_dec(x_6816); +if (lean_is_scalar(x_6815)) { + x_6818 = lean_alloc_ctor(1, 2, 0); +} else { + x_6818 = x_6815; +} +lean_ctor_set(x_6818, 0, x_6813); +lean_ctor_set(x_6818, 1, x_6814); +return x_6818; +} +else +{ +lean_object* x_6819; lean_object* x_6820; lean_object* x_6821; lean_object* x_6822; +lean_dec(x_6815); +lean_dec(x_6814); +x_6819 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19; +x_6820 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6820, 0, x_6819); +x_6821 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18; +x_6822 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6821, x_6820, x_6813); +if (lean_obj_tag(x_6822) == 0) +{ +lean_object* x_6823; lean_object* x_6824; lean_object* x_6825; lean_object* x_6826; +lean_dec(x_6816); +x_6823 = lean_ctor_get(x_6822, 0); +lean_inc(x_6823); +x_6824 = lean_ctor_get(x_6822, 1); +lean_inc(x_6824); +if (lean_is_exclusive(x_6822)) { + lean_ctor_release(x_6822, 0); + lean_ctor_release(x_6822, 1); + x_6825 = x_6822; +} else { + lean_dec_ref(x_6822); + x_6825 = lean_box(0); +} +if (lean_is_scalar(x_6825)) { + x_6826 = lean_alloc_ctor(0, 2, 0); +} else { + x_6826 = x_6825; +} +lean_ctor_set(x_6826, 0, x_6823); +lean_ctor_set(x_6826, 1, x_6824); +return x_6826; +} +else +{ +lean_object* x_6827; lean_object* x_6828; lean_object* x_6829; lean_object* x_6830; uint8_t x_6831; +x_6827 = lean_ctor_get(x_6822, 0); +lean_inc(x_6827); +x_6828 = lean_ctor_get(x_6822, 1); +lean_inc(x_6828); +if (lean_is_exclusive(x_6822)) { + lean_ctor_release(x_6822, 0); + lean_ctor_release(x_6822, 1); + x_6829 = x_6822; +} else { + lean_dec_ref(x_6822); + x_6829 = lean_box(0); +} +x_6830 = lean_ctor_get(x_6827, 1); +lean_inc(x_6830); +x_6831 = lean_nat_dec_eq(x_6816, x_6830); +lean_dec(x_6816); +if (x_6831 == 0) +{ +lean_object* x_6832; +lean_dec(x_6830); +if (lean_is_scalar(x_6829)) { + x_6832 = lean_alloc_ctor(1, 2, 0); +} else { + x_6832 = x_6829; +} +lean_ctor_set(x_6832, 0, x_6827); +lean_ctor_set(x_6832, 1, x_6828); +return x_6832; +} +else +{ +lean_object* x_6833; lean_object* x_6834; lean_object* x_6835; lean_object* x_6836; +lean_dec(x_6829); +lean_dec(x_6828); +x_6833 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22; +x_6834 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6834, 0, x_6833); +x_6835 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21; +x_6836 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6835, x_6834, x_6827); +if (lean_obj_tag(x_6836) == 0) +{ +lean_object* x_6837; lean_object* x_6838; lean_object* x_6839; lean_object* x_6840; +lean_dec(x_6830); +x_6837 = lean_ctor_get(x_6836, 0); +lean_inc(x_6837); +x_6838 = lean_ctor_get(x_6836, 1); +lean_inc(x_6838); +if (lean_is_exclusive(x_6836)) { + lean_ctor_release(x_6836, 0); + lean_ctor_release(x_6836, 1); + x_6839 = x_6836; +} else { + lean_dec_ref(x_6836); + x_6839 = lean_box(0); +} +if (lean_is_scalar(x_6839)) { + x_6840 = lean_alloc_ctor(0, 2, 0); +} else { + x_6840 = x_6839; +} +lean_ctor_set(x_6840, 0, x_6837); +lean_ctor_set(x_6840, 1, x_6838); +return x_6840; +} +else +{ +lean_object* x_6841; lean_object* x_6842; lean_object* x_6843; lean_object* x_6844; uint8_t x_6845; +x_6841 = lean_ctor_get(x_6836, 0); +lean_inc(x_6841); +x_6842 = lean_ctor_get(x_6836, 1); +lean_inc(x_6842); +if (lean_is_exclusive(x_6836)) { + lean_ctor_release(x_6836, 0); + lean_ctor_release(x_6836, 1); + x_6843 = x_6836; +} else { + lean_dec_ref(x_6836); + x_6843 = lean_box(0); +} +x_6844 = lean_ctor_get(x_6841, 1); +lean_inc(x_6844); +x_6845 = lean_nat_dec_eq(x_6830, x_6844); +lean_dec(x_6830); +if (x_6845 == 0) +{ +lean_object* x_6846; +lean_dec(x_6844); +if (lean_is_scalar(x_6843)) { + x_6846 = lean_alloc_ctor(1, 2, 0); +} else { + x_6846 = x_6843; +} +lean_ctor_set(x_6846, 0, x_6841); +lean_ctor_set(x_6846, 1, x_6842); +return x_6846; +} +else +{ +lean_object* x_6847; lean_object* x_6848; lean_object* x_6849; +lean_dec(x_6843); +lean_dec(x_6842); +x_6847 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23; +x_6848 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6848, 0, x_6847); +x_6849 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6835, x_6848, x_6841); +if (lean_obj_tag(x_6849) == 0) +{ +lean_object* x_6850; lean_object* x_6851; lean_object* x_6852; lean_object* x_6853; +lean_dec(x_6844); +x_6850 = lean_ctor_get(x_6849, 0); +lean_inc(x_6850); +x_6851 = lean_ctor_get(x_6849, 1); +lean_inc(x_6851); +if (lean_is_exclusive(x_6849)) { + lean_ctor_release(x_6849, 0); + lean_ctor_release(x_6849, 1); + x_6852 = x_6849; +} else { + lean_dec_ref(x_6849); + x_6852 = lean_box(0); +} +if (lean_is_scalar(x_6852)) { + x_6853 = lean_alloc_ctor(0, 2, 0); +} else { + x_6853 = x_6852; +} +lean_ctor_set(x_6853, 0, x_6850); +lean_ctor_set(x_6853, 1, x_6851); +return x_6853; +} +else +{ +lean_object* x_6854; lean_object* x_6855; lean_object* x_6856; lean_object* x_6857; uint8_t x_6858; +x_6854 = lean_ctor_get(x_6849, 0); +lean_inc(x_6854); +x_6855 = lean_ctor_get(x_6849, 1); +lean_inc(x_6855); +if (lean_is_exclusive(x_6849)) { + lean_ctor_release(x_6849, 0); + lean_ctor_release(x_6849, 1); + x_6856 = x_6849; +} else { + lean_dec_ref(x_6849); + x_6856 = lean_box(0); +} +x_6857 = lean_ctor_get(x_6854, 1); +lean_inc(x_6857); +x_6858 = lean_nat_dec_eq(x_6844, x_6857); +lean_dec(x_6844); +if (x_6858 == 0) +{ +lean_object* x_6859; +lean_dec(x_6857); +if (lean_is_scalar(x_6856)) { + x_6859 = lean_alloc_ctor(1, 2, 0); +} else { + x_6859 = x_6856; +} +lean_ctor_set(x_6859, 0, x_6854); +lean_ctor_set(x_6859, 1, x_6855); +return x_6859; +} +else +{ +lean_object* x_6860; lean_object* x_6861; lean_object* x_6862; lean_object* x_6863; +lean_dec(x_6856); +lean_dec(x_6855); +x_6860 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26; +x_6861 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6861, 0, x_6860); +x_6862 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25; +x_6863 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6862, x_6861, x_6854); +if (lean_obj_tag(x_6863) == 0) +{ +lean_object* x_6864; lean_object* x_6865; lean_object* x_6866; lean_object* x_6867; +lean_dec(x_6857); +x_6864 = lean_ctor_get(x_6863, 0); +lean_inc(x_6864); +x_6865 = lean_ctor_get(x_6863, 1); +lean_inc(x_6865); +if (lean_is_exclusive(x_6863)) { + lean_ctor_release(x_6863, 0); + lean_ctor_release(x_6863, 1); + x_6866 = x_6863; +} else { + lean_dec_ref(x_6863); + x_6866 = lean_box(0); +} +if (lean_is_scalar(x_6866)) { + x_6867 = lean_alloc_ctor(0, 2, 0); +} else { + x_6867 = x_6866; +} +lean_ctor_set(x_6867, 0, x_6864); +lean_ctor_set(x_6867, 1, x_6865); +return x_6867; +} +else +{ +lean_object* x_6868; lean_object* x_6869; lean_object* x_6870; lean_object* x_6871; uint8_t x_6872; +x_6868 = lean_ctor_get(x_6863, 0); +lean_inc(x_6868); +x_6869 = lean_ctor_get(x_6863, 1); +lean_inc(x_6869); +if (lean_is_exclusive(x_6863)) { + lean_ctor_release(x_6863, 0); + lean_ctor_release(x_6863, 1); + x_6870 = x_6863; +} else { + lean_dec_ref(x_6863); + x_6870 = lean_box(0); +} +x_6871 = lean_ctor_get(x_6868, 1); +lean_inc(x_6871); +x_6872 = lean_nat_dec_eq(x_6857, x_6871); +lean_dec(x_6857); +if (x_6872 == 0) +{ +lean_object* x_6873; +lean_dec(x_6871); +if (lean_is_scalar(x_6870)) { + x_6873 = lean_alloc_ctor(1, 2, 0); +} else { + x_6873 = x_6870; +} +lean_ctor_set(x_6873, 0, x_6868); +lean_ctor_set(x_6873, 1, x_6869); +return x_6873; +} +else +{ +lean_object* x_6874; lean_object* x_6875; lean_object* x_6876; lean_object* x_6877; +lean_dec(x_6870); +lean_dec(x_6869); +x_6874 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29; +x_6875 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6875, 0, x_6874); +x_6876 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28; +x_6877 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6876, x_6875, x_6868); +if (lean_obj_tag(x_6877) == 0) +{ +lean_object* x_6878; lean_object* x_6879; lean_object* x_6880; lean_object* x_6881; +lean_dec(x_6871); +x_6878 = lean_ctor_get(x_6877, 0); +lean_inc(x_6878); +x_6879 = lean_ctor_get(x_6877, 1); +lean_inc(x_6879); +if (lean_is_exclusive(x_6877)) { + lean_ctor_release(x_6877, 0); + lean_ctor_release(x_6877, 1); + x_6880 = x_6877; +} else { + lean_dec_ref(x_6877); + x_6880 = lean_box(0); +} +if (lean_is_scalar(x_6880)) { + x_6881 = lean_alloc_ctor(0, 2, 0); +} else { + x_6881 = x_6880; +} +lean_ctor_set(x_6881, 0, x_6878); +lean_ctor_set(x_6881, 1, x_6879); +return x_6881; +} +else +{ +lean_object* x_6882; lean_object* x_6883; lean_object* x_6884; lean_object* x_6885; uint8_t x_6886; +x_6882 = lean_ctor_get(x_6877, 0); +lean_inc(x_6882); +x_6883 = lean_ctor_get(x_6877, 1); +lean_inc(x_6883); +if (lean_is_exclusive(x_6877)) { + lean_ctor_release(x_6877, 0); + lean_ctor_release(x_6877, 1); + x_6884 = x_6877; +} else { + lean_dec_ref(x_6877); + x_6884 = lean_box(0); +} +x_6885 = lean_ctor_get(x_6882, 1); +lean_inc(x_6885); +x_6886 = lean_nat_dec_eq(x_6871, x_6885); +lean_dec(x_6871); +if (x_6886 == 0) +{ +lean_object* x_6887; +lean_dec(x_6885); +if (lean_is_scalar(x_6884)) { + x_6887 = lean_alloc_ctor(1, 2, 0); +} else { + x_6887 = x_6884; +} +lean_ctor_set(x_6887, 0, x_6882); +lean_ctor_set(x_6887, 1, x_6883); +return x_6887; +} +else +{ +lean_object* x_6888; lean_object* x_6889; lean_object* x_6890; lean_object* x_6891; +lean_dec(x_6884); +lean_dec(x_6883); +x_6888 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32; +x_6889 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6889, 0, x_6888); +x_6890 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31; +x_6891 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6890, x_6889, x_6882); +if (lean_obj_tag(x_6891) == 0) +{ +lean_object* x_6892; lean_object* x_6893; lean_object* x_6894; lean_object* x_6895; +lean_dec(x_6885); +x_6892 = lean_ctor_get(x_6891, 0); +lean_inc(x_6892); +x_6893 = lean_ctor_get(x_6891, 1); +lean_inc(x_6893); +if (lean_is_exclusive(x_6891)) { + lean_ctor_release(x_6891, 0); + lean_ctor_release(x_6891, 1); + x_6894 = x_6891; +} else { + lean_dec_ref(x_6891); + x_6894 = lean_box(0); +} +if (lean_is_scalar(x_6894)) { + x_6895 = lean_alloc_ctor(0, 2, 0); +} else { + x_6895 = x_6894; +} +lean_ctor_set(x_6895, 0, x_6892); +lean_ctor_set(x_6895, 1, x_6893); +return x_6895; +} +else +{ +lean_object* x_6896; lean_object* x_6897; lean_object* x_6898; lean_object* x_6899; uint8_t x_6900; +x_6896 = lean_ctor_get(x_6891, 0); +lean_inc(x_6896); +x_6897 = lean_ctor_get(x_6891, 1); +lean_inc(x_6897); +if (lean_is_exclusive(x_6891)) { + lean_ctor_release(x_6891, 0); + lean_ctor_release(x_6891, 1); + x_6898 = x_6891; +} else { + lean_dec_ref(x_6891); + x_6898 = lean_box(0); +} +x_6899 = lean_ctor_get(x_6896, 1); +lean_inc(x_6899); +x_6900 = lean_nat_dec_eq(x_6885, x_6899); +lean_dec(x_6885); +if (x_6900 == 0) +{ +lean_object* x_6901; +lean_dec(x_6899); +if (lean_is_scalar(x_6898)) { + x_6901 = lean_alloc_ctor(1, 2, 0); +} else { + x_6901 = x_6898; +} +lean_ctor_set(x_6901, 0, x_6896); +lean_ctor_set(x_6901, 1, x_6897); +return x_6901; +} +else +{ +lean_object* x_6902; lean_object* x_6903; lean_object* x_6904; lean_object* x_6905; +lean_dec(x_6898); +lean_dec(x_6897); +x_6902 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35; +x_6903 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6903, 0, x_6902); +x_6904 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34; +x_6905 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6904, x_6903, x_6896); +if (lean_obj_tag(x_6905) == 0) +{ +lean_object* x_6906; lean_object* x_6907; lean_object* x_6908; lean_object* x_6909; +lean_dec(x_6899); +x_6906 = lean_ctor_get(x_6905, 0); +lean_inc(x_6906); +x_6907 = lean_ctor_get(x_6905, 1); +lean_inc(x_6907); +if (lean_is_exclusive(x_6905)) { + lean_ctor_release(x_6905, 0); + lean_ctor_release(x_6905, 1); + x_6908 = x_6905; +} else { + lean_dec_ref(x_6905); + x_6908 = lean_box(0); +} +if (lean_is_scalar(x_6908)) { + x_6909 = lean_alloc_ctor(0, 2, 0); +} else { + x_6909 = x_6908; +} +lean_ctor_set(x_6909, 0, x_6906); +lean_ctor_set(x_6909, 1, x_6907); +return x_6909; +} +else +{ +lean_object* x_6910; lean_object* x_6911; lean_object* x_6912; lean_object* x_6913; uint8_t x_6914; +x_6910 = lean_ctor_get(x_6905, 0); +lean_inc(x_6910); +x_6911 = lean_ctor_get(x_6905, 1); +lean_inc(x_6911); +if (lean_is_exclusive(x_6905)) { + lean_ctor_release(x_6905, 0); + lean_ctor_release(x_6905, 1); + x_6912 = x_6905; +} else { + lean_dec_ref(x_6905); + x_6912 = lean_box(0); +} +x_6913 = lean_ctor_get(x_6910, 1); +lean_inc(x_6913); +x_6914 = lean_nat_dec_eq(x_6899, x_6913); +lean_dec(x_6899); +if (x_6914 == 0) +{ +lean_object* x_6915; +lean_dec(x_6913); +if (lean_is_scalar(x_6912)) { + x_6915 = lean_alloc_ctor(1, 2, 0); +} else { + x_6915 = x_6912; +} +lean_ctor_set(x_6915, 0, x_6910); +lean_ctor_set(x_6915, 1, x_6911); +return x_6915; +} +else +{ +lean_object* x_6916; lean_object* x_6917; lean_object* x_6918; +lean_dec(x_6912); +lean_dec(x_6911); +x_6916 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36; +x_6917 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6917, 0, x_6916); +x_6918 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6904, x_6917, x_6910); +if (lean_obj_tag(x_6918) == 0) +{ +lean_object* x_6919; lean_object* x_6920; lean_object* x_6921; lean_object* x_6922; +lean_dec(x_6913); +x_6919 = lean_ctor_get(x_6918, 0); +lean_inc(x_6919); +x_6920 = lean_ctor_get(x_6918, 1); +lean_inc(x_6920); +if (lean_is_exclusive(x_6918)) { + lean_ctor_release(x_6918, 0); + lean_ctor_release(x_6918, 1); + x_6921 = x_6918; +} else { + lean_dec_ref(x_6918); + x_6921 = lean_box(0); +} +if (lean_is_scalar(x_6921)) { + x_6922 = lean_alloc_ctor(0, 2, 0); +} else { + x_6922 = x_6921; +} +lean_ctor_set(x_6922, 0, x_6919); +lean_ctor_set(x_6922, 1, x_6920); +return x_6922; +} +else +{ +lean_object* x_6923; lean_object* x_6924; lean_object* x_6925; lean_object* x_6926; uint8_t x_6927; +x_6923 = lean_ctor_get(x_6918, 0); +lean_inc(x_6923); +x_6924 = lean_ctor_get(x_6918, 1); +lean_inc(x_6924); +if (lean_is_exclusive(x_6918)) { + lean_ctor_release(x_6918, 0); + lean_ctor_release(x_6918, 1); + x_6925 = x_6918; +} else { + lean_dec_ref(x_6918); + x_6925 = lean_box(0); +} +x_6926 = lean_ctor_get(x_6923, 1); +lean_inc(x_6926); +x_6927 = lean_nat_dec_eq(x_6913, x_6926); +lean_dec(x_6913); +if (x_6927 == 0) +{ +lean_object* x_6928; +lean_dec(x_6926); +if (lean_is_scalar(x_6925)) { + x_6928 = lean_alloc_ctor(1, 2, 0); +} else { + x_6928 = x_6925; +} +lean_ctor_set(x_6928, 0, x_6923); +lean_ctor_set(x_6928, 1, x_6924); +return x_6928; +} +else +{ +lean_object* x_6929; lean_object* x_6930; lean_object* x_6931; lean_object* x_6932; +lean_dec(x_6925); +lean_dec(x_6924); +x_6929 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39; +x_6930 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6930, 0, x_6929); +x_6931 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38; +x_6932 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6931, x_6930, x_6923); +if (lean_obj_tag(x_6932) == 0) +{ +lean_object* x_6933; lean_object* x_6934; lean_object* x_6935; lean_object* x_6936; +lean_dec(x_6926); +x_6933 = lean_ctor_get(x_6932, 0); +lean_inc(x_6933); +x_6934 = lean_ctor_get(x_6932, 1); +lean_inc(x_6934); +if (lean_is_exclusive(x_6932)) { + lean_ctor_release(x_6932, 0); + lean_ctor_release(x_6932, 1); + x_6935 = x_6932; +} else { + lean_dec_ref(x_6932); + x_6935 = lean_box(0); +} +if (lean_is_scalar(x_6935)) { + x_6936 = lean_alloc_ctor(0, 2, 0); +} else { + x_6936 = x_6935; +} +lean_ctor_set(x_6936, 0, x_6933); +lean_ctor_set(x_6936, 1, x_6934); +return x_6936; +} +else +{ +lean_object* x_6937; lean_object* x_6938; lean_object* x_6939; lean_object* x_6940; uint8_t x_6941; +x_6937 = lean_ctor_get(x_6932, 0); +lean_inc(x_6937); +x_6938 = lean_ctor_get(x_6932, 1); +lean_inc(x_6938); +if (lean_is_exclusive(x_6932)) { + lean_ctor_release(x_6932, 0); + lean_ctor_release(x_6932, 1); + x_6939 = x_6932; +} else { + lean_dec_ref(x_6932); + x_6939 = lean_box(0); +} +x_6940 = lean_ctor_get(x_6937, 1); +lean_inc(x_6940); +x_6941 = lean_nat_dec_eq(x_6926, x_6940); +lean_dec(x_6926); +if (x_6941 == 0) +{ +lean_object* x_6942; +lean_dec(x_6940); +if (lean_is_scalar(x_6939)) { + x_6942 = lean_alloc_ctor(1, 2, 0); +} else { + x_6942 = x_6939; +} +lean_ctor_set(x_6942, 0, x_6937); +lean_ctor_set(x_6942, 1, x_6938); +return x_6942; +} +else +{ +lean_object* x_6943; lean_object* x_6944; lean_object* x_6945; lean_object* x_6946; +lean_dec(x_6939); +lean_dec(x_6938); +x_6943 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42; +x_6944 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6944, 0, x_6943); +x_6945 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41; +x_6946 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6945, x_6944, x_6937); +if (lean_obj_tag(x_6946) == 0) +{ +lean_object* x_6947; lean_object* x_6948; lean_object* x_6949; lean_object* x_6950; +lean_dec(x_6940); +x_6947 = lean_ctor_get(x_6946, 0); +lean_inc(x_6947); +x_6948 = lean_ctor_get(x_6946, 1); +lean_inc(x_6948); +if (lean_is_exclusive(x_6946)) { + lean_ctor_release(x_6946, 0); + lean_ctor_release(x_6946, 1); + x_6949 = x_6946; +} else { + lean_dec_ref(x_6946); + x_6949 = lean_box(0); +} +if (lean_is_scalar(x_6949)) { + x_6950 = lean_alloc_ctor(0, 2, 0); +} else { + x_6950 = x_6949; +} +lean_ctor_set(x_6950, 0, x_6947); +lean_ctor_set(x_6950, 1, x_6948); +return x_6950; +} +else +{ +lean_object* x_6951; lean_object* x_6952; lean_object* x_6953; lean_object* x_6954; uint8_t x_6955; +x_6951 = lean_ctor_get(x_6946, 0); +lean_inc(x_6951); +x_6952 = lean_ctor_get(x_6946, 1); +lean_inc(x_6952); +if (lean_is_exclusive(x_6946)) { + lean_ctor_release(x_6946, 0); + lean_ctor_release(x_6946, 1); + x_6953 = x_6946; +} else { + lean_dec_ref(x_6946); + x_6953 = lean_box(0); +} +x_6954 = lean_ctor_get(x_6951, 1); +lean_inc(x_6954); +x_6955 = lean_nat_dec_eq(x_6940, x_6954); +lean_dec(x_6940); +if (x_6955 == 0) +{ +lean_object* x_6956; +lean_dec(x_6954); +if (lean_is_scalar(x_6953)) { + x_6956 = lean_alloc_ctor(1, 2, 0); +} else { + x_6956 = x_6953; +} +lean_ctor_set(x_6956, 0, x_6951); +lean_ctor_set(x_6956, 1, x_6952); +return x_6956; +} +else +{ +lean_object* x_6957; lean_object* x_6958; lean_object* x_6959; lean_object* x_6960; +lean_dec(x_6953); +lean_dec(x_6952); +x_6957 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45; +x_6958 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6958, 0, x_6957); +x_6959 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44; +x_6960 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6959, x_6958, x_6951); +if (lean_obj_tag(x_6960) == 0) +{ +lean_object* x_6961; lean_object* x_6962; lean_object* x_6963; lean_object* x_6964; +lean_dec(x_6954); +x_6961 = lean_ctor_get(x_6960, 0); +lean_inc(x_6961); +x_6962 = lean_ctor_get(x_6960, 1); +lean_inc(x_6962); +if (lean_is_exclusive(x_6960)) { + lean_ctor_release(x_6960, 0); + lean_ctor_release(x_6960, 1); + x_6963 = x_6960; +} else { + lean_dec_ref(x_6960); + x_6963 = lean_box(0); +} +if (lean_is_scalar(x_6963)) { + x_6964 = lean_alloc_ctor(0, 2, 0); +} else { + x_6964 = x_6963; +} +lean_ctor_set(x_6964, 0, x_6961); +lean_ctor_set(x_6964, 1, x_6962); +return x_6964; +} +else +{ +lean_object* x_6965; lean_object* x_6966; lean_object* x_6967; lean_object* x_6968; uint8_t x_6969; +x_6965 = lean_ctor_get(x_6960, 0); +lean_inc(x_6965); +x_6966 = lean_ctor_get(x_6960, 1); +lean_inc(x_6966); +if (lean_is_exclusive(x_6960)) { + lean_ctor_release(x_6960, 0); + lean_ctor_release(x_6960, 1); + x_6967 = x_6960; +} else { + lean_dec_ref(x_6960); + x_6967 = lean_box(0); +} +x_6968 = lean_ctor_get(x_6965, 1); +lean_inc(x_6968); +x_6969 = lean_nat_dec_eq(x_6954, x_6968); +lean_dec(x_6954); +if (x_6969 == 0) +{ +lean_object* x_6970; +lean_dec(x_6968); +if (lean_is_scalar(x_6967)) { + x_6970 = lean_alloc_ctor(1, 2, 0); +} else { + x_6970 = x_6967; +} +lean_ctor_set(x_6970, 0, x_6965); +lean_ctor_set(x_6970, 1, x_6966); +return x_6970; +} +else +{ +lean_object* x_6971; lean_object* x_6972; lean_object* x_6973; lean_object* x_6974; +lean_dec(x_6967); +lean_dec(x_6966); +x_6971 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48; +x_6972 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6972, 0, x_6971); +x_6973 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47; +x_6974 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6973, x_6972, x_6965); +if (lean_obj_tag(x_6974) == 0) +{ +lean_object* x_6975; lean_object* x_6976; lean_object* x_6977; lean_object* x_6978; +lean_dec(x_6968); +x_6975 = lean_ctor_get(x_6974, 0); +lean_inc(x_6975); +x_6976 = lean_ctor_get(x_6974, 1); +lean_inc(x_6976); +if (lean_is_exclusive(x_6974)) { + lean_ctor_release(x_6974, 0); + lean_ctor_release(x_6974, 1); + x_6977 = x_6974; +} else { + lean_dec_ref(x_6974); + x_6977 = lean_box(0); +} +if (lean_is_scalar(x_6977)) { + x_6978 = lean_alloc_ctor(0, 2, 0); +} else { + x_6978 = x_6977; +} +lean_ctor_set(x_6978, 0, x_6975); +lean_ctor_set(x_6978, 1, x_6976); +return x_6978; +} +else +{ +lean_object* x_6979; lean_object* x_6980; lean_object* x_6981; lean_object* x_6982; uint8_t x_6983; +x_6979 = lean_ctor_get(x_6974, 0); +lean_inc(x_6979); +x_6980 = lean_ctor_get(x_6974, 1); +lean_inc(x_6980); +if (lean_is_exclusive(x_6974)) { + lean_ctor_release(x_6974, 0); + lean_ctor_release(x_6974, 1); + x_6981 = x_6974; +} else { + lean_dec_ref(x_6974); + x_6981 = lean_box(0); +} +x_6982 = lean_ctor_get(x_6979, 1); +lean_inc(x_6982); +x_6983 = lean_nat_dec_eq(x_6968, x_6982); +lean_dec(x_6968); +if (x_6983 == 0) +{ +lean_object* x_6984; +lean_dec(x_6982); +if (lean_is_scalar(x_6981)) { + x_6984 = lean_alloc_ctor(1, 2, 0); +} else { + x_6984 = x_6981; +} +lean_ctor_set(x_6984, 0, x_6979); +lean_ctor_set(x_6984, 1, x_6980); +return x_6984; +} +else +{ +lean_object* x_6985; lean_object* x_6986; lean_object* x_6987; lean_object* x_6988; +lean_dec(x_6981); +lean_dec(x_6980); +x_6985 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51; +x_6986 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_6986, 0, x_6985); +x_6987 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50; +x_6988 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_6987, x_6986, x_6979); +if (lean_obj_tag(x_6988) == 0) +{ +lean_object* x_6989; lean_object* x_6990; lean_object* x_6991; lean_object* x_6992; +lean_dec(x_6982); +x_6989 = lean_ctor_get(x_6988, 0); +lean_inc(x_6989); +x_6990 = lean_ctor_get(x_6988, 1); +lean_inc(x_6990); +if (lean_is_exclusive(x_6988)) { + lean_ctor_release(x_6988, 0); + lean_ctor_release(x_6988, 1); + x_6991 = x_6988; +} else { + lean_dec_ref(x_6988); + x_6991 = lean_box(0); +} +if (lean_is_scalar(x_6991)) { + x_6992 = lean_alloc_ctor(0, 2, 0); +} else { + x_6992 = x_6991; +} +lean_ctor_set(x_6992, 0, x_6989); +lean_ctor_set(x_6992, 1, x_6990); +return x_6992; +} +else +{ +lean_object* x_6993; lean_object* x_6994; lean_object* x_6995; lean_object* x_6996; uint8_t x_6997; +x_6993 = lean_ctor_get(x_6988, 0); +lean_inc(x_6993); +x_6994 = lean_ctor_get(x_6988, 1); +lean_inc(x_6994); +if (lean_is_exclusive(x_6988)) { + lean_ctor_release(x_6988, 0); + lean_ctor_release(x_6988, 1); + x_6995 = x_6988; +} else { + lean_dec_ref(x_6988); + x_6995 = lean_box(0); +} +x_6996 = lean_ctor_get(x_6993, 1); +lean_inc(x_6996); +x_6997 = lean_nat_dec_eq(x_6982, x_6996); +lean_dec(x_6982); +if (x_6997 == 0) +{ +lean_object* x_6998; +lean_dec(x_6996); +if (lean_is_scalar(x_6995)) { + x_6998 = lean_alloc_ctor(1, 2, 0); +} else { + x_6998 = x_6995; +} +lean_ctor_set(x_6998, 0, x_6993); +lean_ctor_set(x_6998, 1, x_6994); +return x_6998; +} +else +{ +lean_object* x_6999; lean_object* x_7000; lean_object* x_7001; lean_object* x_7002; +lean_dec(x_6995); +lean_dec(x_6994); +x_6999 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54; +x_7000 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7000, 0, x_6999); +x_7001 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53; +x_7002 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7001, x_7000, x_6993); +if (lean_obj_tag(x_7002) == 0) +{ +lean_object* x_7003; lean_object* x_7004; lean_object* x_7005; lean_object* x_7006; +lean_dec(x_6996); +x_7003 = lean_ctor_get(x_7002, 0); +lean_inc(x_7003); +x_7004 = lean_ctor_get(x_7002, 1); +lean_inc(x_7004); +if (lean_is_exclusive(x_7002)) { + lean_ctor_release(x_7002, 0); + lean_ctor_release(x_7002, 1); + x_7005 = x_7002; +} else { + lean_dec_ref(x_7002); + x_7005 = lean_box(0); +} +if (lean_is_scalar(x_7005)) { + x_7006 = lean_alloc_ctor(0, 2, 0); +} else { + x_7006 = x_7005; +} +lean_ctor_set(x_7006, 0, x_7003); +lean_ctor_set(x_7006, 1, x_7004); +return x_7006; +} +else +{ +lean_object* x_7007; lean_object* x_7008; lean_object* x_7009; lean_object* x_7010; uint8_t x_7011; +x_7007 = lean_ctor_get(x_7002, 0); +lean_inc(x_7007); +x_7008 = lean_ctor_get(x_7002, 1); +lean_inc(x_7008); +if (lean_is_exclusive(x_7002)) { + lean_ctor_release(x_7002, 0); + lean_ctor_release(x_7002, 1); + x_7009 = x_7002; +} else { + lean_dec_ref(x_7002); + x_7009 = lean_box(0); +} +x_7010 = lean_ctor_get(x_7007, 1); +lean_inc(x_7010); +x_7011 = lean_nat_dec_eq(x_6996, x_7010); +lean_dec(x_6996); +if (x_7011 == 0) +{ +lean_object* x_7012; +lean_dec(x_7010); +if (lean_is_scalar(x_7009)) { + x_7012 = lean_alloc_ctor(1, 2, 0); +} else { + x_7012 = x_7009; +} +lean_ctor_set(x_7012, 0, x_7007); +lean_ctor_set(x_7012, 1, x_7008); +return x_7012; +} +else +{ +lean_object* x_7013; lean_object* x_7014; lean_object* x_7015; lean_object* x_7016; +lean_dec(x_7009); +lean_dec(x_7008); +x_7013 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57; +x_7014 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7014, 0, x_7013); +x_7015 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56; +x_7016 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7015, x_7014, x_7007); +if (lean_obj_tag(x_7016) == 0) +{ +lean_object* x_7017; lean_object* x_7018; lean_object* x_7019; lean_object* x_7020; +lean_dec(x_7010); +x_7017 = lean_ctor_get(x_7016, 0); +lean_inc(x_7017); +x_7018 = lean_ctor_get(x_7016, 1); +lean_inc(x_7018); +if (lean_is_exclusive(x_7016)) { + lean_ctor_release(x_7016, 0); + lean_ctor_release(x_7016, 1); + x_7019 = x_7016; +} else { + lean_dec_ref(x_7016); + x_7019 = lean_box(0); +} +if (lean_is_scalar(x_7019)) { + x_7020 = lean_alloc_ctor(0, 2, 0); +} else { + x_7020 = x_7019; +} +lean_ctor_set(x_7020, 0, x_7017); +lean_ctor_set(x_7020, 1, x_7018); +return x_7020; +} +else +{ +lean_object* x_7021; lean_object* x_7022; lean_object* x_7023; lean_object* x_7024; uint8_t x_7025; +x_7021 = lean_ctor_get(x_7016, 0); +lean_inc(x_7021); +x_7022 = lean_ctor_get(x_7016, 1); +lean_inc(x_7022); +if (lean_is_exclusive(x_7016)) { + lean_ctor_release(x_7016, 0); + lean_ctor_release(x_7016, 1); + x_7023 = x_7016; +} else { + lean_dec_ref(x_7016); + x_7023 = lean_box(0); +} +x_7024 = lean_ctor_get(x_7021, 1); +lean_inc(x_7024); +x_7025 = lean_nat_dec_eq(x_7010, x_7024); +lean_dec(x_7010); +if (x_7025 == 0) +{ +lean_object* x_7026; +lean_dec(x_7024); +if (lean_is_scalar(x_7023)) { + x_7026 = lean_alloc_ctor(1, 2, 0); +} else { + x_7026 = x_7023; +} +lean_ctor_set(x_7026, 0, x_7021); +lean_ctor_set(x_7026, 1, x_7022); +return x_7026; +} +else +{ +lean_object* x_7027; lean_object* x_7028; lean_object* x_7029; lean_object* x_7030; +lean_dec(x_7023); +lean_dec(x_7022); +x_7027 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60; +x_7028 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7028, 0, x_7027); +x_7029 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59; +x_7030 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7029, x_7028, x_7021); +if (lean_obj_tag(x_7030) == 0) +{ +lean_object* x_7031; lean_object* x_7032; lean_object* x_7033; lean_object* x_7034; +lean_dec(x_7024); +x_7031 = lean_ctor_get(x_7030, 0); +lean_inc(x_7031); +x_7032 = lean_ctor_get(x_7030, 1); +lean_inc(x_7032); +if (lean_is_exclusive(x_7030)) { + lean_ctor_release(x_7030, 0); + lean_ctor_release(x_7030, 1); + x_7033 = x_7030; +} else { + lean_dec_ref(x_7030); + x_7033 = lean_box(0); +} +if (lean_is_scalar(x_7033)) { + x_7034 = lean_alloc_ctor(0, 2, 0); +} else { + x_7034 = x_7033; +} +lean_ctor_set(x_7034, 0, x_7031); +lean_ctor_set(x_7034, 1, x_7032); +return x_7034; +} +else +{ +lean_object* x_7035; lean_object* x_7036; lean_object* x_7037; lean_object* x_7038; uint8_t x_7039; +x_7035 = lean_ctor_get(x_7030, 0); +lean_inc(x_7035); +x_7036 = lean_ctor_get(x_7030, 1); +lean_inc(x_7036); +if (lean_is_exclusive(x_7030)) { + lean_ctor_release(x_7030, 0); + lean_ctor_release(x_7030, 1); + x_7037 = x_7030; +} else { + lean_dec_ref(x_7030); + x_7037 = lean_box(0); +} +x_7038 = lean_ctor_get(x_7035, 1); +lean_inc(x_7038); +x_7039 = lean_nat_dec_eq(x_7024, x_7038); +lean_dec(x_7024); +if (x_7039 == 0) +{ +lean_object* x_7040; +lean_dec(x_7038); +if (lean_is_scalar(x_7037)) { + x_7040 = lean_alloc_ctor(1, 2, 0); +} else { + x_7040 = x_7037; +} +lean_ctor_set(x_7040, 0, x_7035); +lean_ctor_set(x_7040, 1, x_7036); +return x_7040; +} +else +{ +lean_object* x_7041; lean_object* x_7042; lean_object* x_7043; lean_object* x_7044; +lean_dec(x_7037); +lean_dec(x_7036); +x_7041 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63; +x_7042 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7042, 0, x_7041); +x_7043 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62; +x_7044 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7043, x_7042, x_7035); +if (lean_obj_tag(x_7044) == 0) +{ +lean_object* x_7045; lean_object* x_7046; lean_object* x_7047; lean_object* x_7048; +lean_dec(x_7038); +x_7045 = lean_ctor_get(x_7044, 0); +lean_inc(x_7045); +x_7046 = lean_ctor_get(x_7044, 1); +lean_inc(x_7046); +if (lean_is_exclusive(x_7044)) { + lean_ctor_release(x_7044, 0); + lean_ctor_release(x_7044, 1); + x_7047 = x_7044; +} else { + lean_dec_ref(x_7044); + x_7047 = lean_box(0); +} +if (lean_is_scalar(x_7047)) { + x_7048 = lean_alloc_ctor(0, 2, 0); +} else { + x_7048 = x_7047; +} +lean_ctor_set(x_7048, 0, x_7045); +lean_ctor_set(x_7048, 1, x_7046); +return x_7048; +} +else +{ +lean_object* x_7049; lean_object* x_7050; lean_object* x_7051; lean_object* x_7052; uint8_t x_7053; +x_7049 = lean_ctor_get(x_7044, 0); +lean_inc(x_7049); +x_7050 = lean_ctor_get(x_7044, 1); +lean_inc(x_7050); +if (lean_is_exclusive(x_7044)) { + lean_ctor_release(x_7044, 0); + lean_ctor_release(x_7044, 1); + x_7051 = x_7044; +} else { + lean_dec_ref(x_7044); + x_7051 = lean_box(0); +} +x_7052 = lean_ctor_get(x_7049, 1); +lean_inc(x_7052); +x_7053 = lean_nat_dec_eq(x_7038, x_7052); +lean_dec(x_7038); +if (x_7053 == 0) +{ +lean_object* x_7054; +lean_dec(x_7052); +if (lean_is_scalar(x_7051)) { + x_7054 = lean_alloc_ctor(1, 2, 0); +} else { + x_7054 = x_7051; +} +lean_ctor_set(x_7054, 0, x_7049); +lean_ctor_set(x_7054, 1, x_7050); +return x_7054; +} +else +{ +lean_object* x_7055; lean_object* x_7056; lean_object* x_7057; lean_object* x_7058; +lean_dec(x_7051); +lean_dec(x_7050); +x_7055 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66; +x_7056 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7056, 0, x_7055); +x_7057 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65; +x_7058 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7057, x_7056, x_7049); +if (lean_obj_tag(x_7058) == 0) +{ +lean_object* x_7059; lean_object* x_7060; lean_object* x_7061; lean_object* x_7062; +lean_dec(x_7052); +x_7059 = lean_ctor_get(x_7058, 0); +lean_inc(x_7059); +x_7060 = lean_ctor_get(x_7058, 1); +lean_inc(x_7060); +if (lean_is_exclusive(x_7058)) { + lean_ctor_release(x_7058, 0); + lean_ctor_release(x_7058, 1); + x_7061 = x_7058; +} else { + lean_dec_ref(x_7058); + x_7061 = lean_box(0); +} +if (lean_is_scalar(x_7061)) { + x_7062 = lean_alloc_ctor(0, 2, 0); +} else { + x_7062 = x_7061; +} +lean_ctor_set(x_7062, 0, x_7059); +lean_ctor_set(x_7062, 1, x_7060); +return x_7062; +} +else +{ +lean_object* x_7063; lean_object* x_7064; lean_object* x_7065; lean_object* x_7066; uint8_t x_7067; +x_7063 = lean_ctor_get(x_7058, 0); +lean_inc(x_7063); +x_7064 = lean_ctor_get(x_7058, 1); +lean_inc(x_7064); +if (lean_is_exclusive(x_7058)) { + lean_ctor_release(x_7058, 0); + lean_ctor_release(x_7058, 1); + x_7065 = x_7058; +} else { + lean_dec_ref(x_7058); + x_7065 = lean_box(0); +} +x_7066 = lean_ctor_get(x_7063, 1); +lean_inc(x_7066); +x_7067 = lean_nat_dec_eq(x_7052, x_7066); +lean_dec(x_7052); +if (x_7067 == 0) +{ +lean_object* x_7068; +lean_dec(x_7066); +if (lean_is_scalar(x_7065)) { + x_7068 = lean_alloc_ctor(1, 2, 0); +} else { + x_7068 = x_7065; +} +lean_ctor_set(x_7068, 0, x_7063); +lean_ctor_set(x_7068, 1, x_7064); +return x_7068; +} +else +{ +lean_object* x_7069; lean_object* x_7070; lean_object* x_7071; lean_object* x_7072; +lean_dec(x_7065); +lean_dec(x_7064); +x_7069 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69; +x_7070 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7070, 0, x_7069); +x_7071 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68; +x_7072 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7071, x_7070, x_7063); +if (lean_obj_tag(x_7072) == 0) +{ +lean_object* x_7073; lean_object* x_7074; lean_object* x_7075; lean_object* x_7076; +lean_dec(x_7066); +x_7073 = lean_ctor_get(x_7072, 0); +lean_inc(x_7073); +x_7074 = lean_ctor_get(x_7072, 1); +lean_inc(x_7074); +if (lean_is_exclusive(x_7072)) { + lean_ctor_release(x_7072, 0); + lean_ctor_release(x_7072, 1); + x_7075 = x_7072; +} else { + lean_dec_ref(x_7072); + x_7075 = lean_box(0); +} +if (lean_is_scalar(x_7075)) { + x_7076 = lean_alloc_ctor(0, 2, 0); +} else { + x_7076 = x_7075; +} +lean_ctor_set(x_7076, 0, x_7073); +lean_ctor_set(x_7076, 1, x_7074); +return x_7076; +} +else +{ +lean_object* x_7077; lean_object* x_7078; lean_object* x_7079; lean_object* x_7080; uint8_t x_7081; +x_7077 = lean_ctor_get(x_7072, 0); +lean_inc(x_7077); +x_7078 = lean_ctor_get(x_7072, 1); +lean_inc(x_7078); +if (lean_is_exclusive(x_7072)) { + lean_ctor_release(x_7072, 0); + lean_ctor_release(x_7072, 1); + x_7079 = x_7072; +} else { + lean_dec_ref(x_7072); + x_7079 = lean_box(0); +} +x_7080 = lean_ctor_get(x_7077, 1); +lean_inc(x_7080); +x_7081 = lean_nat_dec_eq(x_7066, x_7080); +lean_dec(x_7066); +if (x_7081 == 0) +{ +lean_object* x_7082; +lean_dec(x_7080); +if (lean_is_scalar(x_7079)) { + x_7082 = lean_alloc_ctor(1, 2, 0); +} else { + x_7082 = x_7079; +} +lean_ctor_set(x_7082, 0, x_7077); +lean_ctor_set(x_7082, 1, x_7078); +return x_7082; +} +else +{ +lean_object* x_7083; lean_object* x_7084; lean_object* x_7085; lean_object* x_7086; +lean_dec(x_7079); +lean_dec(x_7078); +x_7083 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72; +x_7084 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7084, 0, x_7083); +x_7085 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71; +x_7086 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7085, x_7084, x_7077); +if (lean_obj_tag(x_7086) == 0) +{ +lean_object* x_7087; lean_object* x_7088; lean_object* x_7089; lean_object* x_7090; +lean_dec(x_7080); +x_7087 = lean_ctor_get(x_7086, 0); +lean_inc(x_7087); +x_7088 = lean_ctor_get(x_7086, 1); +lean_inc(x_7088); +if (lean_is_exclusive(x_7086)) { + lean_ctor_release(x_7086, 0); + lean_ctor_release(x_7086, 1); + x_7089 = x_7086; +} else { + lean_dec_ref(x_7086); + x_7089 = lean_box(0); +} +if (lean_is_scalar(x_7089)) { + x_7090 = lean_alloc_ctor(0, 2, 0); +} else { + x_7090 = x_7089; +} +lean_ctor_set(x_7090, 0, x_7087); +lean_ctor_set(x_7090, 1, x_7088); +return x_7090; +} +else +{ +lean_object* x_7091; lean_object* x_7092; lean_object* x_7093; lean_object* x_7094; uint8_t x_7095; +x_7091 = lean_ctor_get(x_7086, 0); +lean_inc(x_7091); +x_7092 = lean_ctor_get(x_7086, 1); +lean_inc(x_7092); +if (lean_is_exclusive(x_7086)) { + lean_ctor_release(x_7086, 0); + lean_ctor_release(x_7086, 1); + x_7093 = x_7086; +} else { + lean_dec_ref(x_7086); + x_7093 = lean_box(0); +} +x_7094 = lean_ctor_get(x_7091, 1); +lean_inc(x_7094); +x_7095 = lean_nat_dec_eq(x_7080, x_7094); +lean_dec(x_7080); +if (x_7095 == 0) +{ +lean_object* x_7096; +lean_dec(x_7094); +if (lean_is_scalar(x_7093)) { + x_7096 = lean_alloc_ctor(1, 2, 0); +} else { + x_7096 = x_7093; +} +lean_ctor_set(x_7096, 0, x_7091); +lean_ctor_set(x_7096, 1, x_7092); +return x_7096; +} +else +{ +lean_object* x_7097; lean_object* x_7098; lean_object* x_7099; lean_object* x_7100; +lean_dec(x_7093); +lean_dec(x_7092); +x_7097 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73; +x_7098 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7098, 0, x_7097); +x_7099 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74; +x_7100 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7099, x_7098, x_7091); +if (lean_obj_tag(x_7100) == 0) +{ +lean_object* x_7101; lean_object* x_7102; lean_object* x_7103; lean_object* x_7104; +lean_dec(x_7094); +x_7101 = lean_ctor_get(x_7100, 0); +lean_inc(x_7101); +x_7102 = lean_ctor_get(x_7100, 1); +lean_inc(x_7102); +if (lean_is_exclusive(x_7100)) { + lean_ctor_release(x_7100, 0); + lean_ctor_release(x_7100, 1); + x_7103 = x_7100; +} else { + lean_dec_ref(x_7100); + x_7103 = lean_box(0); +} +if (lean_is_scalar(x_7103)) { + x_7104 = lean_alloc_ctor(0, 2, 0); +} else { + x_7104 = x_7103; +} +lean_ctor_set(x_7104, 0, x_7101); +lean_ctor_set(x_7104, 1, x_7102); +return x_7104; +} +else +{ +lean_object* x_7105; lean_object* x_7106; lean_object* x_7107; lean_object* x_7108; uint8_t x_7109; +x_7105 = lean_ctor_get(x_7100, 0); +lean_inc(x_7105); +x_7106 = lean_ctor_get(x_7100, 1); +lean_inc(x_7106); +if (lean_is_exclusive(x_7100)) { + lean_ctor_release(x_7100, 0); + lean_ctor_release(x_7100, 1); + x_7107 = x_7100; +} else { + lean_dec_ref(x_7100); + x_7107 = lean_box(0); +} +x_7108 = lean_ctor_get(x_7105, 1); +lean_inc(x_7108); +x_7109 = lean_nat_dec_eq(x_7094, x_7108); +lean_dec(x_7094); +if (x_7109 == 0) +{ +lean_object* x_7110; +lean_dec(x_7108); +if (lean_is_scalar(x_7107)) { + x_7110 = lean_alloc_ctor(1, 2, 0); +} else { + x_7110 = x_7107; +} +lean_ctor_set(x_7110, 0, x_7105); +lean_ctor_set(x_7110, 1, x_7106); +return x_7110; +} +else +{ +lean_object* x_7111; lean_object* x_7112; lean_object* x_7113; lean_object* x_7114; +lean_dec(x_7107); +lean_dec(x_7106); +x_7111 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77; +x_7112 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7112, 0, x_7111); +x_7113 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76; +x_7114 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7113, x_7112, x_7105); +if (lean_obj_tag(x_7114) == 0) +{ +lean_object* x_7115; lean_object* x_7116; lean_object* x_7117; lean_object* x_7118; +lean_dec(x_7108); +x_7115 = lean_ctor_get(x_7114, 0); +lean_inc(x_7115); +x_7116 = lean_ctor_get(x_7114, 1); +lean_inc(x_7116); +if (lean_is_exclusive(x_7114)) { + lean_ctor_release(x_7114, 0); + lean_ctor_release(x_7114, 1); + x_7117 = x_7114; +} else { + lean_dec_ref(x_7114); + x_7117 = lean_box(0); +} +if (lean_is_scalar(x_7117)) { + x_7118 = lean_alloc_ctor(0, 2, 0); +} else { + x_7118 = x_7117; +} +lean_ctor_set(x_7118, 0, x_7115); +lean_ctor_set(x_7118, 1, x_7116); +return x_7118; +} +else +{ +lean_object* x_7119; lean_object* x_7120; lean_object* x_7121; lean_object* x_7122; uint8_t x_7123; +x_7119 = lean_ctor_get(x_7114, 0); +lean_inc(x_7119); +x_7120 = lean_ctor_get(x_7114, 1); +lean_inc(x_7120); +if (lean_is_exclusive(x_7114)) { + lean_ctor_release(x_7114, 0); + lean_ctor_release(x_7114, 1); + x_7121 = x_7114; +} else { + lean_dec_ref(x_7114); + x_7121 = lean_box(0); +} +x_7122 = lean_ctor_get(x_7119, 1); +lean_inc(x_7122); +x_7123 = lean_nat_dec_eq(x_7108, x_7122); +lean_dec(x_7108); +if (x_7123 == 0) +{ +lean_object* x_7124; +lean_dec(x_7122); +if (lean_is_scalar(x_7121)) { + x_7124 = lean_alloc_ctor(1, 2, 0); +} else { + x_7124 = x_7121; +} +lean_ctor_set(x_7124, 0, x_7119); +lean_ctor_set(x_7124, 1, x_7120); +return x_7124; +} +else +{ +lean_object* x_7125; lean_object* x_7126; lean_object* x_7127; lean_object* x_7128; +lean_dec(x_7121); +lean_dec(x_7120); +x_7125 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80; +x_7126 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7126, 0, x_7125); +x_7127 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79; +x_7128 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7127, x_7126, x_7119); +if (lean_obj_tag(x_7128) == 0) +{ +lean_object* x_7129; lean_object* x_7130; lean_object* x_7131; lean_object* x_7132; +lean_dec(x_7122); +x_7129 = lean_ctor_get(x_7128, 0); +lean_inc(x_7129); +x_7130 = lean_ctor_get(x_7128, 1); +lean_inc(x_7130); +if (lean_is_exclusive(x_7128)) { + lean_ctor_release(x_7128, 0); + lean_ctor_release(x_7128, 1); + x_7131 = x_7128; +} else { + lean_dec_ref(x_7128); + x_7131 = lean_box(0); +} +if (lean_is_scalar(x_7131)) { + x_7132 = lean_alloc_ctor(0, 2, 0); +} else { + x_7132 = x_7131; +} +lean_ctor_set(x_7132, 0, x_7129); +lean_ctor_set(x_7132, 1, x_7130); +return x_7132; +} +else +{ +lean_object* x_7133; lean_object* x_7134; lean_object* x_7135; lean_object* x_7136; uint8_t x_7137; +x_7133 = lean_ctor_get(x_7128, 0); +lean_inc(x_7133); +x_7134 = lean_ctor_get(x_7128, 1); +lean_inc(x_7134); +if (lean_is_exclusive(x_7128)) { + lean_ctor_release(x_7128, 0); + lean_ctor_release(x_7128, 1); + x_7135 = x_7128; +} else { + lean_dec_ref(x_7128); + x_7135 = lean_box(0); +} +x_7136 = lean_ctor_get(x_7133, 1); +lean_inc(x_7136); +x_7137 = lean_nat_dec_eq(x_7122, x_7136); +lean_dec(x_7122); +if (x_7137 == 0) +{ +lean_object* x_7138; +lean_dec(x_7136); +if (lean_is_scalar(x_7135)) { + x_7138 = lean_alloc_ctor(1, 2, 0); +} else { + x_7138 = x_7135; +} +lean_ctor_set(x_7138, 0, x_7133); +lean_ctor_set(x_7138, 1, x_7134); +return x_7138; +} +else +{ +lean_object* x_7139; lean_object* x_7140; lean_object* x_7141; lean_object* x_7142; +lean_dec(x_7135); +lean_dec(x_7134); +x_7139 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83; +x_7140 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7140, 0, x_7139); +x_7141 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82; +x_7142 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7141, x_7140, x_7133); +if (lean_obj_tag(x_7142) == 0) +{ +lean_object* x_7143; lean_object* x_7144; lean_object* x_7145; lean_object* x_7146; +lean_dec(x_7136); +x_7143 = lean_ctor_get(x_7142, 0); +lean_inc(x_7143); +x_7144 = lean_ctor_get(x_7142, 1); +lean_inc(x_7144); +if (lean_is_exclusive(x_7142)) { + lean_ctor_release(x_7142, 0); + lean_ctor_release(x_7142, 1); + x_7145 = x_7142; +} else { + lean_dec_ref(x_7142); + x_7145 = lean_box(0); +} +if (lean_is_scalar(x_7145)) { + x_7146 = lean_alloc_ctor(0, 2, 0); +} else { + x_7146 = x_7145; +} +lean_ctor_set(x_7146, 0, x_7143); +lean_ctor_set(x_7146, 1, x_7144); +return x_7146; +} +else +{ +lean_object* x_7147; lean_object* x_7148; lean_object* x_7149; lean_object* x_7150; uint8_t x_7151; +x_7147 = lean_ctor_get(x_7142, 0); +lean_inc(x_7147); +x_7148 = lean_ctor_get(x_7142, 1); +lean_inc(x_7148); +if (lean_is_exclusive(x_7142)) { + lean_ctor_release(x_7142, 0); + lean_ctor_release(x_7142, 1); + x_7149 = x_7142; +} else { + lean_dec_ref(x_7142); + x_7149 = lean_box(0); +} +x_7150 = lean_ctor_get(x_7147, 1); +lean_inc(x_7150); +x_7151 = lean_nat_dec_eq(x_7136, x_7150); +lean_dec(x_7136); +if (x_7151 == 0) +{ +lean_object* x_7152; +lean_dec(x_7150); +if (lean_is_scalar(x_7149)) { + x_7152 = lean_alloc_ctor(1, 2, 0); +} else { + x_7152 = x_7149; +} +lean_ctor_set(x_7152, 0, x_7147); +lean_ctor_set(x_7152, 1, x_7148); +return x_7152; +} +else +{ +lean_object* x_7153; lean_object* x_7154; lean_object* x_7155; lean_object* x_7156; +lean_dec(x_7149); +lean_dec(x_7148); +x_7153 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86; +x_7154 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7154, 0, x_7153); +x_7155 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85; +x_7156 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7155, x_7154, x_7147); +if (lean_obj_tag(x_7156) == 0) +{ +lean_object* x_7157; lean_object* x_7158; lean_object* x_7159; lean_object* x_7160; +lean_dec(x_7150); +x_7157 = lean_ctor_get(x_7156, 0); +lean_inc(x_7157); +x_7158 = lean_ctor_get(x_7156, 1); +lean_inc(x_7158); +if (lean_is_exclusive(x_7156)) { + lean_ctor_release(x_7156, 0); + lean_ctor_release(x_7156, 1); + x_7159 = x_7156; +} else { + lean_dec_ref(x_7156); + x_7159 = lean_box(0); +} +if (lean_is_scalar(x_7159)) { + x_7160 = lean_alloc_ctor(0, 2, 0); +} else { + x_7160 = x_7159; +} +lean_ctor_set(x_7160, 0, x_7157); +lean_ctor_set(x_7160, 1, x_7158); +return x_7160; +} +else +{ +lean_object* x_7161; lean_object* x_7162; lean_object* x_7163; lean_object* x_7164; uint8_t x_7165; +x_7161 = lean_ctor_get(x_7156, 0); +lean_inc(x_7161); +x_7162 = lean_ctor_get(x_7156, 1); +lean_inc(x_7162); +if (lean_is_exclusive(x_7156)) { + lean_ctor_release(x_7156, 0); + lean_ctor_release(x_7156, 1); + x_7163 = x_7156; +} else { + lean_dec_ref(x_7156); + x_7163 = lean_box(0); +} +x_7164 = lean_ctor_get(x_7161, 1); +lean_inc(x_7164); +x_7165 = lean_nat_dec_eq(x_7150, x_7164); +lean_dec(x_7164); +lean_dec(x_7150); +if (x_7165 == 0) +{ +lean_object* x_7166; +if (lean_is_scalar(x_7163)) { + x_7166 = lean_alloc_ctor(1, 2, 0); +} else { + x_7166 = x_7163; +} +lean_ctor_set(x_7166, 0, x_7161); +lean_ctor_set(x_7166, 1, x_7162); +return x_7166; +} +else +{ +lean_object* x_7167; lean_object* x_7168; lean_object* x_7169; lean_object* x_7170; +lean_dec(x_7163); +lean_dec(x_7162); +x_7167 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89; +x_7168 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_7168, 0, x_7167); +x_7169 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88; +x_7170 = l_Bind_bindLeft___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__3(x_7169, x_7168, x_7161); +return x_7170; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__1(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__10(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__13(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__24(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__25(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__26(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__27(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___lambda__28(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.FormatPart.string", 26, 26); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.FormatPart.modifier", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_unsigned_to_nat(1024u); +x_6 = lean_nat_dec_le(x_5, x_2); +x_7 = l_String_quote(x_4); +lean_dec(x_4); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_7); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3; +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_1); +if (x_6 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_10 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_11 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = 0; +x_13 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*1, x_12); +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_9); +x_17 = 0; +x_18 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_17); +x_19 = l_Repr_addAppParen(x_18, x_2); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +lean_dec(x_1); +x_21 = lean_unsigned_to_nat(1024u); +x_22 = lean_nat_dec_le(x_21, x_2); +x_23 = l_String_quote(x_20); +lean_dec(x_20); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3; +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +if (x_22 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_28 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = 0; +x_30 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set_uint8(x_30, sizeof(void*)*1, x_29); +x_31 = l_Repr_addAppParen(x_30, x_2); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_33 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_26); +x_34 = 0; +x_35 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_34); +x_36 = l_Repr_addAppParen(x_35, x_2); +return x_36; +} +} +} +else +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_unsigned_to_nat(1024u); +x_39 = lean_nat_dec_le(x_38, x_2); +x_40 = l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836_(x_37, x_38); +x_41 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6; +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +if (x_39 == 0) +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; +x_43 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_44 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = 0; +x_46 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*1, x_45); +x_47 = l_Repr_addAppParen(x_46, x_2); +return x_47; +} +else +{ +lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_49 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_42); +x_50 = 0; +x_51 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set_uint8(x_51, sizeof(void*)*1, x_50); +x_52 = l_Repr_addAppParen(x_51, x_2); +return x_52; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprFormatPart___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprFormatPart() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprFormatPart___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instCoeStringFormatPart(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instCoeModifierFormatPart(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instCoeTimeZone(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_3, x_4); +return x_5; +} +else +{ +lean_dec(x_3); +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_Awareness_type_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_5 = lean_int_add(x_1, x_4); +x_6 = lean_int_sub(x_5, x_1); +lean_dec(x_5); +x_7 = lean_int_add(x_6, x_1); +lean_dec(x_6); +x_8 = lean_int_sub(x_1, x_1); +x_9 = lean_int_emod(x_8, x_7); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = lean_int_emod(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = lean_int_add(x_11, x_1); +lean_dec(x_11); +x_13 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1; +x_14 = lean_int_add(x_1, x_13); +x_15 = lean_int_sub(x_14, x_1); +lean_dec(x_14); +x_16 = lean_int_add(x_15, x_1); +lean_dec(x_15); +x_17 = lean_int_emod(x_8, x_16); +lean_dec(x_8); +x_18 = lean_int_add(x_17, x_16); +lean_dec(x_17); +x_19 = lean_int_emod(x_18, x_16); +lean_dec(x_16); +lean_dec(x_18); +x_20 = lean_int_add(x_19, x_1); +lean_dec(x_19); +lean_inc(x_2); +x_21 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_21, 0, x_2); +lean_ctor_set(x_21, 1, x_12); +lean_ctor_set(x_21, 2, x_20); +x_22 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2; +x_23 = lean_int_add(x_2, x_22); +x_24 = lean_int_sub(x_23, x_2); +lean_dec(x_23); +x_25 = lean_int_add(x_24, x_1); +lean_dec(x_24); +x_26 = lean_int_sub(x_2, x_2); +x_27 = lean_int_emod(x_26, x_25); +x_28 = lean_int_add(x_27, x_25); +lean_dec(x_27); +x_29 = lean_int_emod(x_28, x_25); +lean_dec(x_25); +lean_dec(x_28); +x_30 = lean_int_add(x_29, x_2); +lean_dec(x_29); +x_31 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3; +x_32 = lean_int_add(x_2, x_31); +x_33 = lean_int_sub(x_32, x_2); +lean_dec(x_32); +x_34 = lean_int_add(x_33, x_1); +lean_dec(x_33); +x_35 = lean_int_emod(x_26, x_34); +lean_dec(x_26); +x_36 = lean_int_add(x_35, x_34); +lean_dec(x_35); +x_37 = lean_int_emod(x_36, x_34); +lean_dec(x_34); +lean_dec(x_36); +x_38 = lean_int_add(x_37, x_2); +lean_dec(x_37); +x_39 = 0; +x_40 = lean_box(x_39); +lean_inc(x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +x_42 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_42, 0, x_30); +lean_ctor_set(x_42, 1, x_38); +lean_ctor_set(x_42, 2, x_41); +lean_ctor_set(x_42, 3, x_2); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_21); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9; +x_3 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32; +x_3 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33; +x_4 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35; +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Awareness_instInhabitedType___lambda__2___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Awareness_instInhabitedType___closed__2; +x_2 = lean_mk_thunk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = 0; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_4 = 0; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 2, x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___closed__4; +x_2 = l_Std_Time_Awareness_instInhabitedType___closed__5; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*3, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Awareness_instInhabitedType___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_Awareness_instInhabitedType___closed__3; +x_2 = l_Std_Time_Awareness_instInhabitedType___closed__1; +x_3 = l_Std_Time_Awareness_instInhabitedType___closed__6; +x_4 = l_Std_Time_Awareness_instInhabitedType___closed__7; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Awareness_instInhabitedType___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_3); +x_5 = lean_mk_thunk(x_4); +x_6 = l_Std_Time_Awareness_instInhabitedType___closed__1; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = l_Std_Time_Awareness_instInhabitedType___closed__8; +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Awareness_instInhabitedType___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___lambda__2___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Awareness_instInhabitedType___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Awareness_instInhabitedType(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Awareness_getD(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +return x_3; +} +else +{ +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Awareness_getD___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_Awareness_getD(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedGenericFormat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedGenericFormat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instInhabitedGenericFormat(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_5, x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_11, x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_5, x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701_(x_8, x_9); +x_11 = l_List_foldl___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__4(x_2, x_10, x_4); +return x_11; +} +} +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[]", 2, 2); +return x_1; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5; +lean_inc(x_1); +x_5 = l_Std_Format_joinSep___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__3(x_1, x_4); +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_7 = lean_ctor_get(x_1, 1); +lean_dec(x_7); +x_8 = lean_ctor_get(x_1, 0); +lean_dec(x_8); +x_9 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9; +lean_ctor_set_tag(x_1, 5); +lean_ctor_set(x_1, 1, x_5); +lean_ctor_set(x_1, 0, x_9); +x_10 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8; +x_13 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = 0; +x_15 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set_uint8(x_15, sizeof(void*)*1, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +lean_dec(x_1); +x_16 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_5); +x_18 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8; +x_21 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = 0; +x_23 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set_uint8(x_23, sizeof(void*)*1, x_22); +return x_23; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("string", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1(x_1, x_3); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +x_6 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = 0; +x_8 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*1, x_7); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4; +x_10 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10; +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_7); +return x_17; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956_(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956_(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instReprGenericFormat___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprGenericFormat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instReprGenericFormat___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instReprGenericFormat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instReprGenericFormat(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("condition not satisfied", 23, 23); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_string_utf8_byte_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_7 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +uint32_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_string_utf8_get_fast(x_3, x_4); +x_10 = lean_box_uint32(x_9); +x_11 = lean_apply_1(x_1, x_10); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_4); +lean_dec(x_3); +x_13 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_2, 1); +lean_dec(x_16); +x_17 = lean_ctor_get(x_2, 0); +lean_dec(x_17); +x_18 = lean_string_utf8_next_fast(x_3, x_4); +lean_dec(x_4); +lean_ctor_set(x_2, 1, x_18); +x_19 = lean_box_uint32(x_9); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_2); +x_21 = lean_string_utf8_next_fast(x_3, x_4); +lean_dec(x_4); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_3); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_box_uint32(x_9); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1(uint32_t x_1, uint32_t x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; +x_3 = lean_uint32_dec_eq(x_2, x_1); +x_4 = l_instDecidableNot___rarg(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_box_uint32(x_1); +x_5 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1___boxed), 2, 1); +lean_closure_set(x_5, 0, x_4); +lean_inc(x_3); +x_6 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(x_5, x_3); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; uint32_t x_9; lean_object* x_10; +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_unbox_uint32(x_8); +lean_dec(x_8); +x_10 = lean_string_push(x_2, x_9); +x_2 = x_10; +x_3 = x_7; +goto _start; +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_6, 0); +x_14 = lean_ctor_get(x_6, 1); +x_15 = lean_ctor_get(x_3, 1); +lean_inc(x_15); +lean_dec(x_3); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_2); +return x_6; +} +else +{ +lean_dec(x_14); +lean_ctor_set_tag(x_6, 0); +lean_ctor_set(x_6, 1, x_2); +return x_6; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_6, 0); +x_19 = lean_ctor_get(x_6, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_6); +x_20 = lean_ctor_get(x_3, 1); +lean_inc(x_20); +lean_dec(x_3); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +x_22 = lean_nat_dec_eq(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_2); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_19); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_19); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_2); +return x_24; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_apply_1(x_1, x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 1); +x_7 = lean_apply_1(x_2, x_6); +lean_ctor_set(x_4, 1, x_7); +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_4); +x_10 = lean_apply_1(x_2, x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +lean_dec(x_2); +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) +{ +return x_4; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_box_uint32(x_1); +x_5 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1___boxed), 2, 1); +lean_closure_set(x_5, 0, x_4); +lean_inc(x_3); +x_6 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(x_5, x_3); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; uint32_t x_9; lean_object* x_10; +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_unbox_uint32(x_8); +lean_dec(x_8); +x_10 = lean_string_push(x_2, x_9); +x_2 = x_10; +x_3 = x_7; +goto _start; +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_6); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_6, 0); +x_14 = lean_ctor_get(x_6, 1); +x_15 = lean_ctor_get(x_3, 1); +lean_inc(x_15); +lean_dec(x_3); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_2); +return x_6; +} +else +{ +lean_dec(x_14); +lean_ctor_set_tag(x_6, 0); +lean_ctor_set(x_6, 1, x_2); +return x_6; +} +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_6, 0); +x_19 = lean_ctor_get(x_6, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_6); +x_20 = lean_ctor_get(x_3, 1); +lean_inc(x_20); +lean_dec(x_3); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +x_22 = lean_nat_dec_eq(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_2); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_18); +lean_ctor_set(x_23, 1, x_19); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_19); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_2); +return x_24; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_apply_1(x_1, x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 1); +x_7 = lean_apply_1(x_2, x_6); +lean_ctor_set(x_4, 1, x_7); +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_4); +x_10 = lean_apply_1(x_2, x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +lean_dec(x_2); +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) +{ +return x_4; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_any___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_string_utf8_byte_size(x_2); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_3); +lean_dec(x_2); +x_6 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; uint32_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_1, 1); +lean_dec(x_9); +x_10 = lean_ctor_get(x_1, 0); +lean_dec(x_10); +x_11 = lean_string_utf8_get_fast(x_2, x_3); +x_12 = lean_string_utf8_next_fast(x_2, x_3); +lean_dec(x_3); +lean_ctor_set(x_1, 1, x_12); +x_13 = lean_box_uint32(x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +uint32_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_1); +x_15 = lean_string_utf8_get_fast(x_2, x_3); +x_16 = lean_string_utf8_next_fast(x_2, x_3); +lean_dec(x_3); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_2); +lean_ctor_set(x_17, 1, x_16); +x_18 = lean_box_uint32(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("expected: '", 11, 11); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1(uint32_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_string_utf8_byte_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_7 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +uint32_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_string_utf8_get_fast(x_3, x_4); +x_10 = lean_string_utf8_next_fast(x_3, x_4); +lean_dec(x_4); +lean_inc(x_10); +lean_inc(x_3); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_uint32_dec_eq(x_9, x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_3); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_14 = lean_string_push(x_13, x_1); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_16 = lean_string_append(x_15, x_14); +lean_dec(x_14); +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_2); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_2, 1); +lean_dec(x_21); +x_22 = lean_ctor_get(x_2, 0); +lean_dec(x_22); +x_23 = lean_nat_dec_lt(x_10, x_5); +lean_dec(x_5); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_free_object(x_2); +lean_dec(x_10); +lean_dec(x_3); +x_24 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_11); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +else +{ +uint32_t x_26; uint8_t x_27; uint8_t x_28; +x_26 = lean_string_utf8_get_fast(x_3, x_10); +x_27 = lean_uint32_dec_eq(x_26, x_1); +x_28 = l_instDecidableNot___rarg(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_2); +lean_dec(x_10); +lean_dec(x_3); +x_29 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_11); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_11); +x_31 = lean_string_utf8_next_fast(x_3, x_10); +lean_dec(x_10); +lean_ctor_set(x_2, 1, x_31); +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_33 = lean_string_push(x_32, x_26); +x_34 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2(x_1, x_33, x_2); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_36 = lean_ctor_get(x_34, 0); +x_37 = lean_ctor_get(x_34, 1); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +x_40 = lean_string_utf8_byte_size(x_38); +x_41 = lean_nat_dec_lt(x_39, x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +lean_object* x_42; +lean_dec(x_39); +lean_dec(x_38); +lean_dec(x_37); +x_42 = l_Std_Internal_Parsec_unexpectedEndOfInput; +lean_ctor_set_tag(x_34, 1); +lean_ctor_set(x_34, 1, x_42); +return x_34; +} +else +{ +uint32_t x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_string_utf8_get_fast(x_38, x_39); +x_44 = lean_string_utf8_next_fast(x_38, x_39); +lean_dec(x_39); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_38); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_uint32_dec_eq(x_43, x_1); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_45); +lean_dec(x_37); +x_47 = lean_string_push(x_32, x_1); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_49 = lean_string_append(x_48, x_47); +lean_dec(x_47); +x_50 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_51 = lean_string_append(x_49, x_50); +lean_ctor_set_tag(x_34, 1); +lean_ctor_set(x_34, 1, x_51); +return x_34; +} +else +{ +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_45); +return x_34; +} +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_52 = lean_ctor_get(x_34, 0); +x_53 = lean_ctor_get(x_34, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_34); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +x_56 = lean_string_utf8_byte_size(x_54); +x_57 = lean_nat_dec_lt(x_55, x_56); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_53); +x_58 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +else +{ +uint32_t x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_string_utf8_get_fast(x_54, x_55); +x_61 = lean_string_utf8_next_fast(x_54, x_55); +lean_dec(x_55); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_54); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_uint32_dec_eq(x_60, x_1); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_62); +lean_dec(x_53); +x_64 = lean_string_push(x_32, x_1); +x_65 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_66 = lean_string_append(x_65, x_64); +lean_dec(x_64); +x_67 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_68 = lean_string_append(x_66, x_67); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_52); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +else +{ +lean_object* x_70; +lean_dec(x_52); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_62); +lean_ctor_set(x_70, 1, x_53); +return x_70; +} +} +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_34); +if (x_71 == 0) +{ +return x_34; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_34, 0); +x_73 = lean_ctor_get(x_34, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_34); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_2); +x_75 = lean_nat_dec_lt(x_10, x_5); +lean_dec(x_5); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_10); +lean_dec(x_3); +x_76 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_11); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +else +{ +uint32_t x_78; uint8_t x_79; uint8_t x_80; +x_78 = lean_string_utf8_get_fast(x_3, x_10); +x_79 = lean_uint32_dec_eq(x_78, x_1); +x_80 = l_instDecidableNot___rarg(x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +lean_dec(x_10); +lean_dec(x_3); +x_81 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_11); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_11); +x_83 = lean_string_utf8_next_fast(x_3, x_10); +lean_dec(x_10); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_3); +lean_ctor_set(x_84, 1, x_83); +x_85 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_86 = lean_string_push(x_85, x_78); +x_87 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2(x_1, x_86, x_84); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +x_91 = lean_ctor_get(x_88, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_88, 1); +lean_inc(x_92); +x_93 = lean_string_utf8_byte_size(x_91); +x_94 = lean_nat_dec_lt(x_92, x_93); +lean_dec(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_92); +lean_dec(x_91); +lean_dec(x_89); +x_95 = l_Std_Internal_Parsec_unexpectedEndOfInput; +if (lean_is_scalar(x_90)) { + x_96 = lean_alloc_ctor(1, 2, 0); +} else { + x_96 = x_90; + lean_ctor_set_tag(x_96, 1); +} +lean_ctor_set(x_96, 0, x_88); +lean_ctor_set(x_96, 1, x_95); +return x_96; +} +else +{ +uint32_t x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_97 = lean_string_utf8_get_fast(x_91, x_92); +x_98 = lean_string_utf8_next_fast(x_91, x_92); +lean_dec(x_92); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_91); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_uint32_dec_eq(x_97, x_1); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_99); +lean_dec(x_89); +x_101 = lean_string_push(x_85, x_1); +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_103 = lean_string_append(x_102, x_101); +lean_dec(x_101); +x_104 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_105 = lean_string_append(x_103, x_104); +if (lean_is_scalar(x_90)) { + x_106 = lean_alloc_ctor(1, 2, 0); +} else { + x_106 = x_90; + lean_ctor_set_tag(x_106, 1); +} +lean_ctor_set(x_106, 0, x_88); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +else +{ +lean_object* x_107; +lean_dec(x_88); +if (lean_is_scalar(x_90)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_90; +} +lean_ctor_set(x_107, 0, x_99); +lean_ctor_set(x_107, 1, x_89); +return x_107; +} +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_87, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_87, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_110 = x_87; +} else { + lean_dec_ref(x_87); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3(uint32_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_string_utf8_byte_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_7 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +uint32_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_string_utf8_get_fast(x_3, x_4); +x_10 = lean_string_utf8_next_fast(x_3, x_4); +lean_dec(x_4); +lean_inc(x_10); +lean_inc(x_3); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_uint32_dec_eq(x_9, x_1); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_5); +lean_dec(x_3); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_14 = lean_string_push(x_13, x_1); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_16 = lean_string_append(x_15, x_14); +lean_dec(x_14); +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_18 = lean_string_append(x_16, x_17); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_2); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_2, 1); +lean_dec(x_21); +x_22 = lean_ctor_get(x_2, 0); +lean_dec(x_22); +x_23 = lean_nat_dec_lt(x_10, x_5); +lean_dec(x_5); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_free_object(x_2); +lean_dec(x_10); +lean_dec(x_3); +x_24 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_11); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +else +{ +uint32_t x_26; uint8_t x_27; uint8_t x_28; +x_26 = lean_string_utf8_get_fast(x_3, x_10); +x_27 = lean_uint32_dec_eq(x_26, x_1); +x_28 = l_instDecidableNot___rarg(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_2); +lean_dec(x_10); +lean_dec(x_3); +x_29 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_11); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_11); +x_31 = lean_string_utf8_next_fast(x_3, x_10); +lean_dec(x_10); +lean_ctor_set(x_2, 1, x_31); +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_33 = lean_string_push(x_32, x_26); +x_34 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4(x_1, x_33, x_2); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_36 = lean_ctor_get(x_34, 0); +x_37 = lean_ctor_get(x_34, 1); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +x_40 = lean_string_utf8_byte_size(x_38); +x_41 = lean_nat_dec_lt(x_39, x_40); +lean_dec(x_40); +if (x_41 == 0) +{ +lean_object* x_42; +lean_dec(x_39); +lean_dec(x_38); +lean_dec(x_37); +x_42 = l_Std_Internal_Parsec_unexpectedEndOfInput; +lean_ctor_set_tag(x_34, 1); +lean_ctor_set(x_34, 1, x_42); +return x_34; +} +else +{ +uint32_t x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_string_utf8_get_fast(x_38, x_39); +x_44 = lean_string_utf8_next_fast(x_38, x_39); +lean_dec(x_39); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_38); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_uint32_dec_eq(x_43, x_1); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_45); +lean_dec(x_37); +x_47 = lean_string_push(x_32, x_1); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_49 = lean_string_append(x_48, x_47); +lean_dec(x_47); +x_50 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_51 = lean_string_append(x_49, x_50); +lean_ctor_set_tag(x_34, 1); +lean_ctor_set(x_34, 1, x_51); +return x_34; +} +else +{ +lean_dec(x_36); +lean_ctor_set(x_34, 0, x_45); +return x_34; +} +} +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_52 = lean_ctor_get(x_34, 0); +x_53 = lean_ctor_get(x_34, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_34); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_52, 1); +lean_inc(x_55); +x_56 = lean_string_utf8_byte_size(x_54); +x_57 = lean_nat_dec_lt(x_55, x_56); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_53); +x_58 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_52); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +else +{ +uint32_t x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_string_utf8_get_fast(x_54, x_55); +x_61 = lean_string_utf8_next_fast(x_54, x_55); +lean_dec(x_55); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_54); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_uint32_dec_eq(x_60, x_1); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_62); +lean_dec(x_53); +x_64 = lean_string_push(x_32, x_1); +x_65 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_66 = lean_string_append(x_65, x_64); +lean_dec(x_64); +x_67 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_68 = lean_string_append(x_66, x_67); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_52); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +else +{ +lean_object* x_70; +lean_dec(x_52); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_62); +lean_ctor_set(x_70, 1, x_53); +return x_70; +} +} +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_34); +if (x_71 == 0) +{ +return x_34; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_34, 0); +x_73 = lean_ctor_get(x_34, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_34); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +} +} +else +{ +uint8_t x_75; +lean_dec(x_2); +x_75 = lean_nat_dec_lt(x_10, x_5); +lean_dec(x_5); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; +lean_dec(x_10); +lean_dec(x_3); +x_76 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_11); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +else +{ +uint32_t x_78; uint8_t x_79; uint8_t x_80; +x_78 = lean_string_utf8_get_fast(x_3, x_10); +x_79 = lean_uint32_dec_eq(x_78, x_1); +x_80 = l_instDecidableNot___rarg(x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +lean_dec(x_10); +lean_dec(x_3); +x_81 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_82 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_82, 0, x_11); +lean_ctor_set(x_82, 1, x_81); +return x_82; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_11); +x_83 = lean_string_utf8_next_fast(x_3, x_10); +lean_dec(x_10); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_3); +lean_ctor_set(x_84, 1, x_83); +x_85 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_86 = lean_string_push(x_85, x_78); +x_87 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4(x_1, x_86, x_84); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_90 = x_87; +} else { + lean_dec_ref(x_87); + x_90 = lean_box(0); +} +x_91 = lean_ctor_get(x_88, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_88, 1); +lean_inc(x_92); +x_93 = lean_string_utf8_byte_size(x_91); +x_94 = lean_nat_dec_lt(x_92, x_93); +lean_dec(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_92); +lean_dec(x_91); +lean_dec(x_89); +x_95 = l_Std_Internal_Parsec_unexpectedEndOfInput; +if (lean_is_scalar(x_90)) { + x_96 = lean_alloc_ctor(1, 2, 0); +} else { + x_96 = x_90; + lean_ctor_set_tag(x_96, 1); +} +lean_ctor_set(x_96, 0, x_88); +lean_ctor_set(x_96, 1, x_95); +return x_96; +} +else +{ +uint32_t x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_97 = lean_string_utf8_get_fast(x_91, x_92); +x_98 = lean_string_utf8_next_fast(x_91, x_92); +lean_dec(x_92); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_91); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_uint32_dec_eq(x_97, x_1); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_99); +lean_dec(x_89); +x_101 = lean_string_push(x_85, x_1); +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_103 = lean_string_append(x_102, x_101); +lean_dec(x_101); +x_104 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_105 = lean_string_append(x_103, x_104); +if (lean_is_scalar(x_90)) { + x_106 = lean_alloc_ctor(1, 2, 0); +} else { + x_106 = x_90; + lean_ctor_set_tag(x_106, 1); +} +lean_ctor_set(x_106, 0, x_88); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +else +{ +lean_object* x_107; +lean_dec(x_88); +if (lean_is_scalar(x_90)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_90; +} +lean_ctor_set(x_107, 0, x_99); +lean_ctor_set(x_107, 1, x_89); +return x_107; +} +} +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_87, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_87, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_110 = x_87; +} else { + lean_dec_ref(x_87); + x_110 = lean_box(0); +} +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); +} else { + x_111 = x_110; +} +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; +} +} +} +} +} +} +} +} +static uint8_t _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__1() { +_start: +{ +uint8_t x_1; uint8_t x_2; +x_1 = 0; +x_2 = l_instDecidableNot___rarg(x_1); +return x_2; +} +} +static uint8_t _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__2() { +_start: +{ +uint8_t x_1; uint8_t x_2; +x_1 = 1; +x_2 = l_instDecidableNot___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4(uint32_t x_1, uint32_t x_2, uint32_t x_3) { +_start: +{ +uint8_t x_4; uint8_t x_13; uint32_t x_22; uint8_t x_23; +x_22 = 65; +x_23 = lean_uint32_dec_le(x_22, x_3); +if (x_23 == 0) +{ +uint32_t x_24; uint8_t x_25; +x_24 = 97; +x_25 = lean_uint32_dec_le(x_24, x_3); +if (x_25 == 0) +{ +uint8_t x_26; +x_26 = 0; +x_4 = x_26; +goto block_12; +} +else +{ +uint32_t x_27; uint8_t x_28; +x_27 = 122; +x_28 = lean_uint32_dec_le(x_3, x_27); +if (x_28 == 0) +{ +x_4 = x_28; +goto block_12; +} +else +{ +x_13 = x_28; +goto block_21; +} +} +} +else +{ +uint32_t x_29; uint8_t x_30; +x_29 = 90; +x_30 = lean_uint32_dec_le(x_3, x_29); +if (x_30 == 0) +{ +uint32_t x_31; uint8_t x_32; +x_31 = 97; +x_32 = lean_uint32_dec_le(x_31, x_3); +if (x_32 == 0) +{ +uint8_t x_33; +x_33 = 0; +x_4 = x_33; +goto block_12; +} +else +{ +uint32_t x_34; uint8_t x_35; +x_34 = 122; +x_35 = lean_uint32_dec_le(x_3, x_34); +if (x_35 == 0) +{ +x_4 = x_35; +goto block_12; +} +else +{ +x_13 = x_35; +goto block_21; +} +} +} +else +{ +uint8_t x_36; +x_36 = 1; +x_13 = x_36; +goto block_21; +} +} +block_12: +{ +uint8_t x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__1; +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +uint8_t x_7; uint8_t x_8; +x_7 = lean_uint32_dec_eq(x_3, x_1); +x_8 = l_instDecidableNot___rarg(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +x_9 = 0; +return x_9; +} +else +{ +uint8_t x_10; uint8_t x_11; +x_10 = lean_uint32_dec_eq(x_3, x_2); +x_11 = l_instDecidableNot___rarg(x_10); +return x_11; +} +} +} +block_21: +{ +uint8_t x_14; +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__2; +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = 0; +return x_15; +} +else +{ +uint8_t x_16; uint8_t x_17; +x_16 = lean_uint32_dec_eq(x_3, x_1); +x_17 = l_instDecidableNot___rarg(x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +else +{ +uint8_t x_19; uint8_t x_20; +x_19 = lean_uint32_dec_eq(x_3, x_2); +x_20 = l_instDecidableNot___rarg(x_19); +return x_20; +} +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__2), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2() { +_start: +{ +uint32_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 92; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = lean_string_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_2 = lean_alloc_closure((void*)(l_String_push___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5; +x_3 = lean_alloc_closure((void*)(l_Function_comp___rarg), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_any___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__6), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 39; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = 34; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart(lean_object* x_1) { +_start: +{ +lean_object* x_2; +lean_inc(x_1); +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier(x_1); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +lean_dec(x_1); +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_2, 1, x_5); +return x_2; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +lean_inc(x_7); +lean_inc(x_6); +lean_dec(x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_6); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + x_12 = x_2; +} else { + lean_dec_ref(x_2); + x_12 = lean_box(0); +} +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_14); +if (lean_is_scalar(x_12)) { + x_16 = lean_alloc_ctor(1, 2, 0); +} else { + x_16 = x_12; +} +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_11); +return x_16; +} +else +{ +uint32_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +lean_dec(x_11); +x_17 = 92; +x_90 = lean_ctor_get(x_10, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_10, 1); +lean_inc(x_91); +x_92 = lean_string_utf8_byte_size(x_90); +x_93 = lean_nat_dec_lt(x_91, x_92); +lean_dec(x_92); +if (x_93 == 0) +{ +lean_object* x_94; +lean_dec(x_91); +lean_dec(x_90); +x_94 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_18 = x_10; +x_19 = x_94; +goto block_89; +} +else +{ +uint32_t x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_95 = lean_string_utf8_get_fast(x_90, x_91); +x_96 = lean_string_utf8_next_fast(x_90, x_91); +lean_dec(x_91); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_90); +lean_ctor_set(x_97, 1, x_96); +x_98 = lean_uint32_dec_eq(x_95, x_17); +if (x_98 == 0) +{ +lean_object* x_99; +lean_dec(x_97); +x_99 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4; +x_18 = x_10; +x_19 = x_99; +goto block_89; +} +else +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_10); +x_100 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7; +x_101 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6; +x_102 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(x_100, x_101, x_97); +if (lean_obj_tag(x_102) == 0) +{ +uint8_t x_103; +lean_dec(x_14); +lean_dec(x_12); +x_103 = !lean_is_exclusive(x_102); +if (x_103 == 0) +{ +return x_102; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_102, 0); +x_105 = lean_ctor_get(x_102, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_102); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} +else +{ +lean_object* x_107; lean_object* x_108; +x_107 = lean_ctor_get(x_102, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_102, 1); +lean_inc(x_108); +lean_dec(x_102); +x_18 = x_107; +x_19 = x_108; +goto block_89; +} +} +} +block_89: +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +x_21 = lean_nat_dec_eq(x_14, x_20); +lean_dec(x_14); +if (x_21 == 0) +{ +lean_object* x_22; +lean_dec(x_20); +if (lean_is_scalar(x_12)) { + x_22 = lean_alloc_ctor(1, 2, 0); +} else { + x_22 = x_12; +} +lean_ctor_set(x_22, 0, x_18); +lean_ctor_set(x_22, 1, x_19); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_19); +lean_dec(x_12); +x_23 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2; +x_24 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___boxed), 2, 1); +lean_closure_set(x_24, 0, x_23); +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1; +x_26 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(x_24, x_25, x_18); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +lean_dec(x_20); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +return x_26; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +x_35 = lean_nat_dec_eq(x_20, x_34); +lean_dec(x_20); +if (x_35 == 0) +{ +lean_dec(x_34); +return x_26; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_free_object(x_26); +lean_dec(x_33); +x_36 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +x_37 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3___boxed), 2, 1); +lean_closure_set(x_37, 0, x_36); +x_38 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(x_37, x_25, x_32); +if (lean_obj_tag(x_38) == 0) +{ +uint8_t x_39; +lean_dec(x_34); +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +return x_38; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_38, 0); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_38); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_38); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_38, 0); +x_45 = lean_ctor_get(x_38, 1); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +x_47 = lean_nat_dec_eq(x_34, x_46); +lean_dec(x_46); +lean_dec(x_34); +if (x_47 == 0) +{ +return x_38; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_free_object(x_38); +lean_dec(x_45); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +x_49 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2; +x_50 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___boxed), 3, 2); +lean_closure_set(x_50, 0, x_48); +lean_closure_set(x_50, 1, x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1), 2, 1); +lean_closure_set(x_51, 0, x_50); +x_52 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_52, 0, x_51); +x_53 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg(x_52, x_25, x_44); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_54 = lean_ctor_get(x_38, 0); +x_55 = lean_ctor_get(x_38, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_38); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +x_57 = lean_nat_dec_eq(x_34, x_56); +lean_dec(x_56); +lean_dec(x_34); +if (x_57 == 0) +{ +lean_object* x_58; +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_54); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_55); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +x_60 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2; +x_61 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___boxed), 3, 2); +lean_closure_set(x_61, 0, x_59); +lean_closure_set(x_61, 1, x_60); +x_62 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1), 2, 1); +lean_closure_set(x_62, 0, x_61); +x_63 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_63, 0, x_62); +x_64 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg(x_63, x_25, x_54); +return x_64; +} +} +} +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_65 = lean_ctor_get(x_26, 0); +x_66 = lean_ctor_get(x_26, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_26); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +x_68 = lean_nat_dec_eq(x_20, x_67); +lean_dec(x_20); +if (x_68 == 0) +{ +lean_object* x_69; +lean_dec(x_67); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_65); +lean_ctor_set(x_69, 1, x_66); +return x_69; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_dec(x_66); +x_70 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +x_71 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3___boxed), 2, 1); +lean_closure_set(x_71, 0, x_70); +x_72 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__3___rarg(x_71, x_25, x_65); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_67); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; +} else { + lean_dec_ref(x_72); + x_75 = lean_box(0); +} +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_75; +} +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_74); +return x_76; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +x_77 = lean_ctor_get(x_72, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_72, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_79 = x_72; +} else { + lean_dec_ref(x_72); + x_79 = lean_box(0); +} +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +x_81 = lean_nat_dec_eq(x_67, x_80); +lean_dec(x_80); +lean_dec(x_67); +if (x_81 == 0) +{ +lean_object* x_82; +if (lean_is_scalar(x_79)) { + x_82 = lean_alloc_ctor(1, 2, 0); +} else { + x_82 = x_79; +} +lean_ctor_set(x_82, 0, x_77); +lean_ctor_set(x_82, 1, x_78); +return x_82; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_79); +lean_dec(x_78); +x_83 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1; +x_84 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2; +x_85 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___boxed), 3, 2); +lean_closure_set(x_85, 0, x_83); +lean_closure_set(x_85, 1, x_84); +x_86 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1), 2, 1); +lean_closure_set(x_86, 0, x_85); +x_87 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_many1Chars___at___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___spec__1), 2, 1); +lean_closure_set(x_87, 0, x_86); +x_88 = l_Functor_mapRev___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__5___rarg(x_87, x_25, x_77); +return x_88; +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; uint32_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___lambda__1(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__2(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__4(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__3(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; uint32_t x_5; uint32_t x_6; uint8_t x_7; lean_object* x_8; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_6 = lean_unbox_uint32(x_3); +lean_dec(x_3); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4(x_4, x_5, x_6); +x_8 = lean_box(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Format_Basic_0__Std_Time_specParser___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +lean_inc(x_2); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_array_push(x_1, x_5); +x_1 = x_6; +x_2 = x_4; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_3); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +x_13 = lean_nat_dec_eq(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_dec(x_1); +return x_3; +} +else +{ +lean_dec(x_10); +lean_ctor_set_tag(x_3, 0); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_14 = lean_ctor_get(x_3, 0); +x_15 = lean_ctor_get(x_3, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_3); +x_16 = lean_ctor_get(x_2, 1); +lean_inc(x_16); +lean_dec(x_2); +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +x_18 = lean_nat_dec_eq(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_1); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_14); +lean_ctor_set(x_19, 1, x_15); +return x_19; +} +else +{ +lean_object* x_20; +lean_dec(x_15); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_14); +lean_ctor_set(x_20, 1, x_1); +return x_20; +} +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParser(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1; +x_3 = l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Format_Basic_0__Std_Time_specParser___spec__1(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +x_7 = lean_array_to_list(x_6); +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +x_10 = lean_string_utf8_byte_size(x_8); +lean_dec(x_8); +x_11 = lean_nat_dec_lt(x_9, x_10); +lean_dec(x_10); +lean_dec(x_9); +if (x_11 == 0) +{ +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_12; +lean_dec(x_7); +x_12 = l_Std_Internal_Parsec_expectedEndOfInput; +lean_ctor_set_tag(x_3, 1); +lean_ctor_set(x_3, 1, x_12); +return x_3; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_3); +x_15 = lean_array_to_list(x_14); +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_13, 1); +lean_inc(x_17); +x_18 = lean_string_utf8_byte_size(x_16); +lean_dec(x_16); +x_19 = lean_nat_dec_lt(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +if (x_19 == 0) +{ +lean_object* x_20; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_15); +x_21 = l_Std_Internal_Parsec_expectedEndOfInput; +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_13); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_3); +if (x_23 == 0) +{ +return x_3; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_3, 0); +x_25 = lean_ctor_get(x_3, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_3); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_specParser), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_specParse(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1; +x_3 = l_Std_Internal_Parsec_String_Parser_run___rarg(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_2, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_2, x_6); +lean_dec(x_2); +x_8 = lean_string_push(x_3, x_1); +x_2 = x_7; +x_3 = x_8; +goto _start; +} +else +{ +lean_dec(x_2); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(lean_object* x_1, uint32_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_string_length(x_3); +x_5 = lean_nat_sub(x_1, x_4); +lean_dec(x_4); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_7 = l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1(x_2, x_5, x_6); +x_8 = lean_string_append(x_7, x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_leftPad___spec__1(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_leftPad___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_1, x_4, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1(uint32_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_eq(x_2, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_2, x_6); +lean_dec(x_2); +x_8 = lean_string_push(x_3, x_1); +x_2 = x_7; +x_3 = x_8; +goto _start; +} +else +{ +lean_dec(x_2); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(lean_object* x_1, uint32_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_string_length(x_3); +x_5 = lean_nat_sub(x_1, x_4); +lean_dec(x_4); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_7 = l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1(x_2, x_5, x_6); +x_8 = lean_string_append(x_3, x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_5 = l_Nat_repeatTR_loop___at___private_Std_Time_Format_Basic_0__Std_Time_rightPad___spec__1(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightPad___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(x_1, x_4, x_3); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_pad(lean_object* x_1, lean_object* x_2, uint8_t x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_46; uint8_t x_47; +x_46 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_47 = lean_int_dec_lt(x_2, x_46); +if (x_47 == 0) +{ +lean_object* x_48; +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_4 = x_48; +x_5 = x_2; +goto block_45; +} +else +{ +lean_object* x_49; lean_object* x_50; +x_49 = lean_int_neg(x_2); +lean_dec(x_2); +x_50 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_4 = x_50; +x_5 = x_49; +goto block_45; +} +block_45: +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_7 = lean_int_dec_lt(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_nat_abs(x_5); +lean_dec(x_5); +x_9 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_10 = lean_string_length(x_9); +x_11 = lean_nat_dec_lt(x_1, x_10); +if (x_11 == 0) +{ +uint32_t x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_10); +x_12 = 48; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_1, x_12, x_9); +lean_dec(x_9); +x_14 = lean_string_append(x_4, x_13); +lean_dec(x_13); +return x_14; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_15; +lean_dec(x_10); +x_15 = lean_string_append(x_4, x_9); +lean_dec(x_9); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_16 = lean_nat_sub(x_10, x_1); +lean_dec(x_10); +x_17 = lean_string_utf8_byte_size(x_9); +x_18 = lean_unsigned_to_nat(0u); +lean_inc(x_17); +lean_inc(x_9); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_9); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_19, 2, x_17); +x_20 = l_Substring_nextn(x_19, x_16, x_18); +lean_dec(x_19); +x_21 = lean_nat_add(x_18, x_20); +lean_dec(x_20); +x_22 = lean_string_utf8_extract(x_9, x_21, x_17); +lean_dec(x_17); +lean_dec(x_21); +lean_dec(x_9); +x_23 = lean_string_append(x_4, x_22); +lean_dec(x_22); +return x_23; +} +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_24 = lean_nat_abs(x_5); +lean_dec(x_5); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_sub(x_24, x_25); +lean_dec(x_24); +x_27 = lean_nat_add(x_26, x_25); +lean_dec(x_26); +x_28 = l___private_Init_Data_Repr_0__Nat_reprFast(x_27); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_30 = lean_string_append(x_29, x_28); +lean_dec(x_28); +x_31 = lean_string_length(x_30); +x_32 = lean_nat_dec_lt(x_1, x_31); +if (x_32 == 0) +{ +uint32_t x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_31); +x_33 = 48; +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_1, x_33, x_30); +lean_dec(x_30); +x_35 = lean_string_append(x_4, x_34); +lean_dec(x_34); +return x_35; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_36; +lean_dec(x_31); +x_36 = lean_string_append(x_4, x_30); +lean_dec(x_30); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_37 = lean_nat_sub(x_31, x_1); +lean_dec(x_31); +x_38 = lean_string_utf8_byte_size(x_30); +x_39 = lean_unsigned_to_nat(0u); +lean_inc(x_38); +lean_inc(x_30); +x_40 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_40, 0, x_30); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = l_Substring_nextn(x_40, x_37, x_39); +lean_dec(x_40); +x_42 = lean_nat_add(x_39, x_41); +lean_dec(x_41); +x_43 = lean_string_utf8_extract(x_30, x_42, x_38); +lean_dec(x_38); +lean_dec(x_42); +lean_dec(x_30); +x_44 = lean_string_append(x_4, x_43); +lean_dec(x_43); +return x_44; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_pad___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_1, x_2, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate(lean_object* x_1, lean_object* x_2, uint8_t x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_44; uint8_t x_45; +x_44 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_45 = lean_int_dec_lt(x_2, x_44); +if (x_45 == 0) +{ +lean_object* x_46; +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_4 = x_46; +x_5 = x_2; +goto block_43; +} +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_int_neg(x_2); +lean_dec(x_2); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_4 = x_48; +x_5 = x_47; +goto block_43; +} +block_43: +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_7 = lean_int_dec_lt(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_nat_abs(x_5); +lean_dec(x_5); +x_9 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_10 = lean_string_length(x_9); +x_11 = lean_nat_dec_lt(x_1, x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +uint32_t x_12; lean_object* x_13; lean_object* x_14; +x_12 = 48; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(x_1, x_12, x_9); +lean_dec(x_1); +x_14 = lean_string_append(x_4, x_13); +lean_dec(x_13); +return x_14; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_15; +lean_dec(x_1); +x_15 = lean_string_append(x_4, x_9); +lean_dec(x_9); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_string_utf8_byte_size(x_9); +x_17 = lean_unsigned_to_nat(0u); +lean_inc(x_9); +x_18 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set(x_18, 2, x_16); +x_19 = l_Substring_nextn(x_18, x_1, x_17); +lean_dec(x_18); +x_20 = lean_nat_add(x_17, x_19); +lean_dec(x_19); +x_21 = lean_string_utf8_extract(x_9, x_17, x_20); +lean_dec(x_20); +lean_dec(x_9); +x_22 = lean_string_append(x_4, x_21); +lean_dec(x_21); +return x_22; +} +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_23 = lean_nat_abs(x_5); +lean_dec(x_5); +x_24 = lean_unsigned_to_nat(1u); +x_25 = lean_nat_sub(x_23, x_24); +lean_dec(x_23); +x_26 = lean_nat_add(x_25, x_24); +lean_dec(x_25); +x_27 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_28 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_29 = lean_string_append(x_28, x_27); +lean_dec(x_27); +x_30 = lean_string_length(x_29); +x_31 = lean_nat_dec_lt(x_1, x_30); +lean_dec(x_30); +if (x_31 == 0) +{ +uint32_t x_32; lean_object* x_33; lean_object* x_34; +x_32 = 48; +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(x_1, x_32, x_29); +lean_dec(x_1); +x_34 = lean_string_append(x_4, x_33); +lean_dec(x_33); +return x_34; +} +else +{ +if (x_3 == 0) +{ +lean_object* x_35; +lean_dec(x_1); +x_35 = lean_string_append(x_4, x_29); +lean_dec(x_29); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_36 = lean_string_utf8_byte_size(x_29); +x_37 = lean_unsigned_to_nat(0u); +lean_inc(x_29); +x_38 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_38, 0, x_29); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +x_39 = l_Substring_nextn(x_38, x_1, x_37); +lean_dec(x_38); +x_40 = lean_nat_add(x_37, x_39); +lean_dec(x_39); +x_41 = lean_string_utf8_extract(x_29, x_37, x_40); +lean_dec(x_40); +lean_dec(x_29); +x_42 = lean_string_append(x_4, x_41); +lean_dec(x_41); +return x_42; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate(x_1, x_2, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("December", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("November", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("October", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("September", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("August", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("July", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("June", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("May", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("April", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("March", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("February", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("January", 7, 7); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_nat_sub(x_9, x_4); +lean_dec(x_9); +x_12 = lean_nat_dec_eq(x_11, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_nat_sub(x_11, x_4); +lean_dec(x_11); +x_14 = lean_nat_dec_eq(x_13, x_3); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_nat_sub(x_13, x_4); +lean_dec(x_13); +x_16 = lean_nat_dec_eq(x_15, x_3); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_nat_sub(x_15, x_4); +lean_dec(x_15); +x_18 = lean_nat_dec_eq(x_17, x_3); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_nat_sub(x_17, x_4); +lean_dec(x_17); +x_20 = lean_nat_dec_eq(x_19, x_3); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_nat_sub(x_19, x_4); +lean_dec(x_19); +x_22 = lean_nat_dec_eq(x_21, x_3); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_nat_sub(x_21, x_4); +lean_dec(x_21); +x_24 = lean_nat_dec_eq(x_23, x_3); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_nat_sub(x_23, x_4); +lean_dec(x_23); +x_26 = lean_nat_dec_eq(x_25, x_3); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +return x_27; +} +else +{ +lean_object* x_28; +x_28 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +return x_28; +} +} +else +{ +lean_object* x_29; +lean_dec(x_23); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +return x_29; +} +} +else +{ +lean_object* x_30; +lean_dec(x_21); +x_30 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +return x_30; +} +} +else +{ +lean_object* x_31; +lean_dec(x_19); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +return x_31; +} +} +else +{ +lean_object* x_32; +lean_dec(x_17); +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +return x_32; +} +} +else +{ +lean_object* x_33; +lean_dec(x_15); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +return x_33; +} +} +else +{ +lean_object* x_34; +lean_dec(x_13); +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +return x_34; +} +} +else +{ +lean_object* x_35; +lean_dec(x_11); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +return x_35; +} +} +else +{ +lean_object* x_36; +lean_dec(x_9); +x_36 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10; +return x_36; +} +} +else +{ +lean_object* x_37; +lean_dec(x_7); +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11; +return x_37; +} +} +else +{ +lean_object* x_38; +lean_dec(x_5); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12; +return x_38; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Dec", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Nov", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Oct", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sep", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Aug", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Jul", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Jun", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Apr", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Mar", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Feb", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Jan", 3, 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_nat_sub(x_9, x_4); +lean_dec(x_9); +x_12 = lean_nat_dec_eq(x_11, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_nat_sub(x_11, x_4); +lean_dec(x_11); +x_14 = lean_nat_dec_eq(x_13, x_3); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_nat_sub(x_13, x_4); +lean_dec(x_13); +x_16 = lean_nat_dec_eq(x_15, x_3); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_nat_sub(x_15, x_4); +lean_dec(x_15); +x_18 = lean_nat_dec_eq(x_17, x_3); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_nat_sub(x_17, x_4); +lean_dec(x_17); +x_20 = lean_nat_dec_eq(x_19, x_3); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_nat_sub(x_19, x_4); +lean_dec(x_19); +x_22 = lean_nat_dec_eq(x_21, x_3); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_nat_sub(x_21, x_4); +lean_dec(x_21); +x_24 = lean_nat_dec_eq(x_23, x_3); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_nat_sub(x_23, x_4); +lean_dec(x_23); +x_26 = lean_nat_dec_eq(x_25, x_3); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +return x_27; +} +else +{ +lean_object* x_28; +x_28 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +return x_28; +} +} +else +{ +lean_object* x_29; +lean_dec(x_23); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +return x_29; +} +} +else +{ +lean_object* x_30; +lean_dec(x_21); +x_30 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +return x_30; +} +} +else +{ +lean_object* x_31; +lean_dec(x_19); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +return x_31; +} +} +else +{ +lean_object* x_32; +lean_dec(x_17); +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +return x_32; +} +} +else +{ +lean_object* x_33; +lean_dec(x_15); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +return x_33; +} +} +else +{ +lean_object* x_34; +lean_dec(x_13); +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +return x_34; +} +} +else +{ +lean_object* x_35; +lean_dec(x_11); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +return x_35; +} +} +else +{ +lean_object* x_36; +lean_dec(x_9); +x_36 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9; +return x_36; +} +} +else +{ +lean_object* x_37; +lean_dec(x_7); +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10; +return x_37; +} +} +else +{ +lean_object* x_38; +lean_dec(x_5); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11; +return x_38; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("D", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("N", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("O", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("S", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("A", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("J", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("M", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("F", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_nat_sub(x_9, x_4); +lean_dec(x_9); +x_12 = lean_nat_dec_eq(x_11, x_3); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_nat_sub(x_11, x_4); +lean_dec(x_11); +x_14 = lean_nat_dec_eq(x_13, x_3); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_nat_sub(x_13, x_4); +lean_dec(x_13); +x_16 = lean_nat_dec_eq(x_15, x_3); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_nat_sub(x_15, x_4); +lean_dec(x_15); +x_18 = lean_nat_dec_eq(x_17, x_3); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = lean_nat_sub(x_17, x_4); +lean_dec(x_17); +x_20 = lean_nat_dec_eq(x_19, x_3); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_nat_sub(x_19, x_4); +lean_dec(x_19); +x_22 = lean_nat_dec_eq(x_21, x_3); +if (x_22 == 0) +{ +lean_object* x_23; uint8_t x_24; +x_23 = lean_nat_sub(x_21, x_4); +lean_dec(x_21); +x_24 = lean_nat_dec_eq(x_23, x_3); +if (x_24 == 0) +{ +lean_object* x_25; uint8_t x_26; +x_25 = lean_nat_sub(x_23, x_4); +lean_dec(x_23); +x_26 = lean_nat_dec_eq(x_25, x_3); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +return x_27; +} +else +{ +lean_object* x_28; +x_28 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +return x_28; +} +} +else +{ +lean_object* x_29; +lean_dec(x_23); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +return x_29; +} +} +else +{ +lean_object* x_30; +lean_dec(x_21); +x_30 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +return x_30; +} +} +else +{ +lean_object* x_31; +lean_dec(x_19); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +return x_31; +} +} +else +{ +lean_object* x_32; +lean_dec(x_17); +x_32 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6; +return x_32; +} +} +else +{ +lean_object* x_33; +lean_dec(x_15); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6; +return x_33; +} +} +else +{ +lean_object* x_34; +lean_dec(x_13); +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +return x_34; +} +} +else +{ +lean_object* x_35; +lean_dec(x_11); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +return x_35; +} +} +else +{ +lean_object* x_36; +lean_dec(x_9); +x_36 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +return x_36; +} +} +else +{ +lean_object* x_37; +lean_dec(x_7); +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +return x_37; +} +} +else +{ +lean_object* x_38; +lean_dec(x_5); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6; +return x_38; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Monday", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tuesday", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Wednesday", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Thursday", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Friday", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Saturday", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sunday", 6, 6); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1; +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2; +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +return x_5; +} +case 4: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +return x_6; +} +case 5: +{ +lean_object* x_7; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +return x_7; +} +default: +{ +lean_object* x_8; +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Mon", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tue", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Wed", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Thu", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Fri", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sat", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Sun", 3, 3); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1; +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2; +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +return x_4; +} +case 3: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +return x_5; +} +case 4: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +return x_6; +} +case 5: +{ +lean_object* x_7; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +return x_7; +} +default: +{ +lean_object* x_8; +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("T", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("W", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(x_1); +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +return x_3; +} +case 2: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +return x_4; +} +case 4: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +return x_5; +} +case 5: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +return x_6; +} +case 6: +{ +lean_object* x_7; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +return x_7; +} +default: +{ +lean_object* x_8; +lean_dec(x_2); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("BCE", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("CE", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Before Common Era", 17, 17); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Common Era", 10, 10); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("B", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("C", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("4", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("3", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("2", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("1", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +return x_11; +} +else +{ +lean_object* x_12; +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2; +return x_12; +} +} +else +{ +lean_object* x_13; +lean_dec(x_7); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3; +return x_13; +} +} +else +{ +lean_object* x_14; +lean_dec(x_5); +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4; +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Q4", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Q3", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Q2", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Q1", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +return x_11; +} +else +{ +lean_object* x_12; +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2; +return x_12; +} +} +else +{ +lean_object* x_13; +lean_dec(x_7); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3; +return x_13; +} +} +else +{ +lean_object* x_14; +lean_dec(x_5); +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4; +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("4th quarter", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("3rd quarter", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("2nd quarter", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("1st quarter", 11, 11); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_nat_abs(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_unsigned_to_nat(1u); +x_5 = lean_nat_sub(x_2, x_4); +lean_dec(x_2); +x_6 = lean_nat_dec_eq(x_5, x_3); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_nat_sub(x_5, x_4); +lean_dec(x_5); +x_8 = lean_nat_dec_eq(x_7, x_3); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_nat_sub(x_7, x_4); +lean_dec(x_7); +x_10 = lean_nat_dec_eq(x_9, x_3); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +return x_11; +} +else +{ +lean_object* x_12; +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2; +return x_12; +} +} +else +{ +lean_object* x_13; +lean_dec(x_7); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3; +return x_13; +} +} +else +{ +lean_object* x_14; +lean_dec(x_5); +x_14 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4; +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("AM", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PM", 2, 2); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Ante Meridiem", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Post Meridiem", 13, 13); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("P", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow(x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("+", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toSigned(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_nat_abs(x_1); +x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1; +x_7 = lean_string_append(x_6, x_5); +lean_dec(x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_8 = lean_nat_abs(x_1); +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_8, x_9); +lean_dec(x_8); +x_11 = lean_nat_add(x_10, x_9); +lean_dec(x_10); +x_12 = l___private_Init_Data_Repr_0__Nat_reprFast(x_11); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_14 = lean_string_append(x_13, x_12); +lean_dec(x_12); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_toSigned(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(lean_object* x_1, uint8_t x_2, uint8_t x_3, uint8_t x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_127; uint8_t x_128; +x_127 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_128 = lean_int_dec_le(x_127, x_1); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; +x_129 = lean_int_neg(x_1); +lean_dec(x_1); +x_130 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_5 = x_130; +x_6 = x_129; +goto block_126; +} +else +{ +lean_object* x_131; +x_131 = l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1; +x_5 = x_131; +x_6 = x_1; +goto block_126; +} +block_126: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint32_t x_10; lean_object* x_11; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +x_8 = lean_int_mul(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_PlainTime_ofNanoseconds(x_8); +lean_dec(x_8); +x_10 = 48; +x_62 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_63 = lean_string_append(x_62, x_5); +lean_dec(x_5); +x_64 = lean_string_append(x_63, x_62); +x_114 = lean_ctor_get(x_9, 0); +lean_inc(x_114); +x_115 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_116 = lean_int_dec_lt(x_114, x_115); +if (x_116 == 0) +{ +lean_object* x_117; lean_object* x_118; +x_117 = lean_nat_abs(x_114); +lean_dec(x_114); +x_118 = l___private_Init_Data_Repr_0__Nat_reprFast(x_117); +x_65 = x_118; +goto block_113; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_119 = lean_nat_abs(x_114); +lean_dec(x_114); +x_120 = lean_unsigned_to_nat(1u); +x_121 = lean_nat_sub(x_119, x_120); +lean_dec(x_119); +x_122 = lean_nat_add(x_121, x_120); +lean_dec(x_121); +x_123 = l___private_Init_Data_Repr_0__Nat_reprFast(x_122); +x_124 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_125 = lean_string_append(x_124, x_123); +lean_dec(x_123); +x_65 = x_125; +goto block_113; +} +block_61: +{ +if (x_3 == 0) +{ +lean_dec(x_9); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_15 = lean_int_dec_eq(x_13, x_14); +x_16 = l_instDecidableNot___rarg(x_15); +if (x_16 == 0) +{ +lean_dec(x_13); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_18 = lean_string_append(x_17, x_11); +lean_dec(x_11); +x_19 = lean_string_append(x_18, x_17); +if (x_4 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_string_append(x_19, x_17); +x_21 = lean_string_append(x_20, x_17); +x_22 = lean_int_dec_lt(x_13, x_14); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_nat_abs(x_13); +lean_dec(x_13); +x_24 = l___private_Init_Data_Repr_0__Nat_reprFast(x_23); +x_25 = lean_unsigned_to_nat(2u); +x_26 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_25, x_10, x_24); +lean_dec(x_24); +x_27 = lean_string_append(x_21, x_26); +lean_dec(x_26); +x_28 = lean_string_append(x_27, x_17); +return x_28; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = lean_nat_abs(x_13); +lean_dec(x_13); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_sub(x_29, x_30); +lean_dec(x_29); +x_32 = lean_nat_add(x_31, x_30); +lean_dec(x_31); +x_33 = l___private_Init_Data_Repr_0__Nat_reprFast(x_32); +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_35 = lean_string_append(x_34, x_33); +lean_dec(x_33); +x_36 = lean_unsigned_to_nat(2u); +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_36, x_10, x_35); +lean_dec(x_35); +x_38 = lean_string_append(x_21, x_37); +lean_dec(x_37); +x_39 = lean_string_append(x_38, x_17); +return x_39; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2; +x_41 = lean_string_append(x_19, x_40); +x_42 = lean_string_append(x_41, x_17); +x_43 = lean_int_dec_lt(x_13, x_14); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_44 = lean_nat_abs(x_13); +lean_dec(x_13); +x_45 = l___private_Init_Data_Repr_0__Nat_reprFast(x_44); +x_46 = lean_unsigned_to_nat(2u); +x_47 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_46, x_10, x_45); +lean_dec(x_45); +x_48 = lean_string_append(x_42, x_47); +lean_dec(x_47); +x_49 = lean_string_append(x_48, x_17); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_50 = lean_nat_abs(x_13); +lean_dec(x_13); +x_51 = lean_unsigned_to_nat(1u); +x_52 = lean_nat_sub(x_50, x_51); +lean_dec(x_50); +x_53 = lean_nat_add(x_52, x_51); +lean_dec(x_52); +x_54 = l___private_Init_Data_Repr_0__Nat_reprFast(x_53); +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_56 = lean_string_append(x_55, x_54); +lean_dec(x_54); +x_57 = lean_unsigned_to_nat(2u); +x_58 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_57, x_10, x_56); +lean_dec(x_56); +x_59 = lean_string_append(x_42, x_58); +lean_dec(x_58); +x_60 = lean_string_append(x_59, x_17); +return x_60; +} +} +} +} +} +block_113: +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_unsigned_to_nat(2u); +x_67 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_66, x_10, x_65); +lean_dec(x_65); +x_68 = lean_string_append(x_64, x_67); +lean_dec(x_67); +x_69 = lean_string_append(x_68, x_62); +if (x_2 == 0) +{ +x_11 = x_69; +goto block_61; +} +else +{ +lean_object* x_70; lean_object* x_71; +x_70 = lean_string_append(x_62, x_69); +lean_dec(x_69); +x_71 = lean_string_append(x_70, x_62); +if (x_4 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_72 = lean_ctor_get(x_9, 1); +lean_inc(x_72); +x_73 = lean_string_append(x_71, x_62); +x_74 = lean_string_append(x_73, x_62); +x_75 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_76 = lean_int_dec_lt(x_72, x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_nat_abs(x_72); +lean_dec(x_72); +x_78 = l___private_Init_Data_Repr_0__Nat_reprFast(x_77); +x_79 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_66, x_10, x_78); +lean_dec(x_78); +x_80 = lean_string_append(x_74, x_79); +lean_dec(x_79); +x_81 = lean_string_append(x_80, x_62); +x_11 = x_81; +goto block_61; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_82 = lean_nat_abs(x_72); +lean_dec(x_72); +x_83 = lean_unsigned_to_nat(1u); +x_84 = lean_nat_sub(x_82, x_83); +lean_dec(x_82); +x_85 = lean_nat_add(x_84, x_83); +lean_dec(x_84); +x_86 = l___private_Init_Data_Repr_0__Nat_reprFast(x_85); +x_87 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_88 = lean_string_append(x_87, x_86); +lean_dec(x_86); +x_89 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_66, x_10, x_88); +lean_dec(x_88); +x_90 = lean_string_append(x_74, x_89); +lean_dec(x_89); +x_91 = lean_string_append(x_90, x_62); +x_11 = x_91; +goto block_61; +} +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; +x_92 = lean_ctor_get(x_9, 1); +lean_inc(x_92); +x_93 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2; +x_94 = lean_string_append(x_71, x_93); +x_95 = lean_string_append(x_94, x_62); +x_96 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_97 = lean_int_dec_lt(x_92, x_96); +if (x_97 == 0) +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_98 = lean_nat_abs(x_92); +lean_dec(x_92); +x_99 = l___private_Init_Data_Repr_0__Nat_reprFast(x_98); +x_100 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_66, x_10, x_99); +lean_dec(x_99); +x_101 = lean_string_append(x_95, x_100); +lean_dec(x_100); +x_102 = lean_string_append(x_101, x_62); +x_11 = x_102; +goto block_61; +} +else +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_103 = lean_nat_abs(x_92); +lean_dec(x_92); +x_104 = lean_unsigned_to_nat(1u); +x_105 = lean_nat_sub(x_103, x_104); +lean_dec(x_103); +x_106 = lean_nat_add(x_105, x_104); +lean_dec(x_105); +x_107 = l___private_Init_Data_Repr_0__Nat_reprFast(x_106); +x_108 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +x_109 = lean_string_append(x_108, x_107); +lean_dec(x_107); +x_110 = l___private_Std_Time_Format_Basic_0__Std_Time_leftPad(x_66, x_10, x_109); +lean_dec(x_109); +x_111 = lean_string_append(x_95, x_110); +lean_dec(x_110); +x_112 = lean_string_append(x_111, x_62); +x_11 = x_112; +goto block_61; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = lean_unbox(x_3); +lean_dec(x_3); +x_7 = lean_unbox(x_4); +lean_dec(x_4); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_1, x_5, x_6, x_7); +return x_8; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("GMT", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Z", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatWith(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_3; +x_3 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_3) { +case 0: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort(x_4); +return x_5; +} +case 1: +{ +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong(x_6); +return x_7; +} +default: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow(x_8); +return x_9; +} +} +} +case 1: +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_12 = lean_int_dec_le(x_2, x_11); +if (x_12 == 0) +{ +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +x_14 = lean_int_emod(x_2, x_13); +lean_dec(x_2); +x_15 = lean_unsigned_to_nat(2u); +x_16 = 0; +x_17 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_15, x_14, x_16); +return x_17; +} +case 1: +{ +lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_unsigned_to_nat(4u); +x_19 = 0; +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_18, x_2, x_19); +return x_20; +} +default: +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_10, 0); +lean_inc(x_21); +lean_dec(x_10); +x_22 = 0; +x_23 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_21, x_2, x_22); +lean_dec(x_21); +return x_23; +} +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_int_neg(x_2); +lean_dec(x_2); +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_26 = lean_int_add(x_24, x_25); +lean_dec(x_24); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_27 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +x_28 = lean_int_emod(x_26, x_27); +lean_dec(x_26); +x_29 = lean_unsigned_to_nat(2u); +x_30 = 0; +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_29, x_28, x_30); +return x_31; +} +case 1: +{ +lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_32 = lean_unsigned_to_nat(4u); +x_33 = 0; +x_34 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_32, x_26, x_33); +return x_34; +} +default: +{ +lean_object* x_35; uint8_t x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_10, 0); +lean_inc(x_35); +lean_dec(x_10); +x_36 = 0; +x_37 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_35, x_26, x_36); +lean_dec(x_35); +return x_37; +} +} +} +} +case 2: +{ +lean_object* x_38; +x_38 = lean_ctor_get(x_1, 0); +lean_inc(x_38); +lean_dec(x_1); +switch (lean_obj_tag(x_38)) { +case 0: +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; +x_39 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +x_40 = lean_int_emod(x_2, x_39); +lean_dec(x_2); +x_41 = lean_unsigned_to_nat(2u); +x_42 = 0; +x_43 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_41, x_40, x_42); +return x_43; +} +case 1: +{ +lean_object* x_44; uint8_t x_45; lean_object* x_46; +x_44 = lean_unsigned_to_nat(4u); +x_45 = 0; +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_44, x_2, x_45); +return x_46; +} +default: +{ +lean_object* x_47; uint8_t x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_38, 0); +lean_inc(x_47); +lean_dec(x_38); +x_48 = 0; +x_49 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_47, x_2, x_48); +lean_dec(x_47); +return x_49; +} +} +} +case 3: +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_1, 0); +lean_inc(x_50); +lean_dec(x_1); +x_51 = lean_ctor_get(x_2, 1); +lean_inc(x_51); +lean_dec(x_2); +x_52 = 0; +x_53 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_50, x_51, x_52); +lean_dec(x_50); +return x_53; +} +case 4: +{ +lean_object* x_54; +x_54 = lean_ctor_get(x_1, 0); +lean_inc(x_54); +lean_dec(x_1); +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_55; uint8_t x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +lean_dec(x_54); +x_56 = 0; +x_57 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_55, x_2, x_56); +lean_dec(x_55); +return x_57; +} +else +{ +lean_object* x_58; uint8_t x_59; +x_58 = lean_ctor_get(x_54, 0); +lean_inc(x_58); +lean_dec(x_54); +x_59 = lean_unbox(x_58); +lean_dec(x_58); +switch (x_59) { +case 0: +{ +lean_object* x_60; +x_60 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort(x_2); +lean_dec(x_2); +return x_60; +} +case 1: +{ +lean_object* x_61; +x_61 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong(x_2); +lean_dec(x_2); +return x_61; +} +default: +{ +lean_object* x_62; +x_62 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow(x_2); +lean_dec(x_2); +return x_62; +} +} +} +} +case 6: +{ +lean_object* x_63; +x_63 = lean_ctor_get(x_1, 0); +lean_inc(x_63); +lean_dec(x_1); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; uint8_t x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +lean_dec(x_63); +x_65 = 0; +x_66 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_64, x_2, x_65); +lean_dec(x_64); +return x_66; +} +else +{ +lean_object* x_67; uint8_t x_68; +x_67 = lean_ctor_get(x_63, 0); +lean_inc(x_67); +lean_dec(x_63); +x_68 = lean_unbox(x_67); +lean_dec(x_67); +switch (x_68) { +case 0: +{ +lean_object* x_69; +x_69 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort(x_2); +lean_dec(x_2); +return x_69; +} +case 1: +{ +lean_object* x_70; +x_70 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong(x_2); +lean_dec(x_2); +return x_70; +} +default: +{ +lean_object* x_71; +x_71 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber(x_2); +lean_dec(x_2); +return x_71; +} +} +} +} +case 9: +{ +uint8_t x_72; +x_72 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_72) { +case 0: +{ +uint8_t x_73; lean_object* x_74; +x_73 = lean_unbox(x_2); +lean_dec(x_2); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort(x_73); +return x_74; +} +case 1: +{ +uint8_t x_75; lean_object* x_76; +x_75 = lean_unbox(x_2); +lean_dec(x_2); +x_76 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong(x_75); +return x_76; +} +default: +{ +uint8_t x_77; lean_object* x_78; +x_77 = lean_unbox(x_2); +lean_dec(x_2); +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow(x_77); +return x_78; +} +} +} +case 10: +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_1, 0); +lean_inc(x_79); +lean_dec(x_1); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +lean_dec(x_79); +x_81 = lean_unbox(x_2); +lean_dec(x_2); +x_82 = l_Std_Time_Weekday_toOrdinal(x_81); +x_83 = 0; +x_84 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_80, x_82, x_83); +lean_dec(x_80); +return x_84; +} +else +{ +lean_object* x_85; uint8_t x_86; +x_85 = lean_ctor_get(x_79, 0); +lean_inc(x_85); +lean_dec(x_79); +x_86 = lean_unbox(x_85); +lean_dec(x_85); +switch (x_86) { +case 0: +{ +uint8_t x_87; lean_object* x_88; +x_87 = lean_unbox(x_2); +lean_dec(x_2); +x_88 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort(x_87); +return x_88; +} +case 1: +{ +uint8_t x_89; lean_object* x_90; +x_89 = lean_unbox(x_2); +lean_dec(x_2); +x_90 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong(x_89); +return x_90; +} +default: +{ +uint8_t x_91; lean_object* x_92; +x_91 = lean_unbox(x_2); +lean_dec(x_2); +x_92 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow(x_91); +return x_92; +} +} +} +} +case 12: +{ +uint8_t x_93; +x_93 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_93) { +case 0: +{ +uint8_t x_94; lean_object* x_95; +x_94 = lean_unbox(x_2); +lean_dec(x_2); +x_95 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort(x_94); +return x_95; +} +case 1: +{ +uint8_t x_96; lean_object* x_97; +x_96 = lean_unbox(x_2); +lean_dec(x_2); +x_97 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong(x_96); +return x_97; +} +default: +{ +uint8_t x_98; lean_object* x_99; +x_98 = lean_unbox(x_2); +lean_dec(x_2); +x_99 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow(x_98); +return x_99; +} +} +} +case 13: +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; lean_object* x_104; +x_100 = lean_ctor_get(x_1, 0); +lean_inc(x_100); +lean_dec(x_1); +x_101 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_102 = lean_int_emod(x_2, x_101); +lean_dec(x_2); +x_103 = 0; +x_104 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_100, x_102, x_103); +lean_dec(x_100); +return x_104; +} +case 14: +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_object* x_109; +x_105 = lean_ctor_get(x_1, 0); +lean_inc(x_105); +lean_dec(x_1); +x_106 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_107 = lean_int_emod(x_2, x_106); +lean_dec(x_2); +x_108 = 0; +x_109 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_105, x_107, x_108); +lean_dec(x_105); +return x_109; +} +case 18: +{ +lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_1, 0); +lean_inc(x_110); +lean_dec(x_1); +x_111 = lean_ctor_get(x_2, 1); +lean_inc(x_111); +lean_dec(x_2); +x_112 = 0; +x_113 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_110, x_111, x_112); +lean_dec(x_110); +return x_113; +} +case 19: +{ +lean_object* x_114; +x_114 = lean_ctor_get(x_1, 0); +lean_inc(x_114); +lean_dec(x_1); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; uint8_t x_116; lean_object* x_117; +x_115 = lean_unsigned_to_nat(9u); +x_116 = 0; +x_117 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_115, x_2, x_116); +return x_117; +} +else +{ +lean_object* x_118; uint8_t x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_114, 0); +lean_inc(x_118); +lean_dec(x_114); +x_119 = 1; +x_120 = l___private_Std_Time_Format_Basic_0__Std_Time_rightTruncate(x_118, x_2, x_119); +return x_120; +} +} +case 23: +{ +return x_2; +} +case 24: +{ +lean_dec(x_1); +return x_2; +} +case 25: +{ +uint8_t x_121; +x_121 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_122 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3; +x_123 = lean_int_div(x_2, x_122); +lean_dec(x_2); +x_124 = l___private_Std_Time_Format_Basic_0__Std_Time_toSigned(x_123); +lean_dec(x_123); +x_125 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_126 = lean_string_append(x_125, x_124); +lean_dec(x_124); +x_127 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_128 = lean_string_append(x_126, x_127); +return x_128; +} +else +{ +uint8_t x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_129 = 1; +x_130 = 0; +x_131 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_129, x_130, x_129); +x_132 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_133 = lean_string_append(x_132, x_131); +lean_dec(x_131); +x_134 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_135 = lean_string_append(x_133, x_134); +return x_135; +} +} +case 26: +{ +uint8_t x_136; lean_object* x_137; uint8_t x_138; +x_136 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +x_137 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_138 = lean_int_dec_eq(x_2, x_137); +if (x_138 == 0) +{ +switch (x_136) { +case 0: +{ +uint8_t x_139; lean_object* x_140; +x_139 = 0; +x_140 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_139, x_139, x_139); +return x_140; +} +case 1: +{ +uint8_t x_141; uint8_t x_142; lean_object* x_143; +x_141 = 1; +x_142 = 0; +x_143 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_141, x_142, x_142); +return x_143; +} +case 2: +{ +uint8_t x_144; uint8_t x_145; lean_object* x_146; +x_144 = 1; +x_145 = 0; +x_146 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_144, x_145, x_144); +return x_146; +} +case 3: +{ +uint8_t x_147; uint8_t x_148; lean_object* x_149; +x_147 = 1; +x_148 = 0; +x_149 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_147, x_147, x_148); +return x_149; +} +default: +{ +uint8_t x_150; lean_object* x_151; +x_150 = 1; +x_151 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_150, x_150, x_150); +return x_151; +} +} +} +else +{ +lean_object* x_152; +lean_dec(x_2); +x_152 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5; +return x_152; +} +} +case 27: +{ +uint8_t x_153; +x_153 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_153) { +case 0: +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; uint8_t x_159; uint8_t x_160; lean_object* x_161; +x_154 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_155 = lean_int_div(x_2, x_154); +x_156 = lean_int_emod(x_155, x_154); +lean_dec(x_155); +x_157 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_158 = lean_int_dec_eq(x_156, x_157); +lean_dec(x_156); +x_159 = l_instDecidableNot___rarg(x_158); +x_160 = 0; +x_161 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_159, x_160, x_160); +return x_161; +} +case 1: +{ +uint8_t x_162; uint8_t x_163; lean_object* x_164; +x_162 = 1; +x_163 = 0; +x_164 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_162, x_163, x_163); +return x_164; +} +case 2: +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; uint8_t x_169; uint8_t x_170; lean_object* x_171; +x_165 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_166 = lean_int_emod(x_2, x_165); +x_167 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_168 = lean_int_dec_eq(x_166, x_167); +lean_dec(x_166); +x_169 = l_instDecidableNot___rarg(x_168); +x_170 = 1; +x_171 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_170, x_169, x_170); +return x_171; +} +case 3: +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; uint8_t x_178; lean_object* x_179; +x_172 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_173 = lean_int_emod(x_2, x_172); +x_174 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_175 = lean_int_dec_eq(x_173, x_174); +lean_dec(x_173); +x_176 = l_instDecidableNot___rarg(x_175); +x_177 = 1; +x_178 = 0; +x_179 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_177, x_176, x_178); +return x_179; +} +default: +{ +uint8_t x_180; lean_object* x_181; +x_180 = 1; +x_181 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_180, x_180, x_180); +return x_181; +} +} +} +case 28: +{ +uint8_t x_182; +x_182 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_182) { +case 0: +{ +uint8_t x_183; uint8_t x_184; lean_object* x_185; +x_183 = 1; +x_184 = 0; +x_185 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_183, x_184, x_184); +return x_185; +} +case 1: +{ +lean_object* x_186; uint8_t x_187; +x_186 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_187 = lean_int_dec_eq(x_2, x_186); +if (x_187 == 0) +{ +uint8_t x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_188 = 1; +x_189 = 0; +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_188, x_189, x_188); +x_191 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_192 = lean_string_append(x_191, x_190); +lean_dec(x_190); +x_193 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_194 = lean_string_append(x_192, x_193); +return x_194; +} +else +{ +lean_object* x_195; +lean_dec(x_2); +x_195 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +return x_195; +} +} +default: +{ +lean_object* x_196; uint8_t x_197; +x_196 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_197 = lean_int_dec_eq(x_2, x_196); +if (x_197 == 0) +{ +lean_object* x_198; lean_object* x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; lean_object* x_203; +x_198 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_199 = lean_int_emod(x_2, x_198); +x_200 = lean_int_dec_eq(x_199, x_196); +lean_dec(x_199); +x_201 = l_instDecidableNot___rarg(x_200); +x_202 = 1; +x_203 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString(x_2, x_202, x_201, x_202); +return x_203; +} +else +{ +lean_object* x_204; +lean_dec(x_2); +x_204 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5; +return x_204; +} +} +} +} +default: +{ +lean_object* x_205; uint8_t x_206; lean_object* x_207; +x_205 = lean_ctor_get(x_1, 0); +lean_inc(x_205); +lean_dec(x_1); +x_206 = 0; +x_207 = l___private_Std_Time_Format_Basic_0__Std_Time_pad(x_205, x_2, x_206); +lean_dec(x_205); +return x_207; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Std_Time_DateTime_era___rarg(x_3); +x_5 = lean_box(x_4); +return x_5; +} +case 1: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_3, 1); +x_7 = lean_thunk_get_own(x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_dec(x_8); +return x_9; +} +case 2: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_3, 1); +x_11 = lean_thunk_get_own(x_10); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +return x_13; +} +case 3: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_thunk_get_own(x_14); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_19 = lean_int_mod(x_17, x_18); +x_20 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_21 = lean_int_dec_eq(x_19, x_20); +lean_dec(x_19); +x_22 = l_Std_Time_DateTime_dayOfYear___rarg(x_3); +if (x_21 == 0) +{ +uint8_t x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_17); +x_23 = 0; +x_24 = lean_box(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; +x_26 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +x_27 = lean_int_mod(x_17, x_26); +x_28 = lean_int_dec_eq(x_27, x_20); +lean_dec(x_27); +x_29 = l_instDecidableNot___rarg(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2; +x_31 = lean_int_mod(x_17, x_30); +lean_dec(x_17); +x_32 = lean_int_dec_eq(x_31, x_20); +lean_dec(x_31); +if (x_32 == 0) +{ +uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_33 = 0; +x_34 = lean_box(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_22); +return x_35; +} +else +{ +uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_36 = 1; +x_37 = lean_box(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_22); +return x_38; +} +} +else +{ +uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_17); +x_39 = 1; +x_40 = lean_box(x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_22); +return x_41; +} +} +} +case 4: +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_3, 1); +x_43 = lean_thunk_get_own(x_42); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +return x_45; +} +case 5: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_3, 1); +x_47 = lean_thunk_get_own(x_46); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_ctor_get(x_48, 2); +lean_inc(x_49); +lean_dec(x_48); +return x_49; +} +case 6: +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_3, 1); +x_51 = lean_thunk_get_own(x_50); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_Std_Time_PlainDate_quarter(x_52); +lean_dec(x_52); +return x_53; +} +case 7: +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_3, 1); +x_55 = lean_thunk_get_own(x_54); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +lean_dec(x_55); +x_57 = l_Std_Time_PlainDate_weekOfYear(x_56); +return x_57; +} +case 8: +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_3, 1); +x_59 = lean_thunk_get_own(x_58); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_60); +return x_61; +} +case 9: +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; +x_62 = lean_ctor_get(x_3, 1); +x_63 = lean_thunk_get_own(x_62); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +lean_dec(x_63); +x_65 = l_Std_Time_PlainDate_weekday(x_64); +x_66 = lean_box(x_65); +return x_66; +} +case 10: +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; +x_67 = lean_ctor_get(x_3, 1); +x_68 = lean_thunk_get_own(x_67); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_Std_Time_PlainDate_weekday(x_69); +x_71 = lean_box(x_70); +return x_71; +} +case 11: +{ +lean_object* x_72; +x_72 = l_Std_Time_DateTime_weekOfMonth___rarg(x_3); +return x_72; +} +case 12: +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; +x_73 = lean_ctor_get(x_3, 1); +x_74 = lean_thunk_get_own(x_73); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +lean_dec(x_75); +x_77 = l_Std_Time_HourMarker_ofOrdinal(x_76); +lean_dec(x_76); +x_78 = lean_box(x_77); +return x_78; +} +case 13: +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_79 = lean_ctor_get(x_3, 1); +x_80 = lean_thunk_get_own(x_79); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +lean_dec(x_81); +x_83 = l_Std_Time_HourMarker_toRelative(x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +lean_dec(x_83); +return x_84; +} +case 14: +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_85 = lean_ctor_get(x_3, 1); +x_86 = lean_thunk_get_own(x_85); +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +lean_dec(x_86); +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +lean_dec(x_87); +x_89 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_90 = lean_int_emod(x_88, x_89); +lean_dec(x_88); +return x_90; +} +case 15: +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_91 = lean_ctor_get(x_3, 1); +x_92 = lean_thunk_get_own(x_91); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +lean_dec(x_93); +x_95 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_94); +lean_dec(x_94); +return x_95; +} +case 16: +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_3, 1); +x_97 = lean_thunk_get_own(x_96); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +lean_dec(x_98); +return x_99; +} +case 17: +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_100 = lean_ctor_get(x_3, 1); +x_101 = lean_thunk_get_own(x_100); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +lean_dec(x_101); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +return x_103; +} +case 18: +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_3, 1); +x_105 = lean_thunk_get_own(x_104); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_ctor_get(x_106, 2); +lean_inc(x_107); +lean_dec(x_106); +return x_107; +} +case 19: +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_3, 1); +x_109 = lean_thunk_get_own(x_108); +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +lean_dec(x_109); +x_111 = lean_ctor_get(x_110, 3); +lean_inc(x_111); +lean_dec(x_110); +return x_111; +} +case 20: +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_112 = lean_ctor_get(x_3, 1); +x_113 = lean_thunk_get_own(x_112); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +lean_dec(x_113); +x_115 = l_Std_Time_PlainTime_toMilliseconds(x_114); +lean_dec(x_114); +return x_115; +} +case 21: +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_116 = lean_ctor_get(x_3, 1); +x_117 = lean_thunk_get_own(x_116); +x_118 = lean_ctor_get(x_117, 1); +lean_inc(x_118); +lean_dec(x_117); +x_119 = lean_ctor_get(x_118, 3); +lean_inc(x_119); +lean_dec(x_118); +return x_119; +} +case 22: +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_120 = lean_ctor_get(x_3, 1); +x_121 = lean_thunk_get_own(x_120); +x_122 = lean_ctor_get(x_121, 1); +lean_inc(x_122); +lean_dec(x_121); +x_123 = l_Std_Time_PlainTime_toNanoseconds(x_122); +lean_dec(x_122); +return x_123; +} +case 23: +{ +lean_object* x_124; +x_124 = lean_ctor_get(x_1, 1); +lean_inc(x_124); +return x_124; +} +case 24: +{ +uint8_t x_125; +x_125 = lean_ctor_get_uint8(x_2, 0); +if (x_125 == 0) +{ +lean_object* x_126; +x_126 = lean_ctor_get(x_1, 2); +lean_inc(x_126); +return x_126; +} +else +{ +lean_object* x_127; +x_127 = lean_ctor_get(x_1, 1); +lean_inc(x_127); +return x_127; +} +} +default: +{ +lean_object* x_128; +x_128 = lean_ctor_get(x_1, 0); +lean_inc(x_128); +return x_128; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(5u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +lean_dec(x_40); +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +lean_dec(x_40); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +x_54 = lean_nat_dec_eq(x_40, x_53); +lean_dec(x_40); +if (x_54 == 0) +{ +lean_dec(x_53); +return x_43; +} +else +{ +lean_object* x_55; lean_object* x_56; +lean_free_object(x_43); +lean_dec(x_52); +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_56 = l_Std_Internal_Parsec_String_pstring(x_55, x_51); +if (lean_obj_tag(x_56) == 0) +{ +uint8_t x_57; +lean_dec(x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 1); +lean_dec(x_58); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +lean_ctor_set(x_56, 1, x_59); +return x_56; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_56, 0); +lean_inc(x_60); +lean_dec(x_56); +x_61 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_56); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_56, 0); +x_65 = lean_ctor_get(x_56, 1); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +x_67 = lean_nat_dec_eq(x_53, x_66); +lean_dec(x_53); +if (x_67 == 0) +{ +lean_dec(x_66); +return x_56; +} +else +{ +lean_object* x_68; lean_object* x_69; +lean_free_object(x_56); +lean_dec(x_65); +x_68 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_69 = l_Std_Internal_Parsec_String_pstring(x_68, x_64); +if (lean_obj_tag(x_69) == 0) +{ +uint8_t x_70; +lean_dec(x_66); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_69, 1); +lean_dec(x_71); +x_72 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +lean_ctor_set(x_69, 1, x_72); +return x_69; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_69, 0); +lean_inc(x_73); +lean_dec(x_69); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +else +{ +uint8_t x_76; +x_76 = !lean_is_exclusive(x_69); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_77 = lean_ctor_get(x_69, 0); +x_78 = lean_ctor_get(x_69, 1); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +x_80 = lean_nat_dec_eq(x_66, x_79); +lean_dec(x_66); +if (x_80 == 0) +{ +lean_dec(x_79); +return x_69; +} +else +{ +lean_object* x_81; lean_object* x_82; +lean_free_object(x_69); +lean_dec(x_78); +x_81 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_82 = l_Std_Internal_Parsec_String_pstring(x_81, x_77); +if (lean_obj_tag(x_82) == 0) +{ +uint8_t x_83; +lean_dec(x_79); +x_83 = !lean_is_exclusive(x_82); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_82, 1); +lean_dec(x_84); +x_85 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +lean_ctor_set(x_82, 1, x_85); +return x_82; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_82, 0); +lean_inc(x_86); +lean_dec(x_82); +x_87 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; +} +} +else +{ +uint8_t x_89; +x_89 = !lean_is_exclusive(x_82); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_90 = lean_ctor_get(x_82, 0); +x_91 = lean_ctor_get(x_82, 1); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +x_93 = lean_nat_dec_eq(x_79, x_92); +lean_dec(x_79); +if (x_93 == 0) +{ +lean_dec(x_92); +return x_82; +} +else +{ +lean_object* x_94; lean_object* x_95; +lean_free_object(x_82); +lean_dec(x_91); +x_94 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_95 = l_Std_Internal_Parsec_String_pstring(x_94, x_90); +if (lean_obj_tag(x_95) == 0) +{ +uint8_t x_96; +lean_dec(x_92); +x_96 = !lean_is_exclusive(x_95); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_95, 1); +lean_dec(x_97); +x_98 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +lean_ctor_set(x_95, 1, x_98); +return x_95; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_95, 0); +lean_inc(x_99); +lean_dec(x_95); +x_100 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; +} +} +else +{ +uint8_t x_102; +x_102 = !lean_is_exclusive(x_95); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_103 = lean_ctor_get(x_95, 0); +x_104 = lean_ctor_get(x_95, 1); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +x_106 = lean_nat_dec_eq(x_92, x_105); +lean_dec(x_92); +if (x_106 == 0) +{ +lean_dec(x_105); +return x_95; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_free_object(x_95); +lean_dec(x_104); +x_107 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_108 = l_Std_Internal_Parsec_String_pstring(x_107, x_103); +if (lean_obj_tag(x_108) == 0) +{ +uint8_t x_109; +lean_dec(x_105); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_108, 1); +lean_dec(x_110); +x_111 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +lean_ctor_set(x_108, 1, x_111); +return x_108; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_108, 0); +lean_inc(x_112); +lean_dec(x_108); +x_113 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; +} +} +else +{ +uint8_t x_115; +x_115 = !lean_is_exclusive(x_108); +if (x_115 == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_116 = lean_ctor_get(x_108, 0); +x_117 = lean_ctor_get(x_108, 1); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +x_119 = lean_nat_dec_eq(x_105, x_118); +lean_dec(x_105); +if (x_119 == 0) +{ +lean_dec(x_118); +return x_108; +} +else +{ +lean_object* x_120; lean_object* x_121; +lean_free_object(x_108); +lean_dec(x_117); +x_120 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_121 = l_Std_Internal_Parsec_String_pstring(x_120, x_116); +if (lean_obj_tag(x_121) == 0) +{ +uint8_t x_122; +lean_dec(x_118); +x_122 = !lean_is_exclusive(x_121); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_121, 1); +lean_dec(x_123); +x_124 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +lean_ctor_set(x_121, 1, x_124); +return x_121; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_121, 0); +lean_inc(x_125); +lean_dec(x_121); +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +else +{ +uint8_t x_128; +x_128 = !lean_is_exclusive(x_121); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_129 = lean_ctor_get(x_121, 0); +x_130 = lean_ctor_get(x_121, 1); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +x_132 = lean_nat_dec_eq(x_118, x_131); +lean_dec(x_118); +if (x_132 == 0) +{ +lean_dec(x_131); +return x_121; +} +else +{ +lean_object* x_133; lean_object* x_134; +lean_free_object(x_121); +lean_dec(x_130); +x_133 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_134 = l_Std_Internal_Parsec_String_pstring(x_133, x_129); +if (lean_obj_tag(x_134) == 0) +{ +uint8_t x_135; +lean_dec(x_131); +x_135 = !lean_is_exclusive(x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_134, 1); +lean_dec(x_136); +x_137 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +lean_ctor_set(x_134, 1, x_137); +return x_134; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_134, 0); +lean_inc(x_138); +lean_dec(x_134); +x_139 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; +} +} +else +{ +uint8_t x_141; +x_141 = !lean_is_exclusive(x_134); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_142 = lean_ctor_get(x_134, 0); +x_143 = lean_ctor_get(x_134, 1); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +x_145 = lean_nat_dec_eq(x_131, x_144); +lean_dec(x_144); +lean_dec(x_131); +if (x_145 == 0) +{ +return x_134; +} +else +{ +lean_object* x_146; lean_object* x_147; +lean_free_object(x_134); +lean_dec(x_143); +x_146 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_147 = l_Std_Internal_Parsec_String_pstring(x_146, x_142); +if (lean_obj_tag(x_147) == 0) +{ +uint8_t x_148; +x_148 = !lean_is_exclusive(x_147); +if (x_148 == 0) +{ +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_147, 1); +lean_dec(x_149); +x_150 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +lean_ctor_set(x_147, 1, x_150); +return x_147; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_147, 0); +lean_inc(x_151); +lean_dec(x_147); +x_152 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +else +{ +uint8_t x_154; +x_154 = !lean_is_exclusive(x_147); +if (x_154 == 0) +{ +return x_147; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_147, 0); +x_156 = lean_ctor_get(x_147, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_147); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; +} +} +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; +x_158 = lean_ctor_get(x_134, 0); +x_159 = lean_ctor_get(x_134, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_134); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +x_161 = lean_nat_dec_eq(x_131, x_160); +lean_dec(x_160); +lean_dec(x_131); +if (x_161 == 0) +{ +lean_object* x_162; +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_158); +lean_ctor_set(x_162, 1, x_159); +return x_162; +} +else +{ +lean_object* x_163; lean_object* x_164; +lean_dec(x_159); +x_163 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_164 = l_Std_Internal_Parsec_String_pstring(x_163, x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_166 = x_164; +} else { + lean_dec_ref(x_164); + x_166 = lean_box(0); +} +x_167 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_166)) { + x_168 = lean_alloc_ctor(0, 2, 0); +} else { + x_168 = x_166; +} +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_167); +return x_168; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_169 = lean_ctor_get(x_164, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_164, 1); +lean_inc(x_170); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_171 = x_164; +} else { + lean_dec_ref(x_164); + x_171 = lean_box(0); +} +if (lean_is_scalar(x_171)) { + x_172 = lean_alloc_ctor(1, 2, 0); +} else { + x_172 = x_171; +} +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_170); +return x_172; +} +} +} +} +} +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; +x_173 = lean_ctor_get(x_121, 0); +x_174 = lean_ctor_get(x_121, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_121); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +x_176 = lean_nat_dec_eq(x_118, x_175); +lean_dec(x_118); +if (x_176 == 0) +{ +lean_object* x_177; +lean_dec(x_175); +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_174); +return x_177; +} +else +{ +lean_object* x_178; lean_object* x_179; +lean_dec(x_174); +x_178 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_179 = l_Std_Internal_Parsec_String_pstring(x_178, x_173); +if (lean_obj_tag(x_179) == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_175); +x_180 = lean_ctor_get(x_179, 0); +lean_inc(x_180); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_181 = x_179; +} else { + lean_dec_ref(x_179); + x_181 = lean_box(0); +} +x_182 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 2, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_182); +return x_183; +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; +x_184 = lean_ctor_get(x_179, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_179, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_186 = x_179; +} else { + lean_dec_ref(x_179); + x_186 = lean_box(0); +} +x_187 = lean_ctor_get(x_184, 1); +lean_inc(x_187); +x_188 = lean_nat_dec_eq(x_175, x_187); +lean_dec(x_187); +lean_dec(x_175); +if (x_188 == 0) +{ +lean_object* x_189; +if (lean_is_scalar(x_186)) { + x_189 = lean_alloc_ctor(1, 2, 0); +} else { + x_189 = x_186; +} +lean_ctor_set(x_189, 0, x_184); +lean_ctor_set(x_189, 1, x_185); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; +lean_dec(x_186); +lean_dec(x_185); +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_191 = l_Std_Internal_Parsec_String_pstring(x_190, x_184); +if (lean_obj_tag(x_191) == 0) +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_193 = x_191; +} else { + lean_dec_ref(x_191); + x_193 = lean_box(0); +} +x_194 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_193)) { + x_195 = lean_alloc_ctor(0, 2, 0); +} else { + x_195 = x_193; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_194); +return x_195; +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_196 = lean_ctor_get(x_191, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_191, 1); +lean_inc(x_197); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_198 = x_191; +} else { + lean_dec_ref(x_191); + x_198 = lean_box(0); +} +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(1, 2, 0); +} else { + x_199 = x_198; +} +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +return x_199; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; +x_200 = lean_ctor_get(x_108, 0); +x_201 = lean_ctor_get(x_108, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_108); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +x_203 = lean_nat_dec_eq(x_105, x_202); +lean_dec(x_105); +if (x_203 == 0) +{ +lean_object* x_204; +lean_dec(x_202); +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_200); +lean_ctor_set(x_204, 1, x_201); +return x_204; +} +else +{ +lean_object* x_205; lean_object* x_206; +lean_dec(x_201); +x_205 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_206 = l_Std_Internal_Parsec_String_pstring(x_205, x_200); +if (lean_obj_tag(x_206) == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_202); +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_208 = x_206; +} else { + lean_dec_ref(x_206); + x_208 = lean_box(0); +} +x_209 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_208)) { + x_210 = lean_alloc_ctor(0, 2, 0); +} else { + x_210 = x_208; +} +lean_ctor_set(x_210, 0, x_207); +lean_ctor_set(x_210, 1, x_209); +return x_210; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_211 = lean_ctor_get(x_206, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_206, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_213 = x_206; +} else { + lean_dec_ref(x_206); + x_213 = lean_box(0); +} +x_214 = lean_ctor_get(x_211, 1); +lean_inc(x_214); +x_215 = lean_nat_dec_eq(x_202, x_214); +lean_dec(x_202); +if (x_215 == 0) +{ +lean_object* x_216; +lean_dec(x_214); +if (lean_is_scalar(x_213)) { + x_216 = lean_alloc_ctor(1, 2, 0); +} else { + x_216 = x_213; +} +lean_ctor_set(x_216, 0, x_211); +lean_ctor_set(x_216, 1, x_212); +return x_216; +} +else +{ +lean_object* x_217; lean_object* x_218; +lean_dec(x_213); +lean_dec(x_212); +x_217 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_218 = l_Std_Internal_Parsec_String_pstring(x_217, x_211); +if (lean_obj_tag(x_218) == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_214); +x_219 = lean_ctor_get(x_218, 0); +lean_inc(x_219); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_220 = x_218; +} else { + lean_dec_ref(x_218); + x_220 = lean_box(0); +} +x_221 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_220)) { + x_222 = lean_alloc_ctor(0, 2, 0); +} else { + x_222 = x_220; +} +lean_ctor_set(x_222, 0, x_219); +lean_ctor_set(x_222, 1, x_221); +return x_222; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; uint8_t x_227; +x_223 = lean_ctor_get(x_218, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_218, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_225 = x_218; +} else { + lean_dec_ref(x_218); + x_225 = lean_box(0); +} +x_226 = lean_ctor_get(x_223, 1); +lean_inc(x_226); +x_227 = lean_nat_dec_eq(x_214, x_226); +lean_dec(x_226); +lean_dec(x_214); +if (x_227 == 0) +{ +lean_object* x_228; +if (lean_is_scalar(x_225)) { + x_228 = lean_alloc_ctor(1, 2, 0); +} else { + x_228 = x_225; +} +lean_ctor_set(x_228, 0, x_223); +lean_ctor_set(x_228, 1, x_224); +return x_228; +} +else +{ +lean_object* x_229; lean_object* x_230; +lean_dec(x_225); +lean_dec(x_224); +x_229 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_230 = l_Std_Internal_Parsec_String_pstring(x_229, x_223); +if (lean_obj_tag(x_230) == 0) +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_231 = lean_ctor_get(x_230, 0); +lean_inc(x_231); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_232 = x_230; +} else { + lean_dec_ref(x_230); + x_232 = lean_box(0); +} +x_233 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_232)) { + x_234 = lean_alloc_ctor(0, 2, 0); +} else { + x_234 = x_232; +} +lean_ctor_set(x_234, 0, x_231); +lean_ctor_set(x_234, 1, x_233); +return x_234; +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_235 = lean_ctor_get(x_230, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_230, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_237 = x_230; +} else { + lean_dec_ref(x_230); + x_237 = lean_box(0); +} +if (lean_is_scalar(x_237)) { + x_238 = lean_alloc_ctor(1, 2, 0); +} else { + x_238 = x_237; +} +lean_ctor_set(x_238, 0, x_235); +lean_ctor_set(x_238, 1, x_236); +return x_238; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; uint8_t x_242; +x_239 = lean_ctor_get(x_95, 0); +x_240 = lean_ctor_get(x_95, 1); +lean_inc(x_240); +lean_inc(x_239); +lean_dec(x_95); +x_241 = lean_ctor_get(x_239, 1); +lean_inc(x_241); +x_242 = lean_nat_dec_eq(x_92, x_241); +lean_dec(x_92); +if (x_242 == 0) +{ +lean_object* x_243; +lean_dec(x_241); +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_239); +lean_ctor_set(x_243, 1, x_240); +return x_243; +} +else +{ +lean_object* x_244; lean_object* x_245; +lean_dec(x_240); +x_244 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_245 = l_Std_Internal_Parsec_String_pstring(x_244, x_239); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +lean_dec(x_241); +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_247 = x_245; +} else { + lean_dec_ref(x_245); + x_247 = lean_box(0); +} +x_248 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_247)) { + x_249 = lean_alloc_ctor(0, 2, 0); +} else { + x_249 = x_247; +} +lean_ctor_set(x_249, 0, x_246); +lean_ctor_set(x_249, 1, x_248); +return x_249; +} +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; uint8_t x_254; +x_250 = lean_ctor_get(x_245, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_245, 1); +lean_inc(x_251); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_252 = x_245; +} else { + lean_dec_ref(x_245); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 1); +lean_inc(x_253); +x_254 = lean_nat_dec_eq(x_241, x_253); +lean_dec(x_241); +if (x_254 == 0) +{ +lean_object* x_255; +lean_dec(x_253); +if (lean_is_scalar(x_252)) { + x_255 = lean_alloc_ctor(1, 2, 0); +} else { + x_255 = x_252; +} +lean_ctor_set(x_255, 0, x_250); +lean_ctor_set(x_255, 1, x_251); +return x_255; +} +else +{ +lean_object* x_256; lean_object* x_257; +lean_dec(x_252); +lean_dec(x_251); +x_256 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_257 = l_Std_Internal_Parsec_String_pstring(x_256, x_250); +if (lean_obj_tag(x_257) == 0) +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_253); +x_258 = lean_ctor_get(x_257, 0); +lean_inc(x_258); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_259 = x_257; +} else { + lean_dec_ref(x_257); + x_259 = lean_box(0); +} +x_260 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_259)) { + x_261 = lean_alloc_ctor(0, 2, 0); +} else { + x_261 = x_259; +} +lean_ctor_set(x_261, 0, x_258); +lean_ctor_set(x_261, 1, x_260); +return x_261; +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; uint8_t x_266; +x_262 = lean_ctor_get(x_257, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_257, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_264 = x_257; +} else { + lean_dec_ref(x_257); + x_264 = lean_box(0); +} +x_265 = lean_ctor_get(x_262, 1); +lean_inc(x_265); +x_266 = lean_nat_dec_eq(x_253, x_265); +lean_dec(x_253); +if (x_266 == 0) +{ +lean_object* x_267; +lean_dec(x_265); +if (lean_is_scalar(x_264)) { + x_267 = lean_alloc_ctor(1, 2, 0); +} else { + x_267 = x_264; +} +lean_ctor_set(x_267, 0, x_262); +lean_ctor_set(x_267, 1, x_263); +return x_267; +} +else +{ +lean_object* x_268; lean_object* x_269; +lean_dec(x_264); +lean_dec(x_263); +x_268 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_269 = l_Std_Internal_Parsec_String_pstring(x_268, x_262); +if (lean_obj_tag(x_269) == 0) +{ +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_265); +x_270 = lean_ctor_get(x_269, 0); +lean_inc(x_270); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_271 = x_269; +} else { + lean_dec_ref(x_269); + x_271 = lean_box(0); +} +x_272 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_271)) { + x_273 = lean_alloc_ctor(0, 2, 0); +} else { + x_273 = x_271; +} +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_272); +return x_273; +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; uint8_t x_278; +x_274 = lean_ctor_get(x_269, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_269, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_276 = x_269; +} else { + lean_dec_ref(x_269); + x_276 = lean_box(0); +} +x_277 = lean_ctor_get(x_274, 1); +lean_inc(x_277); +x_278 = lean_nat_dec_eq(x_265, x_277); +lean_dec(x_277); +lean_dec(x_265); +if (x_278 == 0) +{ +lean_object* x_279; +if (lean_is_scalar(x_276)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_276; +} +lean_ctor_set(x_279, 0, x_274); +lean_ctor_set(x_279, 1, x_275); +return x_279; +} +else +{ +lean_object* x_280; lean_object* x_281; +lean_dec(x_276); +lean_dec(x_275); +x_280 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_281 = l_Std_Internal_Parsec_String_pstring(x_280, x_274); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_283 = x_281; +} else { + lean_dec_ref(x_281); + x_283 = lean_box(0); +} +x_284 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_283)) { + x_285 = lean_alloc_ctor(0, 2, 0); +} else { + x_285 = x_283; +} +lean_ctor_set(x_285, 0, x_282); +lean_ctor_set(x_285, 1, x_284); +return x_285; +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_286 = lean_ctor_get(x_281, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_281, 1); +lean_inc(x_287); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_288 = x_281; +} else { + lean_dec_ref(x_281); + x_288 = lean_box(0); +} +if (lean_is_scalar(x_288)) { + x_289 = lean_alloc_ctor(1, 2, 0); +} else { + x_289 = x_288; +} +lean_ctor_set(x_289, 0, x_286); +lean_ctor_set(x_289, 1, x_287); +return x_289; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_290; lean_object* x_291; lean_object* x_292; uint8_t x_293; +x_290 = lean_ctor_get(x_82, 0); +x_291 = lean_ctor_get(x_82, 1); +lean_inc(x_291); +lean_inc(x_290); +lean_dec(x_82); +x_292 = lean_ctor_get(x_290, 1); +lean_inc(x_292); +x_293 = lean_nat_dec_eq(x_79, x_292); +lean_dec(x_79); +if (x_293 == 0) +{ +lean_object* x_294; +lean_dec(x_292); +x_294 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_294, 0, x_290); +lean_ctor_set(x_294, 1, x_291); +return x_294; +} +else +{ +lean_object* x_295; lean_object* x_296; +lean_dec(x_291); +x_295 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_296 = l_Std_Internal_Parsec_String_pstring(x_295, x_290); +if (lean_obj_tag(x_296) == 0) +{ +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_292); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +if (lean_is_exclusive(x_296)) { + lean_ctor_release(x_296, 0); + lean_ctor_release(x_296, 1); + x_298 = x_296; +} else { + lean_dec_ref(x_296); + x_298 = lean_box(0); +} +x_299 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_298)) { + x_300 = lean_alloc_ctor(0, 2, 0); +} else { + x_300 = x_298; +} +lean_ctor_set(x_300, 0, x_297); +lean_ctor_set(x_300, 1, x_299); +return x_300; +} +else +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; +x_301 = lean_ctor_get(x_296, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_296, 1); +lean_inc(x_302); +if (lean_is_exclusive(x_296)) { + lean_ctor_release(x_296, 0); + lean_ctor_release(x_296, 1); + x_303 = x_296; +} else { + lean_dec_ref(x_296); + x_303 = lean_box(0); +} +x_304 = lean_ctor_get(x_301, 1); +lean_inc(x_304); +x_305 = lean_nat_dec_eq(x_292, x_304); +lean_dec(x_292); +if (x_305 == 0) +{ +lean_object* x_306; +lean_dec(x_304); +if (lean_is_scalar(x_303)) { + x_306 = lean_alloc_ctor(1, 2, 0); +} else { + x_306 = x_303; +} +lean_ctor_set(x_306, 0, x_301); +lean_ctor_set(x_306, 1, x_302); +return x_306; +} +else +{ +lean_object* x_307; lean_object* x_308; +lean_dec(x_303); +lean_dec(x_302); +x_307 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_308 = l_Std_Internal_Parsec_String_pstring(x_307, x_301); +if (lean_obj_tag(x_308) == 0) +{ +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_304); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_310 = x_308; +} else { + lean_dec_ref(x_308); + x_310 = lean_box(0); +} +x_311 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_310)) { + x_312 = lean_alloc_ctor(0, 2, 0); +} else { + x_312 = x_310; +} +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_311); +return x_312; +} +else +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; uint8_t x_317; +x_313 = lean_ctor_get(x_308, 0); +lean_inc(x_313); +x_314 = lean_ctor_get(x_308, 1); +lean_inc(x_314); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_315 = x_308; +} else { + lean_dec_ref(x_308); + x_315 = lean_box(0); +} +x_316 = lean_ctor_get(x_313, 1); +lean_inc(x_316); +x_317 = lean_nat_dec_eq(x_304, x_316); +lean_dec(x_304); +if (x_317 == 0) +{ +lean_object* x_318; +lean_dec(x_316); +if (lean_is_scalar(x_315)) { + x_318 = lean_alloc_ctor(1, 2, 0); +} else { + x_318 = x_315; +} +lean_ctor_set(x_318, 0, x_313); +lean_ctor_set(x_318, 1, x_314); +return x_318; +} +else +{ +lean_object* x_319; lean_object* x_320; +lean_dec(x_315); +lean_dec(x_314); +x_319 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_320 = l_Std_Internal_Parsec_String_pstring(x_319, x_313); +if (lean_obj_tag(x_320) == 0) +{ +lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_316); +x_321 = lean_ctor_get(x_320, 0); +lean_inc(x_321); +if (lean_is_exclusive(x_320)) { + lean_ctor_release(x_320, 0); + lean_ctor_release(x_320, 1); + x_322 = x_320; +} else { + lean_dec_ref(x_320); + x_322 = lean_box(0); +} +x_323 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_322)) { + x_324 = lean_alloc_ctor(0, 2, 0); +} else { + x_324 = x_322; +} +lean_ctor_set(x_324, 0, x_321); +lean_ctor_set(x_324, 1, x_323); +return x_324; +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; +x_325 = lean_ctor_get(x_320, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_320, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_320)) { + lean_ctor_release(x_320, 0); + lean_ctor_release(x_320, 1); + x_327 = x_320; +} else { + lean_dec_ref(x_320); + x_327 = lean_box(0); +} +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +x_329 = lean_nat_dec_eq(x_316, x_328); +lean_dec(x_316); +if (x_329 == 0) +{ +lean_object* x_330; +lean_dec(x_328); +if (lean_is_scalar(x_327)) { + x_330 = lean_alloc_ctor(1, 2, 0); +} else { + x_330 = x_327; +} +lean_ctor_set(x_330, 0, x_325); +lean_ctor_set(x_330, 1, x_326); +return x_330; +} +else +{ +lean_object* x_331; lean_object* x_332; +lean_dec(x_327); +lean_dec(x_326); +x_331 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_332 = l_Std_Internal_Parsec_String_pstring(x_331, x_325); +if (lean_obj_tag(x_332) == 0) +{ +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; +lean_dec(x_328); +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_334 = x_332; +} else { + lean_dec_ref(x_332); + x_334 = lean_box(0); +} +x_335 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_334)) { + x_336 = lean_alloc_ctor(0, 2, 0); +} else { + x_336 = x_334; +} +lean_ctor_set(x_336, 0, x_333); +lean_ctor_set(x_336, 1, x_335); +return x_336; +} +else +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; uint8_t x_341; +x_337 = lean_ctor_get(x_332, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_332, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_339 = x_332; +} else { + lean_dec_ref(x_332); + x_339 = lean_box(0); +} +x_340 = lean_ctor_get(x_337, 1); +lean_inc(x_340); +x_341 = lean_nat_dec_eq(x_328, x_340); +lean_dec(x_340); +lean_dec(x_328); +if (x_341 == 0) +{ +lean_object* x_342; +if (lean_is_scalar(x_339)) { + x_342 = lean_alloc_ctor(1, 2, 0); +} else { + x_342 = x_339; +} +lean_ctor_set(x_342, 0, x_337); +lean_ctor_set(x_342, 1, x_338); +return x_342; +} +else +{ +lean_object* x_343; lean_object* x_344; +lean_dec(x_339); +lean_dec(x_338); +x_343 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_344 = l_Std_Internal_Parsec_String_pstring(x_343, x_337); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +if (lean_is_exclusive(x_344)) { + lean_ctor_release(x_344, 0); + lean_ctor_release(x_344, 1); + x_346 = x_344; +} else { + lean_dec_ref(x_344); + x_346 = lean_box(0); +} +x_347 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_346)) { + x_348 = lean_alloc_ctor(0, 2, 0); +} else { + x_348 = x_346; +} +lean_ctor_set(x_348, 0, x_345); +lean_ctor_set(x_348, 1, x_347); +return x_348; +} +else +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +x_349 = lean_ctor_get(x_344, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_344, 1); +lean_inc(x_350); +if (lean_is_exclusive(x_344)) { + lean_ctor_release(x_344, 0); + lean_ctor_release(x_344, 1); + x_351 = x_344; +} else { + lean_dec_ref(x_344); + x_351 = lean_box(0); +} +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_351; +} +lean_ctor_set(x_352, 0, x_349); +lean_ctor_set(x_352, 1, x_350); +return x_352; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; uint8_t x_356; +x_353 = lean_ctor_get(x_69, 0); +x_354 = lean_ctor_get(x_69, 1); +lean_inc(x_354); +lean_inc(x_353); +lean_dec(x_69); +x_355 = lean_ctor_get(x_353, 1); +lean_inc(x_355); +x_356 = lean_nat_dec_eq(x_66, x_355); +lean_dec(x_66); +if (x_356 == 0) +{ +lean_object* x_357; +lean_dec(x_355); +x_357 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_357, 0, x_353); +lean_ctor_set(x_357, 1, x_354); +return x_357; +} +else +{ +lean_object* x_358; lean_object* x_359; +lean_dec(x_354); +x_358 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_359 = l_Std_Internal_Parsec_String_pstring(x_358, x_353); +if (lean_obj_tag(x_359) == 0) +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_355); +x_360 = lean_ctor_get(x_359, 0); +lean_inc(x_360); +if (lean_is_exclusive(x_359)) { + lean_ctor_release(x_359, 0); + lean_ctor_release(x_359, 1); + x_361 = x_359; +} else { + lean_dec_ref(x_359); + x_361 = lean_box(0); +} +x_362 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_361)) { + x_363 = lean_alloc_ctor(0, 2, 0); +} else { + x_363 = x_361; +} +lean_ctor_set(x_363, 0, x_360); +lean_ctor_set(x_363, 1, x_362); +return x_363; +} +else +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; uint8_t x_368; +x_364 = lean_ctor_get(x_359, 0); +lean_inc(x_364); +x_365 = lean_ctor_get(x_359, 1); +lean_inc(x_365); +if (lean_is_exclusive(x_359)) { + lean_ctor_release(x_359, 0); + lean_ctor_release(x_359, 1); + x_366 = x_359; +} else { + lean_dec_ref(x_359); + x_366 = lean_box(0); +} +x_367 = lean_ctor_get(x_364, 1); +lean_inc(x_367); +x_368 = lean_nat_dec_eq(x_355, x_367); +lean_dec(x_355); +if (x_368 == 0) +{ +lean_object* x_369; +lean_dec(x_367); +if (lean_is_scalar(x_366)) { + x_369 = lean_alloc_ctor(1, 2, 0); +} else { + x_369 = x_366; +} +lean_ctor_set(x_369, 0, x_364); +lean_ctor_set(x_369, 1, x_365); +return x_369; +} +else +{ +lean_object* x_370; lean_object* x_371; +lean_dec(x_366); +lean_dec(x_365); +x_370 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_371 = l_Std_Internal_Parsec_String_pstring(x_370, x_364); +if (lean_obj_tag(x_371) == 0) +{ +lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_367); +x_372 = lean_ctor_get(x_371, 0); +lean_inc(x_372); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_373 = x_371; +} else { + lean_dec_ref(x_371); + x_373 = lean_box(0); +} +x_374 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_373)) { + x_375 = lean_alloc_ctor(0, 2, 0); +} else { + x_375 = x_373; +} +lean_ctor_set(x_375, 0, x_372); +lean_ctor_set(x_375, 1, x_374); +return x_375; +} +else +{ +lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; uint8_t x_380; +x_376 = lean_ctor_get(x_371, 0); +lean_inc(x_376); +x_377 = lean_ctor_get(x_371, 1); +lean_inc(x_377); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_378 = x_371; +} else { + lean_dec_ref(x_371); + x_378 = lean_box(0); +} +x_379 = lean_ctor_get(x_376, 1); +lean_inc(x_379); +x_380 = lean_nat_dec_eq(x_367, x_379); +lean_dec(x_367); +if (x_380 == 0) +{ +lean_object* x_381; +lean_dec(x_379); +if (lean_is_scalar(x_378)) { + x_381 = lean_alloc_ctor(1, 2, 0); +} else { + x_381 = x_378; +} +lean_ctor_set(x_381, 0, x_376); +lean_ctor_set(x_381, 1, x_377); +return x_381; +} +else +{ +lean_object* x_382; lean_object* x_383; +lean_dec(x_378); +lean_dec(x_377); +x_382 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_383 = l_Std_Internal_Parsec_String_pstring(x_382, x_376); +if (lean_obj_tag(x_383) == 0) +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_379); +x_384 = lean_ctor_get(x_383, 0); +lean_inc(x_384); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_385 = x_383; +} else { + lean_dec_ref(x_383); + x_385 = lean_box(0); +} +x_386 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_385)) { + x_387 = lean_alloc_ctor(0, 2, 0); +} else { + x_387 = x_385; +} +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_386); +return x_387; +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; uint8_t x_392; +x_388 = lean_ctor_get(x_383, 0); +lean_inc(x_388); +x_389 = lean_ctor_get(x_383, 1); +lean_inc(x_389); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_390 = x_383; +} else { + lean_dec_ref(x_383); + x_390 = lean_box(0); +} +x_391 = lean_ctor_get(x_388, 1); +lean_inc(x_391); +x_392 = lean_nat_dec_eq(x_379, x_391); +lean_dec(x_379); +if (x_392 == 0) +{ +lean_object* x_393; +lean_dec(x_391); +if (lean_is_scalar(x_390)) { + x_393 = lean_alloc_ctor(1, 2, 0); +} else { + x_393 = x_390; +} +lean_ctor_set(x_393, 0, x_388); +lean_ctor_set(x_393, 1, x_389); +return x_393; +} +else +{ +lean_object* x_394; lean_object* x_395; +lean_dec(x_390); +lean_dec(x_389); +x_394 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_395 = l_Std_Internal_Parsec_String_pstring(x_394, x_388); +if (lean_obj_tag(x_395) == 0) +{ +lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; +lean_dec(x_391); +x_396 = lean_ctor_get(x_395, 0); +lean_inc(x_396); +if (lean_is_exclusive(x_395)) { + lean_ctor_release(x_395, 0); + lean_ctor_release(x_395, 1); + x_397 = x_395; +} else { + lean_dec_ref(x_395); + x_397 = lean_box(0); +} +x_398 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_397)) { + x_399 = lean_alloc_ctor(0, 2, 0); +} else { + x_399 = x_397; +} +lean_ctor_set(x_399, 0, x_396); +lean_ctor_set(x_399, 1, x_398); +return x_399; +} +else +{ +lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; uint8_t x_404; +x_400 = lean_ctor_get(x_395, 0); +lean_inc(x_400); +x_401 = lean_ctor_get(x_395, 1); +lean_inc(x_401); +if (lean_is_exclusive(x_395)) { + lean_ctor_release(x_395, 0); + lean_ctor_release(x_395, 1); + x_402 = x_395; +} else { + lean_dec_ref(x_395); + x_402 = lean_box(0); +} +x_403 = lean_ctor_get(x_400, 1); +lean_inc(x_403); +x_404 = lean_nat_dec_eq(x_391, x_403); +lean_dec(x_391); +if (x_404 == 0) +{ +lean_object* x_405; +lean_dec(x_403); +if (lean_is_scalar(x_402)) { + x_405 = lean_alloc_ctor(1, 2, 0); +} else { + x_405 = x_402; +} +lean_ctor_set(x_405, 0, x_400); +lean_ctor_set(x_405, 1, x_401); +return x_405; +} +else +{ +lean_object* x_406; lean_object* x_407; +lean_dec(x_402); +lean_dec(x_401); +x_406 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_407 = l_Std_Internal_Parsec_String_pstring(x_406, x_400); +if (lean_obj_tag(x_407) == 0) +{ +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; +lean_dec(x_403); +x_408 = lean_ctor_get(x_407, 0); +lean_inc(x_408); +if (lean_is_exclusive(x_407)) { + lean_ctor_release(x_407, 0); + lean_ctor_release(x_407, 1); + x_409 = x_407; +} else { + lean_dec_ref(x_407); + x_409 = lean_box(0); +} +x_410 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_409)) { + x_411 = lean_alloc_ctor(0, 2, 0); +} else { + x_411 = x_409; +} +lean_ctor_set(x_411, 0, x_408); +lean_ctor_set(x_411, 1, x_410); +return x_411; +} +else +{ +lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; uint8_t x_416; +x_412 = lean_ctor_get(x_407, 0); +lean_inc(x_412); +x_413 = lean_ctor_get(x_407, 1); +lean_inc(x_413); +if (lean_is_exclusive(x_407)) { + lean_ctor_release(x_407, 0); + lean_ctor_release(x_407, 1); + x_414 = x_407; +} else { + lean_dec_ref(x_407); + x_414 = lean_box(0); +} +x_415 = lean_ctor_get(x_412, 1); +lean_inc(x_415); +x_416 = lean_nat_dec_eq(x_403, x_415); +lean_dec(x_415); +lean_dec(x_403); +if (x_416 == 0) +{ +lean_object* x_417; +if (lean_is_scalar(x_414)) { + x_417 = lean_alloc_ctor(1, 2, 0); +} else { + x_417 = x_414; +} +lean_ctor_set(x_417, 0, x_412); +lean_ctor_set(x_417, 1, x_413); +return x_417; +} +else +{ +lean_object* x_418; lean_object* x_419; +lean_dec(x_414); +lean_dec(x_413); +x_418 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_419 = l_Std_Internal_Parsec_String_pstring(x_418, x_412); +if (lean_obj_tag(x_419) == 0) +{ +lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +x_420 = lean_ctor_get(x_419, 0); +lean_inc(x_420); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + lean_ctor_release(x_419, 1); + x_421 = x_419; +} else { + lean_dec_ref(x_419); + x_421 = lean_box(0); +} +x_422 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_421)) { + x_423 = lean_alloc_ctor(0, 2, 0); +} else { + x_423 = x_421; +} +lean_ctor_set(x_423, 0, x_420); +lean_ctor_set(x_423, 1, x_422); +return x_423; +} +else +{ +lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; +x_424 = lean_ctor_get(x_419, 0); +lean_inc(x_424); +x_425 = lean_ctor_get(x_419, 1); +lean_inc(x_425); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + lean_ctor_release(x_419, 1); + x_426 = x_419; +} else { + lean_dec_ref(x_419); + x_426 = lean_box(0); +} +if (lean_is_scalar(x_426)) { + x_427 = lean_alloc_ctor(1, 2, 0); +} else { + x_427 = x_426; +} +lean_ctor_set(x_427, 0, x_424); +lean_ctor_set(x_427, 1, x_425); +return x_427; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_428; lean_object* x_429; lean_object* x_430; uint8_t x_431; +x_428 = lean_ctor_get(x_56, 0); +x_429 = lean_ctor_get(x_56, 1); +lean_inc(x_429); +lean_inc(x_428); +lean_dec(x_56); +x_430 = lean_ctor_get(x_428, 1); +lean_inc(x_430); +x_431 = lean_nat_dec_eq(x_53, x_430); +lean_dec(x_53); +if (x_431 == 0) +{ +lean_object* x_432; +lean_dec(x_430); +x_432 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_432, 0, x_428); +lean_ctor_set(x_432, 1, x_429); +return x_432; +} +else +{ +lean_object* x_433; lean_object* x_434; +lean_dec(x_429); +x_433 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_434 = l_Std_Internal_Parsec_String_pstring(x_433, x_428); +if (lean_obj_tag(x_434) == 0) +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; +lean_dec(x_430); +x_435 = lean_ctor_get(x_434, 0); +lean_inc(x_435); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_436 = x_434; +} else { + lean_dec_ref(x_434); + x_436 = lean_box(0); +} +x_437 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_436)) { + x_438 = lean_alloc_ctor(0, 2, 0); +} else { + x_438 = x_436; +} +lean_ctor_set(x_438, 0, x_435); +lean_ctor_set(x_438, 1, x_437); +return x_438; +} +else +{ +lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; +x_439 = lean_ctor_get(x_434, 0); +lean_inc(x_439); +x_440 = lean_ctor_get(x_434, 1); +lean_inc(x_440); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_441 = x_434; +} else { + lean_dec_ref(x_434); + x_441 = lean_box(0); +} +x_442 = lean_ctor_get(x_439, 1); +lean_inc(x_442); +x_443 = lean_nat_dec_eq(x_430, x_442); +lean_dec(x_430); +if (x_443 == 0) +{ +lean_object* x_444; +lean_dec(x_442); +if (lean_is_scalar(x_441)) { + x_444 = lean_alloc_ctor(1, 2, 0); +} else { + x_444 = x_441; +} +lean_ctor_set(x_444, 0, x_439); +lean_ctor_set(x_444, 1, x_440); +return x_444; +} +else +{ +lean_object* x_445; lean_object* x_446; +lean_dec(x_441); +lean_dec(x_440); +x_445 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_446 = l_Std_Internal_Parsec_String_pstring(x_445, x_439); +if (lean_obj_tag(x_446) == 0) +{ +lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; +lean_dec(x_442); +x_447 = lean_ctor_get(x_446, 0); +lean_inc(x_447); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_448 = x_446; +} else { + lean_dec_ref(x_446); + x_448 = lean_box(0); +} +x_449 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(0, 2, 0); +} else { + x_450 = x_448; +} +lean_ctor_set(x_450, 0, x_447); +lean_ctor_set(x_450, 1, x_449); +return x_450; +} +else +{ +lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; uint8_t x_455; +x_451 = lean_ctor_get(x_446, 0); +lean_inc(x_451); +x_452 = lean_ctor_get(x_446, 1); +lean_inc(x_452); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_453 = x_446; +} else { + lean_dec_ref(x_446); + x_453 = lean_box(0); +} +x_454 = lean_ctor_get(x_451, 1); +lean_inc(x_454); +x_455 = lean_nat_dec_eq(x_442, x_454); +lean_dec(x_442); +if (x_455 == 0) +{ +lean_object* x_456; +lean_dec(x_454); +if (lean_is_scalar(x_453)) { + x_456 = lean_alloc_ctor(1, 2, 0); +} else { + x_456 = x_453; +} +lean_ctor_set(x_456, 0, x_451); +lean_ctor_set(x_456, 1, x_452); +return x_456; +} +else +{ +lean_object* x_457; lean_object* x_458; +lean_dec(x_453); +lean_dec(x_452); +x_457 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_458 = l_Std_Internal_Parsec_String_pstring(x_457, x_451); +if (lean_obj_tag(x_458) == 0) +{ +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; +lean_dec(x_454); +x_459 = lean_ctor_get(x_458, 0); +lean_inc(x_459); +if (lean_is_exclusive(x_458)) { + lean_ctor_release(x_458, 0); + lean_ctor_release(x_458, 1); + x_460 = x_458; +} else { + lean_dec_ref(x_458); + x_460 = lean_box(0); +} +x_461 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_460)) { + x_462 = lean_alloc_ctor(0, 2, 0); +} else { + x_462 = x_460; +} +lean_ctor_set(x_462, 0, x_459); +lean_ctor_set(x_462, 1, x_461); +return x_462; +} +else +{ +lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; uint8_t x_467; +x_463 = lean_ctor_get(x_458, 0); +lean_inc(x_463); +x_464 = lean_ctor_get(x_458, 1); +lean_inc(x_464); +if (lean_is_exclusive(x_458)) { + lean_ctor_release(x_458, 0); + lean_ctor_release(x_458, 1); + x_465 = x_458; +} else { + lean_dec_ref(x_458); + x_465 = lean_box(0); +} +x_466 = lean_ctor_get(x_463, 1); +lean_inc(x_466); +x_467 = lean_nat_dec_eq(x_454, x_466); +lean_dec(x_454); +if (x_467 == 0) +{ +lean_object* x_468; +lean_dec(x_466); +if (lean_is_scalar(x_465)) { + x_468 = lean_alloc_ctor(1, 2, 0); +} else { + x_468 = x_465; +} +lean_ctor_set(x_468, 0, x_463); +lean_ctor_set(x_468, 1, x_464); +return x_468; +} +else +{ +lean_object* x_469; lean_object* x_470; +lean_dec(x_465); +lean_dec(x_464); +x_469 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_470 = l_Std_Internal_Parsec_String_pstring(x_469, x_463); +if (lean_obj_tag(x_470) == 0) +{ +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; +lean_dec(x_466); +x_471 = lean_ctor_get(x_470, 0); +lean_inc(x_471); +if (lean_is_exclusive(x_470)) { + lean_ctor_release(x_470, 0); + lean_ctor_release(x_470, 1); + x_472 = x_470; +} else { + lean_dec_ref(x_470); + x_472 = lean_box(0); +} +x_473 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_472)) { + x_474 = lean_alloc_ctor(0, 2, 0); +} else { + x_474 = x_472; +} +lean_ctor_set(x_474, 0, x_471); +lean_ctor_set(x_474, 1, x_473); +return x_474; +} +else +{ +lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; +x_475 = lean_ctor_get(x_470, 0); +lean_inc(x_475); +x_476 = lean_ctor_get(x_470, 1); +lean_inc(x_476); +if (lean_is_exclusive(x_470)) { + lean_ctor_release(x_470, 0); + lean_ctor_release(x_470, 1); + x_477 = x_470; +} else { + lean_dec_ref(x_470); + x_477 = lean_box(0); +} +x_478 = lean_ctor_get(x_475, 1); +lean_inc(x_478); +x_479 = lean_nat_dec_eq(x_466, x_478); +lean_dec(x_466); +if (x_479 == 0) +{ +lean_object* x_480; +lean_dec(x_478); +if (lean_is_scalar(x_477)) { + x_480 = lean_alloc_ctor(1, 2, 0); +} else { + x_480 = x_477; +} +lean_ctor_set(x_480, 0, x_475); +lean_ctor_set(x_480, 1, x_476); +return x_480; +} +else +{ +lean_object* x_481; lean_object* x_482; +lean_dec(x_477); +lean_dec(x_476); +x_481 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_482 = l_Std_Internal_Parsec_String_pstring(x_481, x_475); +if (lean_obj_tag(x_482) == 0) +{ +lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; +lean_dec(x_478); +x_483 = lean_ctor_get(x_482, 0); +lean_inc(x_483); +if (lean_is_exclusive(x_482)) { + lean_ctor_release(x_482, 0); + lean_ctor_release(x_482, 1); + x_484 = x_482; +} else { + lean_dec_ref(x_482); + x_484 = lean_box(0); +} +x_485 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_484)) { + x_486 = lean_alloc_ctor(0, 2, 0); +} else { + x_486 = x_484; +} +lean_ctor_set(x_486, 0, x_483); +lean_ctor_set(x_486, 1, x_485); +return x_486; +} +else +{ +lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; uint8_t x_491; +x_487 = lean_ctor_get(x_482, 0); +lean_inc(x_487); +x_488 = lean_ctor_get(x_482, 1); +lean_inc(x_488); +if (lean_is_exclusive(x_482)) { + lean_ctor_release(x_482, 0); + lean_ctor_release(x_482, 1); + x_489 = x_482; +} else { + lean_dec_ref(x_482); + x_489 = lean_box(0); +} +x_490 = lean_ctor_get(x_487, 1); +lean_inc(x_490); +x_491 = lean_nat_dec_eq(x_478, x_490); +lean_dec(x_478); +if (x_491 == 0) +{ +lean_object* x_492; +lean_dec(x_490); +if (lean_is_scalar(x_489)) { + x_492 = lean_alloc_ctor(1, 2, 0); +} else { + x_492 = x_489; +} +lean_ctor_set(x_492, 0, x_487); +lean_ctor_set(x_492, 1, x_488); +return x_492; +} +else +{ +lean_object* x_493; lean_object* x_494; +lean_dec(x_489); +lean_dec(x_488); +x_493 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_494 = l_Std_Internal_Parsec_String_pstring(x_493, x_487); +if (lean_obj_tag(x_494) == 0) +{ +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; +lean_dec(x_490); +x_495 = lean_ctor_get(x_494, 0); +lean_inc(x_495); +if (lean_is_exclusive(x_494)) { + lean_ctor_release(x_494, 0); + lean_ctor_release(x_494, 1); + x_496 = x_494; +} else { + lean_dec_ref(x_494); + x_496 = lean_box(0); +} +x_497 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_496)) { + x_498 = lean_alloc_ctor(0, 2, 0); +} else { + x_498 = x_496; +} +lean_ctor_set(x_498, 0, x_495); +lean_ctor_set(x_498, 1, x_497); +return x_498; +} +else +{ +lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; uint8_t x_503; +x_499 = lean_ctor_get(x_494, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_494, 1); +lean_inc(x_500); +if (lean_is_exclusive(x_494)) { + lean_ctor_release(x_494, 0); + lean_ctor_release(x_494, 1); + x_501 = x_494; +} else { + lean_dec_ref(x_494); + x_501 = lean_box(0); +} +x_502 = lean_ctor_get(x_499, 1); +lean_inc(x_502); +x_503 = lean_nat_dec_eq(x_490, x_502); +lean_dec(x_502); +lean_dec(x_490); +if (x_503 == 0) +{ +lean_object* x_504; +if (lean_is_scalar(x_501)) { + x_504 = lean_alloc_ctor(1, 2, 0); +} else { + x_504 = x_501; +} +lean_ctor_set(x_504, 0, x_499); +lean_ctor_set(x_504, 1, x_500); +return x_504; +} +else +{ +lean_object* x_505; lean_object* x_506; +lean_dec(x_501); +lean_dec(x_500); +x_505 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_506 = l_Std_Internal_Parsec_String_pstring(x_505, x_499); +if (lean_obj_tag(x_506) == 0) +{ +lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +x_507 = lean_ctor_get(x_506, 0); +lean_inc(x_507); +if (lean_is_exclusive(x_506)) { + lean_ctor_release(x_506, 0); + lean_ctor_release(x_506, 1); + x_508 = x_506; +} else { + lean_dec_ref(x_506); + x_508 = lean_box(0); +} +x_509 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_508)) { + x_510 = lean_alloc_ctor(0, 2, 0); +} else { + x_510 = x_508; +} +lean_ctor_set(x_510, 0, x_507); +lean_ctor_set(x_510, 1, x_509); +return x_510; +} +else +{ +lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +x_511 = lean_ctor_get(x_506, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_506, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_506)) { + lean_ctor_release(x_506, 0); + lean_ctor_release(x_506, 1); + x_513 = x_506; +} else { + lean_dec_ref(x_506); + x_513 = lean_box(0); +} +if (lean_is_scalar(x_513)) { + x_514 = lean_alloc_ctor(1, 2, 0); +} else { + x_514 = x_513; +} +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_512); +return x_514; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_515; lean_object* x_516; lean_object* x_517; uint8_t x_518; +x_515 = lean_ctor_get(x_43, 0); +x_516 = lean_ctor_get(x_43, 1); +lean_inc(x_516); +lean_inc(x_515); +lean_dec(x_43); +x_517 = lean_ctor_get(x_515, 1); +lean_inc(x_517); +x_518 = lean_nat_dec_eq(x_40, x_517); +lean_dec(x_40); +if (x_518 == 0) +{ +lean_object* x_519; +lean_dec(x_517); +x_519 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_519, 0, x_515); +lean_ctor_set(x_519, 1, x_516); +return x_519; +} +else +{ +lean_object* x_520; lean_object* x_521; +lean_dec(x_516); +x_520 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_521 = l_Std_Internal_Parsec_String_pstring(x_520, x_515); +if (lean_obj_tag(x_521) == 0) +{ +lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +lean_dec(x_517); +x_522 = lean_ctor_get(x_521, 0); +lean_inc(x_522); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_523 = x_521; +} else { + lean_dec_ref(x_521); + x_523 = lean_box(0); +} +x_524 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_523)) { + x_525 = lean_alloc_ctor(0, 2, 0); +} else { + x_525 = x_523; +} +lean_ctor_set(x_525, 0, x_522); +lean_ctor_set(x_525, 1, x_524); +return x_525; +} +else +{ +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; uint8_t x_530; +x_526 = lean_ctor_get(x_521, 0); +lean_inc(x_526); +x_527 = lean_ctor_get(x_521, 1); +lean_inc(x_527); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_528 = x_521; +} else { + lean_dec_ref(x_521); + x_528 = lean_box(0); +} +x_529 = lean_ctor_get(x_526, 1); +lean_inc(x_529); +x_530 = lean_nat_dec_eq(x_517, x_529); +lean_dec(x_517); +if (x_530 == 0) +{ +lean_object* x_531; +lean_dec(x_529); +if (lean_is_scalar(x_528)) { + x_531 = lean_alloc_ctor(1, 2, 0); +} else { + x_531 = x_528; +} +lean_ctor_set(x_531, 0, x_526); +lean_ctor_set(x_531, 1, x_527); +return x_531; +} +else +{ +lean_object* x_532; lean_object* x_533; +lean_dec(x_528); +lean_dec(x_527); +x_532 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_533 = l_Std_Internal_Parsec_String_pstring(x_532, x_526); +if (lean_obj_tag(x_533) == 0) +{ +lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; +lean_dec(x_529); +x_534 = lean_ctor_get(x_533, 0); +lean_inc(x_534); +if (lean_is_exclusive(x_533)) { + lean_ctor_release(x_533, 0); + lean_ctor_release(x_533, 1); + x_535 = x_533; +} else { + lean_dec_ref(x_533); + x_535 = lean_box(0); +} +x_536 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_535)) { + x_537 = lean_alloc_ctor(0, 2, 0); +} else { + x_537 = x_535; +} +lean_ctor_set(x_537, 0, x_534); +lean_ctor_set(x_537, 1, x_536); +return x_537; +} +else +{ +lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; +x_538 = lean_ctor_get(x_533, 0); +lean_inc(x_538); +x_539 = lean_ctor_get(x_533, 1); +lean_inc(x_539); +if (lean_is_exclusive(x_533)) { + lean_ctor_release(x_533, 0); + lean_ctor_release(x_533, 1); + x_540 = x_533; +} else { + lean_dec_ref(x_533); + x_540 = lean_box(0); +} +x_541 = lean_ctor_get(x_538, 1); +lean_inc(x_541); +x_542 = lean_nat_dec_eq(x_529, x_541); +lean_dec(x_529); +if (x_542 == 0) +{ +lean_object* x_543; +lean_dec(x_541); +if (lean_is_scalar(x_540)) { + x_543 = lean_alloc_ctor(1, 2, 0); +} else { + x_543 = x_540; +} +lean_ctor_set(x_543, 0, x_538); +lean_ctor_set(x_543, 1, x_539); +return x_543; +} +else +{ +lean_object* x_544; lean_object* x_545; +lean_dec(x_540); +lean_dec(x_539); +x_544 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_545 = l_Std_Internal_Parsec_String_pstring(x_544, x_538); +if (lean_obj_tag(x_545) == 0) +{ +lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; +lean_dec(x_541); +x_546 = lean_ctor_get(x_545, 0); +lean_inc(x_546); +if (lean_is_exclusive(x_545)) { + lean_ctor_release(x_545, 0); + lean_ctor_release(x_545, 1); + x_547 = x_545; +} else { + lean_dec_ref(x_545); + x_547 = lean_box(0); +} +x_548 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_547)) { + x_549 = lean_alloc_ctor(0, 2, 0); +} else { + x_549 = x_547; +} +lean_ctor_set(x_549, 0, x_546); +lean_ctor_set(x_549, 1, x_548); +return x_549; +} +else +{ +lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; uint8_t x_554; +x_550 = lean_ctor_get(x_545, 0); +lean_inc(x_550); +x_551 = lean_ctor_get(x_545, 1); +lean_inc(x_551); +if (lean_is_exclusive(x_545)) { + lean_ctor_release(x_545, 0); + lean_ctor_release(x_545, 1); + x_552 = x_545; +} else { + lean_dec_ref(x_545); + x_552 = lean_box(0); +} +x_553 = lean_ctor_get(x_550, 1); +lean_inc(x_553); +x_554 = lean_nat_dec_eq(x_541, x_553); +lean_dec(x_541); +if (x_554 == 0) +{ +lean_object* x_555; +lean_dec(x_553); +if (lean_is_scalar(x_552)) { + x_555 = lean_alloc_ctor(1, 2, 0); +} else { + x_555 = x_552; +} +lean_ctor_set(x_555, 0, x_550); +lean_ctor_set(x_555, 1, x_551); +return x_555; +} +else +{ +lean_object* x_556; lean_object* x_557; +lean_dec(x_552); +lean_dec(x_551); +x_556 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_557 = l_Std_Internal_Parsec_String_pstring(x_556, x_550); +if (lean_obj_tag(x_557) == 0) +{ +lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; +lean_dec(x_553); +x_558 = lean_ctor_get(x_557, 0); +lean_inc(x_558); +if (lean_is_exclusive(x_557)) { + lean_ctor_release(x_557, 0); + lean_ctor_release(x_557, 1); + x_559 = x_557; +} else { + lean_dec_ref(x_557); + x_559 = lean_box(0); +} +x_560 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_559)) { + x_561 = lean_alloc_ctor(0, 2, 0); +} else { + x_561 = x_559; +} +lean_ctor_set(x_561, 0, x_558); +lean_ctor_set(x_561, 1, x_560); +return x_561; +} +else +{ +lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; uint8_t x_566; +x_562 = lean_ctor_get(x_557, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_557, 1); +lean_inc(x_563); +if (lean_is_exclusive(x_557)) { + lean_ctor_release(x_557, 0); + lean_ctor_release(x_557, 1); + x_564 = x_557; +} else { + lean_dec_ref(x_557); + x_564 = lean_box(0); +} +x_565 = lean_ctor_get(x_562, 1); +lean_inc(x_565); +x_566 = lean_nat_dec_eq(x_553, x_565); +lean_dec(x_553); +if (x_566 == 0) +{ +lean_object* x_567; +lean_dec(x_565); +if (lean_is_scalar(x_564)) { + x_567 = lean_alloc_ctor(1, 2, 0); +} else { + x_567 = x_564; +} +lean_ctor_set(x_567, 0, x_562); +lean_ctor_set(x_567, 1, x_563); +return x_567; +} +else +{ +lean_object* x_568; lean_object* x_569; +lean_dec(x_564); +lean_dec(x_563); +x_568 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_569 = l_Std_Internal_Parsec_String_pstring(x_568, x_562); +if (lean_obj_tag(x_569) == 0) +{ +lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; +lean_dec(x_565); +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_571 = x_569; +} else { + lean_dec_ref(x_569); + x_571 = lean_box(0); +} +x_572 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_571)) { + x_573 = lean_alloc_ctor(0, 2, 0); +} else { + x_573 = x_571; +} +lean_ctor_set(x_573, 0, x_570); +lean_ctor_set(x_573, 1, x_572); +return x_573; +} +else +{ +lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; uint8_t x_578; +x_574 = lean_ctor_get(x_569, 0); +lean_inc(x_574); +x_575 = lean_ctor_get(x_569, 1); +lean_inc(x_575); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_576 = x_569; +} else { + lean_dec_ref(x_569); + x_576 = lean_box(0); +} +x_577 = lean_ctor_get(x_574, 1); +lean_inc(x_577); +x_578 = lean_nat_dec_eq(x_565, x_577); +lean_dec(x_565); +if (x_578 == 0) +{ +lean_object* x_579; +lean_dec(x_577); +if (lean_is_scalar(x_576)) { + x_579 = lean_alloc_ctor(1, 2, 0); +} else { + x_579 = x_576; +} +lean_ctor_set(x_579, 0, x_574); +lean_ctor_set(x_579, 1, x_575); +return x_579; +} +else +{ +lean_object* x_580; lean_object* x_581; +lean_dec(x_576); +lean_dec(x_575); +x_580 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_581 = l_Std_Internal_Parsec_String_pstring(x_580, x_574); +if (lean_obj_tag(x_581) == 0) +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; +lean_dec(x_577); +x_582 = lean_ctor_get(x_581, 0); +lean_inc(x_582); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_583 = x_581; +} else { + lean_dec_ref(x_581); + x_583 = lean_box(0); +} +x_584 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_583)) { + x_585 = lean_alloc_ctor(0, 2, 0); +} else { + x_585 = x_583; +} +lean_ctor_set(x_585, 0, x_582); +lean_ctor_set(x_585, 1, x_584); +return x_585; +} +else +{ +lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; uint8_t x_590; +x_586 = lean_ctor_get(x_581, 0); +lean_inc(x_586); +x_587 = lean_ctor_get(x_581, 1); +lean_inc(x_587); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_588 = x_581; +} else { + lean_dec_ref(x_581); + x_588 = lean_box(0); +} +x_589 = lean_ctor_get(x_586, 1); +lean_inc(x_589); +x_590 = lean_nat_dec_eq(x_577, x_589); +lean_dec(x_577); +if (x_590 == 0) +{ +lean_object* x_591; +lean_dec(x_589); +if (lean_is_scalar(x_588)) { + x_591 = lean_alloc_ctor(1, 2, 0); +} else { + x_591 = x_588; +} +lean_ctor_set(x_591, 0, x_586); +lean_ctor_set(x_591, 1, x_587); +return x_591; +} +else +{ +lean_object* x_592; lean_object* x_593; +lean_dec(x_588); +lean_dec(x_587); +x_592 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_593 = l_Std_Internal_Parsec_String_pstring(x_592, x_586); +if (lean_obj_tag(x_593) == 0) +{ +lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; +lean_dec(x_589); +x_594 = lean_ctor_get(x_593, 0); +lean_inc(x_594); +if (lean_is_exclusive(x_593)) { + lean_ctor_release(x_593, 0); + lean_ctor_release(x_593, 1); + x_595 = x_593; +} else { + lean_dec_ref(x_593); + x_595 = lean_box(0); +} +x_596 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_595)) { + x_597 = lean_alloc_ctor(0, 2, 0); +} else { + x_597 = x_595; +} +lean_ctor_set(x_597, 0, x_594); +lean_ctor_set(x_597, 1, x_596); +return x_597; +} +else +{ +lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; uint8_t x_602; +x_598 = lean_ctor_get(x_593, 0); +lean_inc(x_598); +x_599 = lean_ctor_get(x_593, 1); +lean_inc(x_599); +if (lean_is_exclusive(x_593)) { + lean_ctor_release(x_593, 0); + lean_ctor_release(x_593, 1); + x_600 = x_593; +} else { + lean_dec_ref(x_593); + x_600 = lean_box(0); +} +x_601 = lean_ctor_get(x_598, 1); +lean_inc(x_601); +x_602 = lean_nat_dec_eq(x_589, x_601); +lean_dec(x_601); +lean_dec(x_589); +if (x_602 == 0) +{ +lean_object* x_603; +if (lean_is_scalar(x_600)) { + x_603 = lean_alloc_ctor(1, 2, 0); +} else { + x_603 = x_600; +} +lean_ctor_set(x_603, 0, x_598); +lean_ctor_set(x_603, 1, x_599); +return x_603; +} +else +{ +lean_object* x_604; lean_object* x_605; +lean_dec(x_600); +lean_dec(x_599); +x_604 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_605 = l_Std_Internal_Parsec_String_pstring(x_604, x_598); +if (lean_obj_tag(x_605) == 0) +{ +lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_606 = lean_ctor_get(x_605, 0); +lean_inc(x_606); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + x_607 = x_605; +} else { + lean_dec_ref(x_605); + x_607 = lean_box(0); +} +x_608 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_607)) { + x_609 = lean_alloc_ctor(0, 2, 0); +} else { + x_609 = x_607; +} +lean_ctor_set(x_609, 0, x_606); +lean_ctor_set(x_609, 1, x_608); +return x_609; +} +else +{ +lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; +x_610 = lean_ctor_get(x_605, 0); +lean_inc(x_610); +x_611 = lean_ctor_get(x_605, 1); +lean_inc(x_611); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + x_612 = x_605; +} else { + lean_dec_ref(x_605); + x_612 = lean_box(0); +} +if (lean_is_scalar(x_612)) { + x_613 = lean_alloc_ctor(1, 2, 0); +} else { + x_613 = x_612; +} +lean_ctor_set(x_613, 0, x_610); +lean_ctor_set(x_613, 1, x_611); +return x_613; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_614; lean_object* x_615; lean_object* x_616; uint8_t x_617; +x_614 = lean_ctor_get(x_30, 0); +x_615 = lean_ctor_get(x_30, 1); +lean_inc(x_615); +lean_inc(x_614); +lean_dec(x_30); +x_616 = lean_ctor_get(x_614, 1); +lean_inc(x_616); +x_617 = lean_nat_dec_eq(x_27, x_616); +lean_dec(x_27); +if (x_617 == 0) +{ +lean_object* x_618; +lean_dec(x_616); +x_618 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_618, 0, x_614); +lean_ctor_set(x_618, 1, x_615); +return x_618; +} +else +{ +lean_object* x_619; lean_object* x_620; +lean_dec(x_615); +x_619 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +x_620 = l_Std_Internal_Parsec_String_pstring(x_619, x_614); +if (lean_obj_tag(x_620) == 0) +{ +lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; +lean_dec(x_616); +x_621 = lean_ctor_get(x_620, 0); +lean_inc(x_621); +if (lean_is_exclusive(x_620)) { + lean_ctor_release(x_620, 0); + lean_ctor_release(x_620, 1); + x_622 = x_620; +} else { + lean_dec_ref(x_620); + x_622 = lean_box(0); +} +x_623 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_622)) { + x_624 = lean_alloc_ctor(0, 2, 0); +} else { + x_624 = x_622; +} +lean_ctor_set(x_624, 0, x_621); +lean_ctor_set(x_624, 1, x_623); +return x_624; +} +else +{ +lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; uint8_t x_629; +x_625 = lean_ctor_get(x_620, 0); +lean_inc(x_625); +x_626 = lean_ctor_get(x_620, 1); +lean_inc(x_626); +if (lean_is_exclusive(x_620)) { + lean_ctor_release(x_620, 0); + lean_ctor_release(x_620, 1); + x_627 = x_620; +} else { + lean_dec_ref(x_620); + x_627 = lean_box(0); +} +x_628 = lean_ctor_get(x_625, 1); +lean_inc(x_628); +x_629 = lean_nat_dec_eq(x_616, x_628); +lean_dec(x_616); +if (x_629 == 0) +{ +lean_object* x_630; +lean_dec(x_628); +if (lean_is_scalar(x_627)) { + x_630 = lean_alloc_ctor(1, 2, 0); +} else { + x_630 = x_627; +} +lean_ctor_set(x_630, 0, x_625); +lean_ctor_set(x_630, 1, x_626); +return x_630; +} +else +{ +lean_object* x_631; lean_object* x_632; +lean_dec(x_627); +lean_dec(x_626); +x_631 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_632 = l_Std_Internal_Parsec_String_pstring(x_631, x_625); +if (lean_obj_tag(x_632) == 0) +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; +lean_dec(x_628); +x_633 = lean_ctor_get(x_632, 0); +lean_inc(x_633); +if (lean_is_exclusive(x_632)) { + lean_ctor_release(x_632, 0); + lean_ctor_release(x_632, 1); + x_634 = x_632; +} else { + lean_dec_ref(x_632); + x_634 = lean_box(0); +} +x_635 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_634)) { + x_636 = lean_alloc_ctor(0, 2, 0); +} else { + x_636 = x_634; +} +lean_ctor_set(x_636, 0, x_633); +lean_ctor_set(x_636, 1, x_635); +return x_636; +} +else +{ +lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; +x_637 = lean_ctor_get(x_632, 0); +lean_inc(x_637); +x_638 = lean_ctor_get(x_632, 1); +lean_inc(x_638); +if (lean_is_exclusive(x_632)) { + lean_ctor_release(x_632, 0); + lean_ctor_release(x_632, 1); + x_639 = x_632; +} else { + lean_dec_ref(x_632); + x_639 = lean_box(0); +} +x_640 = lean_ctor_get(x_637, 1); +lean_inc(x_640); +x_641 = lean_nat_dec_eq(x_628, x_640); +lean_dec(x_628); +if (x_641 == 0) +{ +lean_object* x_642; +lean_dec(x_640); +if (lean_is_scalar(x_639)) { + x_642 = lean_alloc_ctor(1, 2, 0); +} else { + x_642 = x_639; +} +lean_ctor_set(x_642, 0, x_637); +lean_ctor_set(x_642, 1, x_638); +return x_642; +} +else +{ +lean_object* x_643; lean_object* x_644; +lean_dec(x_639); +lean_dec(x_638); +x_643 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_644 = l_Std_Internal_Parsec_String_pstring(x_643, x_637); +if (lean_obj_tag(x_644) == 0) +{ +lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; +lean_dec(x_640); +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +if (lean_is_exclusive(x_644)) { + lean_ctor_release(x_644, 0); + lean_ctor_release(x_644, 1); + x_646 = x_644; +} else { + lean_dec_ref(x_644); + x_646 = lean_box(0); +} +x_647 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_646)) { + x_648 = lean_alloc_ctor(0, 2, 0); +} else { + x_648 = x_646; +} +lean_ctor_set(x_648, 0, x_645); +lean_ctor_set(x_648, 1, x_647); +return x_648; +} +else +{ +lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; uint8_t x_653; +x_649 = lean_ctor_get(x_644, 0); +lean_inc(x_649); +x_650 = lean_ctor_get(x_644, 1); +lean_inc(x_650); +if (lean_is_exclusive(x_644)) { + lean_ctor_release(x_644, 0); + lean_ctor_release(x_644, 1); + x_651 = x_644; +} else { + lean_dec_ref(x_644); + x_651 = lean_box(0); +} +x_652 = lean_ctor_get(x_649, 1); +lean_inc(x_652); +x_653 = lean_nat_dec_eq(x_640, x_652); +lean_dec(x_640); +if (x_653 == 0) +{ +lean_object* x_654; +lean_dec(x_652); +if (lean_is_scalar(x_651)) { + x_654 = lean_alloc_ctor(1, 2, 0); +} else { + x_654 = x_651; +} +lean_ctor_set(x_654, 0, x_649); +lean_ctor_set(x_654, 1, x_650); +return x_654; +} +else +{ +lean_object* x_655; lean_object* x_656; +lean_dec(x_651); +lean_dec(x_650); +x_655 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_656 = l_Std_Internal_Parsec_String_pstring(x_655, x_649); +if (lean_obj_tag(x_656) == 0) +{ +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; +lean_dec(x_652); +x_657 = lean_ctor_get(x_656, 0); +lean_inc(x_657); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_658 = x_656; +} else { + lean_dec_ref(x_656); + x_658 = lean_box(0); +} +x_659 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_658)) { + x_660 = lean_alloc_ctor(0, 2, 0); +} else { + x_660 = x_658; +} +lean_ctor_set(x_660, 0, x_657); +lean_ctor_set(x_660, 1, x_659); +return x_660; +} +else +{ +lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; uint8_t x_665; +x_661 = lean_ctor_get(x_656, 0); +lean_inc(x_661); +x_662 = lean_ctor_get(x_656, 1); +lean_inc(x_662); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_663 = x_656; +} else { + lean_dec_ref(x_656); + x_663 = lean_box(0); +} +x_664 = lean_ctor_get(x_661, 1); +lean_inc(x_664); +x_665 = lean_nat_dec_eq(x_652, x_664); +lean_dec(x_652); +if (x_665 == 0) +{ +lean_object* x_666; +lean_dec(x_664); +if (lean_is_scalar(x_663)) { + x_666 = lean_alloc_ctor(1, 2, 0); +} else { + x_666 = x_663; +} +lean_ctor_set(x_666, 0, x_661); +lean_ctor_set(x_666, 1, x_662); +return x_666; +} +else +{ +lean_object* x_667; lean_object* x_668; +lean_dec(x_663); +lean_dec(x_662); +x_667 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_668 = l_Std_Internal_Parsec_String_pstring(x_667, x_661); +if (lean_obj_tag(x_668) == 0) +{ +lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; +lean_dec(x_664); +x_669 = lean_ctor_get(x_668, 0); +lean_inc(x_669); +if (lean_is_exclusive(x_668)) { + lean_ctor_release(x_668, 0); + lean_ctor_release(x_668, 1); + x_670 = x_668; +} else { + lean_dec_ref(x_668); + x_670 = lean_box(0); +} +x_671 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_670)) { + x_672 = lean_alloc_ctor(0, 2, 0); +} else { + x_672 = x_670; +} +lean_ctor_set(x_672, 0, x_669); +lean_ctor_set(x_672, 1, x_671); +return x_672; +} +else +{ +lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; uint8_t x_677; +x_673 = lean_ctor_get(x_668, 0); +lean_inc(x_673); +x_674 = lean_ctor_get(x_668, 1); +lean_inc(x_674); +if (lean_is_exclusive(x_668)) { + lean_ctor_release(x_668, 0); + lean_ctor_release(x_668, 1); + x_675 = x_668; +} else { + lean_dec_ref(x_668); + x_675 = lean_box(0); +} +x_676 = lean_ctor_get(x_673, 1); +lean_inc(x_676); +x_677 = lean_nat_dec_eq(x_664, x_676); +lean_dec(x_664); +if (x_677 == 0) +{ +lean_object* x_678; +lean_dec(x_676); +if (lean_is_scalar(x_675)) { + x_678 = lean_alloc_ctor(1, 2, 0); +} else { + x_678 = x_675; +} +lean_ctor_set(x_678, 0, x_673); +lean_ctor_set(x_678, 1, x_674); +return x_678; +} +else +{ +lean_object* x_679; lean_object* x_680; +lean_dec(x_675); +lean_dec(x_674); +x_679 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_680 = l_Std_Internal_Parsec_String_pstring(x_679, x_673); +if (lean_obj_tag(x_680) == 0) +{ +lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; +lean_dec(x_676); +x_681 = lean_ctor_get(x_680, 0); +lean_inc(x_681); +if (lean_is_exclusive(x_680)) { + lean_ctor_release(x_680, 0); + lean_ctor_release(x_680, 1); + x_682 = x_680; +} else { + lean_dec_ref(x_680); + x_682 = lean_box(0); +} +x_683 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_682)) { + x_684 = lean_alloc_ctor(0, 2, 0); +} else { + x_684 = x_682; +} +lean_ctor_set(x_684, 0, x_681); +lean_ctor_set(x_684, 1, x_683); +return x_684; +} +else +{ +lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; uint8_t x_689; +x_685 = lean_ctor_get(x_680, 0); +lean_inc(x_685); +x_686 = lean_ctor_get(x_680, 1); +lean_inc(x_686); +if (lean_is_exclusive(x_680)) { + lean_ctor_release(x_680, 0); + lean_ctor_release(x_680, 1); + x_687 = x_680; +} else { + lean_dec_ref(x_680); + x_687 = lean_box(0); +} +x_688 = lean_ctor_get(x_685, 1); +lean_inc(x_688); +x_689 = lean_nat_dec_eq(x_676, x_688); +lean_dec(x_676); +if (x_689 == 0) +{ +lean_object* x_690; +lean_dec(x_688); +if (lean_is_scalar(x_687)) { + x_690 = lean_alloc_ctor(1, 2, 0); +} else { + x_690 = x_687; +} +lean_ctor_set(x_690, 0, x_685); +lean_ctor_set(x_690, 1, x_686); +return x_690; +} +else +{ +lean_object* x_691; lean_object* x_692; +lean_dec(x_687); +lean_dec(x_686); +x_691 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_692 = l_Std_Internal_Parsec_String_pstring(x_691, x_685); +if (lean_obj_tag(x_692) == 0) +{ +lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; +lean_dec(x_688); +x_693 = lean_ctor_get(x_692, 0); +lean_inc(x_693); +if (lean_is_exclusive(x_692)) { + lean_ctor_release(x_692, 0); + lean_ctor_release(x_692, 1); + x_694 = x_692; +} else { + lean_dec_ref(x_692); + x_694 = lean_box(0); +} +x_695 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_694)) { + x_696 = lean_alloc_ctor(0, 2, 0); +} else { + x_696 = x_694; +} +lean_ctor_set(x_696, 0, x_693); +lean_ctor_set(x_696, 1, x_695); +return x_696; +} +else +{ +lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; uint8_t x_701; +x_697 = lean_ctor_get(x_692, 0); +lean_inc(x_697); +x_698 = lean_ctor_get(x_692, 1); +lean_inc(x_698); +if (lean_is_exclusive(x_692)) { + lean_ctor_release(x_692, 0); + lean_ctor_release(x_692, 1); + x_699 = x_692; +} else { + lean_dec_ref(x_692); + x_699 = lean_box(0); +} +x_700 = lean_ctor_get(x_697, 1); +lean_inc(x_700); +x_701 = lean_nat_dec_eq(x_688, x_700); +lean_dec(x_688); +if (x_701 == 0) +{ +lean_object* x_702; +lean_dec(x_700); +if (lean_is_scalar(x_699)) { + x_702 = lean_alloc_ctor(1, 2, 0); +} else { + x_702 = x_699; +} +lean_ctor_set(x_702, 0, x_697); +lean_ctor_set(x_702, 1, x_698); +return x_702; +} +else +{ +lean_object* x_703; lean_object* x_704; +lean_dec(x_699); +lean_dec(x_698); +x_703 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_704 = l_Std_Internal_Parsec_String_pstring(x_703, x_697); +if (lean_obj_tag(x_704) == 0) +{ +lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; +lean_dec(x_700); +x_705 = lean_ctor_get(x_704, 0); +lean_inc(x_705); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_706 = x_704; +} else { + lean_dec_ref(x_704); + x_706 = lean_box(0); +} +x_707 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_706)) { + x_708 = lean_alloc_ctor(0, 2, 0); +} else { + x_708 = x_706; +} +lean_ctor_set(x_708, 0, x_705); +lean_ctor_set(x_708, 1, x_707); +return x_708; +} +else +{ +lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; uint8_t x_713; +x_709 = lean_ctor_get(x_704, 0); +lean_inc(x_709); +x_710 = lean_ctor_get(x_704, 1); +lean_inc(x_710); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_711 = x_704; +} else { + lean_dec_ref(x_704); + x_711 = lean_box(0); +} +x_712 = lean_ctor_get(x_709, 1); +lean_inc(x_712); +x_713 = lean_nat_dec_eq(x_700, x_712); +lean_dec(x_712); +lean_dec(x_700); +if (x_713 == 0) +{ +lean_object* x_714; +if (lean_is_scalar(x_711)) { + x_714 = lean_alloc_ctor(1, 2, 0); +} else { + x_714 = x_711; +} +lean_ctor_set(x_714, 0, x_709); +lean_ctor_set(x_714, 1, x_710); +return x_714; +} +else +{ +lean_object* x_715; lean_object* x_716; +lean_dec(x_711); +lean_dec(x_710); +x_715 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_716 = l_Std_Internal_Parsec_String_pstring(x_715, x_709); +if (lean_obj_tag(x_716) == 0) +{ +lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; +x_717 = lean_ctor_get(x_716, 0); +lean_inc(x_717); +if (lean_is_exclusive(x_716)) { + lean_ctor_release(x_716, 0); + lean_ctor_release(x_716, 1); + x_718 = x_716; +} else { + lean_dec_ref(x_716); + x_718 = lean_box(0); +} +x_719 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_718)) { + x_720 = lean_alloc_ctor(0, 2, 0); +} else { + x_720 = x_718; +} +lean_ctor_set(x_720, 0, x_717); +lean_ctor_set(x_720, 1, x_719); +return x_720; +} +else +{ +lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; +x_721 = lean_ctor_get(x_716, 0); +lean_inc(x_721); +x_722 = lean_ctor_get(x_716, 1); +lean_inc(x_722); +if (lean_is_exclusive(x_716)) { + lean_ctor_release(x_716, 0); + lean_ctor_release(x_716, 1); + x_723 = x_716; +} else { + lean_dec_ref(x_716); + x_723 = lean_box(0); +} +if (lean_is_scalar(x_723)) { + x_724 = lean_alloc_ctor(1, 2, 0); +} else { + x_724 = x_723; +} +lean_ctor_set(x_724, 0, x_721); +lean_ctor_set(x_724, 1, x_722); +return x_724; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_725; lean_object* x_726; lean_object* x_727; uint8_t x_728; +x_725 = lean_ctor_get(x_17, 0); +x_726 = lean_ctor_get(x_17, 1); +lean_inc(x_726); +lean_inc(x_725); +lean_dec(x_17); +x_727 = lean_ctor_get(x_725, 1); +lean_inc(x_727); +x_728 = lean_nat_dec_eq(x_14, x_727); +lean_dec(x_14); +if (x_728 == 0) +{ +lean_object* x_729; +lean_dec(x_727); +x_729 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_729, 0, x_725); +lean_ctor_set(x_729, 1, x_726); +return x_729; +} +else +{ +lean_object* x_730; lean_object* x_731; +lean_dec(x_726); +x_730 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10; +x_731 = l_Std_Internal_Parsec_String_pstring(x_730, x_725); +if (lean_obj_tag(x_731) == 0) +{ +lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +lean_dec(x_727); +x_732 = lean_ctor_get(x_731, 0); +lean_inc(x_732); +if (lean_is_exclusive(x_731)) { + lean_ctor_release(x_731, 0); + lean_ctor_release(x_731, 1); + x_733 = x_731; +} else { + lean_dec_ref(x_731); + x_733 = lean_box(0); +} +x_734 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_733)) { + x_735 = lean_alloc_ctor(0, 2, 0); +} else { + x_735 = x_733; +} +lean_ctor_set(x_735, 0, x_732); +lean_ctor_set(x_735, 1, x_734); +return x_735; +} +else +{ +lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; uint8_t x_740; +x_736 = lean_ctor_get(x_731, 0); +lean_inc(x_736); +x_737 = lean_ctor_get(x_731, 1); +lean_inc(x_737); +if (lean_is_exclusive(x_731)) { + lean_ctor_release(x_731, 0); + lean_ctor_release(x_731, 1); + x_738 = x_731; +} else { + lean_dec_ref(x_731); + x_738 = lean_box(0); +} +x_739 = lean_ctor_get(x_736, 1); +lean_inc(x_739); +x_740 = lean_nat_dec_eq(x_727, x_739); +lean_dec(x_727); +if (x_740 == 0) +{ +lean_object* x_741; +lean_dec(x_739); +if (lean_is_scalar(x_738)) { + x_741 = lean_alloc_ctor(1, 2, 0); +} else { + x_741 = x_738; +} +lean_ctor_set(x_741, 0, x_736); +lean_ctor_set(x_741, 1, x_737); +return x_741; +} +else +{ +lean_object* x_742; lean_object* x_743; +lean_dec(x_738); +lean_dec(x_737); +x_742 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +x_743 = l_Std_Internal_Parsec_String_pstring(x_742, x_736); +if (lean_obj_tag(x_743) == 0) +{ +lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; +lean_dec(x_739); +x_744 = lean_ctor_get(x_743, 0); +lean_inc(x_744); +if (lean_is_exclusive(x_743)) { + lean_ctor_release(x_743, 0); + lean_ctor_release(x_743, 1); + x_745 = x_743; +} else { + lean_dec_ref(x_743); + x_745 = lean_box(0); +} +x_746 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_745)) { + x_747 = lean_alloc_ctor(0, 2, 0); +} else { + x_747 = x_745; +} +lean_ctor_set(x_747, 0, x_744); +lean_ctor_set(x_747, 1, x_746); +return x_747; +} +else +{ +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; uint8_t x_752; +x_748 = lean_ctor_get(x_743, 0); +lean_inc(x_748); +x_749 = lean_ctor_get(x_743, 1); +lean_inc(x_749); +if (lean_is_exclusive(x_743)) { + lean_ctor_release(x_743, 0); + lean_ctor_release(x_743, 1); + x_750 = x_743; +} else { + lean_dec_ref(x_743); + x_750 = lean_box(0); +} +x_751 = lean_ctor_get(x_748, 1); +lean_inc(x_751); +x_752 = lean_nat_dec_eq(x_739, x_751); +lean_dec(x_739); +if (x_752 == 0) +{ +lean_object* x_753; +lean_dec(x_751); +if (lean_is_scalar(x_750)) { + x_753 = lean_alloc_ctor(1, 2, 0); +} else { + x_753 = x_750; +} +lean_ctor_set(x_753, 0, x_748); +lean_ctor_set(x_753, 1, x_749); +return x_753; +} +else +{ +lean_object* x_754; lean_object* x_755; +lean_dec(x_750); +lean_dec(x_749); +x_754 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_755 = l_Std_Internal_Parsec_String_pstring(x_754, x_748); +if (lean_obj_tag(x_755) == 0) +{ +lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; +lean_dec(x_751); +x_756 = lean_ctor_get(x_755, 0); +lean_inc(x_756); +if (lean_is_exclusive(x_755)) { + lean_ctor_release(x_755, 0); + lean_ctor_release(x_755, 1); + x_757 = x_755; +} else { + lean_dec_ref(x_755); + x_757 = lean_box(0); +} +x_758 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_757)) { + x_759 = lean_alloc_ctor(0, 2, 0); +} else { + x_759 = x_757; +} +lean_ctor_set(x_759, 0, x_756); +lean_ctor_set(x_759, 1, x_758); +return x_759; +} +else +{ +lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; uint8_t x_764; +x_760 = lean_ctor_get(x_755, 0); +lean_inc(x_760); +x_761 = lean_ctor_get(x_755, 1); +lean_inc(x_761); +if (lean_is_exclusive(x_755)) { + lean_ctor_release(x_755, 0); + lean_ctor_release(x_755, 1); + x_762 = x_755; +} else { + lean_dec_ref(x_755); + x_762 = lean_box(0); +} +x_763 = lean_ctor_get(x_760, 1); +lean_inc(x_763); +x_764 = lean_nat_dec_eq(x_751, x_763); +lean_dec(x_751); +if (x_764 == 0) +{ +lean_object* x_765; +lean_dec(x_763); +if (lean_is_scalar(x_762)) { + x_765 = lean_alloc_ctor(1, 2, 0); +} else { + x_765 = x_762; +} +lean_ctor_set(x_765, 0, x_760); +lean_ctor_set(x_765, 1, x_761); +return x_765; +} +else +{ +lean_object* x_766; lean_object* x_767; +lean_dec(x_762); +lean_dec(x_761); +x_766 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_767 = l_Std_Internal_Parsec_String_pstring(x_766, x_760); +if (lean_obj_tag(x_767) == 0) +{ +lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; +lean_dec(x_763); +x_768 = lean_ctor_get(x_767, 0); +lean_inc(x_768); +if (lean_is_exclusive(x_767)) { + lean_ctor_release(x_767, 0); + lean_ctor_release(x_767, 1); + x_769 = x_767; +} else { + lean_dec_ref(x_767); + x_769 = lean_box(0); +} +x_770 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_769)) { + x_771 = lean_alloc_ctor(0, 2, 0); +} else { + x_771 = x_769; +} +lean_ctor_set(x_771, 0, x_768); +lean_ctor_set(x_771, 1, x_770); +return x_771; +} +else +{ +lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; uint8_t x_776; +x_772 = lean_ctor_get(x_767, 0); +lean_inc(x_772); +x_773 = lean_ctor_get(x_767, 1); +lean_inc(x_773); +if (lean_is_exclusive(x_767)) { + lean_ctor_release(x_767, 0); + lean_ctor_release(x_767, 1); + x_774 = x_767; +} else { + lean_dec_ref(x_767); + x_774 = lean_box(0); +} +x_775 = lean_ctor_get(x_772, 1); +lean_inc(x_775); +x_776 = lean_nat_dec_eq(x_763, x_775); +lean_dec(x_763); +if (x_776 == 0) +{ +lean_object* x_777; +lean_dec(x_775); +if (lean_is_scalar(x_774)) { + x_777 = lean_alloc_ctor(1, 2, 0); +} else { + x_777 = x_774; +} +lean_ctor_set(x_777, 0, x_772); +lean_ctor_set(x_777, 1, x_773); +return x_777; +} +else +{ +lean_object* x_778; lean_object* x_779; +lean_dec(x_774); +lean_dec(x_773); +x_778 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_779 = l_Std_Internal_Parsec_String_pstring(x_778, x_772); +if (lean_obj_tag(x_779) == 0) +{ +lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; +lean_dec(x_775); +x_780 = lean_ctor_get(x_779, 0); +lean_inc(x_780); +if (lean_is_exclusive(x_779)) { + lean_ctor_release(x_779, 0); + lean_ctor_release(x_779, 1); + x_781 = x_779; +} else { + lean_dec_ref(x_779); + x_781 = lean_box(0); +} +x_782 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_781)) { + x_783 = lean_alloc_ctor(0, 2, 0); +} else { + x_783 = x_781; +} +lean_ctor_set(x_783, 0, x_780); +lean_ctor_set(x_783, 1, x_782); +return x_783; +} +else +{ +lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; uint8_t x_788; +x_784 = lean_ctor_get(x_779, 0); +lean_inc(x_784); +x_785 = lean_ctor_get(x_779, 1); +lean_inc(x_785); +if (lean_is_exclusive(x_779)) { + lean_ctor_release(x_779, 0); + lean_ctor_release(x_779, 1); + x_786 = x_779; +} else { + lean_dec_ref(x_779); + x_786 = lean_box(0); +} +x_787 = lean_ctor_get(x_784, 1); +lean_inc(x_787); +x_788 = lean_nat_dec_eq(x_775, x_787); +lean_dec(x_775); +if (x_788 == 0) +{ +lean_object* x_789; +lean_dec(x_787); +if (lean_is_scalar(x_786)) { + x_789 = lean_alloc_ctor(1, 2, 0); +} else { + x_789 = x_786; +} +lean_ctor_set(x_789, 0, x_784); +lean_ctor_set(x_789, 1, x_785); +return x_789; +} +else +{ +lean_object* x_790; lean_object* x_791; +lean_dec(x_786); +lean_dec(x_785); +x_790 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_791 = l_Std_Internal_Parsec_String_pstring(x_790, x_784); +if (lean_obj_tag(x_791) == 0) +{ +lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +lean_dec(x_787); +x_792 = lean_ctor_get(x_791, 0); +lean_inc(x_792); +if (lean_is_exclusive(x_791)) { + lean_ctor_release(x_791, 0); + lean_ctor_release(x_791, 1); + x_793 = x_791; +} else { + lean_dec_ref(x_791); + x_793 = lean_box(0); +} +x_794 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_793)) { + x_795 = lean_alloc_ctor(0, 2, 0); +} else { + x_795 = x_793; +} +lean_ctor_set(x_795, 0, x_792); +lean_ctor_set(x_795, 1, x_794); +return x_795; +} +else +{ +lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; uint8_t x_800; +x_796 = lean_ctor_get(x_791, 0); +lean_inc(x_796); +x_797 = lean_ctor_get(x_791, 1); +lean_inc(x_797); +if (lean_is_exclusive(x_791)) { + lean_ctor_release(x_791, 0); + lean_ctor_release(x_791, 1); + x_798 = x_791; +} else { + lean_dec_ref(x_791); + x_798 = lean_box(0); +} +x_799 = lean_ctor_get(x_796, 1); +lean_inc(x_799); +x_800 = lean_nat_dec_eq(x_787, x_799); +lean_dec(x_787); +if (x_800 == 0) +{ +lean_object* x_801; +lean_dec(x_799); +if (lean_is_scalar(x_798)) { + x_801 = lean_alloc_ctor(1, 2, 0); +} else { + x_801 = x_798; +} +lean_ctor_set(x_801, 0, x_796); +lean_ctor_set(x_801, 1, x_797); +return x_801; +} +else +{ +lean_object* x_802; lean_object* x_803; +lean_dec(x_798); +lean_dec(x_797); +x_802 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_803 = l_Std_Internal_Parsec_String_pstring(x_802, x_796); +if (lean_obj_tag(x_803) == 0) +{ +lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; +lean_dec(x_799); +x_804 = lean_ctor_get(x_803, 0); +lean_inc(x_804); +if (lean_is_exclusive(x_803)) { + lean_ctor_release(x_803, 0); + lean_ctor_release(x_803, 1); + x_805 = x_803; +} else { + lean_dec_ref(x_803); + x_805 = lean_box(0); +} +x_806 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_805)) { + x_807 = lean_alloc_ctor(0, 2, 0); +} else { + x_807 = x_805; +} +lean_ctor_set(x_807, 0, x_804); +lean_ctor_set(x_807, 1, x_806); +return x_807; +} +else +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; uint8_t x_812; +x_808 = lean_ctor_get(x_803, 0); +lean_inc(x_808); +x_809 = lean_ctor_get(x_803, 1); +lean_inc(x_809); +if (lean_is_exclusive(x_803)) { + lean_ctor_release(x_803, 0); + lean_ctor_release(x_803, 1); + x_810 = x_803; +} else { + lean_dec_ref(x_803); + x_810 = lean_box(0); +} +x_811 = lean_ctor_get(x_808, 1); +lean_inc(x_811); +x_812 = lean_nat_dec_eq(x_799, x_811); +lean_dec(x_799); +if (x_812 == 0) +{ +lean_object* x_813; +lean_dec(x_811); +if (lean_is_scalar(x_810)) { + x_813 = lean_alloc_ctor(1, 2, 0); +} else { + x_813 = x_810; +} +lean_ctor_set(x_813, 0, x_808); +lean_ctor_set(x_813, 1, x_809); +return x_813; +} +else +{ +lean_object* x_814; lean_object* x_815; +lean_dec(x_810); +lean_dec(x_809); +x_814 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_815 = l_Std_Internal_Parsec_String_pstring(x_814, x_808); +if (lean_obj_tag(x_815) == 0) +{ +lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; +lean_dec(x_811); +x_816 = lean_ctor_get(x_815, 0); +lean_inc(x_816); +if (lean_is_exclusive(x_815)) { + lean_ctor_release(x_815, 0); + lean_ctor_release(x_815, 1); + x_817 = x_815; +} else { + lean_dec_ref(x_815); + x_817 = lean_box(0); +} +x_818 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_817)) { + x_819 = lean_alloc_ctor(0, 2, 0); +} else { + x_819 = x_817; +} +lean_ctor_set(x_819, 0, x_816); +lean_ctor_set(x_819, 1, x_818); +return x_819; +} +else +{ +lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; uint8_t x_824; +x_820 = lean_ctor_get(x_815, 0); +lean_inc(x_820); +x_821 = lean_ctor_get(x_815, 1); +lean_inc(x_821); +if (lean_is_exclusive(x_815)) { + lean_ctor_release(x_815, 0); + lean_ctor_release(x_815, 1); + x_822 = x_815; +} else { + lean_dec_ref(x_815); + x_822 = lean_box(0); +} +x_823 = lean_ctor_get(x_820, 1); +lean_inc(x_823); +x_824 = lean_nat_dec_eq(x_811, x_823); +lean_dec(x_811); +if (x_824 == 0) +{ +lean_object* x_825; +lean_dec(x_823); +if (lean_is_scalar(x_822)) { + x_825 = lean_alloc_ctor(1, 2, 0); +} else { + x_825 = x_822; +} +lean_ctor_set(x_825, 0, x_820); +lean_ctor_set(x_825, 1, x_821); +return x_825; +} +else +{ +lean_object* x_826; lean_object* x_827; +lean_dec(x_822); +lean_dec(x_821); +x_826 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_827 = l_Std_Internal_Parsec_String_pstring(x_826, x_820); +if (lean_obj_tag(x_827) == 0) +{ +lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; +lean_dec(x_823); +x_828 = lean_ctor_get(x_827, 0); +lean_inc(x_828); +if (lean_is_exclusive(x_827)) { + lean_ctor_release(x_827, 0); + lean_ctor_release(x_827, 1); + x_829 = x_827; +} else { + lean_dec_ref(x_827); + x_829 = lean_box(0); +} +x_830 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_829)) { + x_831 = lean_alloc_ctor(0, 2, 0); +} else { + x_831 = x_829; +} +lean_ctor_set(x_831, 0, x_828); +lean_ctor_set(x_831, 1, x_830); +return x_831; +} +else +{ +lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; uint8_t x_836; +x_832 = lean_ctor_get(x_827, 0); +lean_inc(x_832); +x_833 = lean_ctor_get(x_827, 1); +lean_inc(x_833); +if (lean_is_exclusive(x_827)) { + lean_ctor_release(x_827, 0); + lean_ctor_release(x_827, 1); + x_834 = x_827; +} else { + lean_dec_ref(x_827); + x_834 = lean_box(0); +} +x_835 = lean_ctor_get(x_832, 1); +lean_inc(x_835); +x_836 = lean_nat_dec_eq(x_823, x_835); +lean_dec(x_835); +lean_dec(x_823); +if (x_836 == 0) +{ +lean_object* x_837; +if (lean_is_scalar(x_834)) { + x_837 = lean_alloc_ctor(1, 2, 0); +} else { + x_837 = x_834; +} +lean_ctor_set(x_837, 0, x_832); +lean_ctor_set(x_837, 1, x_833); +return x_837; +} +else +{ +lean_object* x_838; lean_object* x_839; +lean_dec(x_834); +lean_dec(x_833); +x_838 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_839 = l_Std_Internal_Parsec_String_pstring(x_838, x_832); +if (lean_obj_tag(x_839) == 0) +{ +lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; +x_840 = lean_ctor_get(x_839, 0); +lean_inc(x_840); +if (lean_is_exclusive(x_839)) { + lean_ctor_release(x_839, 0); + lean_ctor_release(x_839, 1); + x_841 = x_839; +} else { + lean_dec_ref(x_839); + x_841 = lean_box(0); +} +x_842 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_841)) { + x_843 = lean_alloc_ctor(0, 2, 0); +} else { + x_843 = x_841; +} +lean_ctor_set(x_843, 0, x_840); +lean_ctor_set(x_843, 1, x_842); +return x_843; +} +else +{ +lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; +x_844 = lean_ctor_get(x_839, 0); +lean_inc(x_844); +x_845 = lean_ctor_get(x_839, 1); +lean_inc(x_845); +if (lean_is_exclusive(x_839)) { + lean_ctor_release(x_839, 0); + lean_ctor_release(x_839, 1); + x_846 = x_839; +} else { + lean_dec_ref(x_839); + x_846 = lean_box(0); +} +if (lean_is_scalar(x_846)) { + x_847 = lean_alloc_ctor(1, 2, 0); +} else { + x_847 = x_846; +} +lean_ctor_set(x_847, 0, x_844); +lean_ctor_set(x_847, 1, x_845); +return x_847; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; uint8_t x_852; +x_848 = lean_ctor_get(x_3, 0); +x_849 = lean_ctor_get(x_3, 1); +lean_inc(x_849); +lean_inc(x_848); +lean_dec(x_3); +x_850 = lean_ctor_get(x_1, 1); +lean_inc(x_850); +lean_dec(x_1); +x_851 = lean_ctor_get(x_848, 1); +lean_inc(x_851); +x_852 = lean_nat_dec_eq(x_850, x_851); +lean_dec(x_850); +if (x_852 == 0) +{ +lean_object* x_853; +lean_dec(x_851); +x_853 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_853, 0, x_848); +lean_ctor_set(x_853, 1, x_849); +return x_853; +} +else +{ +lean_object* x_854; lean_object* x_855; +lean_dec(x_849); +x_854 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11; +x_855 = l_Std_Internal_Parsec_String_pstring(x_854, x_848); +if (lean_obj_tag(x_855) == 0) +{ +lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; +lean_dec(x_851); +x_856 = lean_ctor_get(x_855, 0); +lean_inc(x_856); +if (lean_is_exclusive(x_855)) { + lean_ctor_release(x_855, 0); + lean_ctor_release(x_855, 1); + x_857 = x_855; +} else { + lean_dec_ref(x_855); + x_857 = lean_box(0); +} +x_858 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_857)) { + x_859 = lean_alloc_ctor(0, 2, 0); +} else { + x_859 = x_857; +} +lean_ctor_set(x_859, 0, x_856); +lean_ctor_set(x_859, 1, x_858); +return x_859; +} +else +{ +lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; uint8_t x_864; +x_860 = lean_ctor_get(x_855, 0); +lean_inc(x_860); +x_861 = lean_ctor_get(x_855, 1); +lean_inc(x_861); +if (lean_is_exclusive(x_855)) { + lean_ctor_release(x_855, 0); + lean_ctor_release(x_855, 1); + x_862 = x_855; +} else { + lean_dec_ref(x_855); + x_862 = lean_box(0); +} +x_863 = lean_ctor_get(x_860, 1); +lean_inc(x_863); +x_864 = lean_nat_dec_eq(x_851, x_863); +lean_dec(x_851); +if (x_864 == 0) +{ +lean_object* x_865; +lean_dec(x_863); +if (lean_is_scalar(x_862)) { + x_865 = lean_alloc_ctor(1, 2, 0); +} else { + x_865 = x_862; +} +lean_ctor_set(x_865, 0, x_860); +lean_ctor_set(x_865, 1, x_861); +return x_865; +} +else +{ +lean_object* x_866; lean_object* x_867; +lean_dec(x_862); +lean_dec(x_861); +x_866 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10; +x_867 = l_Std_Internal_Parsec_String_pstring(x_866, x_860); +if (lean_obj_tag(x_867) == 0) +{ +lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; +lean_dec(x_863); +x_868 = lean_ctor_get(x_867, 0); +lean_inc(x_868); +if (lean_is_exclusive(x_867)) { + lean_ctor_release(x_867, 0); + lean_ctor_release(x_867, 1); + x_869 = x_867; +} else { + lean_dec_ref(x_867); + x_869 = lean_box(0); +} +x_870 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_869)) { + x_871 = lean_alloc_ctor(0, 2, 0); +} else { + x_871 = x_869; +} +lean_ctor_set(x_871, 0, x_868); +lean_ctor_set(x_871, 1, x_870); +return x_871; +} +else +{ +lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; uint8_t x_876; +x_872 = lean_ctor_get(x_867, 0); +lean_inc(x_872); +x_873 = lean_ctor_get(x_867, 1); +lean_inc(x_873); +if (lean_is_exclusive(x_867)) { + lean_ctor_release(x_867, 0); + lean_ctor_release(x_867, 1); + x_874 = x_867; +} else { + lean_dec_ref(x_867); + x_874 = lean_box(0); +} +x_875 = lean_ctor_get(x_872, 1); +lean_inc(x_875); +x_876 = lean_nat_dec_eq(x_863, x_875); +lean_dec(x_863); +if (x_876 == 0) +{ +lean_object* x_877; +lean_dec(x_875); +if (lean_is_scalar(x_874)) { + x_877 = lean_alloc_ctor(1, 2, 0); +} else { + x_877 = x_874; +} +lean_ctor_set(x_877, 0, x_872); +lean_ctor_set(x_877, 1, x_873); +return x_877; +} +else +{ +lean_object* x_878; lean_object* x_879; +lean_dec(x_874); +lean_dec(x_873); +x_878 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9; +x_879 = l_Std_Internal_Parsec_String_pstring(x_878, x_872); +if (lean_obj_tag(x_879) == 0) +{ +lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_dec(x_875); +x_880 = lean_ctor_get(x_879, 0); +lean_inc(x_880); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_881 = x_879; +} else { + lean_dec_ref(x_879); + x_881 = lean_box(0); +} +x_882 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_881)) { + x_883 = lean_alloc_ctor(0, 2, 0); +} else { + x_883 = x_881; +} +lean_ctor_set(x_883, 0, x_880); +lean_ctor_set(x_883, 1, x_882); +return x_883; +} +else +{ +lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; uint8_t x_888; +x_884 = lean_ctor_get(x_879, 0); +lean_inc(x_884); +x_885 = lean_ctor_get(x_879, 1); +lean_inc(x_885); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_886 = x_879; +} else { + lean_dec_ref(x_879); + x_886 = lean_box(0); +} +x_887 = lean_ctor_get(x_884, 1); +lean_inc(x_887); +x_888 = lean_nat_dec_eq(x_875, x_887); +lean_dec(x_875); +if (x_888 == 0) +{ +lean_object* x_889; +lean_dec(x_887); +if (lean_is_scalar(x_886)) { + x_889 = lean_alloc_ctor(1, 2, 0); +} else { + x_889 = x_886; +} +lean_ctor_set(x_889, 0, x_884); +lean_ctor_set(x_889, 1, x_885); +return x_889; +} +else +{ +lean_object* x_890; lean_object* x_891; +lean_dec(x_886); +lean_dec(x_885); +x_890 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_891 = l_Std_Internal_Parsec_String_pstring(x_890, x_884); +if (lean_obj_tag(x_891) == 0) +{ +lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; +lean_dec(x_887); +x_892 = lean_ctor_get(x_891, 0); +lean_inc(x_892); +if (lean_is_exclusive(x_891)) { + lean_ctor_release(x_891, 0); + lean_ctor_release(x_891, 1); + x_893 = x_891; +} else { + lean_dec_ref(x_891); + x_893 = lean_box(0); +} +x_894 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_893)) { + x_895 = lean_alloc_ctor(0, 2, 0); +} else { + x_895 = x_893; +} +lean_ctor_set(x_895, 0, x_892); +lean_ctor_set(x_895, 1, x_894); +return x_895; +} +else +{ +lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; uint8_t x_900; +x_896 = lean_ctor_get(x_891, 0); +lean_inc(x_896); +x_897 = lean_ctor_get(x_891, 1); +lean_inc(x_897); +if (lean_is_exclusive(x_891)) { + lean_ctor_release(x_891, 0); + lean_ctor_release(x_891, 1); + x_898 = x_891; +} else { + lean_dec_ref(x_891); + x_898 = lean_box(0); +} +x_899 = lean_ctor_get(x_896, 1); +lean_inc(x_899); +x_900 = lean_nat_dec_eq(x_887, x_899); +lean_dec(x_887); +if (x_900 == 0) +{ +lean_object* x_901; +lean_dec(x_899); +if (lean_is_scalar(x_898)) { + x_901 = lean_alloc_ctor(1, 2, 0); +} else { + x_901 = x_898; +} +lean_ctor_set(x_901, 0, x_896); +lean_ctor_set(x_901, 1, x_897); +return x_901; +} +else +{ +lean_object* x_902; lean_object* x_903; +lean_dec(x_898); +lean_dec(x_897); +x_902 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7; +x_903 = l_Std_Internal_Parsec_String_pstring(x_902, x_896); +if (lean_obj_tag(x_903) == 0) +{ +lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; +lean_dec(x_899); +x_904 = lean_ctor_get(x_903, 0); +lean_inc(x_904); +if (lean_is_exclusive(x_903)) { + lean_ctor_release(x_903, 0); + lean_ctor_release(x_903, 1); + x_905 = x_903; +} else { + lean_dec_ref(x_903); + x_905 = lean_box(0); +} +x_906 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_905)) { + x_907 = lean_alloc_ctor(0, 2, 0); +} else { + x_907 = x_905; +} +lean_ctor_set(x_907, 0, x_904); +lean_ctor_set(x_907, 1, x_906); +return x_907; +} +else +{ +lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; uint8_t x_912; +x_908 = lean_ctor_get(x_903, 0); +lean_inc(x_908); +x_909 = lean_ctor_get(x_903, 1); +lean_inc(x_909); +if (lean_is_exclusive(x_903)) { + lean_ctor_release(x_903, 0); + lean_ctor_release(x_903, 1); + x_910 = x_903; +} else { + lean_dec_ref(x_903); + x_910 = lean_box(0); +} +x_911 = lean_ctor_get(x_908, 1); +lean_inc(x_911); +x_912 = lean_nat_dec_eq(x_899, x_911); +lean_dec(x_899); +if (x_912 == 0) +{ +lean_object* x_913; +lean_dec(x_911); +if (lean_is_scalar(x_910)) { + x_913 = lean_alloc_ctor(1, 2, 0); +} else { + x_913 = x_910; +} +lean_ctor_set(x_913, 0, x_908); +lean_ctor_set(x_913, 1, x_909); +return x_913; +} +else +{ +lean_object* x_914; lean_object* x_915; +lean_dec(x_910); +lean_dec(x_909); +x_914 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6; +x_915 = l_Std_Internal_Parsec_String_pstring(x_914, x_908); +if (lean_obj_tag(x_915) == 0) +{ +lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; +lean_dec(x_911); +x_916 = lean_ctor_get(x_915, 0); +lean_inc(x_916); +if (lean_is_exclusive(x_915)) { + lean_ctor_release(x_915, 0); + lean_ctor_release(x_915, 1); + x_917 = x_915; +} else { + lean_dec_ref(x_915); + x_917 = lean_box(0); +} +x_918 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_917)) { + x_919 = lean_alloc_ctor(0, 2, 0); +} else { + x_919 = x_917; +} +lean_ctor_set(x_919, 0, x_916); +lean_ctor_set(x_919, 1, x_918); +return x_919; +} +else +{ +lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; uint8_t x_924; +x_920 = lean_ctor_get(x_915, 0); +lean_inc(x_920); +x_921 = lean_ctor_get(x_915, 1); +lean_inc(x_921); +if (lean_is_exclusive(x_915)) { + lean_ctor_release(x_915, 0); + lean_ctor_release(x_915, 1); + x_922 = x_915; +} else { + lean_dec_ref(x_915); + x_922 = lean_box(0); +} +x_923 = lean_ctor_get(x_920, 1); +lean_inc(x_923); +x_924 = lean_nat_dec_eq(x_911, x_923); +lean_dec(x_911); +if (x_924 == 0) +{ +lean_object* x_925; +lean_dec(x_923); +if (lean_is_scalar(x_922)) { + x_925 = lean_alloc_ctor(1, 2, 0); +} else { + x_925 = x_922; +} +lean_ctor_set(x_925, 0, x_920); +lean_ctor_set(x_925, 1, x_921); +return x_925; +} +else +{ +lean_object* x_926; lean_object* x_927; +lean_dec(x_922); +lean_dec(x_921); +x_926 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5; +x_927 = l_Std_Internal_Parsec_String_pstring(x_926, x_920); +if (lean_obj_tag(x_927) == 0) +{ +lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +lean_dec(x_923); +x_928 = lean_ctor_get(x_927, 0); +lean_inc(x_928); +if (lean_is_exclusive(x_927)) { + lean_ctor_release(x_927, 0); + lean_ctor_release(x_927, 1); + x_929 = x_927; +} else { + lean_dec_ref(x_927); + x_929 = lean_box(0); +} +x_930 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_929)) { + x_931 = lean_alloc_ctor(0, 2, 0); +} else { + x_931 = x_929; +} +lean_ctor_set(x_931, 0, x_928); +lean_ctor_set(x_931, 1, x_930); +return x_931; +} +else +{ +lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; uint8_t x_936; +x_932 = lean_ctor_get(x_927, 0); +lean_inc(x_932); +x_933 = lean_ctor_get(x_927, 1); +lean_inc(x_933); +if (lean_is_exclusive(x_927)) { + lean_ctor_release(x_927, 0); + lean_ctor_release(x_927, 1); + x_934 = x_927; +} else { + lean_dec_ref(x_927); + x_934 = lean_box(0); +} +x_935 = lean_ctor_get(x_932, 1); +lean_inc(x_935); +x_936 = lean_nat_dec_eq(x_923, x_935); +lean_dec(x_923); +if (x_936 == 0) +{ +lean_object* x_937; +lean_dec(x_935); +if (lean_is_scalar(x_934)) { + x_937 = lean_alloc_ctor(1, 2, 0); +} else { + x_937 = x_934; +} +lean_ctor_set(x_937, 0, x_932); +lean_ctor_set(x_937, 1, x_933); +return x_937; +} +else +{ +lean_object* x_938; lean_object* x_939; +lean_dec(x_934); +lean_dec(x_933); +x_938 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4; +x_939 = l_Std_Internal_Parsec_String_pstring(x_938, x_932); +if (lean_obj_tag(x_939) == 0) +{ +lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; +lean_dec(x_935); +x_940 = lean_ctor_get(x_939, 0); +lean_inc(x_940); +if (lean_is_exclusive(x_939)) { + lean_ctor_release(x_939, 0); + lean_ctor_release(x_939, 1); + x_941 = x_939; +} else { + lean_dec_ref(x_939); + x_941 = lean_box(0); +} +x_942 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_941)) { + x_943 = lean_alloc_ctor(0, 2, 0); +} else { + x_943 = x_941; +} +lean_ctor_set(x_943, 0, x_940); +lean_ctor_set(x_943, 1, x_942); +return x_943; +} +else +{ +lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; uint8_t x_948; +x_944 = lean_ctor_get(x_939, 0); +lean_inc(x_944); +x_945 = lean_ctor_get(x_939, 1); +lean_inc(x_945); +if (lean_is_exclusive(x_939)) { + lean_ctor_release(x_939, 0); + lean_ctor_release(x_939, 1); + x_946 = x_939; +} else { + lean_dec_ref(x_939); + x_946 = lean_box(0); +} +x_947 = lean_ctor_get(x_944, 1); +lean_inc(x_947); +x_948 = lean_nat_dec_eq(x_935, x_947); +lean_dec(x_935); +if (x_948 == 0) +{ +lean_object* x_949; +lean_dec(x_947); +if (lean_is_scalar(x_946)) { + x_949 = lean_alloc_ctor(1, 2, 0); +} else { + x_949 = x_946; +} +lean_ctor_set(x_949, 0, x_944); +lean_ctor_set(x_949, 1, x_945); +return x_949; +} +else +{ +lean_object* x_950; lean_object* x_951; +lean_dec(x_946); +lean_dec(x_945); +x_950 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3; +x_951 = l_Std_Internal_Parsec_String_pstring(x_950, x_944); +if (lean_obj_tag(x_951) == 0) +{ +lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; +lean_dec(x_947); +x_952 = lean_ctor_get(x_951, 0); +lean_inc(x_952); +if (lean_is_exclusive(x_951)) { + lean_ctor_release(x_951, 0); + lean_ctor_release(x_951, 1); + x_953 = x_951; +} else { + lean_dec_ref(x_951); + x_953 = lean_box(0); +} +x_954 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_953)) { + x_955 = lean_alloc_ctor(0, 2, 0); +} else { + x_955 = x_953; +} +lean_ctor_set(x_955, 0, x_952); +lean_ctor_set(x_955, 1, x_954); +return x_955; +} +else +{ +lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; uint8_t x_960; +x_956 = lean_ctor_get(x_951, 0); +lean_inc(x_956); +x_957 = lean_ctor_get(x_951, 1); +lean_inc(x_957); +if (lean_is_exclusive(x_951)) { + lean_ctor_release(x_951, 0); + lean_ctor_release(x_951, 1); + x_958 = x_951; +} else { + lean_dec_ref(x_951); + x_958 = lean_box(0); +} +x_959 = lean_ctor_get(x_956, 1); +lean_inc(x_959); +x_960 = lean_nat_dec_eq(x_947, x_959); +lean_dec(x_947); +if (x_960 == 0) +{ +lean_object* x_961; +lean_dec(x_959); +if (lean_is_scalar(x_958)) { + x_961 = lean_alloc_ctor(1, 2, 0); +} else { + x_961 = x_958; +} +lean_ctor_set(x_961, 0, x_956); +lean_ctor_set(x_961, 1, x_957); +return x_961; +} +else +{ +lean_object* x_962; lean_object* x_963; +lean_dec(x_958); +lean_dec(x_957); +x_962 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2; +x_963 = l_Std_Internal_Parsec_String_pstring(x_962, x_956); +if (lean_obj_tag(x_963) == 0) +{ +lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; +lean_dec(x_959); +x_964 = lean_ctor_get(x_963, 0); +lean_inc(x_964); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_965 = x_963; +} else { + lean_dec_ref(x_963); + x_965 = lean_box(0); +} +x_966 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_965)) { + x_967 = lean_alloc_ctor(0, 2, 0); +} else { + x_967 = x_965; +} +lean_ctor_set(x_967, 0, x_964); +lean_ctor_set(x_967, 1, x_966); +return x_967; +} +else +{ +lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; uint8_t x_972; +x_968 = lean_ctor_get(x_963, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_963, 1); +lean_inc(x_969); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_970 = x_963; +} else { + lean_dec_ref(x_963); + x_970 = lean_box(0); +} +x_971 = lean_ctor_get(x_968, 1); +lean_inc(x_971); +x_972 = lean_nat_dec_eq(x_959, x_971); +lean_dec(x_971); +lean_dec(x_959); +if (x_972 == 0) +{ +lean_object* x_973; +if (lean_is_scalar(x_970)) { + x_973 = lean_alloc_ctor(1, 2, 0); +} else { + x_973 = x_970; +} +lean_ctor_set(x_973, 0, x_968); +lean_ctor_set(x_973, 1, x_969); +return x_973; +} +else +{ +lean_object* x_974; lean_object* x_975; +lean_dec(x_970); +lean_dec(x_969); +x_974 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1; +x_975 = l_Std_Internal_Parsec_String_pstring(x_974, x_968); +if (lean_obj_tag(x_975) == 0) +{ +lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; +x_976 = lean_ctor_get(x_975, 0); +lean_inc(x_976); +if (lean_is_exclusive(x_975)) { + lean_ctor_release(x_975, 0); + lean_ctor_release(x_975, 1); + x_977 = x_975; +} else { + lean_dec_ref(x_975); + x_977 = lean_box(0); +} +x_978 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_977)) { + x_979 = lean_alloc_ctor(0, 2, 0); +} else { + x_979 = x_977; +} +lean_ctor_set(x_979, 0, x_976); +lean_ctor_set(x_979, 1, x_978); +return x_979; +} +else +{ +lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +x_980 = lean_ctor_get(x_975, 0); +lean_inc(x_980); +x_981 = lean_ctor_get(x_975, 1); +lean_inc(x_981); +if (lean_is_exclusive(x_975)) { + lean_ctor_release(x_975, 0); + lean_ctor_release(x_975, 1); + x_982 = x_975; +} else { + lean_dec_ref(x_975); + x_982 = lean_box(0); +} +if (lean_is_scalar(x_982)) { + x_983 = lean_alloc_ctor(1, 2, 0); +} else { + x_983 = x_982; +} +lean_ctor_set(x_983, 0, x_980); +lean_ctor_set(x_983, 1, x_981); +return x_983; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_parseMonthShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +lean_dec(x_40); +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +lean_dec(x_40); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +x_54 = lean_nat_dec_eq(x_40, x_53); +lean_dec(x_40); +if (x_54 == 0) +{ +lean_dec(x_53); +return x_43; +} +else +{ +lean_object* x_55; lean_object* x_56; +lean_free_object(x_43); +lean_dec(x_52); +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_56 = l_Std_Internal_Parsec_String_pstring(x_55, x_51); +if (lean_obj_tag(x_56) == 0) +{ +uint8_t x_57; +lean_dec(x_53); +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 1); +lean_dec(x_58); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +lean_ctor_set(x_56, 1, x_59); +return x_56; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_56, 0); +lean_inc(x_60); +lean_dec(x_56); +x_61 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_56); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_64 = lean_ctor_get(x_56, 0); +x_65 = lean_ctor_get(x_56, 1); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +x_67 = lean_nat_dec_eq(x_53, x_66); +lean_dec(x_53); +if (x_67 == 0) +{ +lean_dec(x_66); +return x_56; +} +else +{ +lean_object* x_68; lean_object* x_69; +lean_free_object(x_56); +lean_dec(x_65); +x_68 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_69 = l_Std_Internal_Parsec_String_pstring(x_68, x_64); +if (lean_obj_tag(x_69) == 0) +{ +uint8_t x_70; +lean_dec(x_66); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_69, 1); +lean_dec(x_71); +x_72 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +lean_ctor_set(x_69, 1, x_72); +return x_69; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_69, 0); +lean_inc(x_73); +lean_dec(x_69); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +else +{ +uint8_t x_76; +x_76 = !lean_is_exclusive(x_69); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_77 = lean_ctor_get(x_69, 0); +x_78 = lean_ctor_get(x_69, 1); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +x_80 = lean_nat_dec_eq(x_66, x_79); +lean_dec(x_66); +if (x_80 == 0) +{ +lean_dec(x_79); +return x_69; +} +else +{ +lean_object* x_81; lean_object* x_82; +lean_free_object(x_69); +lean_dec(x_78); +x_81 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_82 = l_Std_Internal_Parsec_String_pstring(x_81, x_77); +if (lean_obj_tag(x_82) == 0) +{ +uint8_t x_83; +lean_dec(x_79); +x_83 = !lean_is_exclusive(x_82); +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; +x_84 = lean_ctor_get(x_82, 1); +lean_dec(x_84); +x_85 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +lean_ctor_set(x_82, 1, x_85); +return x_82; +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_82, 0); +lean_inc(x_86); +lean_dec(x_82); +x_87 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; +} +} +else +{ +uint8_t x_89; +x_89 = !lean_is_exclusive(x_82); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_90 = lean_ctor_get(x_82, 0); +x_91 = lean_ctor_get(x_82, 1); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +x_93 = lean_nat_dec_eq(x_79, x_92); +lean_dec(x_79); +if (x_93 == 0) +{ +lean_dec(x_92); +return x_82; +} +else +{ +lean_object* x_94; lean_object* x_95; +lean_free_object(x_82); +lean_dec(x_91); +x_94 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_95 = l_Std_Internal_Parsec_String_pstring(x_94, x_90); +if (lean_obj_tag(x_95) == 0) +{ +uint8_t x_96; +lean_dec(x_92); +x_96 = !lean_is_exclusive(x_95); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_95, 1); +lean_dec(x_97); +x_98 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +lean_ctor_set(x_95, 1, x_98); +return x_95; +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_95, 0); +lean_inc(x_99); +lean_dec(x_95); +x_100 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_100); +return x_101; +} +} +else +{ +uint8_t x_102; +x_102 = !lean_is_exclusive(x_95); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_103 = lean_ctor_get(x_95, 0); +x_104 = lean_ctor_get(x_95, 1); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +x_106 = lean_nat_dec_eq(x_92, x_105); +lean_dec(x_92); +if (x_106 == 0) +{ +lean_dec(x_105); +return x_95; +} +else +{ +lean_object* x_107; lean_object* x_108; +lean_free_object(x_95); +lean_dec(x_104); +x_107 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_108 = l_Std_Internal_Parsec_String_pstring(x_107, x_103); +if (lean_obj_tag(x_108) == 0) +{ +uint8_t x_109; +lean_dec(x_105); +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_108, 1); +lean_dec(x_110); +x_111 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +lean_ctor_set(x_108, 1, x_111); +return x_108; +} +else +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_108, 0); +lean_inc(x_112); +lean_dec(x_108); +x_113 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; +} +} +else +{ +uint8_t x_115; +x_115 = !lean_is_exclusive(x_108); +if (x_115 == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; +x_116 = lean_ctor_get(x_108, 0); +x_117 = lean_ctor_get(x_108, 1); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +x_119 = lean_nat_dec_eq(x_105, x_118); +lean_dec(x_105); +if (x_119 == 0) +{ +lean_dec(x_118); +return x_108; +} +else +{ +lean_object* x_120; lean_object* x_121; +lean_free_object(x_108); +lean_dec(x_117); +x_120 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_121 = l_Std_Internal_Parsec_String_pstring(x_120, x_116); +if (lean_obj_tag(x_121) == 0) +{ +uint8_t x_122; +lean_dec(x_118); +x_122 = !lean_is_exclusive(x_121); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_121, 1); +lean_dec(x_123); +x_124 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +lean_ctor_set(x_121, 1, x_124); +return x_121; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_121, 0); +lean_inc(x_125); +lean_dec(x_121); +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +else +{ +uint8_t x_128; +x_128 = !lean_is_exclusive(x_121); +if (x_128 == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_129 = lean_ctor_get(x_121, 0); +x_130 = lean_ctor_get(x_121, 1); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +x_132 = lean_nat_dec_eq(x_118, x_131); +lean_dec(x_118); +if (x_132 == 0) +{ +lean_dec(x_131); +return x_121; +} +else +{ +lean_object* x_133; lean_object* x_134; +lean_free_object(x_121); +lean_dec(x_130); +x_133 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_134 = l_Std_Internal_Parsec_String_pstring(x_133, x_129); +if (lean_obj_tag(x_134) == 0) +{ +uint8_t x_135; +lean_dec(x_131); +x_135 = !lean_is_exclusive(x_134); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_ctor_get(x_134, 1); +lean_dec(x_136); +x_137 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +lean_ctor_set(x_134, 1, x_137); +return x_134; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_134, 0); +lean_inc(x_138); +lean_dec(x_134); +x_139 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_138); +lean_ctor_set(x_140, 1, x_139); +return x_140; +} +} +else +{ +uint8_t x_141; +x_141 = !lean_is_exclusive(x_134); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_142 = lean_ctor_get(x_134, 0); +x_143 = lean_ctor_get(x_134, 1); +x_144 = lean_ctor_get(x_142, 1); +lean_inc(x_144); +x_145 = lean_nat_dec_eq(x_131, x_144); +lean_dec(x_144); +lean_dec(x_131); +if (x_145 == 0) +{ +return x_134; +} +else +{ +lean_object* x_146; lean_object* x_147; +lean_free_object(x_134); +lean_dec(x_143); +x_146 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_147 = l_Std_Internal_Parsec_String_pstring(x_146, x_142); +if (lean_obj_tag(x_147) == 0) +{ +uint8_t x_148; +x_148 = !lean_is_exclusive(x_147); +if (x_148 == 0) +{ +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_147, 1); +lean_dec(x_149); +x_150 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +lean_ctor_set(x_147, 1, x_150); +return x_147; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_147, 0); +lean_inc(x_151); +lean_dec(x_147); +x_152 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +else +{ +uint8_t x_154; +x_154 = !lean_is_exclusive(x_147); +if (x_154 == 0) +{ +return x_147; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_155 = lean_ctor_get(x_147, 0); +x_156 = lean_ctor_get(x_147, 1); +lean_inc(x_156); +lean_inc(x_155); +lean_dec(x_147); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_155); +lean_ctor_set(x_157, 1, x_156); +return x_157; +} +} +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; +x_158 = lean_ctor_get(x_134, 0); +x_159 = lean_ctor_get(x_134, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_134); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +x_161 = lean_nat_dec_eq(x_131, x_160); +lean_dec(x_160); +lean_dec(x_131); +if (x_161 == 0) +{ +lean_object* x_162; +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_158); +lean_ctor_set(x_162, 1, x_159); +return x_162; +} +else +{ +lean_object* x_163; lean_object* x_164; +lean_dec(x_159); +x_163 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_164 = l_Std_Internal_Parsec_String_pstring(x_163, x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_166 = x_164; +} else { + lean_dec_ref(x_164); + x_166 = lean_box(0); +} +x_167 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_166)) { + x_168 = lean_alloc_ctor(0, 2, 0); +} else { + x_168 = x_166; +} +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_167); +return x_168; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_169 = lean_ctor_get(x_164, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_164, 1); +lean_inc(x_170); +if (lean_is_exclusive(x_164)) { + lean_ctor_release(x_164, 0); + lean_ctor_release(x_164, 1); + x_171 = x_164; +} else { + lean_dec_ref(x_164); + x_171 = lean_box(0); +} +if (lean_is_scalar(x_171)) { + x_172 = lean_alloc_ctor(1, 2, 0); +} else { + x_172 = x_171; +} +lean_ctor_set(x_172, 0, x_169); +lean_ctor_set(x_172, 1, x_170); +return x_172; +} +} +} +} +} +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; +x_173 = lean_ctor_get(x_121, 0); +x_174 = lean_ctor_get(x_121, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_121); +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +x_176 = lean_nat_dec_eq(x_118, x_175); +lean_dec(x_118); +if (x_176 == 0) +{ +lean_object* x_177; +lean_dec(x_175); +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_174); +return x_177; +} +else +{ +lean_object* x_178; lean_object* x_179; +lean_dec(x_174); +x_178 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_179 = l_Std_Internal_Parsec_String_pstring(x_178, x_173); +if (lean_obj_tag(x_179) == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +lean_dec(x_175); +x_180 = lean_ctor_get(x_179, 0); +lean_inc(x_180); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_181 = x_179; +} else { + lean_dec_ref(x_179); + x_181 = lean_box(0); +} +x_182 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 2, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_182); +return x_183; +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; uint8_t x_188; +x_184 = lean_ctor_get(x_179, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_179, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_186 = x_179; +} else { + lean_dec_ref(x_179); + x_186 = lean_box(0); +} +x_187 = lean_ctor_get(x_184, 1); +lean_inc(x_187); +x_188 = lean_nat_dec_eq(x_175, x_187); +lean_dec(x_187); +lean_dec(x_175); +if (x_188 == 0) +{ +lean_object* x_189; +if (lean_is_scalar(x_186)) { + x_189 = lean_alloc_ctor(1, 2, 0); +} else { + x_189 = x_186; +} +lean_ctor_set(x_189, 0, x_184); +lean_ctor_set(x_189, 1, x_185); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; +lean_dec(x_186); +lean_dec(x_185); +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_191 = l_Std_Internal_Parsec_String_pstring(x_190, x_184); +if (lean_obj_tag(x_191) == 0) +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_192 = lean_ctor_get(x_191, 0); +lean_inc(x_192); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_193 = x_191; +} else { + lean_dec_ref(x_191); + x_193 = lean_box(0); +} +x_194 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_193)) { + x_195 = lean_alloc_ctor(0, 2, 0); +} else { + x_195 = x_193; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_194); +return x_195; +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_196 = lean_ctor_get(x_191, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_191, 1); +lean_inc(x_197); +if (lean_is_exclusive(x_191)) { + lean_ctor_release(x_191, 0); + lean_ctor_release(x_191, 1); + x_198 = x_191; +} else { + lean_dec_ref(x_191); + x_198 = lean_box(0); +} +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(1, 2, 0); +} else { + x_199 = x_198; +} +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +return x_199; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; +x_200 = lean_ctor_get(x_108, 0); +x_201 = lean_ctor_get(x_108, 1); +lean_inc(x_201); +lean_inc(x_200); +lean_dec(x_108); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +x_203 = lean_nat_dec_eq(x_105, x_202); +lean_dec(x_105); +if (x_203 == 0) +{ +lean_object* x_204; +lean_dec(x_202); +x_204 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_204, 0, x_200); +lean_ctor_set(x_204, 1, x_201); +return x_204; +} +else +{ +lean_object* x_205; lean_object* x_206; +lean_dec(x_201); +x_205 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_206 = l_Std_Internal_Parsec_String_pstring(x_205, x_200); +if (lean_obj_tag(x_206) == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_202); +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_208 = x_206; +} else { + lean_dec_ref(x_206); + x_208 = lean_box(0); +} +x_209 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_208)) { + x_210 = lean_alloc_ctor(0, 2, 0); +} else { + x_210 = x_208; +} +lean_ctor_set(x_210, 0, x_207); +lean_ctor_set(x_210, 1, x_209); +return x_210; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; +x_211 = lean_ctor_get(x_206, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_206, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_213 = x_206; +} else { + lean_dec_ref(x_206); + x_213 = lean_box(0); +} +x_214 = lean_ctor_get(x_211, 1); +lean_inc(x_214); +x_215 = lean_nat_dec_eq(x_202, x_214); +lean_dec(x_202); +if (x_215 == 0) +{ +lean_object* x_216; +lean_dec(x_214); +if (lean_is_scalar(x_213)) { + x_216 = lean_alloc_ctor(1, 2, 0); +} else { + x_216 = x_213; +} +lean_ctor_set(x_216, 0, x_211); +lean_ctor_set(x_216, 1, x_212); +return x_216; +} +else +{ +lean_object* x_217; lean_object* x_218; +lean_dec(x_213); +lean_dec(x_212); +x_217 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_218 = l_Std_Internal_Parsec_String_pstring(x_217, x_211); +if (lean_obj_tag(x_218) == 0) +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_dec(x_214); +x_219 = lean_ctor_get(x_218, 0); +lean_inc(x_219); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_220 = x_218; +} else { + lean_dec_ref(x_218); + x_220 = lean_box(0); +} +x_221 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_220)) { + x_222 = lean_alloc_ctor(0, 2, 0); +} else { + x_222 = x_220; +} +lean_ctor_set(x_222, 0, x_219); +lean_ctor_set(x_222, 1, x_221); +return x_222; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; uint8_t x_227; +x_223 = lean_ctor_get(x_218, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_218, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_225 = x_218; +} else { + lean_dec_ref(x_218); + x_225 = lean_box(0); +} +x_226 = lean_ctor_get(x_223, 1); +lean_inc(x_226); +x_227 = lean_nat_dec_eq(x_214, x_226); +lean_dec(x_226); +lean_dec(x_214); +if (x_227 == 0) +{ +lean_object* x_228; +if (lean_is_scalar(x_225)) { + x_228 = lean_alloc_ctor(1, 2, 0); +} else { + x_228 = x_225; +} +lean_ctor_set(x_228, 0, x_223); +lean_ctor_set(x_228, 1, x_224); +return x_228; +} +else +{ +lean_object* x_229; lean_object* x_230; +lean_dec(x_225); +lean_dec(x_224); +x_229 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_230 = l_Std_Internal_Parsec_String_pstring(x_229, x_223); +if (lean_obj_tag(x_230) == 0) +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_231 = lean_ctor_get(x_230, 0); +lean_inc(x_231); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_232 = x_230; +} else { + lean_dec_ref(x_230); + x_232 = lean_box(0); +} +x_233 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_232)) { + x_234 = lean_alloc_ctor(0, 2, 0); +} else { + x_234 = x_232; +} +lean_ctor_set(x_234, 0, x_231); +lean_ctor_set(x_234, 1, x_233); +return x_234; +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_235 = lean_ctor_get(x_230, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_230, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_230)) { + lean_ctor_release(x_230, 0); + lean_ctor_release(x_230, 1); + x_237 = x_230; +} else { + lean_dec_ref(x_230); + x_237 = lean_box(0); +} +if (lean_is_scalar(x_237)) { + x_238 = lean_alloc_ctor(1, 2, 0); +} else { + x_238 = x_237; +} +lean_ctor_set(x_238, 0, x_235); +lean_ctor_set(x_238, 1, x_236); +return x_238; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; uint8_t x_242; +x_239 = lean_ctor_get(x_95, 0); +x_240 = lean_ctor_get(x_95, 1); +lean_inc(x_240); +lean_inc(x_239); +lean_dec(x_95); +x_241 = lean_ctor_get(x_239, 1); +lean_inc(x_241); +x_242 = lean_nat_dec_eq(x_92, x_241); +lean_dec(x_92); +if (x_242 == 0) +{ +lean_object* x_243; +lean_dec(x_241); +x_243 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_243, 0, x_239); +lean_ctor_set(x_243, 1, x_240); +return x_243; +} +else +{ +lean_object* x_244; lean_object* x_245; +lean_dec(x_240); +x_244 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_245 = l_Std_Internal_Parsec_String_pstring(x_244, x_239); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +lean_dec(x_241); +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_247 = x_245; +} else { + lean_dec_ref(x_245); + x_247 = lean_box(0); +} +x_248 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_247)) { + x_249 = lean_alloc_ctor(0, 2, 0); +} else { + x_249 = x_247; +} +lean_ctor_set(x_249, 0, x_246); +lean_ctor_set(x_249, 1, x_248); +return x_249; +} +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; uint8_t x_254; +x_250 = lean_ctor_get(x_245, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_245, 1); +lean_inc(x_251); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_252 = x_245; +} else { + lean_dec_ref(x_245); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 1); +lean_inc(x_253); +x_254 = lean_nat_dec_eq(x_241, x_253); +lean_dec(x_241); +if (x_254 == 0) +{ +lean_object* x_255; +lean_dec(x_253); +if (lean_is_scalar(x_252)) { + x_255 = lean_alloc_ctor(1, 2, 0); +} else { + x_255 = x_252; +} +lean_ctor_set(x_255, 0, x_250); +lean_ctor_set(x_255, 1, x_251); +return x_255; +} +else +{ +lean_object* x_256; lean_object* x_257; +lean_dec(x_252); +lean_dec(x_251); +x_256 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_257 = l_Std_Internal_Parsec_String_pstring(x_256, x_250); +if (lean_obj_tag(x_257) == 0) +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +lean_dec(x_253); +x_258 = lean_ctor_get(x_257, 0); +lean_inc(x_258); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_259 = x_257; +} else { + lean_dec_ref(x_257); + x_259 = lean_box(0); +} +x_260 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_259)) { + x_261 = lean_alloc_ctor(0, 2, 0); +} else { + x_261 = x_259; +} +lean_ctor_set(x_261, 0, x_258); +lean_ctor_set(x_261, 1, x_260); +return x_261; +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; uint8_t x_266; +x_262 = lean_ctor_get(x_257, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_257, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_257)) { + lean_ctor_release(x_257, 0); + lean_ctor_release(x_257, 1); + x_264 = x_257; +} else { + lean_dec_ref(x_257); + x_264 = lean_box(0); +} +x_265 = lean_ctor_get(x_262, 1); +lean_inc(x_265); +x_266 = lean_nat_dec_eq(x_253, x_265); +lean_dec(x_253); +if (x_266 == 0) +{ +lean_object* x_267; +lean_dec(x_265); +if (lean_is_scalar(x_264)) { + x_267 = lean_alloc_ctor(1, 2, 0); +} else { + x_267 = x_264; +} +lean_ctor_set(x_267, 0, x_262); +lean_ctor_set(x_267, 1, x_263); +return x_267; +} +else +{ +lean_object* x_268; lean_object* x_269; +lean_dec(x_264); +lean_dec(x_263); +x_268 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_269 = l_Std_Internal_Parsec_String_pstring(x_268, x_262); +if (lean_obj_tag(x_269) == 0) +{ +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_265); +x_270 = lean_ctor_get(x_269, 0); +lean_inc(x_270); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_271 = x_269; +} else { + lean_dec_ref(x_269); + x_271 = lean_box(0); +} +x_272 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_271)) { + x_273 = lean_alloc_ctor(0, 2, 0); +} else { + x_273 = x_271; +} +lean_ctor_set(x_273, 0, x_270); +lean_ctor_set(x_273, 1, x_272); +return x_273; +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; uint8_t x_278; +x_274 = lean_ctor_get(x_269, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_269, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_269)) { + lean_ctor_release(x_269, 0); + lean_ctor_release(x_269, 1); + x_276 = x_269; +} else { + lean_dec_ref(x_269); + x_276 = lean_box(0); +} +x_277 = lean_ctor_get(x_274, 1); +lean_inc(x_277); +x_278 = lean_nat_dec_eq(x_265, x_277); +lean_dec(x_277); +lean_dec(x_265); +if (x_278 == 0) +{ +lean_object* x_279; +if (lean_is_scalar(x_276)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_276; +} +lean_ctor_set(x_279, 0, x_274); +lean_ctor_set(x_279, 1, x_275); +return x_279; +} +else +{ +lean_object* x_280; lean_object* x_281; +lean_dec(x_276); +lean_dec(x_275); +x_280 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_281 = l_Std_Internal_Parsec_String_pstring(x_280, x_274); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_283 = x_281; +} else { + lean_dec_ref(x_281); + x_283 = lean_box(0); +} +x_284 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_283)) { + x_285 = lean_alloc_ctor(0, 2, 0); +} else { + x_285 = x_283; +} +lean_ctor_set(x_285, 0, x_282); +lean_ctor_set(x_285, 1, x_284); +return x_285; +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_286 = lean_ctor_get(x_281, 0); +lean_inc(x_286); +x_287 = lean_ctor_get(x_281, 1); +lean_inc(x_287); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_288 = x_281; +} else { + lean_dec_ref(x_281); + x_288 = lean_box(0); +} +if (lean_is_scalar(x_288)) { + x_289 = lean_alloc_ctor(1, 2, 0); +} else { + x_289 = x_288; +} +lean_ctor_set(x_289, 0, x_286); +lean_ctor_set(x_289, 1, x_287); +return x_289; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_290; lean_object* x_291; lean_object* x_292; uint8_t x_293; +x_290 = lean_ctor_get(x_82, 0); +x_291 = lean_ctor_get(x_82, 1); +lean_inc(x_291); +lean_inc(x_290); +lean_dec(x_82); +x_292 = lean_ctor_get(x_290, 1); +lean_inc(x_292); +x_293 = lean_nat_dec_eq(x_79, x_292); +lean_dec(x_79); +if (x_293 == 0) +{ +lean_object* x_294; +lean_dec(x_292); +x_294 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_294, 0, x_290); +lean_ctor_set(x_294, 1, x_291); +return x_294; +} +else +{ +lean_object* x_295; lean_object* x_296; +lean_dec(x_291); +x_295 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_296 = l_Std_Internal_Parsec_String_pstring(x_295, x_290); +if (lean_obj_tag(x_296) == 0) +{ +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_292); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +if (lean_is_exclusive(x_296)) { + lean_ctor_release(x_296, 0); + lean_ctor_release(x_296, 1); + x_298 = x_296; +} else { + lean_dec_ref(x_296); + x_298 = lean_box(0); +} +x_299 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_298)) { + x_300 = lean_alloc_ctor(0, 2, 0); +} else { + x_300 = x_298; +} +lean_ctor_set(x_300, 0, x_297); +lean_ctor_set(x_300, 1, x_299); +return x_300; +} +else +{ +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; +x_301 = lean_ctor_get(x_296, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_296, 1); +lean_inc(x_302); +if (lean_is_exclusive(x_296)) { + lean_ctor_release(x_296, 0); + lean_ctor_release(x_296, 1); + x_303 = x_296; +} else { + lean_dec_ref(x_296); + x_303 = lean_box(0); +} +x_304 = lean_ctor_get(x_301, 1); +lean_inc(x_304); +x_305 = lean_nat_dec_eq(x_292, x_304); +lean_dec(x_292); +if (x_305 == 0) +{ +lean_object* x_306; +lean_dec(x_304); +if (lean_is_scalar(x_303)) { + x_306 = lean_alloc_ctor(1, 2, 0); +} else { + x_306 = x_303; +} +lean_ctor_set(x_306, 0, x_301); +lean_ctor_set(x_306, 1, x_302); +return x_306; +} +else +{ +lean_object* x_307; lean_object* x_308; +lean_dec(x_303); +lean_dec(x_302); +x_307 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_308 = l_Std_Internal_Parsec_String_pstring(x_307, x_301); +if (lean_obj_tag(x_308) == 0) +{ +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_304); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_310 = x_308; +} else { + lean_dec_ref(x_308); + x_310 = lean_box(0); +} +x_311 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_310)) { + x_312 = lean_alloc_ctor(0, 2, 0); +} else { + x_312 = x_310; +} +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_311); +return x_312; +} +else +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; uint8_t x_317; +x_313 = lean_ctor_get(x_308, 0); +lean_inc(x_313); +x_314 = lean_ctor_get(x_308, 1); +lean_inc(x_314); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_315 = x_308; +} else { + lean_dec_ref(x_308); + x_315 = lean_box(0); +} +x_316 = lean_ctor_get(x_313, 1); +lean_inc(x_316); +x_317 = lean_nat_dec_eq(x_304, x_316); +lean_dec(x_304); +if (x_317 == 0) +{ +lean_object* x_318; +lean_dec(x_316); +if (lean_is_scalar(x_315)) { + x_318 = lean_alloc_ctor(1, 2, 0); +} else { + x_318 = x_315; +} +lean_ctor_set(x_318, 0, x_313); +lean_ctor_set(x_318, 1, x_314); +return x_318; +} +else +{ +lean_object* x_319; lean_object* x_320; +lean_dec(x_315); +lean_dec(x_314); +x_319 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_320 = l_Std_Internal_Parsec_String_pstring(x_319, x_313); +if (lean_obj_tag(x_320) == 0) +{ +lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_316); +x_321 = lean_ctor_get(x_320, 0); +lean_inc(x_321); +if (lean_is_exclusive(x_320)) { + lean_ctor_release(x_320, 0); + lean_ctor_release(x_320, 1); + x_322 = x_320; +} else { + lean_dec_ref(x_320); + x_322 = lean_box(0); +} +x_323 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_322)) { + x_324 = lean_alloc_ctor(0, 2, 0); +} else { + x_324 = x_322; +} +lean_ctor_set(x_324, 0, x_321); +lean_ctor_set(x_324, 1, x_323); +return x_324; +} +else +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; +x_325 = lean_ctor_get(x_320, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_320, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_320)) { + lean_ctor_release(x_320, 0); + lean_ctor_release(x_320, 1); + x_327 = x_320; +} else { + lean_dec_ref(x_320); + x_327 = lean_box(0); +} +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +x_329 = lean_nat_dec_eq(x_316, x_328); +lean_dec(x_316); +if (x_329 == 0) +{ +lean_object* x_330; +lean_dec(x_328); +if (lean_is_scalar(x_327)) { + x_330 = lean_alloc_ctor(1, 2, 0); +} else { + x_330 = x_327; +} +lean_ctor_set(x_330, 0, x_325); +lean_ctor_set(x_330, 1, x_326); +return x_330; +} +else +{ +lean_object* x_331; lean_object* x_332; +lean_dec(x_327); +lean_dec(x_326); +x_331 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_332 = l_Std_Internal_Parsec_String_pstring(x_331, x_325); +if (lean_obj_tag(x_332) == 0) +{ +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; +lean_dec(x_328); +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_334 = x_332; +} else { + lean_dec_ref(x_332); + x_334 = lean_box(0); +} +x_335 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_334)) { + x_336 = lean_alloc_ctor(0, 2, 0); +} else { + x_336 = x_334; +} +lean_ctor_set(x_336, 0, x_333); +lean_ctor_set(x_336, 1, x_335); +return x_336; +} +else +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; uint8_t x_341; +x_337 = lean_ctor_get(x_332, 0); +lean_inc(x_337); +x_338 = lean_ctor_get(x_332, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_339 = x_332; +} else { + lean_dec_ref(x_332); + x_339 = lean_box(0); +} +x_340 = lean_ctor_get(x_337, 1); +lean_inc(x_340); +x_341 = lean_nat_dec_eq(x_328, x_340); +lean_dec(x_340); +lean_dec(x_328); +if (x_341 == 0) +{ +lean_object* x_342; +if (lean_is_scalar(x_339)) { + x_342 = lean_alloc_ctor(1, 2, 0); +} else { + x_342 = x_339; +} +lean_ctor_set(x_342, 0, x_337); +lean_ctor_set(x_342, 1, x_338); +return x_342; +} +else +{ +lean_object* x_343; lean_object* x_344; +lean_dec(x_339); +lean_dec(x_338); +x_343 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_344 = l_Std_Internal_Parsec_String_pstring(x_343, x_337); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +if (lean_is_exclusive(x_344)) { + lean_ctor_release(x_344, 0); + lean_ctor_release(x_344, 1); + x_346 = x_344; +} else { + lean_dec_ref(x_344); + x_346 = lean_box(0); +} +x_347 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_346)) { + x_348 = lean_alloc_ctor(0, 2, 0); +} else { + x_348 = x_346; +} +lean_ctor_set(x_348, 0, x_345); +lean_ctor_set(x_348, 1, x_347); +return x_348; +} +else +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +x_349 = lean_ctor_get(x_344, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_344, 1); +lean_inc(x_350); +if (lean_is_exclusive(x_344)) { + lean_ctor_release(x_344, 0); + lean_ctor_release(x_344, 1); + x_351 = x_344; +} else { + lean_dec_ref(x_344); + x_351 = lean_box(0); +} +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(1, 2, 0); +} else { + x_352 = x_351; +} +lean_ctor_set(x_352, 0, x_349); +lean_ctor_set(x_352, 1, x_350); +return x_352; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_353; lean_object* x_354; lean_object* x_355; uint8_t x_356; +x_353 = lean_ctor_get(x_69, 0); +x_354 = lean_ctor_get(x_69, 1); +lean_inc(x_354); +lean_inc(x_353); +lean_dec(x_69); +x_355 = lean_ctor_get(x_353, 1); +lean_inc(x_355); +x_356 = lean_nat_dec_eq(x_66, x_355); +lean_dec(x_66); +if (x_356 == 0) +{ +lean_object* x_357; +lean_dec(x_355); +x_357 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_357, 0, x_353); +lean_ctor_set(x_357, 1, x_354); +return x_357; +} +else +{ +lean_object* x_358; lean_object* x_359; +lean_dec(x_354); +x_358 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_359 = l_Std_Internal_Parsec_String_pstring(x_358, x_353); +if (lean_obj_tag(x_359) == 0) +{ +lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; +lean_dec(x_355); +x_360 = lean_ctor_get(x_359, 0); +lean_inc(x_360); +if (lean_is_exclusive(x_359)) { + lean_ctor_release(x_359, 0); + lean_ctor_release(x_359, 1); + x_361 = x_359; +} else { + lean_dec_ref(x_359); + x_361 = lean_box(0); +} +x_362 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_361)) { + x_363 = lean_alloc_ctor(0, 2, 0); +} else { + x_363 = x_361; +} +lean_ctor_set(x_363, 0, x_360); +lean_ctor_set(x_363, 1, x_362); +return x_363; +} +else +{ +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; uint8_t x_368; +x_364 = lean_ctor_get(x_359, 0); +lean_inc(x_364); +x_365 = lean_ctor_get(x_359, 1); +lean_inc(x_365); +if (lean_is_exclusive(x_359)) { + lean_ctor_release(x_359, 0); + lean_ctor_release(x_359, 1); + x_366 = x_359; +} else { + lean_dec_ref(x_359); + x_366 = lean_box(0); +} +x_367 = lean_ctor_get(x_364, 1); +lean_inc(x_367); +x_368 = lean_nat_dec_eq(x_355, x_367); +lean_dec(x_355); +if (x_368 == 0) +{ +lean_object* x_369; +lean_dec(x_367); +if (lean_is_scalar(x_366)) { + x_369 = lean_alloc_ctor(1, 2, 0); +} else { + x_369 = x_366; +} +lean_ctor_set(x_369, 0, x_364); +lean_ctor_set(x_369, 1, x_365); +return x_369; +} +else +{ +lean_object* x_370; lean_object* x_371; +lean_dec(x_366); +lean_dec(x_365); +x_370 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_371 = l_Std_Internal_Parsec_String_pstring(x_370, x_364); +if (lean_obj_tag(x_371) == 0) +{ +lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; +lean_dec(x_367); +x_372 = lean_ctor_get(x_371, 0); +lean_inc(x_372); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_373 = x_371; +} else { + lean_dec_ref(x_371); + x_373 = lean_box(0); +} +x_374 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_373)) { + x_375 = lean_alloc_ctor(0, 2, 0); +} else { + x_375 = x_373; +} +lean_ctor_set(x_375, 0, x_372); +lean_ctor_set(x_375, 1, x_374); +return x_375; +} +else +{ +lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; uint8_t x_380; +x_376 = lean_ctor_get(x_371, 0); +lean_inc(x_376); +x_377 = lean_ctor_get(x_371, 1); +lean_inc(x_377); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_378 = x_371; +} else { + lean_dec_ref(x_371); + x_378 = lean_box(0); +} +x_379 = lean_ctor_get(x_376, 1); +lean_inc(x_379); +x_380 = lean_nat_dec_eq(x_367, x_379); +lean_dec(x_367); +if (x_380 == 0) +{ +lean_object* x_381; +lean_dec(x_379); +if (lean_is_scalar(x_378)) { + x_381 = lean_alloc_ctor(1, 2, 0); +} else { + x_381 = x_378; +} +lean_ctor_set(x_381, 0, x_376); +lean_ctor_set(x_381, 1, x_377); +return x_381; +} +else +{ +lean_object* x_382; lean_object* x_383; +lean_dec(x_378); +lean_dec(x_377); +x_382 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_383 = l_Std_Internal_Parsec_String_pstring(x_382, x_376); +if (lean_obj_tag(x_383) == 0) +{ +lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_379); +x_384 = lean_ctor_get(x_383, 0); +lean_inc(x_384); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_385 = x_383; +} else { + lean_dec_ref(x_383); + x_385 = lean_box(0); +} +x_386 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_385)) { + x_387 = lean_alloc_ctor(0, 2, 0); +} else { + x_387 = x_385; +} +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_386); +return x_387; +} +else +{ +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; uint8_t x_392; +x_388 = lean_ctor_get(x_383, 0); +lean_inc(x_388); +x_389 = lean_ctor_get(x_383, 1); +lean_inc(x_389); +if (lean_is_exclusive(x_383)) { + lean_ctor_release(x_383, 0); + lean_ctor_release(x_383, 1); + x_390 = x_383; +} else { + lean_dec_ref(x_383); + x_390 = lean_box(0); +} +x_391 = lean_ctor_get(x_388, 1); +lean_inc(x_391); +x_392 = lean_nat_dec_eq(x_379, x_391); +lean_dec(x_379); +if (x_392 == 0) +{ +lean_object* x_393; +lean_dec(x_391); +if (lean_is_scalar(x_390)) { + x_393 = lean_alloc_ctor(1, 2, 0); +} else { + x_393 = x_390; +} +lean_ctor_set(x_393, 0, x_388); +lean_ctor_set(x_393, 1, x_389); +return x_393; +} +else +{ +lean_object* x_394; lean_object* x_395; +lean_dec(x_390); +lean_dec(x_389); +x_394 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_395 = l_Std_Internal_Parsec_String_pstring(x_394, x_388); +if (lean_obj_tag(x_395) == 0) +{ +lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; +lean_dec(x_391); +x_396 = lean_ctor_get(x_395, 0); +lean_inc(x_396); +if (lean_is_exclusive(x_395)) { + lean_ctor_release(x_395, 0); + lean_ctor_release(x_395, 1); + x_397 = x_395; +} else { + lean_dec_ref(x_395); + x_397 = lean_box(0); +} +x_398 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_397)) { + x_399 = lean_alloc_ctor(0, 2, 0); +} else { + x_399 = x_397; +} +lean_ctor_set(x_399, 0, x_396); +lean_ctor_set(x_399, 1, x_398); +return x_399; +} +else +{ +lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; uint8_t x_404; +x_400 = lean_ctor_get(x_395, 0); +lean_inc(x_400); +x_401 = lean_ctor_get(x_395, 1); +lean_inc(x_401); +if (lean_is_exclusive(x_395)) { + lean_ctor_release(x_395, 0); + lean_ctor_release(x_395, 1); + x_402 = x_395; +} else { + lean_dec_ref(x_395); + x_402 = lean_box(0); +} +x_403 = lean_ctor_get(x_400, 1); +lean_inc(x_403); +x_404 = lean_nat_dec_eq(x_391, x_403); +lean_dec(x_391); +if (x_404 == 0) +{ +lean_object* x_405; +lean_dec(x_403); +if (lean_is_scalar(x_402)) { + x_405 = lean_alloc_ctor(1, 2, 0); +} else { + x_405 = x_402; +} +lean_ctor_set(x_405, 0, x_400); +lean_ctor_set(x_405, 1, x_401); +return x_405; +} +else +{ +lean_object* x_406; lean_object* x_407; +lean_dec(x_402); +lean_dec(x_401); +x_406 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_407 = l_Std_Internal_Parsec_String_pstring(x_406, x_400); +if (lean_obj_tag(x_407) == 0) +{ +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; +lean_dec(x_403); +x_408 = lean_ctor_get(x_407, 0); +lean_inc(x_408); +if (lean_is_exclusive(x_407)) { + lean_ctor_release(x_407, 0); + lean_ctor_release(x_407, 1); + x_409 = x_407; +} else { + lean_dec_ref(x_407); + x_409 = lean_box(0); +} +x_410 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_409)) { + x_411 = lean_alloc_ctor(0, 2, 0); +} else { + x_411 = x_409; +} +lean_ctor_set(x_411, 0, x_408); +lean_ctor_set(x_411, 1, x_410); +return x_411; +} +else +{ +lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; uint8_t x_416; +x_412 = lean_ctor_get(x_407, 0); +lean_inc(x_412); +x_413 = lean_ctor_get(x_407, 1); +lean_inc(x_413); +if (lean_is_exclusive(x_407)) { + lean_ctor_release(x_407, 0); + lean_ctor_release(x_407, 1); + x_414 = x_407; +} else { + lean_dec_ref(x_407); + x_414 = lean_box(0); +} +x_415 = lean_ctor_get(x_412, 1); +lean_inc(x_415); +x_416 = lean_nat_dec_eq(x_403, x_415); +lean_dec(x_415); +lean_dec(x_403); +if (x_416 == 0) +{ +lean_object* x_417; +if (lean_is_scalar(x_414)) { + x_417 = lean_alloc_ctor(1, 2, 0); +} else { + x_417 = x_414; +} +lean_ctor_set(x_417, 0, x_412); +lean_ctor_set(x_417, 1, x_413); +return x_417; +} +else +{ +lean_object* x_418; lean_object* x_419; +lean_dec(x_414); +lean_dec(x_413); +x_418 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_419 = l_Std_Internal_Parsec_String_pstring(x_418, x_412); +if (lean_obj_tag(x_419) == 0) +{ +lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; +x_420 = lean_ctor_get(x_419, 0); +lean_inc(x_420); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + lean_ctor_release(x_419, 1); + x_421 = x_419; +} else { + lean_dec_ref(x_419); + x_421 = lean_box(0); +} +x_422 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_421)) { + x_423 = lean_alloc_ctor(0, 2, 0); +} else { + x_423 = x_421; +} +lean_ctor_set(x_423, 0, x_420); +lean_ctor_set(x_423, 1, x_422); +return x_423; +} +else +{ +lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; +x_424 = lean_ctor_get(x_419, 0); +lean_inc(x_424); +x_425 = lean_ctor_get(x_419, 1); +lean_inc(x_425); +if (lean_is_exclusive(x_419)) { + lean_ctor_release(x_419, 0); + lean_ctor_release(x_419, 1); + x_426 = x_419; +} else { + lean_dec_ref(x_419); + x_426 = lean_box(0); +} +if (lean_is_scalar(x_426)) { + x_427 = lean_alloc_ctor(1, 2, 0); +} else { + x_427 = x_426; +} +lean_ctor_set(x_427, 0, x_424); +lean_ctor_set(x_427, 1, x_425); +return x_427; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_428; lean_object* x_429; lean_object* x_430; uint8_t x_431; +x_428 = lean_ctor_get(x_56, 0); +x_429 = lean_ctor_get(x_56, 1); +lean_inc(x_429); +lean_inc(x_428); +lean_dec(x_56); +x_430 = lean_ctor_get(x_428, 1); +lean_inc(x_430); +x_431 = lean_nat_dec_eq(x_53, x_430); +lean_dec(x_53); +if (x_431 == 0) +{ +lean_object* x_432; +lean_dec(x_430); +x_432 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_432, 0, x_428); +lean_ctor_set(x_432, 1, x_429); +return x_432; +} +else +{ +lean_object* x_433; lean_object* x_434; +lean_dec(x_429); +x_433 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_434 = l_Std_Internal_Parsec_String_pstring(x_433, x_428); +if (lean_obj_tag(x_434) == 0) +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; +lean_dec(x_430); +x_435 = lean_ctor_get(x_434, 0); +lean_inc(x_435); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_436 = x_434; +} else { + lean_dec_ref(x_434); + x_436 = lean_box(0); +} +x_437 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_436)) { + x_438 = lean_alloc_ctor(0, 2, 0); +} else { + x_438 = x_436; +} +lean_ctor_set(x_438, 0, x_435); +lean_ctor_set(x_438, 1, x_437); +return x_438; +} +else +{ +lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; +x_439 = lean_ctor_get(x_434, 0); +lean_inc(x_439); +x_440 = lean_ctor_get(x_434, 1); +lean_inc(x_440); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_441 = x_434; +} else { + lean_dec_ref(x_434); + x_441 = lean_box(0); +} +x_442 = lean_ctor_get(x_439, 1); +lean_inc(x_442); +x_443 = lean_nat_dec_eq(x_430, x_442); +lean_dec(x_430); +if (x_443 == 0) +{ +lean_object* x_444; +lean_dec(x_442); +if (lean_is_scalar(x_441)) { + x_444 = lean_alloc_ctor(1, 2, 0); +} else { + x_444 = x_441; +} +lean_ctor_set(x_444, 0, x_439); +lean_ctor_set(x_444, 1, x_440); +return x_444; +} +else +{ +lean_object* x_445; lean_object* x_446; +lean_dec(x_441); +lean_dec(x_440); +x_445 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_446 = l_Std_Internal_Parsec_String_pstring(x_445, x_439); +if (lean_obj_tag(x_446) == 0) +{ +lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; +lean_dec(x_442); +x_447 = lean_ctor_get(x_446, 0); +lean_inc(x_447); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_448 = x_446; +} else { + lean_dec_ref(x_446); + x_448 = lean_box(0); +} +x_449 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(0, 2, 0); +} else { + x_450 = x_448; +} +lean_ctor_set(x_450, 0, x_447); +lean_ctor_set(x_450, 1, x_449); +return x_450; +} +else +{ +lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; uint8_t x_455; +x_451 = lean_ctor_get(x_446, 0); +lean_inc(x_451); +x_452 = lean_ctor_get(x_446, 1); +lean_inc(x_452); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_453 = x_446; +} else { + lean_dec_ref(x_446); + x_453 = lean_box(0); +} +x_454 = lean_ctor_get(x_451, 1); +lean_inc(x_454); +x_455 = lean_nat_dec_eq(x_442, x_454); +lean_dec(x_442); +if (x_455 == 0) +{ +lean_object* x_456; +lean_dec(x_454); +if (lean_is_scalar(x_453)) { + x_456 = lean_alloc_ctor(1, 2, 0); +} else { + x_456 = x_453; +} +lean_ctor_set(x_456, 0, x_451); +lean_ctor_set(x_456, 1, x_452); +return x_456; +} +else +{ +lean_object* x_457; lean_object* x_458; +lean_dec(x_453); +lean_dec(x_452); +x_457 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_458 = l_Std_Internal_Parsec_String_pstring(x_457, x_451); +if (lean_obj_tag(x_458) == 0) +{ +lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; +lean_dec(x_454); +x_459 = lean_ctor_get(x_458, 0); +lean_inc(x_459); +if (lean_is_exclusive(x_458)) { + lean_ctor_release(x_458, 0); + lean_ctor_release(x_458, 1); + x_460 = x_458; +} else { + lean_dec_ref(x_458); + x_460 = lean_box(0); +} +x_461 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_460)) { + x_462 = lean_alloc_ctor(0, 2, 0); +} else { + x_462 = x_460; +} +lean_ctor_set(x_462, 0, x_459); +lean_ctor_set(x_462, 1, x_461); +return x_462; +} +else +{ +lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; uint8_t x_467; +x_463 = lean_ctor_get(x_458, 0); +lean_inc(x_463); +x_464 = lean_ctor_get(x_458, 1); +lean_inc(x_464); +if (lean_is_exclusive(x_458)) { + lean_ctor_release(x_458, 0); + lean_ctor_release(x_458, 1); + x_465 = x_458; +} else { + lean_dec_ref(x_458); + x_465 = lean_box(0); +} +x_466 = lean_ctor_get(x_463, 1); +lean_inc(x_466); +x_467 = lean_nat_dec_eq(x_454, x_466); +lean_dec(x_454); +if (x_467 == 0) +{ +lean_object* x_468; +lean_dec(x_466); +if (lean_is_scalar(x_465)) { + x_468 = lean_alloc_ctor(1, 2, 0); +} else { + x_468 = x_465; +} +lean_ctor_set(x_468, 0, x_463); +lean_ctor_set(x_468, 1, x_464); +return x_468; +} +else +{ +lean_object* x_469; lean_object* x_470; +lean_dec(x_465); +lean_dec(x_464); +x_469 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_470 = l_Std_Internal_Parsec_String_pstring(x_469, x_463); +if (lean_obj_tag(x_470) == 0) +{ +lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; +lean_dec(x_466); +x_471 = lean_ctor_get(x_470, 0); +lean_inc(x_471); +if (lean_is_exclusive(x_470)) { + lean_ctor_release(x_470, 0); + lean_ctor_release(x_470, 1); + x_472 = x_470; +} else { + lean_dec_ref(x_470); + x_472 = lean_box(0); +} +x_473 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_472)) { + x_474 = lean_alloc_ctor(0, 2, 0); +} else { + x_474 = x_472; +} +lean_ctor_set(x_474, 0, x_471); +lean_ctor_set(x_474, 1, x_473); +return x_474; +} +else +{ +lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; uint8_t x_479; +x_475 = lean_ctor_get(x_470, 0); +lean_inc(x_475); +x_476 = lean_ctor_get(x_470, 1); +lean_inc(x_476); +if (lean_is_exclusive(x_470)) { + lean_ctor_release(x_470, 0); + lean_ctor_release(x_470, 1); + x_477 = x_470; +} else { + lean_dec_ref(x_470); + x_477 = lean_box(0); +} +x_478 = lean_ctor_get(x_475, 1); +lean_inc(x_478); +x_479 = lean_nat_dec_eq(x_466, x_478); +lean_dec(x_466); +if (x_479 == 0) +{ +lean_object* x_480; +lean_dec(x_478); +if (lean_is_scalar(x_477)) { + x_480 = lean_alloc_ctor(1, 2, 0); +} else { + x_480 = x_477; +} +lean_ctor_set(x_480, 0, x_475); +lean_ctor_set(x_480, 1, x_476); +return x_480; +} +else +{ +lean_object* x_481; lean_object* x_482; +lean_dec(x_477); +lean_dec(x_476); +x_481 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_482 = l_Std_Internal_Parsec_String_pstring(x_481, x_475); +if (lean_obj_tag(x_482) == 0) +{ +lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; +lean_dec(x_478); +x_483 = lean_ctor_get(x_482, 0); +lean_inc(x_483); +if (lean_is_exclusive(x_482)) { + lean_ctor_release(x_482, 0); + lean_ctor_release(x_482, 1); + x_484 = x_482; +} else { + lean_dec_ref(x_482); + x_484 = lean_box(0); +} +x_485 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_484)) { + x_486 = lean_alloc_ctor(0, 2, 0); +} else { + x_486 = x_484; +} +lean_ctor_set(x_486, 0, x_483); +lean_ctor_set(x_486, 1, x_485); +return x_486; +} +else +{ +lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; uint8_t x_491; +x_487 = lean_ctor_get(x_482, 0); +lean_inc(x_487); +x_488 = lean_ctor_get(x_482, 1); +lean_inc(x_488); +if (lean_is_exclusive(x_482)) { + lean_ctor_release(x_482, 0); + lean_ctor_release(x_482, 1); + x_489 = x_482; +} else { + lean_dec_ref(x_482); + x_489 = lean_box(0); +} +x_490 = lean_ctor_get(x_487, 1); +lean_inc(x_490); +x_491 = lean_nat_dec_eq(x_478, x_490); +lean_dec(x_478); +if (x_491 == 0) +{ +lean_object* x_492; +lean_dec(x_490); +if (lean_is_scalar(x_489)) { + x_492 = lean_alloc_ctor(1, 2, 0); +} else { + x_492 = x_489; +} +lean_ctor_set(x_492, 0, x_487); +lean_ctor_set(x_492, 1, x_488); +return x_492; +} +else +{ +lean_object* x_493; lean_object* x_494; +lean_dec(x_489); +lean_dec(x_488); +x_493 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_494 = l_Std_Internal_Parsec_String_pstring(x_493, x_487); +if (lean_obj_tag(x_494) == 0) +{ +lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; +lean_dec(x_490); +x_495 = lean_ctor_get(x_494, 0); +lean_inc(x_495); +if (lean_is_exclusive(x_494)) { + lean_ctor_release(x_494, 0); + lean_ctor_release(x_494, 1); + x_496 = x_494; +} else { + lean_dec_ref(x_494); + x_496 = lean_box(0); +} +x_497 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_496)) { + x_498 = lean_alloc_ctor(0, 2, 0); +} else { + x_498 = x_496; +} +lean_ctor_set(x_498, 0, x_495); +lean_ctor_set(x_498, 1, x_497); +return x_498; +} +else +{ +lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; uint8_t x_503; +x_499 = lean_ctor_get(x_494, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_494, 1); +lean_inc(x_500); +if (lean_is_exclusive(x_494)) { + lean_ctor_release(x_494, 0); + lean_ctor_release(x_494, 1); + x_501 = x_494; +} else { + lean_dec_ref(x_494); + x_501 = lean_box(0); +} +x_502 = lean_ctor_get(x_499, 1); +lean_inc(x_502); +x_503 = lean_nat_dec_eq(x_490, x_502); +lean_dec(x_502); +lean_dec(x_490); +if (x_503 == 0) +{ +lean_object* x_504; +if (lean_is_scalar(x_501)) { + x_504 = lean_alloc_ctor(1, 2, 0); +} else { + x_504 = x_501; +} +lean_ctor_set(x_504, 0, x_499); +lean_ctor_set(x_504, 1, x_500); +return x_504; +} +else +{ +lean_object* x_505; lean_object* x_506; +lean_dec(x_501); +lean_dec(x_500); +x_505 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_506 = l_Std_Internal_Parsec_String_pstring(x_505, x_499); +if (lean_obj_tag(x_506) == 0) +{ +lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +x_507 = lean_ctor_get(x_506, 0); +lean_inc(x_507); +if (lean_is_exclusive(x_506)) { + lean_ctor_release(x_506, 0); + lean_ctor_release(x_506, 1); + x_508 = x_506; +} else { + lean_dec_ref(x_506); + x_508 = lean_box(0); +} +x_509 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_508)) { + x_510 = lean_alloc_ctor(0, 2, 0); +} else { + x_510 = x_508; +} +lean_ctor_set(x_510, 0, x_507); +lean_ctor_set(x_510, 1, x_509); +return x_510; +} +else +{ +lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +x_511 = lean_ctor_get(x_506, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_506, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_506)) { + lean_ctor_release(x_506, 0); + lean_ctor_release(x_506, 1); + x_513 = x_506; +} else { + lean_dec_ref(x_506); + x_513 = lean_box(0); +} +if (lean_is_scalar(x_513)) { + x_514 = lean_alloc_ctor(1, 2, 0); +} else { + x_514 = x_513; +} +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_512); +return x_514; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_515; lean_object* x_516; lean_object* x_517; uint8_t x_518; +x_515 = lean_ctor_get(x_43, 0); +x_516 = lean_ctor_get(x_43, 1); +lean_inc(x_516); +lean_inc(x_515); +lean_dec(x_43); +x_517 = lean_ctor_get(x_515, 1); +lean_inc(x_517); +x_518 = lean_nat_dec_eq(x_40, x_517); +lean_dec(x_40); +if (x_518 == 0) +{ +lean_object* x_519; +lean_dec(x_517); +x_519 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_519, 0, x_515); +lean_ctor_set(x_519, 1, x_516); +return x_519; +} +else +{ +lean_object* x_520; lean_object* x_521; +lean_dec(x_516); +x_520 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_521 = l_Std_Internal_Parsec_String_pstring(x_520, x_515); +if (lean_obj_tag(x_521) == 0) +{ +lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +lean_dec(x_517); +x_522 = lean_ctor_get(x_521, 0); +lean_inc(x_522); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_523 = x_521; +} else { + lean_dec_ref(x_521); + x_523 = lean_box(0); +} +x_524 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_523)) { + x_525 = lean_alloc_ctor(0, 2, 0); +} else { + x_525 = x_523; +} +lean_ctor_set(x_525, 0, x_522); +lean_ctor_set(x_525, 1, x_524); +return x_525; +} +else +{ +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; uint8_t x_530; +x_526 = lean_ctor_get(x_521, 0); +lean_inc(x_526); +x_527 = lean_ctor_get(x_521, 1); +lean_inc(x_527); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_528 = x_521; +} else { + lean_dec_ref(x_521); + x_528 = lean_box(0); +} +x_529 = lean_ctor_get(x_526, 1); +lean_inc(x_529); +x_530 = lean_nat_dec_eq(x_517, x_529); +lean_dec(x_517); +if (x_530 == 0) +{ +lean_object* x_531; +lean_dec(x_529); +if (lean_is_scalar(x_528)) { + x_531 = lean_alloc_ctor(1, 2, 0); +} else { + x_531 = x_528; +} +lean_ctor_set(x_531, 0, x_526); +lean_ctor_set(x_531, 1, x_527); +return x_531; +} +else +{ +lean_object* x_532; lean_object* x_533; +lean_dec(x_528); +lean_dec(x_527); +x_532 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_533 = l_Std_Internal_Parsec_String_pstring(x_532, x_526); +if (lean_obj_tag(x_533) == 0) +{ +lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; +lean_dec(x_529); +x_534 = lean_ctor_get(x_533, 0); +lean_inc(x_534); +if (lean_is_exclusive(x_533)) { + lean_ctor_release(x_533, 0); + lean_ctor_release(x_533, 1); + x_535 = x_533; +} else { + lean_dec_ref(x_533); + x_535 = lean_box(0); +} +x_536 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_535)) { + x_537 = lean_alloc_ctor(0, 2, 0); +} else { + x_537 = x_535; +} +lean_ctor_set(x_537, 0, x_534); +lean_ctor_set(x_537, 1, x_536); +return x_537; +} +else +{ +lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; uint8_t x_542; +x_538 = lean_ctor_get(x_533, 0); +lean_inc(x_538); +x_539 = lean_ctor_get(x_533, 1); +lean_inc(x_539); +if (lean_is_exclusive(x_533)) { + lean_ctor_release(x_533, 0); + lean_ctor_release(x_533, 1); + x_540 = x_533; +} else { + lean_dec_ref(x_533); + x_540 = lean_box(0); +} +x_541 = lean_ctor_get(x_538, 1); +lean_inc(x_541); +x_542 = lean_nat_dec_eq(x_529, x_541); +lean_dec(x_529); +if (x_542 == 0) +{ +lean_object* x_543; +lean_dec(x_541); +if (lean_is_scalar(x_540)) { + x_543 = lean_alloc_ctor(1, 2, 0); +} else { + x_543 = x_540; +} +lean_ctor_set(x_543, 0, x_538); +lean_ctor_set(x_543, 1, x_539); +return x_543; +} +else +{ +lean_object* x_544; lean_object* x_545; +lean_dec(x_540); +lean_dec(x_539); +x_544 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_545 = l_Std_Internal_Parsec_String_pstring(x_544, x_538); +if (lean_obj_tag(x_545) == 0) +{ +lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; +lean_dec(x_541); +x_546 = lean_ctor_get(x_545, 0); +lean_inc(x_546); +if (lean_is_exclusive(x_545)) { + lean_ctor_release(x_545, 0); + lean_ctor_release(x_545, 1); + x_547 = x_545; +} else { + lean_dec_ref(x_545); + x_547 = lean_box(0); +} +x_548 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_547)) { + x_549 = lean_alloc_ctor(0, 2, 0); +} else { + x_549 = x_547; +} +lean_ctor_set(x_549, 0, x_546); +lean_ctor_set(x_549, 1, x_548); +return x_549; +} +else +{ +lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; uint8_t x_554; +x_550 = lean_ctor_get(x_545, 0); +lean_inc(x_550); +x_551 = lean_ctor_get(x_545, 1); +lean_inc(x_551); +if (lean_is_exclusive(x_545)) { + lean_ctor_release(x_545, 0); + lean_ctor_release(x_545, 1); + x_552 = x_545; +} else { + lean_dec_ref(x_545); + x_552 = lean_box(0); +} +x_553 = lean_ctor_get(x_550, 1); +lean_inc(x_553); +x_554 = lean_nat_dec_eq(x_541, x_553); +lean_dec(x_541); +if (x_554 == 0) +{ +lean_object* x_555; +lean_dec(x_553); +if (lean_is_scalar(x_552)) { + x_555 = lean_alloc_ctor(1, 2, 0); +} else { + x_555 = x_552; +} +lean_ctor_set(x_555, 0, x_550); +lean_ctor_set(x_555, 1, x_551); +return x_555; +} +else +{ +lean_object* x_556; lean_object* x_557; +lean_dec(x_552); +lean_dec(x_551); +x_556 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_557 = l_Std_Internal_Parsec_String_pstring(x_556, x_550); +if (lean_obj_tag(x_557) == 0) +{ +lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; +lean_dec(x_553); +x_558 = lean_ctor_get(x_557, 0); +lean_inc(x_558); +if (lean_is_exclusive(x_557)) { + lean_ctor_release(x_557, 0); + lean_ctor_release(x_557, 1); + x_559 = x_557; +} else { + lean_dec_ref(x_557); + x_559 = lean_box(0); +} +x_560 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_559)) { + x_561 = lean_alloc_ctor(0, 2, 0); +} else { + x_561 = x_559; +} +lean_ctor_set(x_561, 0, x_558); +lean_ctor_set(x_561, 1, x_560); +return x_561; +} +else +{ +lean_object* x_562; lean_object* x_563; lean_object* x_564; lean_object* x_565; uint8_t x_566; +x_562 = lean_ctor_get(x_557, 0); +lean_inc(x_562); +x_563 = lean_ctor_get(x_557, 1); +lean_inc(x_563); +if (lean_is_exclusive(x_557)) { + lean_ctor_release(x_557, 0); + lean_ctor_release(x_557, 1); + x_564 = x_557; +} else { + lean_dec_ref(x_557); + x_564 = lean_box(0); +} +x_565 = lean_ctor_get(x_562, 1); +lean_inc(x_565); +x_566 = lean_nat_dec_eq(x_553, x_565); +lean_dec(x_553); +if (x_566 == 0) +{ +lean_object* x_567; +lean_dec(x_565); +if (lean_is_scalar(x_564)) { + x_567 = lean_alloc_ctor(1, 2, 0); +} else { + x_567 = x_564; +} +lean_ctor_set(x_567, 0, x_562); +lean_ctor_set(x_567, 1, x_563); +return x_567; +} +else +{ +lean_object* x_568; lean_object* x_569; +lean_dec(x_564); +lean_dec(x_563); +x_568 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_569 = l_Std_Internal_Parsec_String_pstring(x_568, x_562); +if (lean_obj_tag(x_569) == 0) +{ +lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; +lean_dec(x_565); +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_571 = x_569; +} else { + lean_dec_ref(x_569); + x_571 = lean_box(0); +} +x_572 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_571)) { + x_573 = lean_alloc_ctor(0, 2, 0); +} else { + x_573 = x_571; +} +lean_ctor_set(x_573, 0, x_570); +lean_ctor_set(x_573, 1, x_572); +return x_573; +} +else +{ +lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; uint8_t x_578; +x_574 = lean_ctor_get(x_569, 0); +lean_inc(x_574); +x_575 = lean_ctor_get(x_569, 1); +lean_inc(x_575); +if (lean_is_exclusive(x_569)) { + lean_ctor_release(x_569, 0); + lean_ctor_release(x_569, 1); + x_576 = x_569; +} else { + lean_dec_ref(x_569); + x_576 = lean_box(0); +} +x_577 = lean_ctor_get(x_574, 1); +lean_inc(x_577); +x_578 = lean_nat_dec_eq(x_565, x_577); +lean_dec(x_565); +if (x_578 == 0) +{ +lean_object* x_579; +lean_dec(x_577); +if (lean_is_scalar(x_576)) { + x_579 = lean_alloc_ctor(1, 2, 0); +} else { + x_579 = x_576; +} +lean_ctor_set(x_579, 0, x_574); +lean_ctor_set(x_579, 1, x_575); +return x_579; +} +else +{ +lean_object* x_580; lean_object* x_581; +lean_dec(x_576); +lean_dec(x_575); +x_580 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_581 = l_Std_Internal_Parsec_String_pstring(x_580, x_574); +if (lean_obj_tag(x_581) == 0) +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; +lean_dec(x_577); +x_582 = lean_ctor_get(x_581, 0); +lean_inc(x_582); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_583 = x_581; +} else { + lean_dec_ref(x_581); + x_583 = lean_box(0); +} +x_584 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_583)) { + x_585 = lean_alloc_ctor(0, 2, 0); +} else { + x_585 = x_583; +} +lean_ctor_set(x_585, 0, x_582); +lean_ctor_set(x_585, 1, x_584); +return x_585; +} +else +{ +lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; uint8_t x_590; +x_586 = lean_ctor_get(x_581, 0); +lean_inc(x_586); +x_587 = lean_ctor_get(x_581, 1); +lean_inc(x_587); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_588 = x_581; +} else { + lean_dec_ref(x_581); + x_588 = lean_box(0); +} +x_589 = lean_ctor_get(x_586, 1); +lean_inc(x_589); +x_590 = lean_nat_dec_eq(x_577, x_589); +lean_dec(x_577); +if (x_590 == 0) +{ +lean_object* x_591; +lean_dec(x_589); +if (lean_is_scalar(x_588)) { + x_591 = lean_alloc_ctor(1, 2, 0); +} else { + x_591 = x_588; +} +lean_ctor_set(x_591, 0, x_586); +lean_ctor_set(x_591, 1, x_587); +return x_591; +} +else +{ +lean_object* x_592; lean_object* x_593; +lean_dec(x_588); +lean_dec(x_587); +x_592 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_593 = l_Std_Internal_Parsec_String_pstring(x_592, x_586); +if (lean_obj_tag(x_593) == 0) +{ +lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; +lean_dec(x_589); +x_594 = lean_ctor_get(x_593, 0); +lean_inc(x_594); +if (lean_is_exclusive(x_593)) { + lean_ctor_release(x_593, 0); + lean_ctor_release(x_593, 1); + x_595 = x_593; +} else { + lean_dec_ref(x_593); + x_595 = lean_box(0); +} +x_596 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_595)) { + x_597 = lean_alloc_ctor(0, 2, 0); +} else { + x_597 = x_595; +} +lean_ctor_set(x_597, 0, x_594); +lean_ctor_set(x_597, 1, x_596); +return x_597; +} +else +{ +lean_object* x_598; lean_object* x_599; lean_object* x_600; lean_object* x_601; uint8_t x_602; +x_598 = lean_ctor_get(x_593, 0); +lean_inc(x_598); +x_599 = lean_ctor_get(x_593, 1); +lean_inc(x_599); +if (lean_is_exclusive(x_593)) { + lean_ctor_release(x_593, 0); + lean_ctor_release(x_593, 1); + x_600 = x_593; +} else { + lean_dec_ref(x_593); + x_600 = lean_box(0); +} +x_601 = lean_ctor_get(x_598, 1); +lean_inc(x_601); +x_602 = lean_nat_dec_eq(x_589, x_601); +lean_dec(x_601); +lean_dec(x_589); +if (x_602 == 0) +{ +lean_object* x_603; +if (lean_is_scalar(x_600)) { + x_603 = lean_alloc_ctor(1, 2, 0); +} else { + x_603 = x_600; +} +lean_ctor_set(x_603, 0, x_598); +lean_ctor_set(x_603, 1, x_599); +return x_603; +} +else +{ +lean_object* x_604; lean_object* x_605; +lean_dec(x_600); +lean_dec(x_599); +x_604 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_605 = l_Std_Internal_Parsec_String_pstring(x_604, x_598); +if (lean_obj_tag(x_605) == 0) +{ +lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; +x_606 = lean_ctor_get(x_605, 0); +lean_inc(x_606); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + x_607 = x_605; +} else { + lean_dec_ref(x_605); + x_607 = lean_box(0); +} +x_608 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_607)) { + x_609 = lean_alloc_ctor(0, 2, 0); +} else { + x_609 = x_607; +} +lean_ctor_set(x_609, 0, x_606); +lean_ctor_set(x_609, 1, x_608); +return x_609; +} +else +{ +lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; +x_610 = lean_ctor_get(x_605, 0); +lean_inc(x_610); +x_611 = lean_ctor_get(x_605, 1); +lean_inc(x_611); +if (lean_is_exclusive(x_605)) { + lean_ctor_release(x_605, 0); + lean_ctor_release(x_605, 1); + x_612 = x_605; +} else { + lean_dec_ref(x_605); + x_612 = lean_box(0); +} +if (lean_is_scalar(x_612)) { + x_613 = lean_alloc_ctor(1, 2, 0); +} else { + x_613 = x_612; +} +lean_ctor_set(x_613, 0, x_610); +lean_ctor_set(x_613, 1, x_611); +return x_613; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_614; lean_object* x_615; lean_object* x_616; uint8_t x_617; +x_614 = lean_ctor_get(x_30, 0); +x_615 = lean_ctor_get(x_30, 1); +lean_inc(x_615); +lean_inc(x_614); +lean_dec(x_30); +x_616 = lean_ctor_get(x_614, 1); +lean_inc(x_616); +x_617 = lean_nat_dec_eq(x_27, x_616); +lean_dec(x_27); +if (x_617 == 0) +{ +lean_object* x_618; +lean_dec(x_616); +x_618 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_618, 0, x_614); +lean_ctor_set(x_618, 1, x_615); +return x_618; +} +else +{ +lean_object* x_619; lean_object* x_620; +lean_dec(x_615); +x_619 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +x_620 = l_Std_Internal_Parsec_String_pstring(x_619, x_614); +if (lean_obj_tag(x_620) == 0) +{ +lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; +lean_dec(x_616); +x_621 = lean_ctor_get(x_620, 0); +lean_inc(x_621); +if (lean_is_exclusive(x_620)) { + lean_ctor_release(x_620, 0); + lean_ctor_release(x_620, 1); + x_622 = x_620; +} else { + lean_dec_ref(x_620); + x_622 = lean_box(0); +} +x_623 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_622)) { + x_624 = lean_alloc_ctor(0, 2, 0); +} else { + x_624 = x_622; +} +lean_ctor_set(x_624, 0, x_621); +lean_ctor_set(x_624, 1, x_623); +return x_624; +} +else +{ +lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; uint8_t x_629; +x_625 = lean_ctor_get(x_620, 0); +lean_inc(x_625); +x_626 = lean_ctor_get(x_620, 1); +lean_inc(x_626); +if (lean_is_exclusive(x_620)) { + lean_ctor_release(x_620, 0); + lean_ctor_release(x_620, 1); + x_627 = x_620; +} else { + lean_dec_ref(x_620); + x_627 = lean_box(0); +} +x_628 = lean_ctor_get(x_625, 1); +lean_inc(x_628); +x_629 = lean_nat_dec_eq(x_616, x_628); +lean_dec(x_616); +if (x_629 == 0) +{ +lean_object* x_630; +lean_dec(x_628); +if (lean_is_scalar(x_627)) { + x_630 = lean_alloc_ctor(1, 2, 0); +} else { + x_630 = x_627; +} +lean_ctor_set(x_630, 0, x_625); +lean_ctor_set(x_630, 1, x_626); +return x_630; +} +else +{ +lean_object* x_631; lean_object* x_632; +lean_dec(x_627); +lean_dec(x_626); +x_631 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_632 = l_Std_Internal_Parsec_String_pstring(x_631, x_625); +if (lean_obj_tag(x_632) == 0) +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; +lean_dec(x_628); +x_633 = lean_ctor_get(x_632, 0); +lean_inc(x_633); +if (lean_is_exclusive(x_632)) { + lean_ctor_release(x_632, 0); + lean_ctor_release(x_632, 1); + x_634 = x_632; +} else { + lean_dec_ref(x_632); + x_634 = lean_box(0); +} +x_635 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_634)) { + x_636 = lean_alloc_ctor(0, 2, 0); +} else { + x_636 = x_634; +} +lean_ctor_set(x_636, 0, x_633); +lean_ctor_set(x_636, 1, x_635); +return x_636; +} +else +{ +lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; uint8_t x_641; +x_637 = lean_ctor_get(x_632, 0); +lean_inc(x_637); +x_638 = lean_ctor_get(x_632, 1); +lean_inc(x_638); +if (lean_is_exclusive(x_632)) { + lean_ctor_release(x_632, 0); + lean_ctor_release(x_632, 1); + x_639 = x_632; +} else { + lean_dec_ref(x_632); + x_639 = lean_box(0); +} +x_640 = lean_ctor_get(x_637, 1); +lean_inc(x_640); +x_641 = lean_nat_dec_eq(x_628, x_640); +lean_dec(x_628); +if (x_641 == 0) +{ +lean_object* x_642; +lean_dec(x_640); +if (lean_is_scalar(x_639)) { + x_642 = lean_alloc_ctor(1, 2, 0); +} else { + x_642 = x_639; +} +lean_ctor_set(x_642, 0, x_637); +lean_ctor_set(x_642, 1, x_638); +return x_642; +} +else +{ +lean_object* x_643; lean_object* x_644; +lean_dec(x_639); +lean_dec(x_638); +x_643 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_644 = l_Std_Internal_Parsec_String_pstring(x_643, x_637); +if (lean_obj_tag(x_644) == 0) +{ +lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; +lean_dec(x_640); +x_645 = lean_ctor_get(x_644, 0); +lean_inc(x_645); +if (lean_is_exclusive(x_644)) { + lean_ctor_release(x_644, 0); + lean_ctor_release(x_644, 1); + x_646 = x_644; +} else { + lean_dec_ref(x_644); + x_646 = lean_box(0); +} +x_647 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_646)) { + x_648 = lean_alloc_ctor(0, 2, 0); +} else { + x_648 = x_646; +} +lean_ctor_set(x_648, 0, x_645); +lean_ctor_set(x_648, 1, x_647); +return x_648; +} +else +{ +lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; uint8_t x_653; +x_649 = lean_ctor_get(x_644, 0); +lean_inc(x_649); +x_650 = lean_ctor_get(x_644, 1); +lean_inc(x_650); +if (lean_is_exclusive(x_644)) { + lean_ctor_release(x_644, 0); + lean_ctor_release(x_644, 1); + x_651 = x_644; +} else { + lean_dec_ref(x_644); + x_651 = lean_box(0); +} +x_652 = lean_ctor_get(x_649, 1); +lean_inc(x_652); +x_653 = lean_nat_dec_eq(x_640, x_652); +lean_dec(x_640); +if (x_653 == 0) +{ +lean_object* x_654; +lean_dec(x_652); +if (lean_is_scalar(x_651)) { + x_654 = lean_alloc_ctor(1, 2, 0); +} else { + x_654 = x_651; +} +lean_ctor_set(x_654, 0, x_649); +lean_ctor_set(x_654, 1, x_650); +return x_654; +} +else +{ +lean_object* x_655; lean_object* x_656; +lean_dec(x_651); +lean_dec(x_650); +x_655 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_656 = l_Std_Internal_Parsec_String_pstring(x_655, x_649); +if (lean_obj_tag(x_656) == 0) +{ +lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; +lean_dec(x_652); +x_657 = lean_ctor_get(x_656, 0); +lean_inc(x_657); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_658 = x_656; +} else { + lean_dec_ref(x_656); + x_658 = lean_box(0); +} +x_659 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_658)) { + x_660 = lean_alloc_ctor(0, 2, 0); +} else { + x_660 = x_658; +} +lean_ctor_set(x_660, 0, x_657); +lean_ctor_set(x_660, 1, x_659); +return x_660; +} +else +{ +lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; uint8_t x_665; +x_661 = lean_ctor_get(x_656, 0); +lean_inc(x_661); +x_662 = lean_ctor_get(x_656, 1); +lean_inc(x_662); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_663 = x_656; +} else { + lean_dec_ref(x_656); + x_663 = lean_box(0); +} +x_664 = lean_ctor_get(x_661, 1); +lean_inc(x_664); +x_665 = lean_nat_dec_eq(x_652, x_664); +lean_dec(x_652); +if (x_665 == 0) +{ +lean_object* x_666; +lean_dec(x_664); +if (lean_is_scalar(x_663)) { + x_666 = lean_alloc_ctor(1, 2, 0); +} else { + x_666 = x_663; +} +lean_ctor_set(x_666, 0, x_661); +lean_ctor_set(x_666, 1, x_662); +return x_666; +} +else +{ +lean_object* x_667; lean_object* x_668; +lean_dec(x_663); +lean_dec(x_662); +x_667 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_668 = l_Std_Internal_Parsec_String_pstring(x_667, x_661); +if (lean_obj_tag(x_668) == 0) +{ +lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; +lean_dec(x_664); +x_669 = lean_ctor_get(x_668, 0); +lean_inc(x_669); +if (lean_is_exclusive(x_668)) { + lean_ctor_release(x_668, 0); + lean_ctor_release(x_668, 1); + x_670 = x_668; +} else { + lean_dec_ref(x_668); + x_670 = lean_box(0); +} +x_671 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_670)) { + x_672 = lean_alloc_ctor(0, 2, 0); +} else { + x_672 = x_670; +} +lean_ctor_set(x_672, 0, x_669); +lean_ctor_set(x_672, 1, x_671); +return x_672; +} +else +{ +lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; uint8_t x_677; +x_673 = lean_ctor_get(x_668, 0); +lean_inc(x_673); +x_674 = lean_ctor_get(x_668, 1); +lean_inc(x_674); +if (lean_is_exclusive(x_668)) { + lean_ctor_release(x_668, 0); + lean_ctor_release(x_668, 1); + x_675 = x_668; +} else { + lean_dec_ref(x_668); + x_675 = lean_box(0); +} +x_676 = lean_ctor_get(x_673, 1); +lean_inc(x_676); +x_677 = lean_nat_dec_eq(x_664, x_676); +lean_dec(x_664); +if (x_677 == 0) +{ +lean_object* x_678; +lean_dec(x_676); +if (lean_is_scalar(x_675)) { + x_678 = lean_alloc_ctor(1, 2, 0); +} else { + x_678 = x_675; +} +lean_ctor_set(x_678, 0, x_673); +lean_ctor_set(x_678, 1, x_674); +return x_678; +} +else +{ +lean_object* x_679; lean_object* x_680; +lean_dec(x_675); +lean_dec(x_674); +x_679 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_680 = l_Std_Internal_Parsec_String_pstring(x_679, x_673); +if (lean_obj_tag(x_680) == 0) +{ +lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; +lean_dec(x_676); +x_681 = lean_ctor_get(x_680, 0); +lean_inc(x_681); +if (lean_is_exclusive(x_680)) { + lean_ctor_release(x_680, 0); + lean_ctor_release(x_680, 1); + x_682 = x_680; +} else { + lean_dec_ref(x_680); + x_682 = lean_box(0); +} +x_683 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_682)) { + x_684 = lean_alloc_ctor(0, 2, 0); +} else { + x_684 = x_682; +} +lean_ctor_set(x_684, 0, x_681); +lean_ctor_set(x_684, 1, x_683); +return x_684; +} +else +{ +lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; uint8_t x_689; +x_685 = lean_ctor_get(x_680, 0); +lean_inc(x_685); +x_686 = lean_ctor_get(x_680, 1); +lean_inc(x_686); +if (lean_is_exclusive(x_680)) { + lean_ctor_release(x_680, 0); + lean_ctor_release(x_680, 1); + x_687 = x_680; +} else { + lean_dec_ref(x_680); + x_687 = lean_box(0); +} +x_688 = lean_ctor_get(x_685, 1); +lean_inc(x_688); +x_689 = lean_nat_dec_eq(x_676, x_688); +lean_dec(x_676); +if (x_689 == 0) +{ +lean_object* x_690; +lean_dec(x_688); +if (lean_is_scalar(x_687)) { + x_690 = lean_alloc_ctor(1, 2, 0); +} else { + x_690 = x_687; +} +lean_ctor_set(x_690, 0, x_685); +lean_ctor_set(x_690, 1, x_686); +return x_690; +} +else +{ +lean_object* x_691; lean_object* x_692; +lean_dec(x_687); +lean_dec(x_686); +x_691 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_692 = l_Std_Internal_Parsec_String_pstring(x_691, x_685); +if (lean_obj_tag(x_692) == 0) +{ +lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; +lean_dec(x_688); +x_693 = lean_ctor_get(x_692, 0); +lean_inc(x_693); +if (lean_is_exclusive(x_692)) { + lean_ctor_release(x_692, 0); + lean_ctor_release(x_692, 1); + x_694 = x_692; +} else { + lean_dec_ref(x_692); + x_694 = lean_box(0); +} +x_695 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_694)) { + x_696 = lean_alloc_ctor(0, 2, 0); +} else { + x_696 = x_694; +} +lean_ctor_set(x_696, 0, x_693); +lean_ctor_set(x_696, 1, x_695); +return x_696; +} +else +{ +lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; uint8_t x_701; +x_697 = lean_ctor_get(x_692, 0); +lean_inc(x_697); +x_698 = lean_ctor_get(x_692, 1); +lean_inc(x_698); +if (lean_is_exclusive(x_692)) { + lean_ctor_release(x_692, 0); + lean_ctor_release(x_692, 1); + x_699 = x_692; +} else { + lean_dec_ref(x_692); + x_699 = lean_box(0); +} +x_700 = lean_ctor_get(x_697, 1); +lean_inc(x_700); +x_701 = lean_nat_dec_eq(x_688, x_700); +lean_dec(x_688); +if (x_701 == 0) +{ +lean_object* x_702; +lean_dec(x_700); +if (lean_is_scalar(x_699)) { + x_702 = lean_alloc_ctor(1, 2, 0); +} else { + x_702 = x_699; +} +lean_ctor_set(x_702, 0, x_697); +lean_ctor_set(x_702, 1, x_698); +return x_702; +} +else +{ +lean_object* x_703; lean_object* x_704; +lean_dec(x_699); +lean_dec(x_698); +x_703 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_704 = l_Std_Internal_Parsec_String_pstring(x_703, x_697); +if (lean_obj_tag(x_704) == 0) +{ +lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; +lean_dec(x_700); +x_705 = lean_ctor_get(x_704, 0); +lean_inc(x_705); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_706 = x_704; +} else { + lean_dec_ref(x_704); + x_706 = lean_box(0); +} +x_707 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_706)) { + x_708 = lean_alloc_ctor(0, 2, 0); +} else { + x_708 = x_706; +} +lean_ctor_set(x_708, 0, x_705); +lean_ctor_set(x_708, 1, x_707); +return x_708; +} +else +{ +lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; uint8_t x_713; +x_709 = lean_ctor_get(x_704, 0); +lean_inc(x_709); +x_710 = lean_ctor_get(x_704, 1); +lean_inc(x_710); +if (lean_is_exclusive(x_704)) { + lean_ctor_release(x_704, 0); + lean_ctor_release(x_704, 1); + x_711 = x_704; +} else { + lean_dec_ref(x_704); + x_711 = lean_box(0); +} +x_712 = lean_ctor_get(x_709, 1); +lean_inc(x_712); +x_713 = lean_nat_dec_eq(x_700, x_712); +lean_dec(x_712); +lean_dec(x_700); +if (x_713 == 0) +{ +lean_object* x_714; +if (lean_is_scalar(x_711)) { + x_714 = lean_alloc_ctor(1, 2, 0); +} else { + x_714 = x_711; +} +lean_ctor_set(x_714, 0, x_709); +lean_ctor_set(x_714, 1, x_710); +return x_714; +} +else +{ +lean_object* x_715; lean_object* x_716; +lean_dec(x_711); +lean_dec(x_710); +x_715 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_716 = l_Std_Internal_Parsec_String_pstring(x_715, x_709); +if (lean_obj_tag(x_716) == 0) +{ +lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; +x_717 = lean_ctor_get(x_716, 0); +lean_inc(x_717); +if (lean_is_exclusive(x_716)) { + lean_ctor_release(x_716, 0); + lean_ctor_release(x_716, 1); + x_718 = x_716; +} else { + lean_dec_ref(x_716); + x_718 = lean_box(0); +} +x_719 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_718)) { + x_720 = lean_alloc_ctor(0, 2, 0); +} else { + x_720 = x_718; +} +lean_ctor_set(x_720, 0, x_717); +lean_ctor_set(x_720, 1, x_719); +return x_720; +} +else +{ +lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; +x_721 = lean_ctor_get(x_716, 0); +lean_inc(x_721); +x_722 = lean_ctor_get(x_716, 1); +lean_inc(x_722); +if (lean_is_exclusive(x_716)) { + lean_ctor_release(x_716, 0); + lean_ctor_release(x_716, 1); + x_723 = x_716; +} else { + lean_dec_ref(x_716); + x_723 = lean_box(0); +} +if (lean_is_scalar(x_723)) { + x_724 = lean_alloc_ctor(1, 2, 0); +} else { + x_724 = x_723; +} +lean_ctor_set(x_724, 0, x_721); +lean_ctor_set(x_724, 1, x_722); +return x_724; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_725; lean_object* x_726; lean_object* x_727; uint8_t x_728; +x_725 = lean_ctor_get(x_17, 0); +x_726 = lean_ctor_get(x_17, 1); +lean_inc(x_726); +lean_inc(x_725); +lean_dec(x_17); +x_727 = lean_ctor_get(x_725, 1); +lean_inc(x_727); +x_728 = lean_nat_dec_eq(x_14, x_727); +lean_dec(x_14); +if (x_728 == 0) +{ +lean_object* x_729; +lean_dec(x_727); +x_729 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_729, 0, x_725); +lean_ctor_set(x_729, 1, x_726); +return x_729; +} +else +{ +lean_object* x_730; lean_object* x_731; +lean_dec(x_726); +x_730 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9; +x_731 = l_Std_Internal_Parsec_String_pstring(x_730, x_725); +if (lean_obj_tag(x_731) == 0) +{ +lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; +lean_dec(x_727); +x_732 = lean_ctor_get(x_731, 0); +lean_inc(x_732); +if (lean_is_exclusive(x_731)) { + lean_ctor_release(x_731, 0); + lean_ctor_release(x_731, 1); + x_733 = x_731; +} else { + lean_dec_ref(x_731); + x_733 = lean_box(0); +} +x_734 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_733)) { + x_735 = lean_alloc_ctor(0, 2, 0); +} else { + x_735 = x_733; +} +lean_ctor_set(x_735, 0, x_732); +lean_ctor_set(x_735, 1, x_734); +return x_735; +} +else +{ +lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; uint8_t x_740; +x_736 = lean_ctor_get(x_731, 0); +lean_inc(x_736); +x_737 = lean_ctor_get(x_731, 1); +lean_inc(x_737); +if (lean_is_exclusive(x_731)) { + lean_ctor_release(x_731, 0); + lean_ctor_release(x_731, 1); + x_738 = x_731; +} else { + lean_dec_ref(x_731); + x_738 = lean_box(0); +} +x_739 = lean_ctor_get(x_736, 1); +lean_inc(x_739); +x_740 = lean_nat_dec_eq(x_727, x_739); +lean_dec(x_727); +if (x_740 == 0) +{ +lean_object* x_741; +lean_dec(x_739); +if (lean_is_scalar(x_738)) { + x_741 = lean_alloc_ctor(1, 2, 0); +} else { + x_741 = x_738; +} +lean_ctor_set(x_741, 0, x_736); +lean_ctor_set(x_741, 1, x_737); +return x_741; +} +else +{ +lean_object* x_742; lean_object* x_743; +lean_dec(x_738); +lean_dec(x_737); +x_742 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +x_743 = l_Std_Internal_Parsec_String_pstring(x_742, x_736); +if (lean_obj_tag(x_743) == 0) +{ +lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; +lean_dec(x_739); +x_744 = lean_ctor_get(x_743, 0); +lean_inc(x_744); +if (lean_is_exclusive(x_743)) { + lean_ctor_release(x_743, 0); + lean_ctor_release(x_743, 1); + x_745 = x_743; +} else { + lean_dec_ref(x_743); + x_745 = lean_box(0); +} +x_746 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_745)) { + x_747 = lean_alloc_ctor(0, 2, 0); +} else { + x_747 = x_745; +} +lean_ctor_set(x_747, 0, x_744); +lean_ctor_set(x_747, 1, x_746); +return x_747; +} +else +{ +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; uint8_t x_752; +x_748 = lean_ctor_get(x_743, 0); +lean_inc(x_748); +x_749 = lean_ctor_get(x_743, 1); +lean_inc(x_749); +if (lean_is_exclusive(x_743)) { + lean_ctor_release(x_743, 0); + lean_ctor_release(x_743, 1); + x_750 = x_743; +} else { + lean_dec_ref(x_743); + x_750 = lean_box(0); +} +x_751 = lean_ctor_get(x_748, 1); +lean_inc(x_751); +x_752 = lean_nat_dec_eq(x_739, x_751); +lean_dec(x_739); +if (x_752 == 0) +{ +lean_object* x_753; +lean_dec(x_751); +if (lean_is_scalar(x_750)) { + x_753 = lean_alloc_ctor(1, 2, 0); +} else { + x_753 = x_750; +} +lean_ctor_set(x_753, 0, x_748); +lean_ctor_set(x_753, 1, x_749); +return x_753; +} +else +{ +lean_object* x_754; lean_object* x_755; +lean_dec(x_750); +lean_dec(x_749); +x_754 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_755 = l_Std_Internal_Parsec_String_pstring(x_754, x_748); +if (lean_obj_tag(x_755) == 0) +{ +lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; +lean_dec(x_751); +x_756 = lean_ctor_get(x_755, 0); +lean_inc(x_756); +if (lean_is_exclusive(x_755)) { + lean_ctor_release(x_755, 0); + lean_ctor_release(x_755, 1); + x_757 = x_755; +} else { + lean_dec_ref(x_755); + x_757 = lean_box(0); +} +x_758 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_757)) { + x_759 = lean_alloc_ctor(0, 2, 0); +} else { + x_759 = x_757; +} +lean_ctor_set(x_759, 0, x_756); +lean_ctor_set(x_759, 1, x_758); +return x_759; +} +else +{ +lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; uint8_t x_764; +x_760 = lean_ctor_get(x_755, 0); +lean_inc(x_760); +x_761 = lean_ctor_get(x_755, 1); +lean_inc(x_761); +if (lean_is_exclusive(x_755)) { + lean_ctor_release(x_755, 0); + lean_ctor_release(x_755, 1); + x_762 = x_755; +} else { + lean_dec_ref(x_755); + x_762 = lean_box(0); +} +x_763 = lean_ctor_get(x_760, 1); +lean_inc(x_763); +x_764 = lean_nat_dec_eq(x_751, x_763); +lean_dec(x_751); +if (x_764 == 0) +{ +lean_object* x_765; +lean_dec(x_763); +if (lean_is_scalar(x_762)) { + x_765 = lean_alloc_ctor(1, 2, 0); +} else { + x_765 = x_762; +} +lean_ctor_set(x_765, 0, x_760); +lean_ctor_set(x_765, 1, x_761); +return x_765; +} +else +{ +lean_object* x_766; lean_object* x_767; +lean_dec(x_762); +lean_dec(x_761); +x_766 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_767 = l_Std_Internal_Parsec_String_pstring(x_766, x_760); +if (lean_obj_tag(x_767) == 0) +{ +lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; +lean_dec(x_763); +x_768 = lean_ctor_get(x_767, 0); +lean_inc(x_768); +if (lean_is_exclusive(x_767)) { + lean_ctor_release(x_767, 0); + lean_ctor_release(x_767, 1); + x_769 = x_767; +} else { + lean_dec_ref(x_767); + x_769 = lean_box(0); +} +x_770 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_769)) { + x_771 = lean_alloc_ctor(0, 2, 0); +} else { + x_771 = x_769; +} +lean_ctor_set(x_771, 0, x_768); +lean_ctor_set(x_771, 1, x_770); +return x_771; +} +else +{ +lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; uint8_t x_776; +x_772 = lean_ctor_get(x_767, 0); +lean_inc(x_772); +x_773 = lean_ctor_get(x_767, 1); +lean_inc(x_773); +if (lean_is_exclusive(x_767)) { + lean_ctor_release(x_767, 0); + lean_ctor_release(x_767, 1); + x_774 = x_767; +} else { + lean_dec_ref(x_767); + x_774 = lean_box(0); +} +x_775 = lean_ctor_get(x_772, 1); +lean_inc(x_775); +x_776 = lean_nat_dec_eq(x_763, x_775); +lean_dec(x_763); +if (x_776 == 0) +{ +lean_object* x_777; +lean_dec(x_775); +if (lean_is_scalar(x_774)) { + x_777 = lean_alloc_ctor(1, 2, 0); +} else { + x_777 = x_774; +} +lean_ctor_set(x_777, 0, x_772); +lean_ctor_set(x_777, 1, x_773); +return x_777; +} +else +{ +lean_object* x_778; lean_object* x_779; +lean_dec(x_774); +lean_dec(x_773); +x_778 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_779 = l_Std_Internal_Parsec_String_pstring(x_778, x_772); +if (lean_obj_tag(x_779) == 0) +{ +lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; +lean_dec(x_775); +x_780 = lean_ctor_get(x_779, 0); +lean_inc(x_780); +if (lean_is_exclusive(x_779)) { + lean_ctor_release(x_779, 0); + lean_ctor_release(x_779, 1); + x_781 = x_779; +} else { + lean_dec_ref(x_779); + x_781 = lean_box(0); +} +x_782 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_781)) { + x_783 = lean_alloc_ctor(0, 2, 0); +} else { + x_783 = x_781; +} +lean_ctor_set(x_783, 0, x_780); +lean_ctor_set(x_783, 1, x_782); +return x_783; +} +else +{ +lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; uint8_t x_788; +x_784 = lean_ctor_get(x_779, 0); +lean_inc(x_784); +x_785 = lean_ctor_get(x_779, 1); +lean_inc(x_785); +if (lean_is_exclusive(x_779)) { + lean_ctor_release(x_779, 0); + lean_ctor_release(x_779, 1); + x_786 = x_779; +} else { + lean_dec_ref(x_779); + x_786 = lean_box(0); +} +x_787 = lean_ctor_get(x_784, 1); +lean_inc(x_787); +x_788 = lean_nat_dec_eq(x_775, x_787); +lean_dec(x_775); +if (x_788 == 0) +{ +lean_object* x_789; +lean_dec(x_787); +if (lean_is_scalar(x_786)) { + x_789 = lean_alloc_ctor(1, 2, 0); +} else { + x_789 = x_786; +} +lean_ctor_set(x_789, 0, x_784); +lean_ctor_set(x_789, 1, x_785); +return x_789; +} +else +{ +lean_object* x_790; lean_object* x_791; +lean_dec(x_786); +lean_dec(x_785); +x_790 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_791 = l_Std_Internal_Parsec_String_pstring(x_790, x_784); +if (lean_obj_tag(x_791) == 0) +{ +lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; +lean_dec(x_787); +x_792 = lean_ctor_get(x_791, 0); +lean_inc(x_792); +if (lean_is_exclusive(x_791)) { + lean_ctor_release(x_791, 0); + lean_ctor_release(x_791, 1); + x_793 = x_791; +} else { + lean_dec_ref(x_791); + x_793 = lean_box(0); +} +x_794 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_793)) { + x_795 = lean_alloc_ctor(0, 2, 0); +} else { + x_795 = x_793; +} +lean_ctor_set(x_795, 0, x_792); +lean_ctor_set(x_795, 1, x_794); +return x_795; +} +else +{ +lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; uint8_t x_800; +x_796 = lean_ctor_get(x_791, 0); +lean_inc(x_796); +x_797 = lean_ctor_get(x_791, 1); +lean_inc(x_797); +if (lean_is_exclusive(x_791)) { + lean_ctor_release(x_791, 0); + lean_ctor_release(x_791, 1); + x_798 = x_791; +} else { + lean_dec_ref(x_791); + x_798 = lean_box(0); +} +x_799 = lean_ctor_get(x_796, 1); +lean_inc(x_799); +x_800 = lean_nat_dec_eq(x_787, x_799); +lean_dec(x_787); +if (x_800 == 0) +{ +lean_object* x_801; +lean_dec(x_799); +if (lean_is_scalar(x_798)) { + x_801 = lean_alloc_ctor(1, 2, 0); +} else { + x_801 = x_798; +} +lean_ctor_set(x_801, 0, x_796); +lean_ctor_set(x_801, 1, x_797); +return x_801; +} +else +{ +lean_object* x_802; lean_object* x_803; +lean_dec(x_798); +lean_dec(x_797); +x_802 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_803 = l_Std_Internal_Parsec_String_pstring(x_802, x_796); +if (lean_obj_tag(x_803) == 0) +{ +lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; +lean_dec(x_799); +x_804 = lean_ctor_get(x_803, 0); +lean_inc(x_804); +if (lean_is_exclusive(x_803)) { + lean_ctor_release(x_803, 0); + lean_ctor_release(x_803, 1); + x_805 = x_803; +} else { + lean_dec_ref(x_803); + x_805 = lean_box(0); +} +x_806 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_805)) { + x_807 = lean_alloc_ctor(0, 2, 0); +} else { + x_807 = x_805; +} +lean_ctor_set(x_807, 0, x_804); +lean_ctor_set(x_807, 1, x_806); +return x_807; +} +else +{ +lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; uint8_t x_812; +x_808 = lean_ctor_get(x_803, 0); +lean_inc(x_808); +x_809 = lean_ctor_get(x_803, 1); +lean_inc(x_809); +if (lean_is_exclusive(x_803)) { + lean_ctor_release(x_803, 0); + lean_ctor_release(x_803, 1); + x_810 = x_803; +} else { + lean_dec_ref(x_803); + x_810 = lean_box(0); +} +x_811 = lean_ctor_get(x_808, 1); +lean_inc(x_811); +x_812 = lean_nat_dec_eq(x_799, x_811); +lean_dec(x_799); +if (x_812 == 0) +{ +lean_object* x_813; +lean_dec(x_811); +if (lean_is_scalar(x_810)) { + x_813 = lean_alloc_ctor(1, 2, 0); +} else { + x_813 = x_810; +} +lean_ctor_set(x_813, 0, x_808); +lean_ctor_set(x_813, 1, x_809); +return x_813; +} +else +{ +lean_object* x_814; lean_object* x_815; +lean_dec(x_810); +lean_dec(x_809); +x_814 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_815 = l_Std_Internal_Parsec_String_pstring(x_814, x_808); +if (lean_obj_tag(x_815) == 0) +{ +lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; +lean_dec(x_811); +x_816 = lean_ctor_get(x_815, 0); +lean_inc(x_816); +if (lean_is_exclusive(x_815)) { + lean_ctor_release(x_815, 0); + lean_ctor_release(x_815, 1); + x_817 = x_815; +} else { + lean_dec_ref(x_815); + x_817 = lean_box(0); +} +x_818 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_817)) { + x_819 = lean_alloc_ctor(0, 2, 0); +} else { + x_819 = x_817; +} +lean_ctor_set(x_819, 0, x_816); +lean_ctor_set(x_819, 1, x_818); +return x_819; +} +else +{ +lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; uint8_t x_824; +x_820 = lean_ctor_get(x_815, 0); +lean_inc(x_820); +x_821 = lean_ctor_get(x_815, 1); +lean_inc(x_821); +if (lean_is_exclusive(x_815)) { + lean_ctor_release(x_815, 0); + lean_ctor_release(x_815, 1); + x_822 = x_815; +} else { + lean_dec_ref(x_815); + x_822 = lean_box(0); +} +x_823 = lean_ctor_get(x_820, 1); +lean_inc(x_823); +x_824 = lean_nat_dec_eq(x_811, x_823); +lean_dec(x_811); +if (x_824 == 0) +{ +lean_object* x_825; +lean_dec(x_823); +if (lean_is_scalar(x_822)) { + x_825 = lean_alloc_ctor(1, 2, 0); +} else { + x_825 = x_822; +} +lean_ctor_set(x_825, 0, x_820); +lean_ctor_set(x_825, 1, x_821); +return x_825; +} +else +{ +lean_object* x_826; lean_object* x_827; +lean_dec(x_822); +lean_dec(x_821); +x_826 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_827 = l_Std_Internal_Parsec_String_pstring(x_826, x_820); +if (lean_obj_tag(x_827) == 0) +{ +lean_object* x_828; lean_object* x_829; lean_object* x_830; lean_object* x_831; +lean_dec(x_823); +x_828 = lean_ctor_get(x_827, 0); +lean_inc(x_828); +if (lean_is_exclusive(x_827)) { + lean_ctor_release(x_827, 0); + lean_ctor_release(x_827, 1); + x_829 = x_827; +} else { + lean_dec_ref(x_827); + x_829 = lean_box(0); +} +x_830 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_829)) { + x_831 = lean_alloc_ctor(0, 2, 0); +} else { + x_831 = x_829; +} +lean_ctor_set(x_831, 0, x_828); +lean_ctor_set(x_831, 1, x_830); +return x_831; +} +else +{ +lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; uint8_t x_836; +x_832 = lean_ctor_get(x_827, 0); +lean_inc(x_832); +x_833 = lean_ctor_get(x_827, 1); +lean_inc(x_833); +if (lean_is_exclusive(x_827)) { + lean_ctor_release(x_827, 0); + lean_ctor_release(x_827, 1); + x_834 = x_827; +} else { + lean_dec_ref(x_827); + x_834 = lean_box(0); +} +x_835 = lean_ctor_get(x_832, 1); +lean_inc(x_835); +x_836 = lean_nat_dec_eq(x_823, x_835); +lean_dec(x_835); +lean_dec(x_823); +if (x_836 == 0) +{ +lean_object* x_837; +if (lean_is_scalar(x_834)) { + x_837 = lean_alloc_ctor(1, 2, 0); +} else { + x_837 = x_834; +} +lean_ctor_set(x_837, 0, x_832); +lean_ctor_set(x_837, 1, x_833); +return x_837; +} +else +{ +lean_object* x_838; lean_object* x_839; +lean_dec(x_834); +lean_dec(x_833); +x_838 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_839 = l_Std_Internal_Parsec_String_pstring(x_838, x_832); +if (lean_obj_tag(x_839) == 0) +{ +lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; +x_840 = lean_ctor_get(x_839, 0); +lean_inc(x_840); +if (lean_is_exclusive(x_839)) { + lean_ctor_release(x_839, 0); + lean_ctor_release(x_839, 1); + x_841 = x_839; +} else { + lean_dec_ref(x_839); + x_841 = lean_box(0); +} +x_842 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_841)) { + x_843 = lean_alloc_ctor(0, 2, 0); +} else { + x_843 = x_841; +} +lean_ctor_set(x_843, 0, x_840); +lean_ctor_set(x_843, 1, x_842); +return x_843; +} +else +{ +lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; +x_844 = lean_ctor_get(x_839, 0); +lean_inc(x_844); +x_845 = lean_ctor_get(x_839, 1); +lean_inc(x_845); +if (lean_is_exclusive(x_839)) { + lean_ctor_release(x_839, 0); + lean_ctor_release(x_839, 1); + x_846 = x_839; +} else { + lean_dec_ref(x_839); + x_846 = lean_box(0); +} +if (lean_is_scalar(x_846)) { + x_847 = lean_alloc_ctor(1, 2, 0); +} else { + x_847 = x_846; +} +lean_ctor_set(x_847, 0, x_844); +lean_ctor_set(x_847, 1, x_845); +return x_847; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; uint8_t x_852; +x_848 = lean_ctor_get(x_3, 0); +x_849 = lean_ctor_get(x_3, 1); +lean_inc(x_849); +lean_inc(x_848); +lean_dec(x_3); +x_850 = lean_ctor_get(x_1, 1); +lean_inc(x_850); +lean_dec(x_1); +x_851 = lean_ctor_get(x_848, 1); +lean_inc(x_851); +x_852 = lean_nat_dec_eq(x_850, x_851); +lean_dec(x_850); +if (x_852 == 0) +{ +lean_object* x_853; +lean_dec(x_851); +x_853 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_853, 0, x_848); +lean_ctor_set(x_853, 1, x_849); +return x_853; +} +else +{ +lean_object* x_854; lean_object* x_855; +lean_dec(x_849); +x_854 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10; +x_855 = l_Std_Internal_Parsec_String_pstring(x_854, x_848); +if (lean_obj_tag(x_855) == 0) +{ +lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; +lean_dec(x_851); +x_856 = lean_ctor_get(x_855, 0); +lean_inc(x_856); +if (lean_is_exclusive(x_855)) { + lean_ctor_release(x_855, 0); + lean_ctor_release(x_855, 1); + x_857 = x_855; +} else { + lean_dec_ref(x_855); + x_857 = lean_box(0); +} +x_858 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_857)) { + x_859 = lean_alloc_ctor(0, 2, 0); +} else { + x_859 = x_857; +} +lean_ctor_set(x_859, 0, x_856); +lean_ctor_set(x_859, 1, x_858); +return x_859; +} +else +{ +lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; uint8_t x_864; +x_860 = lean_ctor_get(x_855, 0); +lean_inc(x_860); +x_861 = lean_ctor_get(x_855, 1); +lean_inc(x_861); +if (lean_is_exclusive(x_855)) { + lean_ctor_release(x_855, 0); + lean_ctor_release(x_855, 1); + x_862 = x_855; +} else { + lean_dec_ref(x_855); + x_862 = lean_box(0); +} +x_863 = lean_ctor_get(x_860, 1); +lean_inc(x_863); +x_864 = lean_nat_dec_eq(x_851, x_863); +lean_dec(x_851); +if (x_864 == 0) +{ +lean_object* x_865; +lean_dec(x_863); +if (lean_is_scalar(x_862)) { + x_865 = lean_alloc_ctor(1, 2, 0); +} else { + x_865 = x_862; +} +lean_ctor_set(x_865, 0, x_860); +lean_ctor_set(x_865, 1, x_861); +return x_865; +} +else +{ +lean_object* x_866; lean_object* x_867; +lean_dec(x_862); +lean_dec(x_861); +x_866 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9; +x_867 = l_Std_Internal_Parsec_String_pstring(x_866, x_860); +if (lean_obj_tag(x_867) == 0) +{ +lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; +lean_dec(x_863); +x_868 = lean_ctor_get(x_867, 0); +lean_inc(x_868); +if (lean_is_exclusive(x_867)) { + lean_ctor_release(x_867, 0); + lean_ctor_release(x_867, 1); + x_869 = x_867; +} else { + lean_dec_ref(x_867); + x_869 = lean_box(0); +} +x_870 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_869)) { + x_871 = lean_alloc_ctor(0, 2, 0); +} else { + x_871 = x_869; +} +lean_ctor_set(x_871, 0, x_868); +lean_ctor_set(x_871, 1, x_870); +return x_871; +} +else +{ +lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; uint8_t x_876; +x_872 = lean_ctor_get(x_867, 0); +lean_inc(x_872); +x_873 = lean_ctor_get(x_867, 1); +lean_inc(x_873); +if (lean_is_exclusive(x_867)) { + lean_ctor_release(x_867, 0); + lean_ctor_release(x_867, 1); + x_874 = x_867; +} else { + lean_dec_ref(x_867); + x_874 = lean_box(0); +} +x_875 = lean_ctor_get(x_872, 1); +lean_inc(x_875); +x_876 = lean_nat_dec_eq(x_863, x_875); +lean_dec(x_863); +if (x_876 == 0) +{ +lean_object* x_877; +lean_dec(x_875); +if (lean_is_scalar(x_874)) { + x_877 = lean_alloc_ctor(1, 2, 0); +} else { + x_877 = x_874; +} +lean_ctor_set(x_877, 0, x_872); +lean_ctor_set(x_877, 1, x_873); +return x_877; +} +else +{ +lean_object* x_878; lean_object* x_879; +lean_dec(x_874); +lean_dec(x_873); +x_878 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8; +x_879 = l_Std_Internal_Parsec_String_pstring(x_878, x_872); +if (lean_obj_tag(x_879) == 0) +{ +lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_dec(x_875); +x_880 = lean_ctor_get(x_879, 0); +lean_inc(x_880); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_881 = x_879; +} else { + lean_dec_ref(x_879); + x_881 = lean_box(0); +} +x_882 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_881)) { + x_883 = lean_alloc_ctor(0, 2, 0); +} else { + x_883 = x_881; +} +lean_ctor_set(x_883, 0, x_880); +lean_ctor_set(x_883, 1, x_882); +return x_883; +} +else +{ +lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; uint8_t x_888; +x_884 = lean_ctor_get(x_879, 0); +lean_inc(x_884); +x_885 = lean_ctor_get(x_879, 1); +lean_inc(x_885); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_886 = x_879; +} else { + lean_dec_ref(x_879); + x_886 = lean_box(0); +} +x_887 = lean_ctor_get(x_884, 1); +lean_inc(x_887); +x_888 = lean_nat_dec_eq(x_875, x_887); +lean_dec(x_875); +if (x_888 == 0) +{ +lean_object* x_889; +lean_dec(x_887); +if (lean_is_scalar(x_886)) { + x_889 = lean_alloc_ctor(1, 2, 0); +} else { + x_889 = x_886; +} +lean_ctor_set(x_889, 0, x_884); +lean_ctor_set(x_889, 1, x_885); +return x_889; +} +else +{ +lean_object* x_890; lean_object* x_891; +lean_dec(x_886); +lean_dec(x_885); +x_890 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8; +x_891 = l_Std_Internal_Parsec_String_pstring(x_890, x_884); +if (lean_obj_tag(x_891) == 0) +{ +lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; +lean_dec(x_887); +x_892 = lean_ctor_get(x_891, 0); +lean_inc(x_892); +if (lean_is_exclusive(x_891)) { + lean_ctor_release(x_891, 0); + lean_ctor_release(x_891, 1); + x_893 = x_891; +} else { + lean_dec_ref(x_891); + x_893 = lean_box(0); +} +x_894 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_893)) { + x_895 = lean_alloc_ctor(0, 2, 0); +} else { + x_895 = x_893; +} +lean_ctor_set(x_895, 0, x_892); +lean_ctor_set(x_895, 1, x_894); +return x_895; +} +else +{ +lean_object* x_896; lean_object* x_897; lean_object* x_898; lean_object* x_899; uint8_t x_900; +x_896 = lean_ctor_get(x_891, 0); +lean_inc(x_896); +x_897 = lean_ctor_get(x_891, 1); +lean_inc(x_897); +if (lean_is_exclusive(x_891)) { + lean_ctor_release(x_891, 0); + lean_ctor_release(x_891, 1); + x_898 = x_891; +} else { + lean_dec_ref(x_891); + x_898 = lean_box(0); +} +x_899 = lean_ctor_get(x_896, 1); +lean_inc(x_899); +x_900 = lean_nat_dec_eq(x_887, x_899); +lean_dec(x_887); +if (x_900 == 0) +{ +lean_object* x_901; +lean_dec(x_899); +if (lean_is_scalar(x_898)) { + x_901 = lean_alloc_ctor(1, 2, 0); +} else { + x_901 = x_898; +} +lean_ctor_set(x_901, 0, x_896); +lean_ctor_set(x_901, 1, x_897); +return x_901; +} +else +{ +lean_object* x_902; lean_object* x_903; +lean_dec(x_898); +lean_dec(x_897); +x_902 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7; +x_903 = l_Std_Internal_Parsec_String_pstring(x_902, x_896); +if (lean_obj_tag(x_903) == 0) +{ +lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; +lean_dec(x_899); +x_904 = lean_ctor_get(x_903, 0); +lean_inc(x_904); +if (lean_is_exclusive(x_903)) { + lean_ctor_release(x_903, 0); + lean_ctor_release(x_903, 1); + x_905 = x_903; +} else { + lean_dec_ref(x_903); + x_905 = lean_box(0); +} +x_906 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_905)) { + x_907 = lean_alloc_ctor(0, 2, 0); +} else { + x_907 = x_905; +} +lean_ctor_set(x_907, 0, x_904); +lean_ctor_set(x_907, 1, x_906); +return x_907; +} +else +{ +lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; uint8_t x_912; +x_908 = lean_ctor_get(x_903, 0); +lean_inc(x_908); +x_909 = lean_ctor_get(x_903, 1); +lean_inc(x_909); +if (lean_is_exclusive(x_903)) { + lean_ctor_release(x_903, 0); + lean_ctor_release(x_903, 1); + x_910 = x_903; +} else { + lean_dec_ref(x_903); + x_910 = lean_box(0); +} +x_911 = lean_ctor_get(x_908, 1); +lean_inc(x_911); +x_912 = lean_nat_dec_eq(x_899, x_911); +lean_dec(x_899); +if (x_912 == 0) +{ +lean_object* x_913; +lean_dec(x_911); +if (lean_is_scalar(x_910)) { + x_913 = lean_alloc_ctor(1, 2, 0); +} else { + x_913 = x_910; +} +lean_ctor_set(x_913, 0, x_908); +lean_ctor_set(x_913, 1, x_909); +return x_913; +} +else +{ +lean_object* x_914; lean_object* x_915; +lean_dec(x_910); +lean_dec(x_909); +x_914 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6; +x_915 = l_Std_Internal_Parsec_String_pstring(x_914, x_908); +if (lean_obj_tag(x_915) == 0) +{ +lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; +lean_dec(x_911); +x_916 = lean_ctor_get(x_915, 0); +lean_inc(x_916); +if (lean_is_exclusive(x_915)) { + lean_ctor_release(x_915, 0); + lean_ctor_release(x_915, 1); + x_917 = x_915; +} else { + lean_dec_ref(x_915); + x_917 = lean_box(0); +} +x_918 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_917)) { + x_919 = lean_alloc_ctor(0, 2, 0); +} else { + x_919 = x_917; +} +lean_ctor_set(x_919, 0, x_916); +lean_ctor_set(x_919, 1, x_918); +return x_919; +} +else +{ +lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; uint8_t x_924; +x_920 = lean_ctor_get(x_915, 0); +lean_inc(x_920); +x_921 = lean_ctor_get(x_915, 1); +lean_inc(x_921); +if (lean_is_exclusive(x_915)) { + lean_ctor_release(x_915, 0); + lean_ctor_release(x_915, 1); + x_922 = x_915; +} else { + lean_dec_ref(x_915); + x_922 = lean_box(0); +} +x_923 = lean_ctor_get(x_920, 1); +lean_inc(x_923); +x_924 = lean_nat_dec_eq(x_911, x_923); +lean_dec(x_911); +if (x_924 == 0) +{ +lean_object* x_925; +lean_dec(x_923); +if (lean_is_scalar(x_922)) { + x_925 = lean_alloc_ctor(1, 2, 0); +} else { + x_925 = x_922; +} +lean_ctor_set(x_925, 0, x_920); +lean_ctor_set(x_925, 1, x_921); +return x_925; +} +else +{ +lean_object* x_926; lean_object* x_927; +lean_dec(x_922); +lean_dec(x_921); +x_926 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5; +x_927 = l_Std_Internal_Parsec_String_pstring(x_926, x_920); +if (lean_obj_tag(x_927) == 0) +{ +lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +lean_dec(x_923); +x_928 = lean_ctor_get(x_927, 0); +lean_inc(x_928); +if (lean_is_exclusive(x_927)) { + lean_ctor_release(x_927, 0); + lean_ctor_release(x_927, 1); + x_929 = x_927; +} else { + lean_dec_ref(x_927); + x_929 = lean_box(0); +} +x_930 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_929)) { + x_931 = lean_alloc_ctor(0, 2, 0); +} else { + x_931 = x_929; +} +lean_ctor_set(x_931, 0, x_928); +lean_ctor_set(x_931, 1, x_930); +return x_931; +} +else +{ +lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; uint8_t x_936; +x_932 = lean_ctor_get(x_927, 0); +lean_inc(x_932); +x_933 = lean_ctor_get(x_927, 1); +lean_inc(x_933); +if (lean_is_exclusive(x_927)) { + lean_ctor_release(x_927, 0); + lean_ctor_release(x_927, 1); + x_934 = x_927; +} else { + lean_dec_ref(x_927); + x_934 = lean_box(0); +} +x_935 = lean_ctor_get(x_932, 1); +lean_inc(x_935); +x_936 = lean_nat_dec_eq(x_923, x_935); +lean_dec(x_923); +if (x_936 == 0) +{ +lean_object* x_937; +lean_dec(x_935); +if (lean_is_scalar(x_934)) { + x_937 = lean_alloc_ctor(1, 2, 0); +} else { + x_937 = x_934; +} +lean_ctor_set(x_937, 0, x_932); +lean_ctor_set(x_937, 1, x_933); +return x_937; +} +else +{ +lean_object* x_938; lean_object* x_939; +lean_dec(x_934); +lean_dec(x_933); +x_938 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4; +x_939 = l_Std_Internal_Parsec_String_pstring(x_938, x_932); +if (lean_obj_tag(x_939) == 0) +{ +lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; +lean_dec(x_935); +x_940 = lean_ctor_get(x_939, 0); +lean_inc(x_940); +if (lean_is_exclusive(x_939)) { + lean_ctor_release(x_939, 0); + lean_ctor_release(x_939, 1); + x_941 = x_939; +} else { + lean_dec_ref(x_939); + x_941 = lean_box(0); +} +x_942 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_941)) { + x_943 = lean_alloc_ctor(0, 2, 0); +} else { + x_943 = x_941; +} +lean_ctor_set(x_943, 0, x_940); +lean_ctor_set(x_943, 1, x_942); +return x_943; +} +else +{ +lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; uint8_t x_948; +x_944 = lean_ctor_get(x_939, 0); +lean_inc(x_944); +x_945 = lean_ctor_get(x_939, 1); +lean_inc(x_945); +if (lean_is_exclusive(x_939)) { + lean_ctor_release(x_939, 0); + lean_ctor_release(x_939, 1); + x_946 = x_939; +} else { + lean_dec_ref(x_939); + x_946 = lean_box(0); +} +x_947 = lean_ctor_get(x_944, 1); +lean_inc(x_947); +x_948 = lean_nat_dec_eq(x_935, x_947); +lean_dec(x_935); +if (x_948 == 0) +{ +lean_object* x_949; +lean_dec(x_947); +if (lean_is_scalar(x_946)) { + x_949 = lean_alloc_ctor(1, 2, 0); +} else { + x_949 = x_946; +} +lean_ctor_set(x_949, 0, x_944); +lean_ctor_set(x_949, 1, x_945); +return x_949; +} +else +{ +lean_object* x_950; lean_object* x_951; +lean_dec(x_946); +lean_dec(x_945); +x_950 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3; +x_951 = l_Std_Internal_Parsec_String_pstring(x_950, x_944); +if (lean_obj_tag(x_951) == 0) +{ +lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; +lean_dec(x_947); +x_952 = lean_ctor_get(x_951, 0); +lean_inc(x_952); +if (lean_is_exclusive(x_951)) { + lean_ctor_release(x_951, 0); + lean_ctor_release(x_951, 1); + x_953 = x_951; +} else { + lean_dec_ref(x_951); + x_953 = lean_box(0); +} +x_954 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_953)) { + x_955 = lean_alloc_ctor(0, 2, 0); +} else { + x_955 = x_953; +} +lean_ctor_set(x_955, 0, x_952); +lean_ctor_set(x_955, 1, x_954); +return x_955; +} +else +{ +lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; uint8_t x_960; +x_956 = lean_ctor_get(x_951, 0); +lean_inc(x_956); +x_957 = lean_ctor_get(x_951, 1); +lean_inc(x_957); +if (lean_is_exclusive(x_951)) { + lean_ctor_release(x_951, 0); + lean_ctor_release(x_951, 1); + x_958 = x_951; +} else { + lean_dec_ref(x_951); + x_958 = lean_box(0); +} +x_959 = lean_ctor_get(x_956, 1); +lean_inc(x_959); +x_960 = lean_nat_dec_eq(x_947, x_959); +lean_dec(x_947); +if (x_960 == 0) +{ +lean_object* x_961; +lean_dec(x_959); +if (lean_is_scalar(x_958)) { + x_961 = lean_alloc_ctor(1, 2, 0); +} else { + x_961 = x_958; +} +lean_ctor_set(x_961, 0, x_956); +lean_ctor_set(x_961, 1, x_957); +return x_961; +} +else +{ +lean_object* x_962; lean_object* x_963; +lean_dec(x_958); +lean_dec(x_957); +x_962 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2; +x_963 = l_Std_Internal_Parsec_String_pstring(x_962, x_956); +if (lean_obj_tag(x_963) == 0) +{ +lean_object* x_964; lean_object* x_965; lean_object* x_966; lean_object* x_967; +lean_dec(x_959); +x_964 = lean_ctor_get(x_963, 0); +lean_inc(x_964); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_965 = x_963; +} else { + lean_dec_ref(x_963); + x_965 = lean_box(0); +} +x_966 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_965)) { + x_967 = lean_alloc_ctor(0, 2, 0); +} else { + x_967 = x_965; +} +lean_ctor_set(x_967, 0, x_964); +lean_ctor_set(x_967, 1, x_966); +return x_967; +} +else +{ +lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; uint8_t x_972; +x_968 = lean_ctor_get(x_963, 0); +lean_inc(x_968); +x_969 = lean_ctor_get(x_963, 1); +lean_inc(x_969); +if (lean_is_exclusive(x_963)) { + lean_ctor_release(x_963, 0); + lean_ctor_release(x_963, 1); + x_970 = x_963; +} else { + lean_dec_ref(x_963); + x_970 = lean_box(0); +} +x_971 = lean_ctor_get(x_968, 1); +lean_inc(x_971); +x_972 = lean_nat_dec_eq(x_959, x_971); +lean_dec(x_971); +lean_dec(x_959); +if (x_972 == 0) +{ +lean_object* x_973; +if (lean_is_scalar(x_970)) { + x_973 = lean_alloc_ctor(1, 2, 0); +} else { + x_973 = x_970; +} +lean_ctor_set(x_973, 0, x_968); +lean_ctor_set(x_973, 1, x_969); +return x_973; +} +else +{ +lean_object* x_974; lean_object* x_975; +lean_dec(x_970); +lean_dec(x_969); +x_974 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1; +x_975 = l_Std_Internal_Parsec_String_pstring(x_974, x_968); +if (lean_obj_tag(x_975) == 0) +{ +lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; +x_976 = lean_ctor_get(x_975, 0); +lean_inc(x_976); +if (lean_is_exclusive(x_975)) { + lean_ctor_release(x_975, 0); + lean_ctor_release(x_975, 1); + x_977 = x_975; +} else { + lean_dec_ref(x_975); + x_977 = lean_box(0); +} +x_978 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_977)) { + x_979 = lean_alloc_ctor(0, 2, 0); +} else { + x_979 = x_977; +} +lean_ctor_set(x_979, 0, x_976); +lean_ctor_set(x_979, 1, x_978); +return x_979; +} +else +{ +lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +x_980 = lean_ctor_get(x_975, 0); +lean_inc(x_980); +x_981 = lean_ctor_get(x_975, 1); +lean_inc(x_981); +if (lean_is_exclusive(x_975)) { + lean_ctor_release(x_975, 0); + lean_ctor_release(x_975, 1); + x_982 = x_975; +} else { + lean_dec_ref(x_975); + x_982 = lean_box(0); +} +if (lean_is_scalar(x_982)) { + x_983 = lean_alloc_ctor(1, 2, 0); +} else { + x_983 = x_982; +} +lean_ctor_set(x_983, 0, x_980); +lean_ctor_set(x_983, 1, x_981); +return x_983; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthNarrow(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +lean_dec(x_40); +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +lean_dec(x_40); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +x_54 = lean_nat_dec_eq(x_40, x_53); +lean_dec(x_40); +if (x_54 == 0) +{ +lean_dec(x_53); +return x_43; +} +else +{ +lean_object* x_55; +lean_free_object(x_43); +lean_dec(x_52); +x_55 = l_Std_Internal_Parsec_String_pstring(x_29, x_51); +if (lean_obj_tag(x_55) == 0) +{ +uint8_t x_56; +lean_dec(x_53); +x_56 = !lean_is_exclusive(x_55); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_55, 1); +lean_dec(x_57); +x_58 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +lean_ctor_set(x_55, 1, x_58); +return x_55; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_55, 0); +lean_inc(x_59); +lean_dec(x_55); +x_60 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +else +{ +uint8_t x_62; +x_62 = !lean_is_exclusive(x_55); +if (x_62 == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_63 = lean_ctor_get(x_55, 0); +x_64 = lean_ctor_get(x_55, 1); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +x_66 = lean_nat_dec_eq(x_53, x_65); +lean_dec(x_53); +if (x_66 == 0) +{ +lean_dec(x_65); +return x_55; +} +else +{ +lean_object* x_67; +lean_free_object(x_55); +lean_dec(x_64); +x_67 = l_Std_Internal_Parsec_String_pstring(x_2, x_63); +if (lean_obj_tag(x_67) == 0) +{ +uint8_t x_68; +lean_dec(x_65); +x_68 = !lean_is_exclusive(x_67); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_67, 1); +lean_dec(x_69); +x_70 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +lean_ctor_set(x_67, 1, x_70); +return x_67; +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_67, 0); +lean_inc(x_71); +lean_dec(x_67); +x_72 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; +} +} +else +{ +uint8_t x_74; +x_74 = !lean_is_exclusive(x_67); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_75 = lean_ctor_get(x_67, 0); +x_76 = lean_ctor_get(x_67, 1); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +x_78 = lean_nat_dec_eq(x_65, x_77); +lean_dec(x_65); +if (x_78 == 0) +{ +lean_dec(x_77); +return x_67; +} +else +{ +lean_object* x_79; +lean_free_object(x_67); +lean_dec(x_76); +x_79 = l_Std_Internal_Parsec_String_pstring(x_2, x_75); +if (lean_obj_tag(x_79) == 0) +{ +uint8_t x_80; +lean_dec(x_77); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_79, 1); +lean_dec(x_81); +x_82 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +lean_ctor_set(x_79, 1, x_82); +return x_79; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_79, 0); +lean_inc(x_83); +lean_dec(x_79); +x_84 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +else +{ +uint8_t x_86; +x_86 = !lean_is_exclusive(x_79); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_87 = lean_ctor_get(x_79, 0); +x_88 = lean_ctor_get(x_79, 1); +x_89 = lean_ctor_get(x_87, 1); +lean_inc(x_89); +x_90 = lean_nat_dec_eq(x_77, x_89); +lean_dec(x_77); +if (x_90 == 0) +{ +lean_dec(x_89); +return x_79; +} +else +{ +lean_object* x_91; +lean_free_object(x_79); +lean_dec(x_88); +x_91 = l_Std_Internal_Parsec_String_pstring(x_42, x_87); +if (lean_obj_tag(x_91) == 0) +{ +uint8_t x_92; +lean_dec(x_89); +x_92 = !lean_is_exclusive(x_91); +if (x_92 == 0) +{ +lean_object* x_93; lean_object* x_94; +x_93 = lean_ctor_get(x_91, 1); +lean_dec(x_93); +x_94 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +lean_ctor_set(x_91, 1, x_94); +return x_91; +} +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_95 = lean_ctor_get(x_91, 0); +lean_inc(x_95); +lean_dec(x_91); +x_96 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +else +{ +uint8_t x_98; +x_98 = !lean_is_exclusive(x_91); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_99 = lean_ctor_get(x_91, 0); +x_100 = lean_ctor_get(x_91, 1); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +x_102 = lean_nat_dec_eq(x_89, x_101); +lean_dec(x_89); +if (x_102 == 0) +{ +lean_dec(x_101); +return x_91; +} +else +{ +lean_object* x_103; lean_object* x_104; +lean_free_object(x_91); +lean_dec(x_100); +x_103 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_104 = l_Std_Internal_Parsec_String_pstring(x_103, x_99); +if (lean_obj_tag(x_104) == 0) +{ +uint8_t x_105; +lean_dec(x_101); +x_105 = !lean_is_exclusive(x_104); +if (x_105 == 0) +{ +lean_object* x_106; lean_object* x_107; +x_106 = lean_ctor_get(x_104, 1); +lean_dec(x_106); +x_107 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +lean_ctor_set(x_104, 1, x_107); +return x_104; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_104, 0); +lean_inc(x_108); +lean_dec(x_104); +x_109 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +x_110 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; +} +} +else +{ +uint8_t x_111; +x_111 = !lean_is_exclusive(x_104); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_112 = lean_ctor_get(x_104, 0); +x_113 = lean_ctor_get(x_104, 1); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +x_115 = lean_nat_dec_eq(x_101, x_114); +lean_dec(x_101); +if (x_115 == 0) +{ +lean_dec(x_114); +return x_104; +} +else +{ +lean_object* x_116; lean_object* x_117; +lean_free_object(x_104); +lean_dec(x_113); +x_116 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_117 = l_Std_Internal_Parsec_String_pstring(x_116, x_112); +if (lean_obj_tag(x_117) == 0) +{ +uint8_t x_118; +lean_dec(x_114); +x_118 = !lean_is_exclusive(x_117); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; +x_119 = lean_ctor_get(x_117, 1); +lean_dec(x_119); +x_120 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +lean_ctor_set(x_117, 1, x_120); +return x_117; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_117, 0); +lean_inc(x_121); +lean_dec(x_117); +x_122 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +else +{ +uint8_t x_124; +x_124 = !lean_is_exclusive(x_117); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_125 = lean_ctor_get(x_117, 0); +x_126 = lean_ctor_get(x_117, 1); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +x_128 = lean_nat_dec_eq(x_114, x_127); +lean_dec(x_114); +if (x_128 == 0) +{ +lean_dec(x_127); +return x_117; +} +else +{ +lean_object* x_129; lean_object* x_130; +lean_free_object(x_117); +lean_dec(x_126); +x_129 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_130 = l_Std_Internal_Parsec_String_pstring(x_129, x_125); +if (lean_obj_tag(x_130) == 0) +{ +uint8_t x_131; +lean_dec(x_127); +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; +x_132 = lean_ctor_get(x_130, 1); +lean_dec(x_132); +x_133 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +lean_ctor_set(x_130, 1, x_133); +return x_130; +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_130, 0); +lean_inc(x_134); +lean_dec(x_130); +x_135 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +return x_136; +} +} +else +{ +uint8_t x_137; +x_137 = !lean_is_exclusive(x_130); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_138 = lean_ctor_get(x_130, 0); +x_139 = lean_ctor_get(x_130, 1); +x_140 = lean_ctor_get(x_138, 1); +lean_inc(x_140); +x_141 = lean_nat_dec_eq(x_127, x_140); +lean_dec(x_140); +lean_dec(x_127); +if (x_141 == 0) +{ +return x_130; +} +else +{ +lean_object* x_142; lean_object* x_143; +lean_free_object(x_130); +lean_dec(x_139); +x_142 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_143 = l_Std_Internal_Parsec_String_pstring(x_142, x_138); +if (lean_obj_tag(x_143) == 0) +{ +uint8_t x_144; +x_144 = !lean_is_exclusive(x_143); +if (x_144 == 0) +{ +lean_object* x_145; lean_object* x_146; +x_145 = lean_ctor_get(x_143, 1); +lean_dec(x_145); +x_146 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +lean_ctor_set(x_143, 1, x_146); +return x_143; +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; +x_147 = lean_ctor_get(x_143, 0); +lean_inc(x_147); +lean_dec(x_143); +x_148 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_147); +lean_ctor_set(x_149, 1, x_148); +return x_149; +} +} +else +{ +uint8_t x_150; +x_150 = !lean_is_exclusive(x_143); +if (x_150 == 0) +{ +return x_143; +} +else +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_151 = lean_ctor_get(x_143, 0); +x_152 = lean_ctor_get(x_143, 1); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_143); +x_153 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_153, 0, x_151); +lean_ctor_set(x_153, 1, x_152); +return x_153; +} +} +} +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; +x_154 = lean_ctor_get(x_130, 0); +x_155 = lean_ctor_get(x_130, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_130); +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +x_157 = lean_nat_dec_eq(x_127, x_156); +lean_dec(x_156); +lean_dec(x_127); +if (x_157 == 0) +{ +lean_object* x_158; +x_158 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_158, 0, x_154); +lean_ctor_set(x_158, 1, x_155); +return x_158; +} +else +{ +lean_object* x_159; lean_object* x_160; +lean_dec(x_155); +x_159 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_160 = l_Std_Internal_Parsec_String_pstring(x_159, x_154); +if (lean_obj_tag(x_160) == 0) +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_162 = x_160; +} else { + lean_dec_ref(x_160); + x_162 = lean_box(0); +} +x_163 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_162)) { + x_164 = lean_alloc_ctor(0, 2, 0); +} else { + x_164 = x_162; +} +lean_ctor_set(x_164, 0, x_161); +lean_ctor_set(x_164, 1, x_163); +return x_164; +} +else +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_165 = lean_ctor_get(x_160, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_160, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_160)) { + lean_ctor_release(x_160, 0); + lean_ctor_release(x_160, 1); + x_167 = x_160; +} else { + lean_dec_ref(x_160); + x_167 = lean_box(0); +} +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); +} else { + x_168 = x_167; +} +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; +} +} +} +} +} +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; +x_169 = lean_ctor_get(x_117, 0); +x_170 = lean_ctor_get(x_117, 1); +lean_inc(x_170); +lean_inc(x_169); +lean_dec(x_117); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +x_172 = lean_nat_dec_eq(x_114, x_171); +lean_dec(x_114); +if (x_172 == 0) +{ +lean_object* x_173; +lean_dec(x_171); +x_173 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_173, 0, x_169); +lean_ctor_set(x_173, 1, x_170); +return x_173; +} +else +{ +lean_object* x_174; lean_object* x_175; +lean_dec(x_170); +x_174 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_175 = l_Std_Internal_Parsec_String_pstring(x_174, x_169); +if (lean_obj_tag(x_175) == 0) +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_171); +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_177 = x_175; +} else { + lean_dec_ref(x_175); + x_177 = lean_box(0); +} +x_178 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_177)) { + x_179 = lean_alloc_ctor(0, 2, 0); +} else { + x_179 = x_177; +} +lean_ctor_set(x_179, 0, x_176); +lean_ctor_set(x_179, 1, x_178); +return x_179; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; +x_180 = lean_ctor_get(x_175, 0); +lean_inc(x_180); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_182 = x_175; +} else { + lean_dec_ref(x_175); + x_182 = lean_box(0); +} +x_183 = lean_ctor_get(x_180, 1); +lean_inc(x_183); +x_184 = lean_nat_dec_eq(x_171, x_183); +lean_dec(x_183); +lean_dec(x_171); +if (x_184 == 0) +{ +lean_object* x_185; +if (lean_is_scalar(x_182)) { + x_185 = lean_alloc_ctor(1, 2, 0); +} else { + x_185 = x_182; +} +lean_ctor_set(x_185, 0, x_180); +lean_ctor_set(x_185, 1, x_181); +return x_185; +} +else +{ +lean_object* x_186; lean_object* x_187; +lean_dec(x_182); +lean_dec(x_181); +x_186 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_187 = l_Std_Internal_Parsec_String_pstring(x_186, x_180); +if (lean_obj_tag(x_187) == 0) +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_189 = x_187; +} else { + lean_dec_ref(x_187); + x_189 = lean_box(0); +} +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_189)) { + x_191 = lean_alloc_ctor(0, 2, 0); +} else { + x_191 = x_189; +} +lean_ctor_set(x_191, 0, x_188); +lean_ctor_set(x_191, 1, x_190); +return x_191; +} +else +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_192 = lean_ctor_get(x_187, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_187, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_194 = x_187; +} else { + lean_dec_ref(x_187); + x_194 = lean_box(0); +} +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); +} else { + x_195 = x_194; +} +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; +x_196 = lean_ctor_get(x_104, 0); +x_197 = lean_ctor_get(x_104, 1); +lean_inc(x_197); +lean_inc(x_196); +lean_dec(x_104); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +x_199 = lean_nat_dec_eq(x_101, x_198); +lean_dec(x_101); +if (x_199 == 0) +{ +lean_object* x_200; +lean_dec(x_198); +x_200 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_200, 0, x_196); +lean_ctor_set(x_200, 1, x_197); +return x_200; +} +else +{ +lean_object* x_201; lean_object* x_202; +lean_dec(x_197); +x_201 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_202 = l_Std_Internal_Parsec_String_pstring(x_201, x_196); +if (lean_obj_tag(x_202) == 0) +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +lean_dec(x_198); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_204 = x_202; +} else { + lean_dec_ref(x_202); + x_204 = lean_box(0); +} +x_205 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_204)) { + x_206 = lean_alloc_ctor(0, 2, 0); +} else { + x_206 = x_204; +} +lean_ctor_set(x_206, 0, x_203); +lean_ctor_set(x_206, 1, x_205); +return x_206; +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; uint8_t x_211; +x_207 = lean_ctor_get(x_202, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_202, 1); +lean_inc(x_208); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_209 = x_202; +} else { + lean_dec_ref(x_202); + x_209 = lean_box(0); +} +x_210 = lean_ctor_get(x_207, 1); +lean_inc(x_210); +x_211 = lean_nat_dec_eq(x_198, x_210); +lean_dec(x_198); +if (x_211 == 0) +{ +lean_object* x_212; +lean_dec(x_210); +if (lean_is_scalar(x_209)) { + x_212 = lean_alloc_ctor(1, 2, 0); +} else { + x_212 = x_209; +} +lean_ctor_set(x_212, 0, x_207); +lean_ctor_set(x_212, 1, x_208); +return x_212; +} +else +{ +lean_object* x_213; lean_object* x_214; +lean_dec(x_209); +lean_dec(x_208); +x_213 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_214 = l_Std_Internal_Parsec_String_pstring(x_213, x_207); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_210); +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_216 = x_214; +} else { + lean_dec_ref(x_214); + x_216 = lean_box(0); +} +x_217 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_216)) { + x_218 = lean_alloc_ctor(0, 2, 0); +} else { + x_218 = x_216; +} +lean_ctor_set(x_218, 0, x_215); +lean_ctor_set(x_218, 1, x_217); +return x_218; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_219 = lean_ctor_get(x_214, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_214, 1); +lean_inc(x_220); +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_221 = x_214; +} else { + lean_dec_ref(x_214); + x_221 = lean_box(0); +} +x_222 = lean_ctor_get(x_219, 1); +lean_inc(x_222); +x_223 = lean_nat_dec_eq(x_210, x_222); +lean_dec(x_222); +lean_dec(x_210); +if (x_223 == 0) +{ +lean_object* x_224; +if (lean_is_scalar(x_221)) { + x_224 = lean_alloc_ctor(1, 2, 0); +} else { + x_224 = x_221; +} +lean_ctor_set(x_224, 0, x_219); +lean_ctor_set(x_224, 1, x_220); +return x_224; +} +else +{ +lean_object* x_225; lean_object* x_226; +lean_dec(x_221); +lean_dec(x_220); +x_225 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_226 = l_Std_Internal_Parsec_String_pstring(x_225, x_219); +if (lean_obj_tag(x_226) == 0) +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_228 = x_226; +} else { + lean_dec_ref(x_226); + x_228 = lean_box(0); +} +x_229 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_228)) { + x_230 = lean_alloc_ctor(0, 2, 0); +} else { + x_230 = x_228; +} +lean_ctor_set(x_230, 0, x_227); +lean_ctor_set(x_230, 1, x_229); +return x_230; +} +else +{ +lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_231 = lean_ctor_get(x_226, 0); +lean_inc(x_231); +x_232 = lean_ctor_get(x_226, 1); +lean_inc(x_232); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_233 = x_226; +} else { + lean_dec_ref(x_226); + x_233 = lean_box(0); +} +if (lean_is_scalar(x_233)) { + x_234 = lean_alloc_ctor(1, 2, 0); +} else { + x_234 = x_233; +} +lean_ctor_set(x_234, 0, x_231); +lean_ctor_set(x_234, 1, x_232); +return x_234; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; +x_235 = lean_ctor_get(x_91, 0); +x_236 = lean_ctor_get(x_91, 1); +lean_inc(x_236); +lean_inc(x_235); +lean_dec(x_91); +x_237 = lean_ctor_get(x_235, 1); +lean_inc(x_237); +x_238 = lean_nat_dec_eq(x_89, x_237); +lean_dec(x_89); +if (x_238 == 0) +{ +lean_object* x_239; +lean_dec(x_237); +x_239 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_239, 0, x_235); +lean_ctor_set(x_239, 1, x_236); +return x_239; +} +else +{ +lean_object* x_240; lean_object* x_241; +lean_dec(x_236); +x_240 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_241 = l_Std_Internal_Parsec_String_pstring(x_240, x_235); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_dec(x_237); +x_242 = lean_ctor_get(x_241, 0); +lean_inc(x_242); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + lean_ctor_release(x_241, 1); + x_243 = x_241; +} else { + lean_dec_ref(x_241); + x_243 = lean_box(0); +} +x_244 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_243)) { + x_245 = lean_alloc_ctor(0, 2, 0); +} else { + x_245 = x_243; +} +lean_ctor_set(x_245, 0, x_242); +lean_ctor_set(x_245, 1, x_244); +return x_245; +} +else +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; uint8_t x_250; +x_246 = lean_ctor_get(x_241, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_241, 1); +lean_inc(x_247); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + lean_ctor_release(x_241, 1); + x_248 = x_241; +} else { + lean_dec_ref(x_241); + x_248 = lean_box(0); +} +x_249 = lean_ctor_get(x_246, 1); +lean_inc(x_249); +x_250 = lean_nat_dec_eq(x_237, x_249); +lean_dec(x_237); +if (x_250 == 0) +{ +lean_object* x_251; +lean_dec(x_249); +if (lean_is_scalar(x_248)) { + x_251 = lean_alloc_ctor(1, 2, 0); +} else { + x_251 = x_248; +} +lean_ctor_set(x_251, 0, x_246); +lean_ctor_set(x_251, 1, x_247); +return x_251; +} +else +{ +lean_object* x_252; lean_object* x_253; +lean_dec(x_248); +lean_dec(x_247); +x_252 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_253 = l_Std_Internal_Parsec_String_pstring(x_252, x_246); +if (lean_obj_tag(x_253) == 0) +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +lean_dec(x_249); +x_254 = lean_ctor_get(x_253, 0); +lean_inc(x_254); +if (lean_is_exclusive(x_253)) { + lean_ctor_release(x_253, 0); + lean_ctor_release(x_253, 1); + x_255 = x_253; +} else { + lean_dec_ref(x_253); + x_255 = lean_box(0); +} +x_256 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_255)) { + x_257 = lean_alloc_ctor(0, 2, 0); +} else { + x_257 = x_255; +} +lean_ctor_set(x_257, 0, x_254); +lean_ctor_set(x_257, 1, x_256); +return x_257; +} +else +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; uint8_t x_262; +x_258 = lean_ctor_get(x_253, 0); +lean_inc(x_258); +x_259 = lean_ctor_get(x_253, 1); +lean_inc(x_259); +if (lean_is_exclusive(x_253)) { + lean_ctor_release(x_253, 0); + lean_ctor_release(x_253, 1); + x_260 = x_253; +} else { + lean_dec_ref(x_253); + x_260 = lean_box(0); +} +x_261 = lean_ctor_get(x_258, 1); +lean_inc(x_261); +x_262 = lean_nat_dec_eq(x_249, x_261); +lean_dec(x_249); +if (x_262 == 0) +{ +lean_object* x_263; +lean_dec(x_261); +if (lean_is_scalar(x_260)) { + x_263 = lean_alloc_ctor(1, 2, 0); +} else { + x_263 = x_260; +} +lean_ctor_set(x_263, 0, x_258); +lean_ctor_set(x_263, 1, x_259); +return x_263; +} +else +{ +lean_object* x_264; lean_object* x_265; +lean_dec(x_260); +lean_dec(x_259); +x_264 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_265 = l_Std_Internal_Parsec_String_pstring(x_264, x_258); +if (lean_obj_tag(x_265) == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_261); +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_267 = x_265; +} else { + lean_dec_ref(x_265); + x_267 = lean_box(0); +} +x_268 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_267)) { + x_269 = lean_alloc_ctor(0, 2, 0); +} else { + x_269 = x_267; +} +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set(x_269, 1, x_268); +return x_269; +} +else +{ +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; uint8_t x_274; +x_270 = lean_ctor_get(x_265, 0); +lean_inc(x_270); +x_271 = lean_ctor_get(x_265, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_272 = x_265; +} else { + lean_dec_ref(x_265); + x_272 = lean_box(0); +} +x_273 = lean_ctor_get(x_270, 1); +lean_inc(x_273); +x_274 = lean_nat_dec_eq(x_261, x_273); +lean_dec(x_273); +lean_dec(x_261); +if (x_274 == 0) +{ +lean_object* x_275; +if (lean_is_scalar(x_272)) { + x_275 = lean_alloc_ctor(1, 2, 0); +} else { + x_275 = x_272; +} +lean_ctor_set(x_275, 0, x_270); +lean_ctor_set(x_275, 1, x_271); +return x_275; +} +else +{ +lean_object* x_276; lean_object* x_277; +lean_dec(x_272); +lean_dec(x_271); +x_276 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_277 = l_Std_Internal_Parsec_String_pstring(x_276, x_270); +if (lean_obj_tag(x_277) == 0) +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; +x_278 = lean_ctor_get(x_277, 0); +lean_inc(x_278); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_279 = x_277; +} else { + lean_dec_ref(x_277); + x_279 = lean_box(0); +} +x_280 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_279)) { + x_281 = lean_alloc_ctor(0, 2, 0); +} else { + x_281 = x_279; +} +lean_ctor_set(x_281, 0, x_278); +lean_ctor_set(x_281, 1, x_280); +return x_281; +} +else +{ +lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_282 = lean_ctor_get(x_277, 0); +lean_inc(x_282); +x_283 = lean_ctor_get(x_277, 1); +lean_inc(x_283); +if (lean_is_exclusive(x_277)) { + lean_ctor_release(x_277, 0); + lean_ctor_release(x_277, 1); + x_284 = x_277; +} else { + lean_dec_ref(x_277); + x_284 = lean_box(0); +} +if (lean_is_scalar(x_284)) { + x_285 = lean_alloc_ctor(1, 2, 0); +} else { + x_285 = x_284; +} +lean_ctor_set(x_285, 0, x_282); +lean_ctor_set(x_285, 1, x_283); +return x_285; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_286; lean_object* x_287; lean_object* x_288; uint8_t x_289; +x_286 = lean_ctor_get(x_79, 0); +x_287 = lean_ctor_get(x_79, 1); +lean_inc(x_287); +lean_inc(x_286); +lean_dec(x_79); +x_288 = lean_ctor_get(x_286, 1); +lean_inc(x_288); +x_289 = lean_nat_dec_eq(x_77, x_288); +lean_dec(x_77); +if (x_289 == 0) +{ +lean_object* x_290; +lean_dec(x_288); +x_290 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_290, 0, x_286); +lean_ctor_set(x_290, 1, x_287); +return x_290; +} +else +{ +lean_object* x_291; +lean_dec(x_287); +x_291 = l_Std_Internal_Parsec_String_pstring(x_42, x_286); +if (lean_obj_tag(x_291) == 0) +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; +lean_dec(x_288); +x_292 = lean_ctor_get(x_291, 0); +lean_inc(x_292); +if (lean_is_exclusive(x_291)) { + lean_ctor_release(x_291, 0); + lean_ctor_release(x_291, 1); + x_293 = x_291; +} else { + lean_dec_ref(x_291); + x_293 = lean_box(0); +} +x_294 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_293)) { + x_295 = lean_alloc_ctor(0, 2, 0); +} else { + x_295 = x_293; +} +lean_ctor_set(x_295, 0, x_292); +lean_ctor_set(x_295, 1, x_294); +return x_295; +} +else +{ +lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; uint8_t x_300; +x_296 = lean_ctor_get(x_291, 0); +lean_inc(x_296); +x_297 = lean_ctor_get(x_291, 1); +lean_inc(x_297); +if (lean_is_exclusive(x_291)) { + lean_ctor_release(x_291, 0); + lean_ctor_release(x_291, 1); + x_298 = x_291; +} else { + lean_dec_ref(x_291); + x_298 = lean_box(0); +} +x_299 = lean_ctor_get(x_296, 1); +lean_inc(x_299); +x_300 = lean_nat_dec_eq(x_288, x_299); +lean_dec(x_288); +if (x_300 == 0) +{ +lean_object* x_301; +lean_dec(x_299); +if (lean_is_scalar(x_298)) { + x_301 = lean_alloc_ctor(1, 2, 0); +} else { + x_301 = x_298; +} +lean_ctor_set(x_301, 0, x_296); +lean_ctor_set(x_301, 1, x_297); +return x_301; +} +else +{ +lean_object* x_302; lean_object* x_303; +lean_dec(x_298); +lean_dec(x_297); +x_302 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_303 = l_Std_Internal_Parsec_String_pstring(x_302, x_296); +if (lean_obj_tag(x_303) == 0) +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_299); +x_304 = lean_ctor_get(x_303, 0); +lean_inc(x_304); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_305 = x_303; +} else { + lean_dec_ref(x_303); + x_305 = lean_box(0); +} +x_306 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_305)) { + x_307 = lean_alloc_ctor(0, 2, 0); +} else { + x_307 = x_305; +} +lean_ctor_set(x_307, 0, x_304); +lean_ctor_set(x_307, 1, x_306); +return x_307; +} +else +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; uint8_t x_312; +x_308 = lean_ctor_get(x_303, 0); +lean_inc(x_308); +x_309 = lean_ctor_get(x_303, 1); +lean_inc(x_309); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_310 = x_303; +} else { + lean_dec_ref(x_303); + x_310 = lean_box(0); +} +x_311 = lean_ctor_get(x_308, 1); +lean_inc(x_311); +x_312 = lean_nat_dec_eq(x_299, x_311); +lean_dec(x_299); +if (x_312 == 0) +{ +lean_object* x_313; +lean_dec(x_311); +if (lean_is_scalar(x_310)) { + x_313 = lean_alloc_ctor(1, 2, 0); +} else { + x_313 = x_310; +} +lean_ctor_set(x_313, 0, x_308); +lean_ctor_set(x_313, 1, x_309); +return x_313; +} +else +{ +lean_object* x_314; lean_object* x_315; +lean_dec(x_310); +lean_dec(x_309); +x_314 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_315 = l_Std_Internal_Parsec_String_pstring(x_314, x_308); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; +lean_dec(x_311); +x_316 = lean_ctor_get(x_315, 0); +lean_inc(x_316); +if (lean_is_exclusive(x_315)) { + lean_ctor_release(x_315, 0); + lean_ctor_release(x_315, 1); + x_317 = x_315; +} else { + lean_dec_ref(x_315); + x_317 = lean_box(0); +} +x_318 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_317)) { + x_319 = lean_alloc_ctor(0, 2, 0); +} else { + x_319 = x_317; +} +lean_ctor_set(x_319, 0, x_316); +lean_ctor_set(x_319, 1, x_318); +return x_319; +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; uint8_t x_324; +x_320 = lean_ctor_get(x_315, 0); +lean_inc(x_320); +x_321 = lean_ctor_get(x_315, 1); +lean_inc(x_321); +if (lean_is_exclusive(x_315)) { + lean_ctor_release(x_315, 0); + lean_ctor_release(x_315, 1); + x_322 = x_315; +} else { + lean_dec_ref(x_315); + x_322 = lean_box(0); +} +x_323 = lean_ctor_get(x_320, 1); +lean_inc(x_323); +x_324 = lean_nat_dec_eq(x_311, x_323); +lean_dec(x_311); +if (x_324 == 0) +{ +lean_object* x_325; +lean_dec(x_323); +if (lean_is_scalar(x_322)) { + x_325 = lean_alloc_ctor(1, 2, 0); +} else { + x_325 = x_322; +} +lean_ctor_set(x_325, 0, x_320); +lean_ctor_set(x_325, 1, x_321); +return x_325; +} +else +{ +lean_object* x_326; lean_object* x_327; +lean_dec(x_322); +lean_dec(x_321); +x_326 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_327 = l_Std_Internal_Parsec_String_pstring(x_326, x_320); +if (lean_obj_tag(x_327) == 0) +{ +lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +lean_dec(x_323); +x_328 = lean_ctor_get(x_327, 0); +lean_inc(x_328); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + lean_ctor_release(x_327, 1); + x_329 = x_327; +} else { + lean_dec_ref(x_327); + x_329 = lean_box(0); +} +x_330 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_329)) { + x_331 = lean_alloc_ctor(0, 2, 0); +} else { + x_331 = x_329; +} +lean_ctor_set(x_331, 0, x_328); +lean_ctor_set(x_331, 1, x_330); +return x_331; +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; uint8_t x_336; +x_332 = lean_ctor_get(x_327, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_327, 1); +lean_inc(x_333); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + lean_ctor_release(x_327, 1); + x_334 = x_327; +} else { + lean_dec_ref(x_327); + x_334 = lean_box(0); +} +x_335 = lean_ctor_get(x_332, 1); +lean_inc(x_335); +x_336 = lean_nat_dec_eq(x_323, x_335); +lean_dec(x_335); +lean_dec(x_323); +if (x_336 == 0) +{ +lean_object* x_337; +if (lean_is_scalar(x_334)) { + x_337 = lean_alloc_ctor(1, 2, 0); +} else { + x_337 = x_334; +} +lean_ctor_set(x_337, 0, x_332); +lean_ctor_set(x_337, 1, x_333); +return x_337; +} +else +{ +lean_object* x_338; lean_object* x_339; +lean_dec(x_334); +lean_dec(x_333); +x_338 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_339 = l_Std_Internal_Parsec_String_pstring(x_338, x_332); +if (lean_obj_tag(x_339) == 0) +{ +lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; +x_340 = lean_ctor_get(x_339, 0); +lean_inc(x_340); +if (lean_is_exclusive(x_339)) { + lean_ctor_release(x_339, 0); + lean_ctor_release(x_339, 1); + x_341 = x_339; +} else { + lean_dec_ref(x_339); + x_341 = lean_box(0); +} +x_342 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_341)) { + x_343 = lean_alloc_ctor(0, 2, 0); +} else { + x_343 = x_341; +} +lean_ctor_set(x_343, 0, x_340); +lean_ctor_set(x_343, 1, x_342); +return x_343; +} +else +{ +lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +x_344 = lean_ctor_get(x_339, 0); +lean_inc(x_344); +x_345 = lean_ctor_get(x_339, 1); +lean_inc(x_345); +if (lean_is_exclusive(x_339)) { + lean_ctor_release(x_339, 0); + lean_ctor_release(x_339, 1); + x_346 = x_339; +} else { + lean_dec_ref(x_339); + x_346 = lean_box(0); +} +if (lean_is_scalar(x_346)) { + x_347 = lean_alloc_ctor(1, 2, 0); +} else { + x_347 = x_346; +} +lean_ctor_set(x_347, 0, x_344); +lean_ctor_set(x_347, 1, x_345); +return x_347; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_348; lean_object* x_349; lean_object* x_350; uint8_t x_351; +x_348 = lean_ctor_get(x_67, 0); +x_349 = lean_ctor_get(x_67, 1); +lean_inc(x_349); +lean_inc(x_348); +lean_dec(x_67); +x_350 = lean_ctor_get(x_348, 1); +lean_inc(x_350); +x_351 = lean_nat_dec_eq(x_65, x_350); +lean_dec(x_65); +if (x_351 == 0) +{ +lean_object* x_352; +lean_dec(x_350); +x_352 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_352, 0, x_348); +lean_ctor_set(x_352, 1, x_349); +return x_352; +} +else +{ +lean_object* x_353; +lean_dec(x_349); +x_353 = l_Std_Internal_Parsec_String_pstring(x_2, x_348); +if (lean_obj_tag(x_353) == 0) +{ +lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; +lean_dec(x_350); +x_354 = lean_ctor_get(x_353, 0); +lean_inc(x_354); +if (lean_is_exclusive(x_353)) { + lean_ctor_release(x_353, 0); + lean_ctor_release(x_353, 1); + x_355 = x_353; +} else { + lean_dec_ref(x_353); + x_355 = lean_box(0); +} +x_356 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_355)) { + x_357 = lean_alloc_ctor(0, 2, 0); +} else { + x_357 = x_355; +} +lean_ctor_set(x_357, 0, x_354); +lean_ctor_set(x_357, 1, x_356); +return x_357; +} +else +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; +x_358 = lean_ctor_get(x_353, 0); +lean_inc(x_358); +x_359 = lean_ctor_get(x_353, 1); +lean_inc(x_359); +if (lean_is_exclusive(x_353)) { + lean_ctor_release(x_353, 0); + lean_ctor_release(x_353, 1); + x_360 = x_353; +} else { + lean_dec_ref(x_353); + x_360 = lean_box(0); +} +x_361 = lean_ctor_get(x_358, 1); +lean_inc(x_361); +x_362 = lean_nat_dec_eq(x_350, x_361); +lean_dec(x_350); +if (x_362 == 0) +{ +lean_object* x_363; +lean_dec(x_361); +if (lean_is_scalar(x_360)) { + x_363 = lean_alloc_ctor(1, 2, 0); +} else { + x_363 = x_360; +} +lean_ctor_set(x_363, 0, x_358); +lean_ctor_set(x_363, 1, x_359); +return x_363; +} +else +{ +lean_object* x_364; +lean_dec(x_360); +lean_dec(x_359); +x_364 = l_Std_Internal_Parsec_String_pstring(x_42, x_358); +if (lean_obj_tag(x_364) == 0) +{ +lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_361); +x_365 = lean_ctor_get(x_364, 0); +lean_inc(x_365); +if (lean_is_exclusive(x_364)) { + lean_ctor_release(x_364, 0); + lean_ctor_release(x_364, 1); + x_366 = x_364; +} else { + lean_dec_ref(x_364); + x_366 = lean_box(0); +} +x_367 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_366)) { + x_368 = lean_alloc_ctor(0, 2, 0); +} else { + x_368 = x_366; +} +lean_ctor_set(x_368, 0, x_365); +lean_ctor_set(x_368, 1, x_367); +return x_368; +} +else +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; uint8_t x_373; +x_369 = lean_ctor_get(x_364, 0); +lean_inc(x_369); +x_370 = lean_ctor_get(x_364, 1); +lean_inc(x_370); +if (lean_is_exclusive(x_364)) { + lean_ctor_release(x_364, 0); + lean_ctor_release(x_364, 1); + x_371 = x_364; +} else { + lean_dec_ref(x_364); + x_371 = lean_box(0); +} +x_372 = lean_ctor_get(x_369, 1); +lean_inc(x_372); +x_373 = lean_nat_dec_eq(x_361, x_372); +lean_dec(x_361); +if (x_373 == 0) +{ +lean_object* x_374; +lean_dec(x_372); +if (lean_is_scalar(x_371)) { + x_374 = lean_alloc_ctor(1, 2, 0); +} else { + x_374 = x_371; +} +lean_ctor_set(x_374, 0, x_369); +lean_ctor_set(x_374, 1, x_370); +return x_374; +} +else +{ +lean_object* x_375; lean_object* x_376; +lean_dec(x_371); +lean_dec(x_370); +x_375 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_376 = l_Std_Internal_Parsec_String_pstring(x_375, x_369); +if (lean_obj_tag(x_376) == 0) +{ +lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; +lean_dec(x_372); +x_377 = lean_ctor_get(x_376, 0); +lean_inc(x_377); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_378 = x_376; +} else { + lean_dec_ref(x_376); + x_378 = lean_box(0); +} +x_379 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_378)) { + x_380 = lean_alloc_ctor(0, 2, 0); +} else { + x_380 = x_378; +} +lean_ctor_set(x_380, 0, x_377); +lean_ctor_set(x_380, 1, x_379); +return x_380; +} +else +{ +lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; uint8_t x_385; +x_381 = lean_ctor_get(x_376, 0); +lean_inc(x_381); +x_382 = lean_ctor_get(x_376, 1); +lean_inc(x_382); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_383 = x_376; +} else { + lean_dec_ref(x_376); + x_383 = lean_box(0); +} +x_384 = lean_ctor_get(x_381, 1); +lean_inc(x_384); +x_385 = lean_nat_dec_eq(x_372, x_384); +lean_dec(x_372); +if (x_385 == 0) +{ +lean_object* x_386; +lean_dec(x_384); +if (lean_is_scalar(x_383)) { + x_386 = lean_alloc_ctor(1, 2, 0); +} else { + x_386 = x_383; +} +lean_ctor_set(x_386, 0, x_381); +lean_ctor_set(x_386, 1, x_382); +return x_386; +} +else +{ +lean_object* x_387; lean_object* x_388; +lean_dec(x_383); +lean_dec(x_382); +x_387 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_388 = l_Std_Internal_Parsec_String_pstring(x_387, x_381); +if (lean_obj_tag(x_388) == 0) +{ +lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; +lean_dec(x_384); +x_389 = lean_ctor_get(x_388, 0); +lean_inc(x_389); +if (lean_is_exclusive(x_388)) { + lean_ctor_release(x_388, 0); + lean_ctor_release(x_388, 1); + x_390 = x_388; +} else { + lean_dec_ref(x_388); + x_390 = lean_box(0); +} +x_391 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_390)) { + x_392 = lean_alloc_ctor(0, 2, 0); +} else { + x_392 = x_390; +} +lean_ctor_set(x_392, 0, x_389); +lean_ctor_set(x_392, 1, x_391); +return x_392; +} +else +{ +lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; uint8_t x_397; +x_393 = lean_ctor_get(x_388, 0); +lean_inc(x_393); +x_394 = lean_ctor_get(x_388, 1); +lean_inc(x_394); +if (lean_is_exclusive(x_388)) { + lean_ctor_release(x_388, 0); + lean_ctor_release(x_388, 1); + x_395 = x_388; +} else { + lean_dec_ref(x_388); + x_395 = lean_box(0); +} +x_396 = lean_ctor_get(x_393, 1); +lean_inc(x_396); +x_397 = lean_nat_dec_eq(x_384, x_396); +lean_dec(x_384); +if (x_397 == 0) +{ +lean_object* x_398; +lean_dec(x_396); +if (lean_is_scalar(x_395)) { + x_398 = lean_alloc_ctor(1, 2, 0); +} else { + x_398 = x_395; +} +lean_ctor_set(x_398, 0, x_393); +lean_ctor_set(x_398, 1, x_394); +return x_398; +} +else +{ +lean_object* x_399; lean_object* x_400; +lean_dec(x_395); +lean_dec(x_394); +x_399 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_400 = l_Std_Internal_Parsec_String_pstring(x_399, x_393); +if (lean_obj_tag(x_400) == 0) +{ +lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; +lean_dec(x_396); +x_401 = lean_ctor_get(x_400, 0); +lean_inc(x_401); +if (lean_is_exclusive(x_400)) { + lean_ctor_release(x_400, 0); + lean_ctor_release(x_400, 1); + x_402 = x_400; +} else { + lean_dec_ref(x_400); + x_402 = lean_box(0); +} +x_403 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_402)) { + x_404 = lean_alloc_ctor(0, 2, 0); +} else { + x_404 = x_402; +} +lean_ctor_set(x_404, 0, x_401); +lean_ctor_set(x_404, 1, x_403); +return x_404; +} +else +{ +lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; uint8_t x_409; +x_405 = lean_ctor_get(x_400, 0); +lean_inc(x_405); +x_406 = lean_ctor_get(x_400, 1); +lean_inc(x_406); +if (lean_is_exclusive(x_400)) { + lean_ctor_release(x_400, 0); + lean_ctor_release(x_400, 1); + x_407 = x_400; +} else { + lean_dec_ref(x_400); + x_407 = lean_box(0); +} +x_408 = lean_ctor_get(x_405, 1); +lean_inc(x_408); +x_409 = lean_nat_dec_eq(x_396, x_408); +lean_dec(x_408); +lean_dec(x_396); +if (x_409 == 0) +{ +lean_object* x_410; +if (lean_is_scalar(x_407)) { + x_410 = lean_alloc_ctor(1, 2, 0); +} else { + x_410 = x_407; +} +lean_ctor_set(x_410, 0, x_405); +lean_ctor_set(x_410, 1, x_406); +return x_410; +} +else +{ +lean_object* x_411; lean_object* x_412; +lean_dec(x_407); +lean_dec(x_406); +x_411 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_412 = l_Std_Internal_Parsec_String_pstring(x_411, x_405); +if (lean_obj_tag(x_412) == 0) +{ +lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; +x_413 = lean_ctor_get(x_412, 0); +lean_inc(x_413); +if (lean_is_exclusive(x_412)) { + lean_ctor_release(x_412, 0); + lean_ctor_release(x_412, 1); + x_414 = x_412; +} else { + lean_dec_ref(x_412); + x_414 = lean_box(0); +} +x_415 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_414)) { + x_416 = lean_alloc_ctor(0, 2, 0); +} else { + x_416 = x_414; +} +lean_ctor_set(x_416, 0, x_413); +lean_ctor_set(x_416, 1, x_415); +return x_416; +} +else +{ +lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_417 = lean_ctor_get(x_412, 0); +lean_inc(x_417); +x_418 = lean_ctor_get(x_412, 1); +lean_inc(x_418); +if (lean_is_exclusive(x_412)) { + lean_ctor_release(x_412, 0); + lean_ctor_release(x_412, 1); + x_419 = x_412; +} else { + lean_dec_ref(x_412); + x_419 = lean_box(0); +} +if (lean_is_scalar(x_419)) { + x_420 = lean_alloc_ctor(1, 2, 0); +} else { + x_420 = x_419; +} +lean_ctor_set(x_420, 0, x_417); +lean_ctor_set(x_420, 1, x_418); +return x_420; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_421; lean_object* x_422; lean_object* x_423; uint8_t x_424; +x_421 = lean_ctor_get(x_55, 0); +x_422 = lean_ctor_get(x_55, 1); +lean_inc(x_422); +lean_inc(x_421); +lean_dec(x_55); +x_423 = lean_ctor_get(x_421, 1); +lean_inc(x_423); +x_424 = lean_nat_dec_eq(x_53, x_423); +lean_dec(x_53); +if (x_424 == 0) +{ +lean_object* x_425; +lean_dec(x_423); +x_425 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_425, 0, x_421); +lean_ctor_set(x_425, 1, x_422); +return x_425; +} +else +{ +lean_object* x_426; +lean_dec(x_422); +x_426 = l_Std_Internal_Parsec_String_pstring(x_2, x_421); +if (lean_obj_tag(x_426) == 0) +{ +lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +lean_dec(x_423); +x_427 = lean_ctor_get(x_426, 0); +lean_inc(x_427); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + x_428 = x_426; +} else { + lean_dec_ref(x_426); + x_428 = lean_box(0); +} +x_429 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_428)) { + x_430 = lean_alloc_ctor(0, 2, 0); +} else { + x_430 = x_428; +} +lean_ctor_set(x_430, 0, x_427); +lean_ctor_set(x_430, 1, x_429); +return x_430; +} +else +{ +lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; uint8_t x_435; +x_431 = lean_ctor_get(x_426, 0); +lean_inc(x_431); +x_432 = lean_ctor_get(x_426, 1); +lean_inc(x_432); +if (lean_is_exclusive(x_426)) { + lean_ctor_release(x_426, 0); + lean_ctor_release(x_426, 1); + x_433 = x_426; +} else { + lean_dec_ref(x_426); + x_433 = lean_box(0); +} +x_434 = lean_ctor_get(x_431, 1); +lean_inc(x_434); +x_435 = lean_nat_dec_eq(x_423, x_434); +lean_dec(x_423); +if (x_435 == 0) +{ +lean_object* x_436; +lean_dec(x_434); +if (lean_is_scalar(x_433)) { + x_436 = lean_alloc_ctor(1, 2, 0); +} else { + x_436 = x_433; +} +lean_ctor_set(x_436, 0, x_431); +lean_ctor_set(x_436, 1, x_432); +return x_436; +} +else +{ +lean_object* x_437; +lean_dec(x_433); +lean_dec(x_432); +x_437 = l_Std_Internal_Parsec_String_pstring(x_2, x_431); +if (lean_obj_tag(x_437) == 0) +{ +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; +lean_dec(x_434); +x_438 = lean_ctor_get(x_437, 0); +lean_inc(x_438); +if (lean_is_exclusive(x_437)) { + lean_ctor_release(x_437, 0); + lean_ctor_release(x_437, 1); + x_439 = x_437; +} else { + lean_dec_ref(x_437); + x_439 = lean_box(0); +} +x_440 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_439)) { + x_441 = lean_alloc_ctor(0, 2, 0); +} else { + x_441 = x_439; +} +lean_ctor_set(x_441, 0, x_438); +lean_ctor_set(x_441, 1, x_440); +return x_441; +} +else +{ +lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; uint8_t x_446; +x_442 = lean_ctor_get(x_437, 0); +lean_inc(x_442); +x_443 = lean_ctor_get(x_437, 1); +lean_inc(x_443); +if (lean_is_exclusive(x_437)) { + lean_ctor_release(x_437, 0); + lean_ctor_release(x_437, 1); + x_444 = x_437; +} else { + lean_dec_ref(x_437); + x_444 = lean_box(0); +} +x_445 = lean_ctor_get(x_442, 1); +lean_inc(x_445); +x_446 = lean_nat_dec_eq(x_434, x_445); +lean_dec(x_434); +if (x_446 == 0) +{ +lean_object* x_447; +lean_dec(x_445); +if (lean_is_scalar(x_444)) { + x_447 = lean_alloc_ctor(1, 2, 0); +} else { + x_447 = x_444; +} +lean_ctor_set(x_447, 0, x_442); +lean_ctor_set(x_447, 1, x_443); +return x_447; +} +else +{ +lean_object* x_448; +lean_dec(x_444); +lean_dec(x_443); +x_448 = l_Std_Internal_Parsec_String_pstring(x_42, x_442); +if (lean_obj_tag(x_448) == 0) +{ +lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; +lean_dec(x_445); +x_449 = lean_ctor_get(x_448, 0); +lean_inc(x_449); +if (lean_is_exclusive(x_448)) { + lean_ctor_release(x_448, 0); + lean_ctor_release(x_448, 1); + x_450 = x_448; +} else { + lean_dec_ref(x_448); + x_450 = lean_box(0); +} +x_451 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_450)) { + x_452 = lean_alloc_ctor(0, 2, 0); +} else { + x_452 = x_450; +} +lean_ctor_set(x_452, 0, x_449); +lean_ctor_set(x_452, 1, x_451); +return x_452; +} +else +{ +lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; uint8_t x_457; +x_453 = lean_ctor_get(x_448, 0); +lean_inc(x_453); +x_454 = lean_ctor_get(x_448, 1); +lean_inc(x_454); +if (lean_is_exclusive(x_448)) { + lean_ctor_release(x_448, 0); + lean_ctor_release(x_448, 1); + x_455 = x_448; +} else { + lean_dec_ref(x_448); + x_455 = lean_box(0); +} +x_456 = lean_ctor_get(x_453, 1); +lean_inc(x_456); +x_457 = lean_nat_dec_eq(x_445, x_456); +lean_dec(x_445); +if (x_457 == 0) +{ +lean_object* x_458; +lean_dec(x_456); +if (lean_is_scalar(x_455)) { + x_458 = lean_alloc_ctor(1, 2, 0); +} else { + x_458 = x_455; +} +lean_ctor_set(x_458, 0, x_453); +lean_ctor_set(x_458, 1, x_454); +return x_458; +} +else +{ +lean_object* x_459; lean_object* x_460; +lean_dec(x_455); +lean_dec(x_454); +x_459 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_460 = l_Std_Internal_Parsec_String_pstring(x_459, x_453); +if (lean_obj_tag(x_460) == 0) +{ +lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; +lean_dec(x_456); +x_461 = lean_ctor_get(x_460, 0); +lean_inc(x_461); +if (lean_is_exclusive(x_460)) { + lean_ctor_release(x_460, 0); + lean_ctor_release(x_460, 1); + x_462 = x_460; +} else { + lean_dec_ref(x_460); + x_462 = lean_box(0); +} +x_463 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_462)) { + x_464 = lean_alloc_ctor(0, 2, 0); +} else { + x_464 = x_462; +} +lean_ctor_set(x_464, 0, x_461); +lean_ctor_set(x_464, 1, x_463); +return x_464; +} +else +{ +lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; uint8_t x_469; +x_465 = lean_ctor_get(x_460, 0); +lean_inc(x_465); +x_466 = lean_ctor_get(x_460, 1); +lean_inc(x_466); +if (lean_is_exclusive(x_460)) { + lean_ctor_release(x_460, 0); + lean_ctor_release(x_460, 1); + x_467 = x_460; +} else { + lean_dec_ref(x_460); + x_467 = lean_box(0); +} +x_468 = lean_ctor_get(x_465, 1); +lean_inc(x_468); +x_469 = lean_nat_dec_eq(x_456, x_468); +lean_dec(x_456); +if (x_469 == 0) +{ +lean_object* x_470; +lean_dec(x_468); +if (lean_is_scalar(x_467)) { + x_470 = lean_alloc_ctor(1, 2, 0); +} else { + x_470 = x_467; +} +lean_ctor_set(x_470, 0, x_465); +lean_ctor_set(x_470, 1, x_466); +return x_470; +} +else +{ +lean_object* x_471; lean_object* x_472; +lean_dec(x_467); +lean_dec(x_466); +x_471 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_472 = l_Std_Internal_Parsec_String_pstring(x_471, x_465); +if (lean_obj_tag(x_472) == 0) +{ +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +lean_dec(x_468); +x_473 = lean_ctor_get(x_472, 0); +lean_inc(x_473); +if (lean_is_exclusive(x_472)) { + lean_ctor_release(x_472, 0); + lean_ctor_release(x_472, 1); + x_474 = x_472; +} else { + lean_dec_ref(x_472); + x_474 = lean_box(0); +} +x_475 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_474)) { + x_476 = lean_alloc_ctor(0, 2, 0); +} else { + x_476 = x_474; +} +lean_ctor_set(x_476, 0, x_473); +lean_ctor_set(x_476, 1, x_475); +return x_476; +} +else +{ +lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; uint8_t x_481; +x_477 = lean_ctor_get(x_472, 0); +lean_inc(x_477); +x_478 = lean_ctor_get(x_472, 1); +lean_inc(x_478); +if (lean_is_exclusive(x_472)) { + lean_ctor_release(x_472, 0); + lean_ctor_release(x_472, 1); + x_479 = x_472; +} else { + lean_dec_ref(x_472); + x_479 = lean_box(0); +} +x_480 = lean_ctor_get(x_477, 1); +lean_inc(x_480); +x_481 = lean_nat_dec_eq(x_468, x_480); +lean_dec(x_468); +if (x_481 == 0) +{ +lean_object* x_482; +lean_dec(x_480); +if (lean_is_scalar(x_479)) { + x_482 = lean_alloc_ctor(1, 2, 0); +} else { + x_482 = x_479; +} +lean_ctor_set(x_482, 0, x_477); +lean_ctor_set(x_482, 1, x_478); +return x_482; +} +else +{ +lean_object* x_483; lean_object* x_484; +lean_dec(x_479); +lean_dec(x_478); +x_483 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_484 = l_Std_Internal_Parsec_String_pstring(x_483, x_477); +if (lean_obj_tag(x_484) == 0) +{ +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; +lean_dec(x_480); +x_485 = lean_ctor_get(x_484, 0); +lean_inc(x_485); +if (lean_is_exclusive(x_484)) { + lean_ctor_release(x_484, 0); + lean_ctor_release(x_484, 1); + x_486 = x_484; +} else { + lean_dec_ref(x_484); + x_486 = lean_box(0); +} +x_487 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_486)) { + x_488 = lean_alloc_ctor(0, 2, 0); +} else { + x_488 = x_486; +} +lean_ctor_set(x_488, 0, x_485); +lean_ctor_set(x_488, 1, x_487); +return x_488; +} +else +{ +lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; uint8_t x_493; +x_489 = lean_ctor_get(x_484, 0); +lean_inc(x_489); +x_490 = lean_ctor_get(x_484, 1); +lean_inc(x_490); +if (lean_is_exclusive(x_484)) { + lean_ctor_release(x_484, 0); + lean_ctor_release(x_484, 1); + x_491 = x_484; +} else { + lean_dec_ref(x_484); + x_491 = lean_box(0); +} +x_492 = lean_ctor_get(x_489, 1); +lean_inc(x_492); +x_493 = lean_nat_dec_eq(x_480, x_492); +lean_dec(x_492); +lean_dec(x_480); +if (x_493 == 0) +{ +lean_object* x_494; +if (lean_is_scalar(x_491)) { + x_494 = lean_alloc_ctor(1, 2, 0); +} else { + x_494 = x_491; +} +lean_ctor_set(x_494, 0, x_489); +lean_ctor_set(x_494, 1, x_490); +return x_494; +} +else +{ +lean_object* x_495; lean_object* x_496; +lean_dec(x_491); +lean_dec(x_490); +x_495 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_496 = l_Std_Internal_Parsec_String_pstring(x_495, x_489); +if (lean_obj_tag(x_496) == 0) +{ +lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; +x_497 = lean_ctor_get(x_496, 0); +lean_inc(x_497); +if (lean_is_exclusive(x_496)) { + lean_ctor_release(x_496, 0); + lean_ctor_release(x_496, 1); + x_498 = x_496; +} else { + lean_dec_ref(x_496); + x_498 = lean_box(0); +} +x_499 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_498)) { + x_500 = lean_alloc_ctor(0, 2, 0); +} else { + x_500 = x_498; +} +lean_ctor_set(x_500, 0, x_497); +lean_ctor_set(x_500, 1, x_499); +return x_500; +} +else +{ +lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; +x_501 = lean_ctor_get(x_496, 0); +lean_inc(x_501); +x_502 = lean_ctor_get(x_496, 1); +lean_inc(x_502); +if (lean_is_exclusive(x_496)) { + lean_ctor_release(x_496, 0); + lean_ctor_release(x_496, 1); + x_503 = x_496; +} else { + lean_dec_ref(x_496); + x_503 = lean_box(0); +} +if (lean_is_scalar(x_503)) { + x_504 = lean_alloc_ctor(1, 2, 0); +} else { + x_504 = x_503; +} +lean_ctor_set(x_504, 0, x_501); +lean_ctor_set(x_504, 1, x_502); +return x_504; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_505; lean_object* x_506; lean_object* x_507; uint8_t x_508; +x_505 = lean_ctor_get(x_43, 0); +x_506 = lean_ctor_get(x_43, 1); +lean_inc(x_506); +lean_inc(x_505); +lean_dec(x_43); +x_507 = lean_ctor_get(x_505, 1); +lean_inc(x_507); +x_508 = lean_nat_dec_eq(x_40, x_507); +lean_dec(x_40); +if (x_508 == 0) +{ +lean_object* x_509; +lean_dec(x_507); +x_509 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_509, 0, x_505); +lean_ctor_set(x_509, 1, x_506); +return x_509; +} +else +{ +lean_object* x_510; +lean_dec(x_506); +x_510 = l_Std_Internal_Parsec_String_pstring(x_29, x_505); +if (lean_obj_tag(x_510) == 0) +{ +lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +lean_dec(x_507); +x_511 = lean_ctor_get(x_510, 0); +lean_inc(x_511); +if (lean_is_exclusive(x_510)) { + lean_ctor_release(x_510, 0); + lean_ctor_release(x_510, 1); + x_512 = x_510; +} else { + lean_dec_ref(x_510); + x_512 = lean_box(0); +} +x_513 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_512)) { + x_514 = lean_alloc_ctor(0, 2, 0); +} else { + x_514 = x_512; +} +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_513); +return x_514; +} +else +{ +lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; uint8_t x_519; +x_515 = lean_ctor_get(x_510, 0); +lean_inc(x_515); +x_516 = lean_ctor_get(x_510, 1); +lean_inc(x_516); +if (lean_is_exclusive(x_510)) { + lean_ctor_release(x_510, 0); + lean_ctor_release(x_510, 1); + x_517 = x_510; +} else { + lean_dec_ref(x_510); + x_517 = lean_box(0); +} +x_518 = lean_ctor_get(x_515, 1); +lean_inc(x_518); +x_519 = lean_nat_dec_eq(x_507, x_518); +lean_dec(x_507); +if (x_519 == 0) +{ +lean_object* x_520; +lean_dec(x_518); +if (lean_is_scalar(x_517)) { + x_520 = lean_alloc_ctor(1, 2, 0); +} else { + x_520 = x_517; +} +lean_ctor_set(x_520, 0, x_515); +lean_ctor_set(x_520, 1, x_516); +return x_520; +} +else +{ +lean_object* x_521; +lean_dec(x_517); +lean_dec(x_516); +x_521 = l_Std_Internal_Parsec_String_pstring(x_2, x_515); +if (lean_obj_tag(x_521) == 0) +{ +lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; +lean_dec(x_518); +x_522 = lean_ctor_get(x_521, 0); +lean_inc(x_522); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_523 = x_521; +} else { + lean_dec_ref(x_521); + x_523 = lean_box(0); +} +x_524 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_523)) { + x_525 = lean_alloc_ctor(0, 2, 0); +} else { + x_525 = x_523; +} +lean_ctor_set(x_525, 0, x_522); +lean_ctor_set(x_525, 1, x_524); +return x_525; +} +else +{ +lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; uint8_t x_530; +x_526 = lean_ctor_get(x_521, 0); +lean_inc(x_526); +x_527 = lean_ctor_get(x_521, 1); +lean_inc(x_527); +if (lean_is_exclusive(x_521)) { + lean_ctor_release(x_521, 0); + lean_ctor_release(x_521, 1); + x_528 = x_521; +} else { + lean_dec_ref(x_521); + x_528 = lean_box(0); +} +x_529 = lean_ctor_get(x_526, 1); +lean_inc(x_529); +x_530 = lean_nat_dec_eq(x_518, x_529); +lean_dec(x_518); +if (x_530 == 0) +{ +lean_object* x_531; +lean_dec(x_529); +if (lean_is_scalar(x_528)) { + x_531 = lean_alloc_ctor(1, 2, 0); +} else { + x_531 = x_528; +} +lean_ctor_set(x_531, 0, x_526); +lean_ctor_set(x_531, 1, x_527); +return x_531; +} +else +{ +lean_object* x_532; +lean_dec(x_528); +lean_dec(x_527); +x_532 = l_Std_Internal_Parsec_String_pstring(x_2, x_526); +if (lean_obj_tag(x_532) == 0) +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; +lean_dec(x_529); +x_533 = lean_ctor_get(x_532, 0); +lean_inc(x_533); +if (lean_is_exclusive(x_532)) { + lean_ctor_release(x_532, 0); + lean_ctor_release(x_532, 1); + x_534 = x_532; +} else { + lean_dec_ref(x_532); + x_534 = lean_box(0); +} +x_535 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_534)) { + x_536 = lean_alloc_ctor(0, 2, 0); +} else { + x_536 = x_534; +} +lean_ctor_set(x_536, 0, x_533); +lean_ctor_set(x_536, 1, x_535); +return x_536; +} +else +{ +lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; uint8_t x_541; +x_537 = lean_ctor_get(x_532, 0); +lean_inc(x_537); +x_538 = lean_ctor_get(x_532, 1); +lean_inc(x_538); +if (lean_is_exclusive(x_532)) { + lean_ctor_release(x_532, 0); + lean_ctor_release(x_532, 1); + x_539 = x_532; +} else { + lean_dec_ref(x_532); + x_539 = lean_box(0); +} +x_540 = lean_ctor_get(x_537, 1); +lean_inc(x_540); +x_541 = lean_nat_dec_eq(x_529, x_540); +lean_dec(x_529); +if (x_541 == 0) +{ +lean_object* x_542; +lean_dec(x_540); +if (lean_is_scalar(x_539)) { + x_542 = lean_alloc_ctor(1, 2, 0); +} else { + x_542 = x_539; +} +lean_ctor_set(x_542, 0, x_537); +lean_ctor_set(x_542, 1, x_538); +return x_542; +} +else +{ +lean_object* x_543; +lean_dec(x_539); +lean_dec(x_538); +x_543 = l_Std_Internal_Parsec_String_pstring(x_42, x_537); +if (lean_obj_tag(x_543) == 0) +{ +lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; +lean_dec(x_540); +x_544 = lean_ctor_get(x_543, 0); +lean_inc(x_544); +if (lean_is_exclusive(x_543)) { + lean_ctor_release(x_543, 0); + lean_ctor_release(x_543, 1); + x_545 = x_543; +} else { + lean_dec_ref(x_543); + x_545 = lean_box(0); +} +x_546 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_545)) { + x_547 = lean_alloc_ctor(0, 2, 0); +} else { + x_547 = x_545; +} +lean_ctor_set(x_547, 0, x_544); +lean_ctor_set(x_547, 1, x_546); +return x_547; +} +else +{ +lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; uint8_t x_552; +x_548 = lean_ctor_get(x_543, 0); +lean_inc(x_548); +x_549 = lean_ctor_get(x_543, 1); +lean_inc(x_549); +if (lean_is_exclusive(x_543)) { + lean_ctor_release(x_543, 0); + lean_ctor_release(x_543, 1); + x_550 = x_543; +} else { + lean_dec_ref(x_543); + x_550 = lean_box(0); +} +x_551 = lean_ctor_get(x_548, 1); +lean_inc(x_551); +x_552 = lean_nat_dec_eq(x_540, x_551); +lean_dec(x_540); +if (x_552 == 0) +{ +lean_object* x_553; +lean_dec(x_551); +if (lean_is_scalar(x_550)) { + x_553 = lean_alloc_ctor(1, 2, 0); +} else { + x_553 = x_550; +} +lean_ctor_set(x_553, 0, x_548); +lean_ctor_set(x_553, 1, x_549); +return x_553; +} +else +{ +lean_object* x_554; lean_object* x_555; +lean_dec(x_550); +lean_dec(x_549); +x_554 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_555 = l_Std_Internal_Parsec_String_pstring(x_554, x_548); +if (lean_obj_tag(x_555) == 0) +{ +lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; +lean_dec(x_551); +x_556 = lean_ctor_get(x_555, 0); +lean_inc(x_556); +if (lean_is_exclusive(x_555)) { + lean_ctor_release(x_555, 0); + lean_ctor_release(x_555, 1); + x_557 = x_555; +} else { + lean_dec_ref(x_555); + x_557 = lean_box(0); +} +x_558 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_557)) { + x_559 = lean_alloc_ctor(0, 2, 0); +} else { + x_559 = x_557; +} +lean_ctor_set(x_559, 0, x_556); +lean_ctor_set(x_559, 1, x_558); +return x_559; +} +else +{ +lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; uint8_t x_564; +x_560 = lean_ctor_get(x_555, 0); +lean_inc(x_560); +x_561 = lean_ctor_get(x_555, 1); +lean_inc(x_561); +if (lean_is_exclusive(x_555)) { + lean_ctor_release(x_555, 0); + lean_ctor_release(x_555, 1); + x_562 = x_555; +} else { + lean_dec_ref(x_555); + x_562 = lean_box(0); +} +x_563 = lean_ctor_get(x_560, 1); +lean_inc(x_563); +x_564 = lean_nat_dec_eq(x_551, x_563); +lean_dec(x_551); +if (x_564 == 0) +{ +lean_object* x_565; +lean_dec(x_563); +if (lean_is_scalar(x_562)) { + x_565 = lean_alloc_ctor(1, 2, 0); +} else { + x_565 = x_562; +} +lean_ctor_set(x_565, 0, x_560); +lean_ctor_set(x_565, 1, x_561); +return x_565; +} +else +{ +lean_object* x_566; lean_object* x_567; +lean_dec(x_562); +lean_dec(x_561); +x_566 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_567 = l_Std_Internal_Parsec_String_pstring(x_566, x_560); +if (lean_obj_tag(x_567) == 0) +{ +lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; +lean_dec(x_563); +x_568 = lean_ctor_get(x_567, 0); +lean_inc(x_568); +if (lean_is_exclusive(x_567)) { + lean_ctor_release(x_567, 0); + lean_ctor_release(x_567, 1); + x_569 = x_567; +} else { + lean_dec_ref(x_567); + x_569 = lean_box(0); +} +x_570 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_569)) { + x_571 = lean_alloc_ctor(0, 2, 0); +} else { + x_571 = x_569; +} +lean_ctor_set(x_571, 0, x_568); +lean_ctor_set(x_571, 1, x_570); +return x_571; +} +else +{ +lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; uint8_t x_576; +x_572 = lean_ctor_get(x_567, 0); +lean_inc(x_572); +x_573 = lean_ctor_get(x_567, 1); +lean_inc(x_573); +if (lean_is_exclusive(x_567)) { + lean_ctor_release(x_567, 0); + lean_ctor_release(x_567, 1); + x_574 = x_567; +} else { + lean_dec_ref(x_567); + x_574 = lean_box(0); +} +x_575 = lean_ctor_get(x_572, 1); +lean_inc(x_575); +x_576 = lean_nat_dec_eq(x_563, x_575); +lean_dec(x_563); +if (x_576 == 0) +{ +lean_object* x_577; +lean_dec(x_575); +if (lean_is_scalar(x_574)) { + x_577 = lean_alloc_ctor(1, 2, 0); +} else { + x_577 = x_574; +} +lean_ctor_set(x_577, 0, x_572); +lean_ctor_set(x_577, 1, x_573); +return x_577; +} +else +{ +lean_object* x_578; lean_object* x_579; +lean_dec(x_574); +lean_dec(x_573); +x_578 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_579 = l_Std_Internal_Parsec_String_pstring(x_578, x_572); +if (lean_obj_tag(x_579) == 0) +{ +lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; +lean_dec(x_575); +x_580 = lean_ctor_get(x_579, 0); +lean_inc(x_580); +if (lean_is_exclusive(x_579)) { + lean_ctor_release(x_579, 0); + lean_ctor_release(x_579, 1); + x_581 = x_579; +} else { + lean_dec_ref(x_579); + x_581 = lean_box(0); +} +x_582 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_581)) { + x_583 = lean_alloc_ctor(0, 2, 0); +} else { + x_583 = x_581; +} +lean_ctor_set(x_583, 0, x_580); +lean_ctor_set(x_583, 1, x_582); +return x_583; +} +else +{ +lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; uint8_t x_588; +x_584 = lean_ctor_get(x_579, 0); +lean_inc(x_584); +x_585 = lean_ctor_get(x_579, 1); +lean_inc(x_585); +if (lean_is_exclusive(x_579)) { + lean_ctor_release(x_579, 0); + lean_ctor_release(x_579, 1); + x_586 = x_579; +} else { + lean_dec_ref(x_579); + x_586 = lean_box(0); +} +x_587 = lean_ctor_get(x_584, 1); +lean_inc(x_587); +x_588 = lean_nat_dec_eq(x_575, x_587); +lean_dec(x_587); +lean_dec(x_575); +if (x_588 == 0) +{ +lean_object* x_589; +if (lean_is_scalar(x_586)) { + x_589 = lean_alloc_ctor(1, 2, 0); +} else { + x_589 = x_586; +} +lean_ctor_set(x_589, 0, x_584); +lean_ctor_set(x_589, 1, x_585); +return x_589; +} +else +{ +lean_object* x_590; lean_object* x_591; +lean_dec(x_586); +lean_dec(x_585); +x_590 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_591 = l_Std_Internal_Parsec_String_pstring(x_590, x_584); +if (lean_obj_tag(x_591) == 0) +{ +lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; +x_592 = lean_ctor_get(x_591, 0); +lean_inc(x_592); +if (lean_is_exclusive(x_591)) { + lean_ctor_release(x_591, 0); + lean_ctor_release(x_591, 1); + x_593 = x_591; +} else { + lean_dec_ref(x_591); + x_593 = lean_box(0); +} +x_594 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_593)) { + x_595 = lean_alloc_ctor(0, 2, 0); +} else { + x_595 = x_593; +} +lean_ctor_set(x_595, 0, x_592); +lean_ctor_set(x_595, 1, x_594); +return x_595; +} +else +{ +lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; +x_596 = lean_ctor_get(x_591, 0); +lean_inc(x_596); +x_597 = lean_ctor_get(x_591, 1); +lean_inc(x_597); +if (lean_is_exclusive(x_591)) { + lean_ctor_release(x_591, 0); + lean_ctor_release(x_591, 1); + x_598 = x_591; +} else { + lean_dec_ref(x_591); + x_598 = lean_box(0); +} +if (lean_is_scalar(x_598)) { + x_599 = lean_alloc_ctor(1, 2, 0); +} else { + x_599 = x_598; +} +lean_ctor_set(x_599, 0, x_596); +lean_ctor_set(x_599, 1, x_597); +return x_599; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_600; lean_object* x_601; lean_object* x_602; uint8_t x_603; +x_600 = lean_ctor_get(x_30, 0); +x_601 = lean_ctor_get(x_30, 1); +lean_inc(x_601); +lean_inc(x_600); +lean_dec(x_30); +x_602 = lean_ctor_get(x_600, 1); +lean_inc(x_602); +x_603 = lean_nat_dec_eq(x_27, x_602); +lean_dec(x_27); +if (x_603 == 0) +{ +lean_object* x_604; +lean_dec(x_602); +x_604 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_604, 0, x_600); +lean_ctor_set(x_604, 1, x_601); +return x_604; +} +else +{ +lean_object* x_605; lean_object* x_606; +lean_dec(x_601); +x_605 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +x_606 = l_Std_Internal_Parsec_String_pstring(x_605, x_600); +if (lean_obj_tag(x_606) == 0) +{ +lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; +lean_dec(x_602); +x_607 = lean_ctor_get(x_606, 0); +lean_inc(x_607); +if (lean_is_exclusive(x_606)) { + lean_ctor_release(x_606, 0); + lean_ctor_release(x_606, 1); + x_608 = x_606; +} else { + lean_dec_ref(x_606); + x_608 = lean_box(0); +} +x_609 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_608)) { + x_610 = lean_alloc_ctor(0, 2, 0); +} else { + x_610 = x_608; +} +lean_ctor_set(x_610, 0, x_607); +lean_ctor_set(x_610, 1, x_609); +return x_610; +} +else +{ +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; uint8_t x_615; +x_611 = lean_ctor_get(x_606, 0); +lean_inc(x_611); +x_612 = lean_ctor_get(x_606, 1); +lean_inc(x_612); +if (lean_is_exclusive(x_606)) { + lean_ctor_release(x_606, 0); + lean_ctor_release(x_606, 1); + x_613 = x_606; +} else { + lean_dec_ref(x_606); + x_613 = lean_box(0); +} +x_614 = lean_ctor_get(x_611, 1); +lean_inc(x_614); +x_615 = lean_nat_dec_eq(x_602, x_614); +lean_dec(x_602); +if (x_615 == 0) +{ +lean_object* x_616; +lean_dec(x_614); +if (lean_is_scalar(x_613)) { + x_616 = lean_alloc_ctor(1, 2, 0); +} else { + x_616 = x_613; +} +lean_ctor_set(x_616, 0, x_611); +lean_ctor_set(x_616, 1, x_612); +return x_616; +} +else +{ +lean_object* x_617; +lean_dec(x_613); +lean_dec(x_612); +x_617 = l_Std_Internal_Parsec_String_pstring(x_29, x_611); +if (lean_obj_tag(x_617) == 0) +{ +lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; +lean_dec(x_614); +x_618 = lean_ctor_get(x_617, 0); +lean_inc(x_618); +if (lean_is_exclusive(x_617)) { + lean_ctor_release(x_617, 0); + lean_ctor_release(x_617, 1); + x_619 = x_617; +} else { + lean_dec_ref(x_617); + x_619 = lean_box(0); +} +x_620 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_619)) { + x_621 = lean_alloc_ctor(0, 2, 0); +} else { + x_621 = x_619; +} +lean_ctor_set(x_621, 0, x_618); +lean_ctor_set(x_621, 1, x_620); +return x_621; +} +else +{ +lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; uint8_t x_626; +x_622 = lean_ctor_get(x_617, 0); +lean_inc(x_622); +x_623 = lean_ctor_get(x_617, 1); +lean_inc(x_623); +if (lean_is_exclusive(x_617)) { + lean_ctor_release(x_617, 0); + lean_ctor_release(x_617, 1); + x_624 = x_617; +} else { + lean_dec_ref(x_617); + x_624 = lean_box(0); +} +x_625 = lean_ctor_get(x_622, 1); +lean_inc(x_625); +x_626 = lean_nat_dec_eq(x_614, x_625); +lean_dec(x_614); +if (x_626 == 0) +{ +lean_object* x_627; +lean_dec(x_625); +if (lean_is_scalar(x_624)) { + x_627 = lean_alloc_ctor(1, 2, 0); +} else { + x_627 = x_624; +} +lean_ctor_set(x_627, 0, x_622); +lean_ctor_set(x_627, 1, x_623); +return x_627; +} +else +{ +lean_object* x_628; +lean_dec(x_624); +lean_dec(x_623); +x_628 = l_Std_Internal_Parsec_String_pstring(x_2, x_622); +if (lean_obj_tag(x_628) == 0) +{ +lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; +lean_dec(x_625); +x_629 = lean_ctor_get(x_628, 0); +lean_inc(x_629); +if (lean_is_exclusive(x_628)) { + lean_ctor_release(x_628, 0); + lean_ctor_release(x_628, 1); + x_630 = x_628; +} else { + lean_dec_ref(x_628); + x_630 = lean_box(0); +} +x_631 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_630)) { + x_632 = lean_alloc_ctor(0, 2, 0); +} else { + x_632 = x_630; +} +lean_ctor_set(x_632, 0, x_629); +lean_ctor_set(x_632, 1, x_631); +return x_632; +} +else +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; uint8_t x_637; +x_633 = lean_ctor_get(x_628, 0); +lean_inc(x_633); +x_634 = lean_ctor_get(x_628, 1); +lean_inc(x_634); +if (lean_is_exclusive(x_628)) { + lean_ctor_release(x_628, 0); + lean_ctor_release(x_628, 1); + x_635 = x_628; +} else { + lean_dec_ref(x_628); + x_635 = lean_box(0); +} +x_636 = lean_ctor_get(x_633, 1); +lean_inc(x_636); +x_637 = lean_nat_dec_eq(x_625, x_636); +lean_dec(x_625); +if (x_637 == 0) +{ +lean_object* x_638; +lean_dec(x_636); +if (lean_is_scalar(x_635)) { + x_638 = lean_alloc_ctor(1, 2, 0); +} else { + x_638 = x_635; +} +lean_ctor_set(x_638, 0, x_633); +lean_ctor_set(x_638, 1, x_634); +return x_638; +} +else +{ +lean_object* x_639; +lean_dec(x_635); +lean_dec(x_634); +x_639 = l_Std_Internal_Parsec_String_pstring(x_2, x_633); +if (lean_obj_tag(x_639) == 0) +{ +lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; +lean_dec(x_636); +x_640 = lean_ctor_get(x_639, 0); +lean_inc(x_640); +if (lean_is_exclusive(x_639)) { + lean_ctor_release(x_639, 0); + lean_ctor_release(x_639, 1); + x_641 = x_639; +} else { + lean_dec_ref(x_639); + x_641 = lean_box(0); +} +x_642 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_641)) { + x_643 = lean_alloc_ctor(0, 2, 0); +} else { + x_643 = x_641; +} +lean_ctor_set(x_643, 0, x_640); +lean_ctor_set(x_643, 1, x_642); +return x_643; +} +else +{ +lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; uint8_t x_648; +x_644 = lean_ctor_get(x_639, 0); +lean_inc(x_644); +x_645 = lean_ctor_get(x_639, 1); +lean_inc(x_645); +if (lean_is_exclusive(x_639)) { + lean_ctor_release(x_639, 0); + lean_ctor_release(x_639, 1); + x_646 = x_639; +} else { + lean_dec_ref(x_639); + x_646 = lean_box(0); +} +x_647 = lean_ctor_get(x_644, 1); +lean_inc(x_647); +x_648 = lean_nat_dec_eq(x_636, x_647); +lean_dec(x_636); +if (x_648 == 0) +{ +lean_object* x_649; +lean_dec(x_647); +if (lean_is_scalar(x_646)) { + x_649 = lean_alloc_ctor(1, 2, 0); +} else { + x_649 = x_646; +} +lean_ctor_set(x_649, 0, x_644); +lean_ctor_set(x_649, 1, x_645); +return x_649; +} +else +{ +lean_object* x_650; +lean_dec(x_646); +lean_dec(x_645); +x_650 = l_Std_Internal_Parsec_String_pstring(x_605, x_644); +if (lean_obj_tag(x_650) == 0) +{ +lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; +lean_dec(x_647); +x_651 = lean_ctor_get(x_650, 0); +lean_inc(x_651); +if (lean_is_exclusive(x_650)) { + lean_ctor_release(x_650, 0); + lean_ctor_release(x_650, 1); + x_652 = x_650; +} else { + lean_dec_ref(x_650); + x_652 = lean_box(0); +} +x_653 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_652)) { + x_654 = lean_alloc_ctor(0, 2, 0); +} else { + x_654 = x_652; +} +lean_ctor_set(x_654, 0, x_651); +lean_ctor_set(x_654, 1, x_653); +return x_654; +} +else +{ +lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; uint8_t x_659; +x_655 = lean_ctor_get(x_650, 0); +lean_inc(x_655); +x_656 = lean_ctor_get(x_650, 1); +lean_inc(x_656); +if (lean_is_exclusive(x_650)) { + lean_ctor_release(x_650, 0); + lean_ctor_release(x_650, 1); + x_657 = x_650; +} else { + lean_dec_ref(x_650); + x_657 = lean_box(0); +} +x_658 = lean_ctor_get(x_655, 1); +lean_inc(x_658); +x_659 = lean_nat_dec_eq(x_647, x_658); +lean_dec(x_647); +if (x_659 == 0) +{ +lean_object* x_660; +lean_dec(x_658); +if (lean_is_scalar(x_657)) { + x_660 = lean_alloc_ctor(1, 2, 0); +} else { + x_660 = x_657; +} +lean_ctor_set(x_660, 0, x_655); +lean_ctor_set(x_660, 1, x_656); +return x_660; +} +else +{ +lean_object* x_661; lean_object* x_662; +lean_dec(x_657); +lean_dec(x_656); +x_661 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_662 = l_Std_Internal_Parsec_String_pstring(x_661, x_655); +if (lean_obj_tag(x_662) == 0) +{ +lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; +lean_dec(x_658); +x_663 = lean_ctor_get(x_662, 0); +lean_inc(x_663); +if (lean_is_exclusive(x_662)) { + lean_ctor_release(x_662, 0); + lean_ctor_release(x_662, 1); + x_664 = x_662; +} else { + lean_dec_ref(x_662); + x_664 = lean_box(0); +} +x_665 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_664)) { + x_666 = lean_alloc_ctor(0, 2, 0); +} else { + x_666 = x_664; +} +lean_ctor_set(x_666, 0, x_663); +lean_ctor_set(x_666, 1, x_665); +return x_666; +} +else +{ +lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; uint8_t x_671; +x_667 = lean_ctor_get(x_662, 0); +lean_inc(x_667); +x_668 = lean_ctor_get(x_662, 1); +lean_inc(x_668); +if (lean_is_exclusive(x_662)) { + lean_ctor_release(x_662, 0); + lean_ctor_release(x_662, 1); + x_669 = x_662; +} else { + lean_dec_ref(x_662); + x_669 = lean_box(0); +} +x_670 = lean_ctor_get(x_667, 1); +lean_inc(x_670); +x_671 = lean_nat_dec_eq(x_658, x_670); +lean_dec(x_658); +if (x_671 == 0) +{ +lean_object* x_672; +lean_dec(x_670); +if (lean_is_scalar(x_669)) { + x_672 = lean_alloc_ctor(1, 2, 0); +} else { + x_672 = x_669; +} +lean_ctor_set(x_672, 0, x_667); +lean_ctor_set(x_672, 1, x_668); +return x_672; +} +else +{ +lean_object* x_673; lean_object* x_674; +lean_dec(x_669); +lean_dec(x_668); +x_673 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_674 = l_Std_Internal_Parsec_String_pstring(x_673, x_667); +if (lean_obj_tag(x_674) == 0) +{ +lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; +lean_dec(x_670); +x_675 = lean_ctor_get(x_674, 0); +lean_inc(x_675); +if (lean_is_exclusive(x_674)) { + lean_ctor_release(x_674, 0); + lean_ctor_release(x_674, 1); + x_676 = x_674; +} else { + lean_dec_ref(x_674); + x_676 = lean_box(0); +} +x_677 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_676)) { + x_678 = lean_alloc_ctor(0, 2, 0); +} else { + x_678 = x_676; +} +lean_ctor_set(x_678, 0, x_675); +lean_ctor_set(x_678, 1, x_677); +return x_678; +} +else +{ +lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; uint8_t x_683; +x_679 = lean_ctor_get(x_674, 0); +lean_inc(x_679); +x_680 = lean_ctor_get(x_674, 1); +lean_inc(x_680); +if (lean_is_exclusive(x_674)) { + lean_ctor_release(x_674, 0); + lean_ctor_release(x_674, 1); + x_681 = x_674; +} else { + lean_dec_ref(x_674); + x_681 = lean_box(0); +} +x_682 = lean_ctor_get(x_679, 1); +lean_inc(x_682); +x_683 = lean_nat_dec_eq(x_670, x_682); +lean_dec(x_670); +if (x_683 == 0) +{ +lean_object* x_684; +lean_dec(x_682); +if (lean_is_scalar(x_681)) { + x_684 = lean_alloc_ctor(1, 2, 0); +} else { + x_684 = x_681; +} +lean_ctor_set(x_684, 0, x_679); +lean_ctor_set(x_684, 1, x_680); +return x_684; +} +else +{ +lean_object* x_685; lean_object* x_686; +lean_dec(x_681); +lean_dec(x_680); +x_685 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_686 = l_Std_Internal_Parsec_String_pstring(x_685, x_679); +if (lean_obj_tag(x_686) == 0) +{ +lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; +lean_dec(x_682); +x_687 = lean_ctor_get(x_686, 0); +lean_inc(x_687); +if (lean_is_exclusive(x_686)) { + lean_ctor_release(x_686, 0); + lean_ctor_release(x_686, 1); + x_688 = x_686; +} else { + lean_dec_ref(x_686); + x_688 = lean_box(0); +} +x_689 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_688)) { + x_690 = lean_alloc_ctor(0, 2, 0); +} else { + x_690 = x_688; +} +lean_ctor_set(x_690, 0, x_687); +lean_ctor_set(x_690, 1, x_689); +return x_690; +} +else +{ +lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; uint8_t x_695; +x_691 = lean_ctor_get(x_686, 0); +lean_inc(x_691); +x_692 = lean_ctor_get(x_686, 1); +lean_inc(x_692); +if (lean_is_exclusive(x_686)) { + lean_ctor_release(x_686, 0); + lean_ctor_release(x_686, 1); + x_693 = x_686; +} else { + lean_dec_ref(x_686); + x_693 = lean_box(0); +} +x_694 = lean_ctor_get(x_691, 1); +lean_inc(x_694); +x_695 = lean_nat_dec_eq(x_682, x_694); +lean_dec(x_694); +lean_dec(x_682); +if (x_695 == 0) +{ +lean_object* x_696; +if (lean_is_scalar(x_693)) { + x_696 = lean_alloc_ctor(1, 2, 0); +} else { + x_696 = x_693; +} +lean_ctor_set(x_696, 0, x_691); +lean_ctor_set(x_696, 1, x_692); +return x_696; +} +else +{ +lean_object* x_697; lean_object* x_698; +lean_dec(x_693); +lean_dec(x_692); +x_697 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_698 = l_Std_Internal_Parsec_String_pstring(x_697, x_691); +if (lean_obj_tag(x_698) == 0) +{ +lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; +x_699 = lean_ctor_get(x_698, 0); +lean_inc(x_699); +if (lean_is_exclusive(x_698)) { + lean_ctor_release(x_698, 0); + lean_ctor_release(x_698, 1); + x_700 = x_698; +} else { + lean_dec_ref(x_698); + x_700 = lean_box(0); +} +x_701 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_700)) { + x_702 = lean_alloc_ctor(0, 2, 0); +} else { + x_702 = x_700; +} +lean_ctor_set(x_702, 0, x_699); +lean_ctor_set(x_702, 1, x_701); +return x_702; +} +else +{ +lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; +x_703 = lean_ctor_get(x_698, 0); +lean_inc(x_703); +x_704 = lean_ctor_get(x_698, 1); +lean_inc(x_704); +if (lean_is_exclusive(x_698)) { + lean_ctor_release(x_698, 0); + lean_ctor_release(x_698, 1); + x_705 = x_698; +} else { + lean_dec_ref(x_698); + x_705 = lean_box(0); +} +if (lean_is_scalar(x_705)) { + x_706 = lean_alloc_ctor(1, 2, 0); +} else { + x_706 = x_705; +} +lean_ctor_set(x_706, 0, x_703); +lean_ctor_set(x_706, 1, x_704); +return x_706; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_707; lean_object* x_708; lean_object* x_709; uint8_t x_710; +x_707 = lean_ctor_get(x_17, 0); +x_708 = lean_ctor_get(x_17, 1); +lean_inc(x_708); +lean_inc(x_707); +lean_dec(x_17); +x_709 = lean_ctor_get(x_707, 1); +lean_inc(x_709); +x_710 = lean_nat_dec_eq(x_14, x_709); +lean_dec(x_14); +if (x_710 == 0) +{ +lean_object* x_711; +lean_dec(x_709); +x_711 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_711, 0, x_707); +lean_ctor_set(x_711, 1, x_708); +return x_711; +} +else +{ +lean_object* x_712; lean_object* x_713; +lean_dec(x_708); +x_712 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +x_713 = l_Std_Internal_Parsec_String_pstring(x_712, x_707); +if (lean_obj_tag(x_713) == 0) +{ +lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; +lean_dec(x_709); +x_714 = lean_ctor_get(x_713, 0); +lean_inc(x_714); +if (lean_is_exclusive(x_713)) { + lean_ctor_release(x_713, 0); + lean_ctor_release(x_713, 1); + x_715 = x_713; +} else { + lean_dec_ref(x_713); + x_715 = lean_box(0); +} +x_716 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_715)) { + x_717 = lean_alloc_ctor(0, 2, 0); +} else { + x_717 = x_715; +} +lean_ctor_set(x_717, 0, x_714); +lean_ctor_set(x_717, 1, x_716); +return x_717; +} +else +{ +lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; uint8_t x_722; +x_718 = lean_ctor_get(x_713, 0); +lean_inc(x_718); +x_719 = lean_ctor_get(x_713, 1); +lean_inc(x_719); +if (lean_is_exclusive(x_713)) { + lean_ctor_release(x_713, 0); + lean_ctor_release(x_713, 1); + x_720 = x_713; +} else { + lean_dec_ref(x_713); + x_720 = lean_box(0); +} +x_721 = lean_ctor_get(x_718, 1); +lean_inc(x_721); +x_722 = lean_nat_dec_eq(x_709, x_721); +lean_dec(x_709); +if (x_722 == 0) +{ +lean_object* x_723; +lean_dec(x_721); +if (lean_is_scalar(x_720)) { + x_723 = lean_alloc_ctor(1, 2, 0); +} else { + x_723 = x_720; +} +lean_ctor_set(x_723, 0, x_718); +lean_ctor_set(x_723, 1, x_719); +return x_723; +} +else +{ +lean_object* x_724; lean_object* x_725; +lean_dec(x_720); +lean_dec(x_719); +x_724 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +x_725 = l_Std_Internal_Parsec_String_pstring(x_724, x_718); +if (lean_obj_tag(x_725) == 0) +{ +lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; +lean_dec(x_721); +x_726 = lean_ctor_get(x_725, 0); +lean_inc(x_726); +if (lean_is_exclusive(x_725)) { + lean_ctor_release(x_725, 0); + lean_ctor_release(x_725, 1); + x_727 = x_725; +} else { + lean_dec_ref(x_725); + x_727 = lean_box(0); +} +x_728 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_727)) { + x_729 = lean_alloc_ctor(0, 2, 0); +} else { + x_729 = x_727; +} +lean_ctor_set(x_729, 0, x_726); +lean_ctor_set(x_729, 1, x_728); +return x_729; +} +else +{ +lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; uint8_t x_734; +x_730 = lean_ctor_get(x_725, 0); +lean_inc(x_730); +x_731 = lean_ctor_get(x_725, 1); +lean_inc(x_731); +if (lean_is_exclusive(x_725)) { + lean_ctor_release(x_725, 0); + lean_ctor_release(x_725, 1); + x_732 = x_725; +} else { + lean_dec_ref(x_725); + x_732 = lean_box(0); +} +x_733 = lean_ctor_get(x_730, 1); +lean_inc(x_733); +x_734 = lean_nat_dec_eq(x_721, x_733); +lean_dec(x_721); +if (x_734 == 0) +{ +lean_object* x_735; +lean_dec(x_733); +if (lean_is_scalar(x_732)) { + x_735 = lean_alloc_ctor(1, 2, 0); +} else { + x_735 = x_732; +} +lean_ctor_set(x_735, 0, x_730); +lean_ctor_set(x_735, 1, x_731); +return x_735; +} +else +{ +lean_object* x_736; +lean_dec(x_732); +lean_dec(x_731); +x_736 = l_Std_Internal_Parsec_String_pstring(x_712, x_730); +if (lean_obj_tag(x_736) == 0) +{ +lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; +lean_dec(x_733); +x_737 = lean_ctor_get(x_736, 0); +lean_inc(x_737); +if (lean_is_exclusive(x_736)) { + lean_ctor_release(x_736, 0); + lean_ctor_release(x_736, 1); + x_738 = x_736; +} else { + lean_dec_ref(x_736); + x_738 = lean_box(0); +} +x_739 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_738)) { + x_740 = lean_alloc_ctor(0, 2, 0); +} else { + x_740 = x_738; +} +lean_ctor_set(x_740, 0, x_737); +lean_ctor_set(x_740, 1, x_739); +return x_740; +} +else +{ +lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; uint8_t x_745; +x_741 = lean_ctor_get(x_736, 0); +lean_inc(x_741); +x_742 = lean_ctor_get(x_736, 1); +lean_inc(x_742); +if (lean_is_exclusive(x_736)) { + lean_ctor_release(x_736, 0); + lean_ctor_release(x_736, 1); + x_743 = x_736; +} else { + lean_dec_ref(x_736); + x_743 = lean_box(0); +} +x_744 = lean_ctor_get(x_741, 1); +lean_inc(x_744); +x_745 = lean_nat_dec_eq(x_733, x_744); +lean_dec(x_733); +if (x_745 == 0) +{ +lean_object* x_746; +lean_dec(x_744); +if (lean_is_scalar(x_743)) { + x_746 = lean_alloc_ctor(1, 2, 0); +} else { + x_746 = x_743; +} +lean_ctor_set(x_746, 0, x_741); +lean_ctor_set(x_746, 1, x_742); +return x_746; +} +else +{ +lean_object* x_747; +lean_dec(x_743); +lean_dec(x_742); +x_747 = l_Std_Internal_Parsec_String_pstring(x_2, x_741); +if (lean_obj_tag(x_747) == 0) +{ +lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; +lean_dec(x_744); +x_748 = lean_ctor_get(x_747, 0); +lean_inc(x_748); +if (lean_is_exclusive(x_747)) { + lean_ctor_release(x_747, 0); + lean_ctor_release(x_747, 1); + x_749 = x_747; +} else { + lean_dec_ref(x_747); + x_749 = lean_box(0); +} +x_750 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_749)) { + x_751 = lean_alloc_ctor(0, 2, 0); +} else { + x_751 = x_749; +} +lean_ctor_set(x_751, 0, x_748); +lean_ctor_set(x_751, 1, x_750); +return x_751; +} +else +{ +lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; uint8_t x_756; +x_752 = lean_ctor_get(x_747, 0); +lean_inc(x_752); +x_753 = lean_ctor_get(x_747, 1); +lean_inc(x_753); +if (lean_is_exclusive(x_747)) { + lean_ctor_release(x_747, 0); + lean_ctor_release(x_747, 1); + x_754 = x_747; +} else { + lean_dec_ref(x_747); + x_754 = lean_box(0); +} +x_755 = lean_ctor_get(x_752, 1); +lean_inc(x_755); +x_756 = lean_nat_dec_eq(x_744, x_755); +lean_dec(x_744); +if (x_756 == 0) +{ +lean_object* x_757; +lean_dec(x_755); +if (lean_is_scalar(x_754)) { + x_757 = lean_alloc_ctor(1, 2, 0); +} else { + x_757 = x_754; +} +lean_ctor_set(x_757, 0, x_752); +lean_ctor_set(x_757, 1, x_753); +return x_757; +} +else +{ +lean_object* x_758; +lean_dec(x_754); +lean_dec(x_753); +x_758 = l_Std_Internal_Parsec_String_pstring(x_2, x_752); +if (lean_obj_tag(x_758) == 0) +{ +lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; +lean_dec(x_755); +x_759 = lean_ctor_get(x_758, 0); +lean_inc(x_759); +if (lean_is_exclusive(x_758)) { + lean_ctor_release(x_758, 0); + lean_ctor_release(x_758, 1); + x_760 = x_758; +} else { + lean_dec_ref(x_758); + x_760 = lean_box(0); +} +x_761 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_760)) { + x_762 = lean_alloc_ctor(0, 2, 0); +} else { + x_762 = x_760; +} +lean_ctor_set(x_762, 0, x_759); +lean_ctor_set(x_762, 1, x_761); +return x_762; +} +else +{ +lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; uint8_t x_767; +x_763 = lean_ctor_get(x_758, 0); +lean_inc(x_763); +x_764 = lean_ctor_get(x_758, 1); +lean_inc(x_764); +if (lean_is_exclusive(x_758)) { + lean_ctor_release(x_758, 0); + lean_ctor_release(x_758, 1); + x_765 = x_758; +} else { + lean_dec_ref(x_758); + x_765 = lean_box(0); +} +x_766 = lean_ctor_get(x_763, 1); +lean_inc(x_766); +x_767 = lean_nat_dec_eq(x_755, x_766); +lean_dec(x_755); +if (x_767 == 0) +{ +lean_object* x_768; +lean_dec(x_766); +if (lean_is_scalar(x_765)) { + x_768 = lean_alloc_ctor(1, 2, 0); +} else { + x_768 = x_765; +} +lean_ctor_set(x_768, 0, x_763); +lean_ctor_set(x_768, 1, x_764); +return x_768; +} +else +{ +lean_object* x_769; +lean_dec(x_765); +lean_dec(x_764); +x_769 = l_Std_Internal_Parsec_String_pstring(x_724, x_763); +if (lean_obj_tag(x_769) == 0) +{ +lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; +lean_dec(x_766); +x_770 = lean_ctor_get(x_769, 0); +lean_inc(x_770); +if (lean_is_exclusive(x_769)) { + lean_ctor_release(x_769, 0); + lean_ctor_release(x_769, 1); + x_771 = x_769; +} else { + lean_dec_ref(x_769); + x_771 = lean_box(0); +} +x_772 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_771)) { + x_773 = lean_alloc_ctor(0, 2, 0); +} else { + x_773 = x_771; +} +lean_ctor_set(x_773, 0, x_770); +lean_ctor_set(x_773, 1, x_772); +return x_773; +} +else +{ +lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; uint8_t x_778; +x_774 = lean_ctor_get(x_769, 0); +lean_inc(x_774); +x_775 = lean_ctor_get(x_769, 1); +lean_inc(x_775); +if (lean_is_exclusive(x_769)) { + lean_ctor_release(x_769, 0); + lean_ctor_release(x_769, 1); + x_776 = x_769; +} else { + lean_dec_ref(x_769); + x_776 = lean_box(0); +} +x_777 = lean_ctor_get(x_774, 1); +lean_inc(x_777); +x_778 = lean_nat_dec_eq(x_766, x_777); +lean_dec(x_766); +if (x_778 == 0) +{ +lean_object* x_779; +lean_dec(x_777); +if (lean_is_scalar(x_776)) { + x_779 = lean_alloc_ctor(1, 2, 0); +} else { + x_779 = x_776; +} +lean_ctor_set(x_779, 0, x_774); +lean_ctor_set(x_779, 1, x_775); +return x_779; +} +else +{ +lean_object* x_780; lean_object* x_781; +lean_dec(x_776); +lean_dec(x_775); +x_780 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_781 = l_Std_Internal_Parsec_String_pstring(x_780, x_774); +if (lean_obj_tag(x_781) == 0) +{ +lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; +lean_dec(x_777); +x_782 = lean_ctor_get(x_781, 0); +lean_inc(x_782); +if (lean_is_exclusive(x_781)) { + lean_ctor_release(x_781, 0); + lean_ctor_release(x_781, 1); + x_783 = x_781; +} else { + lean_dec_ref(x_781); + x_783 = lean_box(0); +} +x_784 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_783)) { + x_785 = lean_alloc_ctor(0, 2, 0); +} else { + x_785 = x_783; +} +lean_ctor_set(x_785, 0, x_782); +lean_ctor_set(x_785, 1, x_784); +return x_785; +} +else +{ +lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; uint8_t x_790; +x_786 = lean_ctor_get(x_781, 0); +lean_inc(x_786); +x_787 = lean_ctor_get(x_781, 1); +lean_inc(x_787); +if (lean_is_exclusive(x_781)) { + lean_ctor_release(x_781, 0); + lean_ctor_release(x_781, 1); + x_788 = x_781; +} else { + lean_dec_ref(x_781); + x_788 = lean_box(0); +} +x_789 = lean_ctor_get(x_786, 1); +lean_inc(x_789); +x_790 = lean_nat_dec_eq(x_777, x_789); +lean_dec(x_777); +if (x_790 == 0) +{ +lean_object* x_791; +lean_dec(x_789); +if (lean_is_scalar(x_788)) { + x_791 = lean_alloc_ctor(1, 2, 0); +} else { + x_791 = x_788; +} +lean_ctor_set(x_791, 0, x_786); +lean_ctor_set(x_791, 1, x_787); +return x_791; +} +else +{ +lean_object* x_792; lean_object* x_793; +lean_dec(x_788); +lean_dec(x_787); +x_792 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_793 = l_Std_Internal_Parsec_String_pstring(x_792, x_786); +if (lean_obj_tag(x_793) == 0) +{ +lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; +lean_dec(x_789); +x_794 = lean_ctor_get(x_793, 0); +lean_inc(x_794); +if (lean_is_exclusive(x_793)) { + lean_ctor_release(x_793, 0); + lean_ctor_release(x_793, 1); + x_795 = x_793; +} else { + lean_dec_ref(x_793); + x_795 = lean_box(0); +} +x_796 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_795)) { + x_797 = lean_alloc_ctor(0, 2, 0); +} else { + x_797 = x_795; +} +lean_ctor_set(x_797, 0, x_794); +lean_ctor_set(x_797, 1, x_796); +return x_797; +} +else +{ +lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; uint8_t x_802; +x_798 = lean_ctor_get(x_793, 0); +lean_inc(x_798); +x_799 = lean_ctor_get(x_793, 1); +lean_inc(x_799); +if (lean_is_exclusive(x_793)) { + lean_ctor_release(x_793, 0); + lean_ctor_release(x_793, 1); + x_800 = x_793; +} else { + lean_dec_ref(x_793); + x_800 = lean_box(0); +} +x_801 = lean_ctor_get(x_798, 1); +lean_inc(x_801); +x_802 = lean_nat_dec_eq(x_789, x_801); +lean_dec(x_789); +if (x_802 == 0) +{ +lean_object* x_803; +lean_dec(x_801); +if (lean_is_scalar(x_800)) { + x_803 = lean_alloc_ctor(1, 2, 0); +} else { + x_803 = x_800; +} +lean_ctor_set(x_803, 0, x_798); +lean_ctor_set(x_803, 1, x_799); +return x_803; +} +else +{ +lean_object* x_804; lean_object* x_805; +lean_dec(x_800); +lean_dec(x_799); +x_804 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_805 = l_Std_Internal_Parsec_String_pstring(x_804, x_798); +if (lean_obj_tag(x_805) == 0) +{ +lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; +lean_dec(x_801); +x_806 = lean_ctor_get(x_805, 0); +lean_inc(x_806); +if (lean_is_exclusive(x_805)) { + lean_ctor_release(x_805, 0); + lean_ctor_release(x_805, 1); + x_807 = x_805; +} else { + lean_dec_ref(x_805); + x_807 = lean_box(0); +} +x_808 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_807)) { + x_809 = lean_alloc_ctor(0, 2, 0); +} else { + x_809 = x_807; +} +lean_ctor_set(x_809, 0, x_806); +lean_ctor_set(x_809, 1, x_808); +return x_809; +} +else +{ +lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; uint8_t x_814; +x_810 = lean_ctor_get(x_805, 0); +lean_inc(x_810); +x_811 = lean_ctor_get(x_805, 1); +lean_inc(x_811); +if (lean_is_exclusive(x_805)) { + lean_ctor_release(x_805, 0); + lean_ctor_release(x_805, 1); + x_812 = x_805; +} else { + lean_dec_ref(x_805); + x_812 = lean_box(0); +} +x_813 = lean_ctor_get(x_810, 1); +lean_inc(x_813); +x_814 = lean_nat_dec_eq(x_801, x_813); +lean_dec(x_813); +lean_dec(x_801); +if (x_814 == 0) +{ +lean_object* x_815; +if (lean_is_scalar(x_812)) { + x_815 = lean_alloc_ctor(1, 2, 0); +} else { + x_815 = x_812; +} +lean_ctor_set(x_815, 0, x_810); +lean_ctor_set(x_815, 1, x_811); +return x_815; +} +else +{ +lean_object* x_816; lean_object* x_817; +lean_dec(x_812); +lean_dec(x_811); +x_816 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_817 = l_Std_Internal_Parsec_String_pstring(x_816, x_810); +if (lean_obj_tag(x_817) == 0) +{ +lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; +x_818 = lean_ctor_get(x_817, 0); +lean_inc(x_818); +if (lean_is_exclusive(x_817)) { + lean_ctor_release(x_817, 0); + lean_ctor_release(x_817, 1); + x_819 = x_817; +} else { + lean_dec_ref(x_817); + x_819 = lean_box(0); +} +x_820 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_819)) { + x_821 = lean_alloc_ctor(0, 2, 0); +} else { + x_821 = x_819; +} +lean_ctor_set(x_821, 0, x_818); +lean_ctor_set(x_821, 1, x_820); +return x_821; +} +else +{ +lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; +x_822 = lean_ctor_get(x_817, 0); +lean_inc(x_822); +x_823 = lean_ctor_get(x_817, 1); +lean_inc(x_823); +if (lean_is_exclusive(x_817)) { + lean_ctor_release(x_817, 0); + lean_ctor_release(x_817, 1); + x_824 = x_817; +} else { + lean_dec_ref(x_817); + x_824 = lean_box(0); +} +if (lean_is_scalar(x_824)) { + x_825 = lean_alloc_ctor(1, 2, 0); +} else { + x_825 = x_824; +} +lean_ctor_set(x_825, 0, x_822); +lean_ctor_set(x_825, 1, x_823); +return x_825; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; uint8_t x_830; +x_826 = lean_ctor_get(x_3, 0); +x_827 = lean_ctor_get(x_3, 1); +lean_inc(x_827); +lean_inc(x_826); +lean_dec(x_3); +x_828 = lean_ctor_get(x_1, 1); +lean_inc(x_828); +lean_dec(x_1); +x_829 = lean_ctor_get(x_826, 1); +lean_inc(x_829); +x_830 = lean_nat_dec_eq(x_828, x_829); +lean_dec(x_828); +if (x_830 == 0) +{ +lean_object* x_831; +lean_dec(x_829); +x_831 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_831, 0, x_826); +lean_ctor_set(x_831, 1, x_827); +return x_831; +} +else +{ +lean_object* x_832; lean_object* x_833; +lean_dec(x_827); +x_832 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_833 = l_Std_Internal_Parsec_String_pstring(x_832, x_826); +if (lean_obj_tag(x_833) == 0) +{ +lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; +lean_dec(x_829); +x_834 = lean_ctor_get(x_833, 0); +lean_inc(x_834); +if (lean_is_exclusive(x_833)) { + lean_ctor_release(x_833, 0); + lean_ctor_release(x_833, 1); + x_835 = x_833; +} else { + lean_dec_ref(x_833); + x_835 = lean_box(0); +} +x_836 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_835)) { + x_837 = lean_alloc_ctor(0, 2, 0); +} else { + x_837 = x_835; +} +lean_ctor_set(x_837, 0, x_834); +lean_ctor_set(x_837, 1, x_836); +return x_837; +} +else +{ +lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; uint8_t x_842; +x_838 = lean_ctor_get(x_833, 0); +lean_inc(x_838); +x_839 = lean_ctor_get(x_833, 1); +lean_inc(x_839); +if (lean_is_exclusive(x_833)) { + lean_ctor_release(x_833, 0); + lean_ctor_release(x_833, 1); + x_840 = x_833; +} else { + lean_dec_ref(x_833); + x_840 = lean_box(0); +} +x_841 = lean_ctor_get(x_838, 1); +lean_inc(x_841); +x_842 = lean_nat_dec_eq(x_829, x_841); +lean_dec(x_829); +if (x_842 == 0) +{ +lean_object* x_843; +lean_dec(x_841); +if (lean_is_scalar(x_840)) { + x_843 = lean_alloc_ctor(1, 2, 0); +} else { + x_843 = x_840; +} +lean_ctor_set(x_843, 0, x_838); +lean_ctor_set(x_843, 1, x_839); +return x_843; +} +else +{ +lean_object* x_844; lean_object* x_845; +lean_dec(x_840); +lean_dec(x_839); +x_844 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +x_845 = l_Std_Internal_Parsec_String_pstring(x_844, x_838); +if (lean_obj_tag(x_845) == 0) +{ +lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; +lean_dec(x_841); +x_846 = lean_ctor_get(x_845, 0); +lean_inc(x_846); +if (lean_is_exclusive(x_845)) { + lean_ctor_release(x_845, 0); + lean_ctor_release(x_845, 1); + x_847 = x_845; +} else { + lean_dec_ref(x_845); + x_847 = lean_box(0); +} +x_848 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_847)) { + x_849 = lean_alloc_ctor(0, 2, 0); +} else { + x_849 = x_847; +} +lean_ctor_set(x_849, 0, x_846); +lean_ctor_set(x_849, 1, x_848); +return x_849; +} +else +{ +lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; uint8_t x_854; +x_850 = lean_ctor_get(x_845, 0); +lean_inc(x_850); +x_851 = lean_ctor_get(x_845, 1); +lean_inc(x_851); +if (lean_is_exclusive(x_845)) { + lean_ctor_release(x_845, 0); + lean_ctor_release(x_845, 1); + x_852 = x_845; +} else { + lean_dec_ref(x_845); + x_852 = lean_box(0); +} +x_853 = lean_ctor_get(x_850, 1); +lean_inc(x_853); +x_854 = lean_nat_dec_eq(x_841, x_853); +lean_dec(x_841); +if (x_854 == 0) +{ +lean_object* x_855; +lean_dec(x_853); +if (lean_is_scalar(x_852)) { + x_855 = lean_alloc_ctor(1, 2, 0); +} else { + x_855 = x_852; +} +lean_ctor_set(x_855, 0, x_850); +lean_ctor_set(x_855, 1, x_851); +return x_855; +} +else +{ +lean_object* x_856; lean_object* x_857; +lean_dec(x_852); +lean_dec(x_851); +x_856 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +x_857 = l_Std_Internal_Parsec_String_pstring(x_856, x_850); +if (lean_obj_tag(x_857) == 0) +{ +lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; +lean_dec(x_853); +x_858 = lean_ctor_get(x_857, 0); +lean_inc(x_858); +if (lean_is_exclusive(x_857)) { + lean_ctor_release(x_857, 0); + lean_ctor_release(x_857, 1); + x_859 = x_857; +} else { + lean_dec_ref(x_857); + x_859 = lean_box(0); +} +x_860 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_859)) { + x_861 = lean_alloc_ctor(0, 2, 0); +} else { + x_861 = x_859; +} +lean_ctor_set(x_861, 0, x_858); +lean_ctor_set(x_861, 1, x_860); +return x_861; +} +else +{ +lean_object* x_862; lean_object* x_863; lean_object* x_864; lean_object* x_865; uint8_t x_866; +x_862 = lean_ctor_get(x_857, 0); +lean_inc(x_862); +x_863 = lean_ctor_get(x_857, 1); +lean_inc(x_863); +if (lean_is_exclusive(x_857)) { + lean_ctor_release(x_857, 0); + lean_ctor_release(x_857, 1); + x_864 = x_857; +} else { + lean_dec_ref(x_857); + x_864 = lean_box(0); +} +x_865 = lean_ctor_get(x_862, 1); +lean_inc(x_865); +x_866 = lean_nat_dec_eq(x_853, x_865); +lean_dec(x_853); +if (x_866 == 0) +{ +lean_object* x_867; +lean_dec(x_865); +if (lean_is_scalar(x_864)) { + x_867 = lean_alloc_ctor(1, 2, 0); +} else { + x_867 = x_864; +} +lean_ctor_set(x_867, 0, x_862); +lean_ctor_set(x_867, 1, x_863); +return x_867; +} +else +{ +lean_object* x_868; +lean_dec(x_864); +lean_dec(x_863); +x_868 = l_Std_Internal_Parsec_String_pstring(x_844, x_862); +if (lean_obj_tag(x_868) == 0) +{ +lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; +lean_dec(x_865); +x_869 = lean_ctor_get(x_868, 0); +lean_inc(x_869); +if (lean_is_exclusive(x_868)) { + lean_ctor_release(x_868, 0); + lean_ctor_release(x_868, 1); + x_870 = x_868; +} else { + lean_dec_ref(x_868); + x_870 = lean_box(0); +} +x_871 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2; +if (lean_is_scalar(x_870)) { + x_872 = lean_alloc_ctor(0, 2, 0); +} else { + x_872 = x_870; +} +lean_ctor_set(x_872, 0, x_869); +lean_ctor_set(x_872, 1, x_871); +return x_872; +} +else +{ +lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; uint8_t x_877; +x_873 = lean_ctor_get(x_868, 0); +lean_inc(x_873); +x_874 = lean_ctor_get(x_868, 1); +lean_inc(x_874); +if (lean_is_exclusive(x_868)) { + lean_ctor_release(x_868, 0); + lean_ctor_release(x_868, 1); + x_875 = x_868; +} else { + lean_dec_ref(x_868); + x_875 = lean_box(0); +} +x_876 = lean_ctor_get(x_873, 1); +lean_inc(x_876); +x_877 = lean_nat_dec_eq(x_865, x_876); +lean_dec(x_865); +if (x_877 == 0) +{ +lean_object* x_878; +lean_dec(x_876); +if (lean_is_scalar(x_875)) { + x_878 = lean_alloc_ctor(1, 2, 0); +} else { + x_878 = x_875; +} +lean_ctor_set(x_878, 0, x_873); +lean_ctor_set(x_878, 1, x_874); +return x_878; +} +else +{ +lean_object* x_879; +lean_dec(x_875); +lean_dec(x_874); +x_879 = l_Std_Internal_Parsec_String_pstring(x_2, x_873); +if (lean_obj_tag(x_879) == 0) +{ +lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; +lean_dec(x_876); +x_880 = lean_ctor_get(x_879, 0); +lean_inc(x_880); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_881 = x_879; +} else { + lean_dec_ref(x_879); + x_881 = lean_box(0); +} +x_882 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3; +if (lean_is_scalar(x_881)) { + x_883 = lean_alloc_ctor(0, 2, 0); +} else { + x_883 = x_881; +} +lean_ctor_set(x_883, 0, x_880); +lean_ctor_set(x_883, 1, x_882); +return x_883; +} +else +{ +lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; uint8_t x_888; +x_884 = lean_ctor_get(x_879, 0); +lean_inc(x_884); +x_885 = lean_ctor_get(x_879, 1); +lean_inc(x_885); +if (lean_is_exclusive(x_879)) { + lean_ctor_release(x_879, 0); + lean_ctor_release(x_879, 1); + x_886 = x_879; +} else { + lean_dec_ref(x_879); + x_886 = lean_box(0); +} +x_887 = lean_ctor_get(x_884, 1); +lean_inc(x_887); +x_888 = lean_nat_dec_eq(x_876, x_887); +lean_dec(x_876); +if (x_888 == 0) +{ +lean_object* x_889; +lean_dec(x_887); +if (lean_is_scalar(x_886)) { + x_889 = lean_alloc_ctor(1, 2, 0); +} else { + x_889 = x_886; +} +lean_ctor_set(x_889, 0, x_884); +lean_ctor_set(x_889, 1, x_885); +return x_889; +} +else +{ +lean_object* x_890; +lean_dec(x_886); +lean_dec(x_885); +x_890 = l_Std_Internal_Parsec_String_pstring(x_2, x_884); +if (lean_obj_tag(x_890) == 0) +{ +lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; +lean_dec(x_887); +x_891 = lean_ctor_get(x_890, 0); +lean_inc(x_891); +if (lean_is_exclusive(x_890)) { + lean_ctor_release(x_890, 0); + lean_ctor_release(x_890, 1); + x_892 = x_890; +} else { + lean_dec_ref(x_890); + x_892 = lean_box(0); +} +x_893 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4; +if (lean_is_scalar(x_892)) { + x_894 = lean_alloc_ctor(0, 2, 0); +} else { + x_894 = x_892; +} +lean_ctor_set(x_894, 0, x_891); +lean_ctor_set(x_894, 1, x_893); +return x_894; +} +else +{ +lean_object* x_895; lean_object* x_896; lean_object* x_897; lean_object* x_898; uint8_t x_899; +x_895 = lean_ctor_get(x_890, 0); +lean_inc(x_895); +x_896 = lean_ctor_get(x_890, 1); +lean_inc(x_896); +if (lean_is_exclusive(x_890)) { + lean_ctor_release(x_890, 0); + lean_ctor_release(x_890, 1); + x_897 = x_890; +} else { + lean_dec_ref(x_890); + x_897 = lean_box(0); +} +x_898 = lean_ctor_get(x_895, 1); +lean_inc(x_898); +x_899 = lean_nat_dec_eq(x_887, x_898); +lean_dec(x_887); +if (x_899 == 0) +{ +lean_object* x_900; +lean_dec(x_898); +if (lean_is_scalar(x_897)) { + x_900 = lean_alloc_ctor(1, 2, 0); +} else { + x_900 = x_897; +} +lean_ctor_set(x_900, 0, x_895); +lean_ctor_set(x_900, 1, x_896); +return x_900; +} +else +{ +lean_object* x_901; +lean_dec(x_897); +lean_dec(x_896); +x_901 = l_Std_Internal_Parsec_String_pstring(x_856, x_895); +if (lean_obj_tag(x_901) == 0) +{ +lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; +lean_dec(x_898); +x_902 = lean_ctor_get(x_901, 0); +lean_inc(x_902); +if (lean_is_exclusive(x_901)) { + lean_ctor_release(x_901, 0); + lean_ctor_release(x_901, 1); + x_903 = x_901; +} else { + lean_dec_ref(x_901); + x_903 = lean_box(0); +} +x_904 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5; +if (lean_is_scalar(x_903)) { + x_905 = lean_alloc_ctor(0, 2, 0); +} else { + x_905 = x_903; +} +lean_ctor_set(x_905, 0, x_902); +lean_ctor_set(x_905, 1, x_904); +return x_905; +} +else +{ +lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; uint8_t x_910; +x_906 = lean_ctor_get(x_901, 0); +lean_inc(x_906); +x_907 = lean_ctor_get(x_901, 1); +lean_inc(x_907); +if (lean_is_exclusive(x_901)) { + lean_ctor_release(x_901, 0); + lean_ctor_release(x_901, 1); + x_908 = x_901; +} else { + lean_dec_ref(x_901); + x_908 = lean_box(0); +} +x_909 = lean_ctor_get(x_906, 1); +lean_inc(x_909); +x_910 = lean_nat_dec_eq(x_898, x_909); +lean_dec(x_898); +if (x_910 == 0) +{ +lean_object* x_911; +lean_dec(x_909); +if (lean_is_scalar(x_908)) { + x_911 = lean_alloc_ctor(1, 2, 0); +} else { + x_911 = x_908; +} +lean_ctor_set(x_911, 0, x_906); +lean_ctor_set(x_911, 1, x_907); +return x_911; +} +else +{ +lean_object* x_912; lean_object* x_913; +lean_dec(x_908); +lean_dec(x_907); +x_912 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +x_913 = l_Std_Internal_Parsec_String_pstring(x_912, x_906); +if (lean_obj_tag(x_913) == 0) +{ +lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; +lean_dec(x_909); +x_914 = lean_ctor_get(x_913, 0); +lean_inc(x_914); +if (lean_is_exclusive(x_913)) { + lean_ctor_release(x_913, 0); + lean_ctor_release(x_913, 1); + x_915 = x_913; +} else { + lean_dec_ref(x_913); + x_915 = lean_box(0); +} +x_916 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6; +if (lean_is_scalar(x_915)) { + x_917 = lean_alloc_ctor(0, 2, 0); +} else { + x_917 = x_915; +} +lean_ctor_set(x_917, 0, x_914); +lean_ctor_set(x_917, 1, x_916); +return x_917; +} +else +{ +lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; uint8_t x_922; +x_918 = lean_ctor_get(x_913, 0); +lean_inc(x_918); +x_919 = lean_ctor_get(x_913, 1); +lean_inc(x_919); +if (lean_is_exclusive(x_913)) { + lean_ctor_release(x_913, 0); + lean_ctor_release(x_913, 1); + x_920 = x_913; +} else { + lean_dec_ref(x_913); + x_920 = lean_box(0); +} +x_921 = lean_ctor_get(x_918, 1); +lean_inc(x_921); +x_922 = lean_nat_dec_eq(x_909, x_921); +lean_dec(x_909); +if (x_922 == 0) +{ +lean_object* x_923; +lean_dec(x_921); +if (lean_is_scalar(x_920)) { + x_923 = lean_alloc_ctor(1, 2, 0); +} else { + x_923 = x_920; +} +lean_ctor_set(x_923, 0, x_918); +lean_ctor_set(x_923, 1, x_919); +return x_923; +} +else +{ +lean_object* x_924; lean_object* x_925; +lean_dec(x_920); +lean_dec(x_919); +x_924 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3; +x_925 = l_Std_Internal_Parsec_String_pstring(x_924, x_918); +if (lean_obj_tag(x_925) == 0) +{ +lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; +lean_dec(x_921); +x_926 = lean_ctor_get(x_925, 0); +lean_inc(x_926); +if (lean_is_exclusive(x_925)) { + lean_ctor_release(x_925, 0); + lean_ctor_release(x_925, 1); + x_927 = x_925; +} else { + lean_dec_ref(x_925); + x_927 = lean_box(0); +} +x_928 = l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5; +if (lean_is_scalar(x_927)) { + x_929 = lean_alloc_ctor(0, 2, 0); +} else { + x_929 = x_927; +} +lean_ctor_set(x_929, 0, x_926); +lean_ctor_set(x_929, 1, x_928); +return x_929; +} +else +{ +lean_object* x_930; lean_object* x_931; lean_object* x_932; lean_object* x_933; uint8_t x_934; +x_930 = lean_ctor_get(x_925, 0); +lean_inc(x_930); +x_931 = lean_ctor_get(x_925, 1); +lean_inc(x_931); +if (lean_is_exclusive(x_925)) { + lean_ctor_release(x_925, 0); + lean_ctor_release(x_925, 1); + x_932 = x_925; +} else { + lean_dec_ref(x_925); + x_932 = lean_box(0); +} +x_933 = lean_ctor_get(x_930, 1); +lean_inc(x_933); +x_934 = lean_nat_dec_eq(x_921, x_933); +lean_dec(x_921); +if (x_934 == 0) +{ +lean_object* x_935; +lean_dec(x_933); +if (lean_is_scalar(x_932)) { + x_935 = lean_alloc_ctor(1, 2, 0); +} else { + x_935 = x_932; +} +lean_ctor_set(x_935, 0, x_930); +lean_ctor_set(x_935, 1, x_931); +return x_935; +} +else +{ +lean_object* x_936; lean_object* x_937; +lean_dec(x_932); +lean_dec(x_931); +x_936 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2; +x_937 = l_Std_Internal_Parsec_String_pstring(x_936, x_930); +if (lean_obj_tag(x_937) == 0) +{ +lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; +lean_dec(x_933); +x_938 = lean_ctor_get(x_937, 0); +lean_inc(x_938); +if (lean_is_exclusive(x_937)) { + lean_ctor_release(x_937, 0); + lean_ctor_release(x_937, 1); + x_939 = x_937; +} else { + lean_dec_ref(x_937); + x_939 = lean_box(0); +} +x_940 = l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7; +if (lean_is_scalar(x_939)) { + x_941 = lean_alloc_ctor(0, 2, 0); +} else { + x_941 = x_939; +} +lean_ctor_set(x_941, 0, x_938); +lean_ctor_set(x_941, 1, x_940); +return x_941; +} +else +{ +lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; uint8_t x_946; +x_942 = lean_ctor_get(x_937, 0); +lean_inc(x_942); +x_943 = lean_ctor_get(x_937, 1); +lean_inc(x_943); +if (lean_is_exclusive(x_937)) { + lean_ctor_release(x_937, 0); + lean_ctor_release(x_937, 1); + x_944 = x_937; +} else { + lean_dec_ref(x_937); + x_944 = lean_box(0); +} +x_945 = lean_ctor_get(x_942, 1); +lean_inc(x_945); +x_946 = lean_nat_dec_eq(x_933, x_945); +lean_dec(x_945); +lean_dec(x_933); +if (x_946 == 0) +{ +lean_object* x_947; +if (lean_is_scalar(x_944)) { + x_947 = lean_alloc_ctor(1, 2, 0); +} else { + x_947 = x_944; +} +lean_ctor_set(x_947, 0, x_942); +lean_ctor_set(x_947, 1, x_943); +return x_947; +} +else +{ +lean_object* x_948; lean_object* x_949; +lean_dec(x_944); +lean_dec(x_943); +x_948 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1; +x_949 = l_Std_Internal_Parsec_String_pstring(x_948, x_942); +if (lean_obj_tag(x_949) == 0) +{ +lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; +x_950 = lean_ctor_get(x_949, 0); +lean_inc(x_950); +if (lean_is_exclusive(x_949)) { + lean_ctor_release(x_949, 0); + lean_ctor_release(x_949, 1); + x_951 = x_949; +} else { + lean_dec_ref(x_949); + x_951 = lean_box(0); +} +x_952 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2; +if (lean_is_scalar(x_951)) { + x_953 = lean_alloc_ctor(0, 2, 0); +} else { + x_953 = x_951; +} +lean_ctor_set(x_953, 0, x_950); +lean_ctor_set(x_953, 1, x_952); +return x_953; +} +else +{ +lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; +x_954 = lean_ctor_get(x_949, 0); +lean_inc(x_954); +x_955 = lean_ctor_get(x_949, 1); +lean_inc(x_955); +if (lean_is_exclusive(x_949)) { + lean_ctor_release(x_949, 0); + lean_ctor_release(x_949, 1); + x_956 = x_949; +} else { + lean_dec_ref(x_949); + x_956 = lean_box(0); +} +if (lean_is_scalar(x_956)) { + x_957 = lean_alloc_ctor(1, 2, 0); +} else { + x_957 = x_956; +} +lean_ctor_set(x_957, 0, x_954); +lean_ctor_set(x_957, 1, x_955); +return x_957; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 6; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 6; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_16); +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 0; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +x_32 = lean_nat_dec_eq(x_16, x_31); +lean_dec(x_16); +if (x_32 == 0) +{ +lean_dec(x_31); +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_free_object(x_19); +lean_dec(x_30); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2; +x_34 = l_Std_Internal_Parsec_String_pstring(x_33, x_29); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +lean_dec(x_31); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_34, 1); +lean_dec(x_36); +x_37 = 1; +x_38 = lean_box(x_37); +lean_ctor_set(x_34, 1, x_38); +return x_34; +} +else +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_34, 0); +lean_inc(x_39); +lean_dec(x_34); +x_40 = 1; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +x_47 = lean_nat_dec_eq(x_31, x_46); +lean_dec(x_31); +if (x_47 == 0) +{ +lean_dec(x_46); +return x_34; +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_free_object(x_34); +lean_dec(x_45); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +x_49 = l_Std_Internal_Parsec_String_pstring(x_48, x_44); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +lean_dec(x_46); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_49, 1); +lean_dec(x_51); +x_52 = 2; +x_53 = lean_box(x_52); +lean_ctor_set(x_49, 1, x_53); +return x_49; +} +else +{ +lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_49, 0); +lean_inc(x_54); +lean_dec(x_49); +x_55 = 2; +x_56 = lean_box(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +else +{ +uint8_t x_58; +x_58 = !lean_is_exclusive(x_49); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_59 = lean_ctor_get(x_49, 0); +x_60 = lean_ctor_get(x_49, 1); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +x_62 = lean_nat_dec_eq(x_46, x_61); +lean_dec(x_46); +if (x_62 == 0) +{ +lean_dec(x_61); +return x_49; +} +else +{ +lean_object* x_63; lean_object* x_64; +lean_free_object(x_49); +lean_dec(x_60); +x_63 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +x_64 = l_Std_Internal_Parsec_String_pstring(x_63, x_59); +if (lean_obj_tag(x_64) == 0) +{ +uint8_t x_65; +lean_dec(x_61); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; uint8_t x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_64, 1); +lean_dec(x_66); +x_67 = 3; +x_68 = lean_box(x_67); +lean_ctor_set(x_64, 1, x_68); +return x_64; +} +else +{ +lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_64, 0); +lean_inc(x_69); +lean_dec(x_64); +x_70 = 3; +x_71 = lean_box(x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} +} +else +{ +uint8_t x_73; +x_73 = !lean_is_exclusive(x_64); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = lean_ctor_get(x_64, 0); +x_75 = lean_ctor_get(x_64, 1); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +x_77 = lean_nat_dec_eq(x_61, x_76); +lean_dec(x_61); +if (x_77 == 0) +{ +lean_dec(x_76); +return x_64; +} +else +{ +lean_object* x_78; lean_object* x_79; +lean_free_object(x_64); +lean_dec(x_75); +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_79 = l_Std_Internal_Parsec_String_pstring(x_78, x_74); +if (lean_obj_tag(x_79) == 0) +{ +uint8_t x_80; +lean_dec(x_76); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) +{ +lean_object* x_81; uint8_t x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_79, 1); +lean_dec(x_81); +x_82 = 4; +x_83 = lean_box(x_82); +lean_ctor_set(x_79, 1, x_83); +return x_79; +} +else +{ +lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +lean_dec(x_79); +x_85 = 4; +x_86 = lean_box(x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_79); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_79, 0); +x_90 = lean_ctor_get(x_79, 1); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +x_92 = lean_nat_dec_eq(x_76, x_91); +lean_dec(x_91); +lean_dec(x_76); +if (x_92 == 0) +{ +return x_79; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_free_object(x_79); +lean_dec(x_90); +x_93 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_94 = l_Std_Internal_Parsec_String_pstring(x_93, x_89); +if (lean_obj_tag(x_94) == 0) +{ +uint8_t x_95; +x_95 = !lean_is_exclusive(x_94); +if (x_95 == 0) +{ +lean_object* x_96; uint8_t x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_94, 1); +lean_dec(x_96); +x_97 = 5; +x_98 = lean_box(x_97); +lean_ctor_set(x_94, 1, x_98); +return x_94; +} +else +{ +lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_94, 0); +lean_inc(x_99); +lean_dec(x_94); +x_100 = 5; +x_101 = lean_box(x_100); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_101); +return x_102; +} +} +else +{ +uint8_t x_103; +x_103 = !lean_is_exclusive(x_94); +if (x_103 == 0) +{ +return x_94; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_94, 0); +x_105 = lean_ctor_get(x_94, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_94); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_107 = lean_ctor_get(x_79, 0); +x_108 = lean_ctor_get(x_79, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_79); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +x_110 = lean_nat_dec_eq(x_76, x_109); +lean_dec(x_109); +lean_dec(x_76); +if (x_110 == 0) +{ +lean_object* x_111; +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_107); +lean_ctor_set(x_111, 1, x_108); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; +lean_dec(x_108); +x_112 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_113 = l_Std_Internal_Parsec_String_pstring(x_112, x_107); +if (lean_obj_tag(x_113) == 0) +{ +lean_object* x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; lean_object* x_118; +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_115 = x_113; +} else { + lean_dec_ref(x_113); + x_115 = lean_box(0); +} +x_116 = 5; +x_117 = lean_box(x_116); +if (lean_is_scalar(x_115)) { + x_118 = lean_alloc_ctor(0, 2, 0); +} else { + x_118 = x_115; +} +lean_ctor_set(x_118, 0, x_114); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_119 = lean_ctor_get(x_113, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_113, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_121 = x_113; +} else { + lean_dec_ref(x_113); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +} +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_123 = lean_ctor_get(x_64, 0); +x_124 = lean_ctor_get(x_64, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_64); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +x_126 = lean_nat_dec_eq(x_61, x_125); +lean_dec(x_61); +if (x_126 == 0) +{ +lean_object* x_127; +lean_dec(x_125); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_123); +lean_ctor_set(x_127, 1, x_124); +return x_127; +} +else +{ +lean_object* x_128; lean_object* x_129; +lean_dec(x_124); +x_128 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_129 = l_Std_Internal_Parsec_String_pstring(x_128, x_123); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_125); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_131 = x_129; +} else { + lean_dec_ref(x_129); + x_131 = lean_box(0); +} +x_132 = 4; +x_133 = lean_box(x_132); +if (lean_is_scalar(x_131)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_131; +} +lean_ctor_set(x_134, 0, x_130); +lean_ctor_set(x_134, 1, x_133); +return x_134; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_135 = lean_ctor_get(x_129, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_129, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_137 = x_129; +} else { + lean_dec_ref(x_129); + x_137 = lean_box(0); +} +x_138 = lean_ctor_get(x_135, 1); +lean_inc(x_138); +x_139 = lean_nat_dec_eq(x_125, x_138); +lean_dec(x_138); +lean_dec(x_125); +if (x_139 == 0) +{ +lean_object* x_140; +if (lean_is_scalar(x_137)) { + x_140 = lean_alloc_ctor(1, 2, 0); +} else { + x_140 = x_137; +} +lean_ctor_set(x_140, 0, x_135); +lean_ctor_set(x_140, 1, x_136); +return x_140; +} +else +{ +lean_object* x_141; lean_object* x_142; +lean_dec(x_137); +lean_dec(x_136); +x_141 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_142 = l_Std_Internal_Parsec_String_pstring(x_141, x_135); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_144 = x_142; +} else { + lean_dec_ref(x_142); + x_144 = lean_box(0); +} +x_145 = 5; +x_146 = lean_box(x_145); +if (lean_is_scalar(x_144)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_144; +} +lean_ctor_set(x_147, 0, x_143); +lean_ctor_set(x_147, 1, x_146); +return x_147; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_148 = lean_ctor_get(x_142, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_142, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_150 = x_142; +} else { + lean_dec_ref(x_142); + x_150 = lean_box(0); +} +if (lean_is_scalar(x_150)) { + x_151 = lean_alloc_ctor(1, 2, 0); +} else { + x_151 = x_150; +} +lean_ctor_set(x_151, 0, x_148); +lean_ctor_set(x_151, 1, x_149); +return x_151; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_152 = lean_ctor_get(x_49, 0); +x_153 = lean_ctor_get(x_49, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_49); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +x_155 = lean_nat_dec_eq(x_46, x_154); +lean_dec(x_46); +if (x_155 == 0) +{ +lean_object* x_156; +lean_dec(x_154); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_152); +lean_ctor_set(x_156, 1, x_153); +return x_156; +} +else +{ +lean_object* x_157; lean_object* x_158; +lean_dec(x_153); +x_157 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +x_158 = l_Std_Internal_Parsec_String_pstring(x_157, x_152); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_154); +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_160 = x_158; +} else { + lean_dec_ref(x_158); + x_160 = lean_box(0); +} +x_161 = 3; +x_162 = lean_box(x_161); +if (lean_is_scalar(x_160)) { + x_163 = lean_alloc_ctor(0, 2, 0); +} else { + x_163 = x_160; +} +lean_ctor_set(x_163, 0, x_159); +lean_ctor_set(x_163, 1, x_162); +return x_163; +} +else +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +x_164 = lean_ctor_get(x_158, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_158, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_166 = x_158; +} else { + lean_dec_ref(x_158); + x_166 = lean_box(0); +} +x_167 = lean_ctor_get(x_164, 1); +lean_inc(x_167); +x_168 = lean_nat_dec_eq(x_154, x_167); +lean_dec(x_154); +if (x_168 == 0) +{ +lean_object* x_169; +lean_dec(x_167); +if (lean_is_scalar(x_166)) { + x_169 = lean_alloc_ctor(1, 2, 0); +} else { + x_169 = x_166; +} +lean_ctor_set(x_169, 0, x_164); +lean_ctor_set(x_169, 1, x_165); +return x_169; +} +else +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_166); +lean_dec(x_165); +x_170 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_171 = l_Std_Internal_Parsec_String_pstring(x_170, x_164); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_167); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_173 = x_171; +} else { + lean_dec_ref(x_171); + x_173 = lean_box(0); +} +x_174 = 4; +x_175 = lean_box(x_174); +if (lean_is_scalar(x_173)) { + x_176 = lean_alloc_ctor(0, 2, 0); +} else { + x_176 = x_173; +} +lean_ctor_set(x_176, 0, x_172); +lean_ctor_set(x_176, 1, x_175); +return x_176; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; +x_177 = lean_ctor_get(x_171, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_171, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_179 = x_171; +} else { + lean_dec_ref(x_171); + x_179 = lean_box(0); +} +x_180 = lean_ctor_get(x_177, 1); +lean_inc(x_180); +x_181 = lean_nat_dec_eq(x_167, x_180); +lean_dec(x_180); +lean_dec(x_167); +if (x_181 == 0) +{ +lean_object* x_182; +if (lean_is_scalar(x_179)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_179; +} +lean_ctor_set(x_182, 0, x_177); +lean_ctor_set(x_182, 1, x_178); +return x_182; +} +else +{ +lean_object* x_183; lean_object* x_184; +lean_dec(x_179); +lean_dec(x_178); +x_183 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_184 = l_Std_Internal_Parsec_String_pstring(x_183, x_177); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = 5; +x_188 = lean_box(x_187); +if (lean_is_scalar(x_186)) { + x_189 = lean_alloc_ctor(0, 2, 0); +} else { + x_189 = x_186; +} +lean_ctor_set(x_189, 0, x_185); +lean_ctor_set(x_189, 1, x_188); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_190 = lean_ctor_get(x_184, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_184, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_192 = x_184; +} else { + lean_dec_ref(x_184); + x_192 = lean_box(0); +} +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); +} else { + x_193 = x_192; +} +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; +x_194 = lean_ctor_get(x_34, 0); +x_195 = lean_ctor_get(x_34, 1); +lean_inc(x_195); +lean_inc(x_194); +lean_dec(x_34); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +x_197 = lean_nat_dec_eq(x_31, x_196); +lean_dec(x_31); +if (x_197 == 0) +{ +lean_object* x_198; +lean_dec(x_196); +x_198 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_198, 0, x_194); +lean_ctor_set(x_198, 1, x_195); +return x_198; +} +else +{ +lean_object* x_199; lean_object* x_200; +lean_dec(x_195); +x_199 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +x_200 = l_Std_Internal_Parsec_String_pstring(x_199, x_194); +if (lean_obj_tag(x_200) == 0) +{ +lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_196); +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_202 = x_200; +} else { + lean_dec_ref(x_200); + x_202 = lean_box(0); +} +x_203 = 2; +x_204 = lean_box(x_203); +if (lean_is_scalar(x_202)) { + x_205 = lean_alloc_ctor(0, 2, 0); +} else { + x_205 = x_202; +} +lean_ctor_set(x_205, 0, x_201); +lean_ctor_set(x_205, 1, x_204); +return x_205; +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; uint8_t x_210; +x_206 = lean_ctor_get(x_200, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_200, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_208 = x_200; +} else { + lean_dec_ref(x_200); + x_208 = lean_box(0); +} +x_209 = lean_ctor_get(x_206, 1); +lean_inc(x_209); +x_210 = lean_nat_dec_eq(x_196, x_209); +lean_dec(x_196); +if (x_210 == 0) +{ +lean_object* x_211; +lean_dec(x_209); +if (lean_is_scalar(x_208)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_208; +} +lean_ctor_set(x_211, 0, x_206); +lean_ctor_set(x_211, 1, x_207); +return x_211; +} +else +{ +lean_object* x_212; lean_object* x_213; +lean_dec(x_208); +lean_dec(x_207); +x_212 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +x_213 = l_Std_Internal_Parsec_String_pstring(x_212, x_206); +if (lean_obj_tag(x_213) == 0) +{ +lean_object* x_214; lean_object* x_215; uint8_t x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_209); +x_214 = lean_ctor_get(x_213, 0); +lean_inc(x_214); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_215 = x_213; +} else { + lean_dec_ref(x_213); + x_215 = lean_box(0); +} +x_216 = 3; +x_217 = lean_box(x_216); +if (lean_is_scalar(x_215)) { + x_218 = lean_alloc_ctor(0, 2, 0); +} else { + x_218 = x_215; +} +lean_ctor_set(x_218, 0, x_214); +lean_ctor_set(x_218, 1, x_217); +return x_218; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_219 = lean_ctor_get(x_213, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_213, 1); +lean_inc(x_220); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_221 = x_213; +} else { + lean_dec_ref(x_213); + x_221 = lean_box(0); +} +x_222 = lean_ctor_get(x_219, 1); +lean_inc(x_222); +x_223 = lean_nat_dec_eq(x_209, x_222); +lean_dec(x_209); +if (x_223 == 0) +{ +lean_object* x_224; +lean_dec(x_222); +if (lean_is_scalar(x_221)) { + x_224 = lean_alloc_ctor(1, 2, 0); +} else { + x_224 = x_221; +} +lean_ctor_set(x_224, 0, x_219); +lean_ctor_set(x_224, 1, x_220); +return x_224; +} +else +{ +lean_object* x_225; lean_object* x_226; +lean_dec(x_221); +lean_dec(x_220); +x_225 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_226 = l_Std_Internal_Parsec_String_pstring(x_225, x_219); +if (lean_obj_tag(x_226) == 0) +{ +lean_object* x_227; lean_object* x_228; uint8_t x_229; lean_object* x_230; lean_object* x_231; +lean_dec(x_222); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_228 = x_226; +} else { + lean_dec_ref(x_226); + x_228 = lean_box(0); +} +x_229 = 4; +x_230 = lean_box(x_229); +if (lean_is_scalar(x_228)) { + x_231 = lean_alloc_ctor(0, 2, 0); +} else { + x_231 = x_228; +} +lean_ctor_set(x_231, 0, x_227); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; uint8_t x_236; +x_232 = lean_ctor_get(x_226, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_226, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_234 = x_226; +} else { + lean_dec_ref(x_226); + x_234 = lean_box(0); +} +x_235 = lean_ctor_get(x_232, 1); +lean_inc(x_235); +x_236 = lean_nat_dec_eq(x_222, x_235); +lean_dec(x_235); +lean_dec(x_222); +if (x_236 == 0) +{ +lean_object* x_237; +if (lean_is_scalar(x_234)) { + x_237 = lean_alloc_ctor(1, 2, 0); +} else { + x_237 = x_234; +} +lean_ctor_set(x_237, 0, x_232); +lean_ctor_set(x_237, 1, x_233); +return x_237; +} +else +{ +lean_object* x_238; lean_object* x_239; +lean_dec(x_234); +lean_dec(x_233); +x_238 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_239 = l_Std_Internal_Parsec_String_pstring(x_238, x_232); +if (lean_obj_tag(x_239) == 0) +{ +lean_object* x_240; lean_object* x_241; uint8_t x_242; lean_object* x_243; lean_object* x_244; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_241 = x_239; +} else { + lean_dec_ref(x_239); + x_241 = lean_box(0); +} +x_242 = 5; +x_243 = lean_box(x_242); +if (lean_is_scalar(x_241)) { + x_244 = lean_alloc_ctor(0, 2, 0); +} else { + x_244 = x_241; +} +lean_ctor_set(x_244, 0, x_240); +lean_ctor_set(x_244, 1, x_243); +return x_244; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_245 = lean_ctor_get(x_239, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_239, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_247 = x_239; +} else { + lean_dec_ref(x_239); + x_247 = lean_box(0); +} +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); +} else { + x_248 = x_247; +} +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; +x_249 = lean_ctor_get(x_19, 0); +x_250 = lean_ctor_get(x_19, 1); +lean_inc(x_250); +lean_inc(x_249); +lean_dec(x_19); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +x_252 = lean_nat_dec_eq(x_16, x_251); +lean_dec(x_16); +if (x_252 == 0) +{ +lean_object* x_253; +lean_dec(x_251); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_249); +lean_ctor_set(x_253, 1, x_250); +return x_253; +} +else +{ +lean_object* x_254; lean_object* x_255; +lean_dec(x_250); +x_254 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2; +x_255 = l_Std_Internal_Parsec_String_pstring(x_254, x_249); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; uint8_t x_258; lean_object* x_259; lean_object* x_260; +lean_dec(x_251); +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_257 = x_255; +} else { + lean_dec_ref(x_255); + x_257 = lean_box(0); +} +x_258 = 1; +x_259 = lean_box(x_258); +if (lean_is_scalar(x_257)) { + x_260 = lean_alloc_ctor(0, 2, 0); +} else { + x_260 = x_257; +} +lean_ctor_set(x_260, 0, x_256); +lean_ctor_set(x_260, 1, x_259); +return x_260; +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; uint8_t x_265; +x_261 = lean_ctor_get(x_255, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_255, 1); +lean_inc(x_262); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_263 = x_255; +} else { + lean_dec_ref(x_255); + x_263 = lean_box(0); +} +x_264 = lean_ctor_get(x_261, 1); +lean_inc(x_264); +x_265 = lean_nat_dec_eq(x_251, x_264); +lean_dec(x_251); +if (x_265 == 0) +{ +lean_object* x_266; +lean_dec(x_264); +if (lean_is_scalar(x_263)) { + x_266 = lean_alloc_ctor(1, 2, 0); +} else { + x_266 = x_263; +} +lean_ctor_set(x_266, 0, x_261); +lean_ctor_set(x_266, 1, x_262); +return x_266; +} +else +{ +lean_object* x_267; lean_object* x_268; +lean_dec(x_263); +lean_dec(x_262); +x_267 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +x_268 = l_Std_Internal_Parsec_String_pstring(x_267, x_261); +if (lean_obj_tag(x_268) == 0) +{ +lean_object* x_269; lean_object* x_270; uint8_t x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_264); +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + x_270 = x_268; +} else { + lean_dec_ref(x_268); + x_270 = lean_box(0); +} +x_271 = 2; +x_272 = lean_box(x_271); +if (lean_is_scalar(x_270)) { + x_273 = lean_alloc_ctor(0, 2, 0); +} else { + x_273 = x_270; +} +lean_ctor_set(x_273, 0, x_269); +lean_ctor_set(x_273, 1, x_272); +return x_273; +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; uint8_t x_278; +x_274 = lean_ctor_get(x_268, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_268, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + x_276 = x_268; +} else { + lean_dec_ref(x_268); + x_276 = lean_box(0); +} +x_277 = lean_ctor_get(x_274, 1); +lean_inc(x_277); +x_278 = lean_nat_dec_eq(x_264, x_277); +lean_dec(x_264); +if (x_278 == 0) +{ +lean_object* x_279; +lean_dec(x_277); +if (lean_is_scalar(x_276)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_276; +} +lean_ctor_set(x_279, 0, x_274); +lean_ctor_set(x_279, 1, x_275); +return x_279; +} +else +{ +lean_object* x_280; lean_object* x_281; +lean_dec(x_276); +lean_dec(x_275); +x_280 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +x_281 = l_Std_Internal_Parsec_String_pstring(x_280, x_274); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; lean_object* x_286; +lean_dec(x_277); +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_283 = x_281; +} else { + lean_dec_ref(x_281); + x_283 = lean_box(0); +} +x_284 = 3; +x_285 = lean_box(x_284); +if (lean_is_scalar(x_283)) { + x_286 = lean_alloc_ctor(0, 2, 0); +} else { + x_286 = x_283; +} +lean_ctor_set(x_286, 0, x_282); +lean_ctor_set(x_286, 1, x_285); +return x_286; +} +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; +x_287 = lean_ctor_get(x_281, 0); +lean_inc(x_287); +x_288 = lean_ctor_get(x_281, 1); +lean_inc(x_288); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_289 = x_281; +} else { + lean_dec_ref(x_281); + x_289 = lean_box(0); +} +x_290 = lean_ctor_get(x_287, 1); +lean_inc(x_290); +x_291 = lean_nat_dec_eq(x_277, x_290); +lean_dec(x_277); +if (x_291 == 0) +{ +lean_object* x_292; +lean_dec(x_290); +if (lean_is_scalar(x_289)) { + x_292 = lean_alloc_ctor(1, 2, 0); +} else { + x_292 = x_289; +} +lean_ctor_set(x_292, 0, x_287); +lean_ctor_set(x_292, 1, x_288); +return x_292; +} +else +{ +lean_object* x_293; lean_object* x_294; +lean_dec(x_289); +lean_dec(x_288); +x_293 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_294 = l_Std_Internal_Parsec_String_pstring(x_293, x_287); +if (lean_obj_tag(x_294) == 0) +{ +lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_298; lean_object* x_299; +lean_dec(x_290); +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_296 = x_294; +} else { + lean_dec_ref(x_294); + x_296 = lean_box(0); +} +x_297 = 4; +x_298 = lean_box(x_297); +if (lean_is_scalar(x_296)) { + x_299 = lean_alloc_ctor(0, 2, 0); +} else { + x_299 = x_296; +} +lean_ctor_set(x_299, 0, x_295); +lean_ctor_set(x_299, 1, x_298); +return x_299; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; +x_300 = lean_ctor_get(x_294, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_294, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_302 = x_294; +} else { + lean_dec_ref(x_294); + x_302 = lean_box(0); +} +x_303 = lean_ctor_get(x_300, 1); +lean_inc(x_303); +x_304 = lean_nat_dec_eq(x_290, x_303); +lean_dec(x_303); +lean_dec(x_290); +if (x_304 == 0) +{ +lean_object* x_305; +if (lean_is_scalar(x_302)) { + x_305 = lean_alloc_ctor(1, 2, 0); +} else { + x_305 = x_302; +} +lean_ctor_set(x_305, 0, x_300); +lean_ctor_set(x_305, 1, x_301); +return x_305; +} +else +{ +lean_object* x_306; lean_object* x_307; +lean_dec(x_302); +lean_dec(x_301); +x_306 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_307 = l_Std_Internal_Parsec_String_pstring(x_306, x_300); +if (lean_obj_tag(x_307) == 0) +{ +lean_object* x_308; lean_object* x_309; uint8_t x_310; lean_object* x_311; lean_object* x_312; +x_308 = lean_ctor_get(x_307, 0); +lean_inc(x_308); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_309 = x_307; +} else { + lean_dec_ref(x_307); + x_309 = lean_box(0); +} +x_310 = 5; +x_311 = lean_box(x_310); +if (lean_is_scalar(x_309)) { + x_312 = lean_alloc_ctor(0, 2, 0); +} else { + x_312 = x_309; +} +lean_ctor_set(x_312, 0, x_308); +lean_ctor_set(x_312, 1, x_311); +return x_312; +} +else +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +x_313 = lean_ctor_get(x_307, 0); +lean_inc(x_313); +x_314 = lean_ctor_get(x_307, 1); +lean_inc(x_314); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_315 = x_307; +} else { + lean_dec_ref(x_307); + x_315 = lean_box(0); +} +if (lean_is_scalar(x_315)) { + x_316 = lean_alloc_ctor(1, 2, 0); +} else { + x_316 = x_315; +} +lean_ctor_set(x_316, 0, x_313); +lean_ctor_set(x_316, 1, x_314); +return x_316; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; uint8_t x_321; +x_317 = lean_ctor_get(x_3, 0); +x_318 = lean_ctor_get(x_3, 1); +lean_inc(x_318); +lean_inc(x_317); +lean_dec(x_3); +x_319 = lean_ctor_get(x_1, 1); +lean_inc(x_319); +lean_dec(x_1); +x_320 = lean_ctor_get(x_317, 1); +lean_inc(x_320); +x_321 = lean_nat_dec_eq(x_319, x_320); +lean_dec(x_319); +if (x_321 == 0) +{ +lean_object* x_322; +lean_dec(x_320); +x_322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_322, 0, x_317); +lean_ctor_set(x_322, 1, x_318); +return x_322; +} +else +{ +lean_object* x_323; lean_object* x_324; +lean_dec(x_318); +x_323 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1; +x_324 = l_Std_Internal_Parsec_String_pstring(x_323, x_317); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; uint8_t x_327; lean_object* x_328; lean_object* x_329; +lean_dec(x_320); +x_325 = lean_ctor_get(x_324, 0); +lean_inc(x_325); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_326 = x_324; +} else { + lean_dec_ref(x_324); + x_326 = lean_box(0); +} +x_327 = 0; +x_328 = lean_box(x_327); +if (lean_is_scalar(x_326)) { + x_329 = lean_alloc_ctor(0, 2, 0); +} else { + x_329 = x_326; +} +lean_ctor_set(x_329, 0, x_325); +lean_ctor_set(x_329, 1, x_328); +return x_329; +} +else +{ +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; uint8_t x_334; +x_330 = lean_ctor_get(x_324, 0); +lean_inc(x_330); +x_331 = lean_ctor_get(x_324, 1); +lean_inc(x_331); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_332 = x_324; +} else { + lean_dec_ref(x_324); + x_332 = lean_box(0); +} +x_333 = lean_ctor_get(x_330, 1); +lean_inc(x_333); +x_334 = lean_nat_dec_eq(x_320, x_333); +lean_dec(x_320); +if (x_334 == 0) +{ +lean_object* x_335; +lean_dec(x_333); +if (lean_is_scalar(x_332)) { + x_335 = lean_alloc_ctor(1, 2, 0); +} else { + x_335 = x_332; +} +lean_ctor_set(x_335, 0, x_330); +lean_ctor_set(x_335, 1, x_331); +return x_335; +} +else +{ +lean_object* x_336; lean_object* x_337; +lean_dec(x_332); +lean_dec(x_331); +x_336 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2; +x_337 = l_Std_Internal_Parsec_String_pstring(x_336, x_330); +if (lean_obj_tag(x_337) == 0) +{ +lean_object* x_338; lean_object* x_339; uint8_t x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_333); +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_339 = x_337; +} else { + lean_dec_ref(x_337); + x_339 = lean_box(0); +} +x_340 = 1; +x_341 = lean_box(x_340); +if (lean_is_scalar(x_339)) { + x_342 = lean_alloc_ctor(0, 2, 0); +} else { + x_342 = x_339; +} +lean_ctor_set(x_342, 0, x_338); +lean_ctor_set(x_342, 1, x_341); +return x_342; +} +else +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; uint8_t x_347; +x_343 = lean_ctor_get(x_337, 0); +lean_inc(x_343); +x_344 = lean_ctor_get(x_337, 1); +lean_inc(x_344); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_345 = x_337; +} else { + lean_dec_ref(x_337); + x_345 = lean_box(0); +} +x_346 = lean_ctor_get(x_343, 1); +lean_inc(x_346); +x_347 = lean_nat_dec_eq(x_333, x_346); +lean_dec(x_333); +if (x_347 == 0) +{ +lean_object* x_348; +lean_dec(x_346); +if (lean_is_scalar(x_345)) { + x_348 = lean_alloc_ctor(1, 2, 0); +} else { + x_348 = x_345; +} +lean_ctor_set(x_348, 0, x_343); +lean_ctor_set(x_348, 1, x_344); +return x_348; +} +else +{ +lean_object* x_349; lean_object* x_350; +lean_dec(x_345); +lean_dec(x_344); +x_349 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3; +x_350 = l_Std_Internal_Parsec_String_pstring(x_349, x_343); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; uint8_t x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_346); +x_351 = lean_ctor_get(x_350, 0); +lean_inc(x_351); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + x_352 = x_350; +} else { + lean_dec_ref(x_350); + x_352 = lean_box(0); +} +x_353 = 2; +x_354 = lean_box(x_353); +if (lean_is_scalar(x_352)) { + x_355 = lean_alloc_ctor(0, 2, 0); +} else { + x_355 = x_352; +} +lean_ctor_set(x_355, 0, x_351); +lean_ctor_set(x_355, 1, x_354); +return x_355; +} +else +{ +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; +x_356 = lean_ctor_get(x_350, 0); +lean_inc(x_356); +x_357 = lean_ctor_get(x_350, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + x_358 = x_350; +} else { + lean_dec_ref(x_350); + x_358 = lean_box(0); +} +x_359 = lean_ctor_get(x_356, 1); +lean_inc(x_359); +x_360 = lean_nat_dec_eq(x_346, x_359); +lean_dec(x_346); +if (x_360 == 0) +{ +lean_object* x_361; +lean_dec(x_359); +if (lean_is_scalar(x_358)) { + x_361 = lean_alloc_ctor(1, 2, 0); +} else { + x_361 = x_358; +} +lean_ctor_set(x_361, 0, x_356); +lean_ctor_set(x_361, 1, x_357); +return x_361; +} +else +{ +lean_object* x_362; lean_object* x_363; +lean_dec(x_358); +lean_dec(x_357); +x_362 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4; +x_363 = l_Std_Internal_Parsec_String_pstring(x_362, x_356); +if (lean_obj_tag(x_363) == 0) +{ +lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_359); +x_364 = lean_ctor_get(x_363, 0); +lean_inc(x_364); +if (lean_is_exclusive(x_363)) { + lean_ctor_release(x_363, 0); + lean_ctor_release(x_363, 1); + x_365 = x_363; +} else { + lean_dec_ref(x_363); + x_365 = lean_box(0); +} +x_366 = 3; +x_367 = lean_box(x_366); +if (lean_is_scalar(x_365)) { + x_368 = lean_alloc_ctor(0, 2, 0); +} else { + x_368 = x_365; +} +lean_ctor_set(x_368, 0, x_364); +lean_ctor_set(x_368, 1, x_367); +return x_368; +} +else +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; uint8_t x_373; +x_369 = lean_ctor_get(x_363, 0); +lean_inc(x_369); +x_370 = lean_ctor_get(x_363, 1); +lean_inc(x_370); +if (lean_is_exclusive(x_363)) { + lean_ctor_release(x_363, 0); + lean_ctor_release(x_363, 1); + x_371 = x_363; +} else { + lean_dec_ref(x_363); + x_371 = lean_box(0); +} +x_372 = lean_ctor_get(x_369, 1); +lean_inc(x_372); +x_373 = lean_nat_dec_eq(x_359, x_372); +lean_dec(x_359); +if (x_373 == 0) +{ +lean_object* x_374; +lean_dec(x_372); +if (lean_is_scalar(x_371)) { + x_374 = lean_alloc_ctor(1, 2, 0); +} else { + x_374 = x_371; +} +lean_ctor_set(x_374, 0, x_369); +lean_ctor_set(x_374, 1, x_370); +return x_374; +} +else +{ +lean_object* x_375; lean_object* x_376; +lean_dec(x_371); +lean_dec(x_370); +x_375 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5; +x_376 = l_Std_Internal_Parsec_String_pstring(x_375, x_369); +if (lean_obj_tag(x_376) == 0) +{ +lean_object* x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_372); +x_377 = lean_ctor_get(x_376, 0); +lean_inc(x_377); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_378 = x_376; +} else { + lean_dec_ref(x_376); + x_378 = lean_box(0); +} +x_379 = 4; +x_380 = lean_box(x_379); +if (lean_is_scalar(x_378)) { + x_381 = lean_alloc_ctor(0, 2, 0); +} else { + x_381 = x_378; +} +lean_ctor_set(x_381, 0, x_377); +lean_ctor_set(x_381, 1, x_380); +return x_381; +} +else +{ +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; uint8_t x_386; +x_382 = lean_ctor_get(x_376, 0); +lean_inc(x_382); +x_383 = lean_ctor_get(x_376, 1); +lean_inc(x_383); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_384 = x_376; +} else { + lean_dec_ref(x_376); + x_384 = lean_box(0); +} +x_385 = lean_ctor_get(x_382, 1); +lean_inc(x_385); +x_386 = lean_nat_dec_eq(x_372, x_385); +lean_dec(x_385); +lean_dec(x_372); +if (x_386 == 0) +{ +lean_object* x_387; +if (lean_is_scalar(x_384)) { + x_387 = lean_alloc_ctor(1, 2, 0); +} else { + x_387 = x_384; +} +lean_ctor_set(x_387, 0, x_382); +lean_ctor_set(x_387, 1, x_383); +return x_387; +} +else +{ +lean_object* x_388; lean_object* x_389; +lean_dec(x_384); +lean_dec(x_383); +x_388 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6; +x_389 = l_Std_Internal_Parsec_String_pstring(x_388, x_382); +if (lean_obj_tag(x_389) == 0) +{ +lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; lean_object* x_394; +x_390 = lean_ctor_get(x_389, 0); +lean_inc(x_390); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_391 = x_389; +} else { + lean_dec_ref(x_389); + x_391 = lean_box(0); +} +x_392 = 5; +x_393 = lean_box(x_392); +if (lean_is_scalar(x_391)) { + x_394 = lean_alloc_ctor(0, 2, 0); +} else { + x_394 = x_391; +} +lean_ctor_set(x_394, 0, x_390); +lean_ctor_set(x_394, 1, x_393); +return x_394; +} +else +{ +lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; +x_395 = lean_ctor_get(x_389, 0); +lean_inc(x_395); +x_396 = lean_ctor_get(x_389, 1); +lean_inc(x_396); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_397 = x_389; +} else { + lean_dec_ref(x_389); + x_397 = lean_box(0); +} +if (lean_is_scalar(x_397)) { + x_398 = lean_alloc_ctor(1, 2, 0); +} else { + x_398 = x_397; +} +lean_ctor_set(x_398, 0, x_395); +lean_ctor_set(x_398, 1, x_396); +return x_398; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 6; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 6; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_16); +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 0; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +x_32 = lean_nat_dec_eq(x_16, x_31); +lean_dec(x_16); +if (x_32 == 0) +{ +lean_dec(x_31); +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_free_object(x_19); +lean_dec(x_30); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2; +x_34 = l_Std_Internal_Parsec_String_pstring(x_33, x_29); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +lean_dec(x_31); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_34, 1); +lean_dec(x_36); +x_37 = 1; +x_38 = lean_box(x_37); +lean_ctor_set(x_34, 1, x_38); +return x_34; +} +else +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_34, 0); +lean_inc(x_39); +lean_dec(x_34); +x_40 = 1; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +x_47 = lean_nat_dec_eq(x_31, x_46); +lean_dec(x_31); +if (x_47 == 0) +{ +lean_dec(x_46); +return x_34; +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_free_object(x_34); +lean_dec(x_45); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +x_49 = l_Std_Internal_Parsec_String_pstring(x_48, x_44); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +lean_dec(x_46); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_49, 1); +lean_dec(x_51); +x_52 = 2; +x_53 = lean_box(x_52); +lean_ctor_set(x_49, 1, x_53); +return x_49; +} +else +{ +lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_49, 0); +lean_inc(x_54); +lean_dec(x_49); +x_55 = 2; +x_56 = lean_box(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +else +{ +uint8_t x_58; +x_58 = !lean_is_exclusive(x_49); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_59 = lean_ctor_get(x_49, 0); +x_60 = lean_ctor_get(x_49, 1); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +x_62 = lean_nat_dec_eq(x_46, x_61); +lean_dec(x_46); +if (x_62 == 0) +{ +lean_dec(x_61); +return x_49; +} +else +{ +lean_object* x_63; lean_object* x_64; +lean_free_object(x_49); +lean_dec(x_60); +x_63 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +x_64 = l_Std_Internal_Parsec_String_pstring(x_63, x_59); +if (lean_obj_tag(x_64) == 0) +{ +uint8_t x_65; +lean_dec(x_61); +x_65 = !lean_is_exclusive(x_64); +if (x_65 == 0) +{ +lean_object* x_66; uint8_t x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_64, 1); +lean_dec(x_66); +x_67 = 3; +x_68 = lean_box(x_67); +lean_ctor_set(x_64, 1, x_68); +return x_64; +} +else +{ +lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_64, 0); +lean_inc(x_69); +lean_dec(x_64); +x_70 = 3; +x_71 = lean_box(x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_69); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} +} +else +{ +uint8_t x_73; +x_73 = !lean_is_exclusive(x_64); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = lean_ctor_get(x_64, 0); +x_75 = lean_ctor_get(x_64, 1); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +x_77 = lean_nat_dec_eq(x_61, x_76); +lean_dec(x_61); +if (x_77 == 0) +{ +lean_dec(x_76); +return x_64; +} +else +{ +lean_object* x_78; lean_object* x_79; +lean_free_object(x_64); +lean_dec(x_75); +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_79 = l_Std_Internal_Parsec_String_pstring(x_78, x_74); +if (lean_obj_tag(x_79) == 0) +{ +uint8_t x_80; +lean_dec(x_76); +x_80 = !lean_is_exclusive(x_79); +if (x_80 == 0) +{ +lean_object* x_81; uint8_t x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_79, 1); +lean_dec(x_81); +x_82 = 4; +x_83 = lean_box(x_82); +lean_ctor_set(x_79, 1, x_83); +return x_79; +} +else +{ +lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_79, 0); +lean_inc(x_84); +lean_dec(x_79); +x_85 = 4; +x_86 = lean_box(x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_79); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; +x_89 = lean_ctor_get(x_79, 0); +x_90 = lean_ctor_get(x_79, 1); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +x_92 = lean_nat_dec_eq(x_76, x_91); +lean_dec(x_91); +lean_dec(x_76); +if (x_92 == 0) +{ +return x_79; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_free_object(x_79); +lean_dec(x_90); +x_93 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_94 = l_Std_Internal_Parsec_String_pstring(x_93, x_89); +if (lean_obj_tag(x_94) == 0) +{ +uint8_t x_95; +x_95 = !lean_is_exclusive(x_94); +if (x_95 == 0) +{ +lean_object* x_96; uint8_t x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_94, 1); +lean_dec(x_96); +x_97 = 5; +x_98 = lean_box(x_97); +lean_ctor_set(x_94, 1, x_98); +return x_94; +} +else +{ +lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_94, 0); +lean_inc(x_99); +lean_dec(x_94); +x_100 = 5; +x_101 = lean_box(x_100); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_101); +return x_102; +} +} +else +{ +uint8_t x_103; +x_103 = !lean_is_exclusive(x_94); +if (x_103 == 0) +{ +return x_94; +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_94, 0); +x_105 = lean_ctor_get(x_94, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_94); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +return x_106; +} +} +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_107 = lean_ctor_get(x_79, 0); +x_108 = lean_ctor_get(x_79, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_79); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +x_110 = lean_nat_dec_eq(x_76, x_109); +lean_dec(x_109); +lean_dec(x_76); +if (x_110 == 0) +{ +lean_object* x_111; +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_107); +lean_ctor_set(x_111, 1, x_108); +return x_111; +} +else +{ +lean_object* x_112; lean_object* x_113; +lean_dec(x_108); +x_112 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_113 = l_Std_Internal_Parsec_String_pstring(x_112, x_107); +if (lean_obj_tag(x_113) == 0) +{ +lean_object* x_114; lean_object* x_115; uint8_t x_116; lean_object* x_117; lean_object* x_118; +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_115 = x_113; +} else { + lean_dec_ref(x_113); + x_115 = lean_box(0); +} +x_116 = 5; +x_117 = lean_box(x_116); +if (lean_is_scalar(x_115)) { + x_118 = lean_alloc_ctor(0, 2, 0); +} else { + x_118 = x_115; +} +lean_ctor_set(x_118, 0, x_114); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +else +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_119 = lean_ctor_get(x_113, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_113, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + x_121 = x_113; +} else { + lean_dec_ref(x_113); + x_121 = lean_box(0); +} +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); +} else { + x_122 = x_121; +} +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; +} +} +} +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_123 = lean_ctor_get(x_64, 0); +x_124 = lean_ctor_get(x_64, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_64); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +x_126 = lean_nat_dec_eq(x_61, x_125); +lean_dec(x_61); +if (x_126 == 0) +{ +lean_object* x_127; +lean_dec(x_125); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_123); +lean_ctor_set(x_127, 1, x_124); +return x_127; +} +else +{ +lean_object* x_128; lean_object* x_129; +lean_dec(x_124); +x_128 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_129 = l_Std_Internal_Parsec_String_pstring(x_128, x_123); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_125); +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_131 = x_129; +} else { + lean_dec_ref(x_129); + x_131 = lean_box(0); +} +x_132 = 4; +x_133 = lean_box(x_132); +if (lean_is_scalar(x_131)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_131; +} +lean_ctor_set(x_134, 0, x_130); +lean_ctor_set(x_134, 1, x_133); +return x_134; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; +x_135 = lean_ctor_get(x_129, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_129, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_137 = x_129; +} else { + lean_dec_ref(x_129); + x_137 = lean_box(0); +} +x_138 = lean_ctor_get(x_135, 1); +lean_inc(x_138); +x_139 = lean_nat_dec_eq(x_125, x_138); +lean_dec(x_138); +lean_dec(x_125); +if (x_139 == 0) +{ +lean_object* x_140; +if (lean_is_scalar(x_137)) { + x_140 = lean_alloc_ctor(1, 2, 0); +} else { + x_140 = x_137; +} +lean_ctor_set(x_140, 0, x_135); +lean_ctor_set(x_140, 1, x_136); +return x_140; +} +else +{ +lean_object* x_141; lean_object* x_142; +lean_dec(x_137); +lean_dec(x_136); +x_141 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_142 = l_Std_Internal_Parsec_String_pstring(x_141, x_135); +if (lean_obj_tag(x_142) == 0) +{ +lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_144 = x_142; +} else { + lean_dec_ref(x_142); + x_144 = lean_box(0); +} +x_145 = 5; +x_146 = lean_box(x_145); +if (lean_is_scalar(x_144)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_144; +} +lean_ctor_set(x_147, 0, x_143); +lean_ctor_set(x_147, 1, x_146); +return x_147; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_148 = lean_ctor_get(x_142, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_142, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_142)) { + lean_ctor_release(x_142, 0); + lean_ctor_release(x_142, 1); + x_150 = x_142; +} else { + lean_dec_ref(x_142); + x_150 = lean_box(0); +} +if (lean_is_scalar(x_150)) { + x_151 = lean_alloc_ctor(1, 2, 0); +} else { + x_151 = x_150; +} +lean_ctor_set(x_151, 0, x_148); +lean_ctor_set(x_151, 1, x_149); +return x_151; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_152 = lean_ctor_get(x_49, 0); +x_153 = lean_ctor_get(x_49, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_49); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +x_155 = lean_nat_dec_eq(x_46, x_154); +lean_dec(x_46); +if (x_155 == 0) +{ +lean_object* x_156; +lean_dec(x_154); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_152); +lean_ctor_set(x_156, 1, x_153); +return x_156; +} +else +{ +lean_object* x_157; lean_object* x_158; +lean_dec(x_153); +x_157 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +x_158 = l_Std_Internal_Parsec_String_pstring(x_157, x_152); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_154); +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_160 = x_158; +} else { + lean_dec_ref(x_158); + x_160 = lean_box(0); +} +x_161 = 3; +x_162 = lean_box(x_161); +if (lean_is_scalar(x_160)) { + x_163 = lean_alloc_ctor(0, 2, 0); +} else { + x_163 = x_160; +} +lean_ctor_set(x_163, 0, x_159); +lean_ctor_set(x_163, 1, x_162); +return x_163; +} +else +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; +x_164 = lean_ctor_get(x_158, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_158, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_166 = x_158; +} else { + lean_dec_ref(x_158); + x_166 = lean_box(0); +} +x_167 = lean_ctor_get(x_164, 1); +lean_inc(x_167); +x_168 = lean_nat_dec_eq(x_154, x_167); +lean_dec(x_154); +if (x_168 == 0) +{ +lean_object* x_169; +lean_dec(x_167); +if (lean_is_scalar(x_166)) { + x_169 = lean_alloc_ctor(1, 2, 0); +} else { + x_169 = x_166; +} +lean_ctor_set(x_169, 0, x_164); +lean_ctor_set(x_169, 1, x_165); +return x_169; +} +else +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_166); +lean_dec(x_165); +x_170 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_171 = l_Std_Internal_Parsec_String_pstring(x_170, x_164); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; +lean_dec(x_167); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_173 = x_171; +} else { + lean_dec_ref(x_171); + x_173 = lean_box(0); +} +x_174 = 4; +x_175 = lean_box(x_174); +if (lean_is_scalar(x_173)) { + x_176 = lean_alloc_ctor(0, 2, 0); +} else { + x_176 = x_173; +} +lean_ctor_set(x_176, 0, x_172); +lean_ctor_set(x_176, 1, x_175); +return x_176; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; +x_177 = lean_ctor_get(x_171, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_171, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + lean_ctor_release(x_171, 1); + x_179 = x_171; +} else { + lean_dec_ref(x_171); + x_179 = lean_box(0); +} +x_180 = lean_ctor_get(x_177, 1); +lean_inc(x_180); +x_181 = lean_nat_dec_eq(x_167, x_180); +lean_dec(x_180); +lean_dec(x_167); +if (x_181 == 0) +{ +lean_object* x_182; +if (lean_is_scalar(x_179)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_179; +} +lean_ctor_set(x_182, 0, x_177); +lean_ctor_set(x_182, 1, x_178); +return x_182; +} +else +{ +lean_object* x_183; lean_object* x_184; +lean_dec(x_179); +lean_dec(x_178); +x_183 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_184 = l_Std_Internal_Parsec_String_pstring(x_183, x_177); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; uint8_t x_187; lean_object* x_188; lean_object* x_189; +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = 5; +x_188 = lean_box(x_187); +if (lean_is_scalar(x_186)) { + x_189 = lean_alloc_ctor(0, 2, 0); +} else { + x_189 = x_186; +} +lean_ctor_set(x_189, 0, x_185); +lean_ctor_set(x_189, 1, x_188); +return x_189; +} +else +{ +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_190 = lean_ctor_get(x_184, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_184, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_192 = x_184; +} else { + lean_dec_ref(x_184); + x_192 = lean_box(0); +} +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); +} else { + x_193 = x_192; +} +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; +x_194 = lean_ctor_get(x_34, 0); +x_195 = lean_ctor_get(x_34, 1); +lean_inc(x_195); +lean_inc(x_194); +lean_dec(x_34); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +x_197 = lean_nat_dec_eq(x_31, x_196); +lean_dec(x_31); +if (x_197 == 0) +{ +lean_object* x_198; +lean_dec(x_196); +x_198 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_198, 0, x_194); +lean_ctor_set(x_198, 1, x_195); +return x_198; +} +else +{ +lean_object* x_199; lean_object* x_200; +lean_dec(x_195); +x_199 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +x_200 = l_Std_Internal_Parsec_String_pstring(x_199, x_194); +if (lean_obj_tag(x_200) == 0) +{ +lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_196); +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_202 = x_200; +} else { + lean_dec_ref(x_200); + x_202 = lean_box(0); +} +x_203 = 2; +x_204 = lean_box(x_203); +if (lean_is_scalar(x_202)) { + x_205 = lean_alloc_ctor(0, 2, 0); +} else { + x_205 = x_202; +} +lean_ctor_set(x_205, 0, x_201); +lean_ctor_set(x_205, 1, x_204); +return x_205; +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; uint8_t x_210; +x_206 = lean_ctor_get(x_200, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_200, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_208 = x_200; +} else { + lean_dec_ref(x_200); + x_208 = lean_box(0); +} +x_209 = lean_ctor_get(x_206, 1); +lean_inc(x_209); +x_210 = lean_nat_dec_eq(x_196, x_209); +lean_dec(x_196); +if (x_210 == 0) +{ +lean_object* x_211; +lean_dec(x_209); +if (lean_is_scalar(x_208)) { + x_211 = lean_alloc_ctor(1, 2, 0); +} else { + x_211 = x_208; +} +lean_ctor_set(x_211, 0, x_206); +lean_ctor_set(x_211, 1, x_207); +return x_211; +} +else +{ +lean_object* x_212; lean_object* x_213; +lean_dec(x_208); +lean_dec(x_207); +x_212 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +x_213 = l_Std_Internal_Parsec_String_pstring(x_212, x_206); +if (lean_obj_tag(x_213) == 0) +{ +lean_object* x_214; lean_object* x_215; uint8_t x_216; lean_object* x_217; lean_object* x_218; +lean_dec(x_209); +x_214 = lean_ctor_get(x_213, 0); +lean_inc(x_214); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_215 = x_213; +} else { + lean_dec_ref(x_213); + x_215 = lean_box(0); +} +x_216 = 3; +x_217 = lean_box(x_216); +if (lean_is_scalar(x_215)) { + x_218 = lean_alloc_ctor(0, 2, 0); +} else { + x_218 = x_215; +} +lean_ctor_set(x_218, 0, x_214); +lean_ctor_set(x_218, 1, x_217); +return x_218; +} +else +{ +lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; uint8_t x_223; +x_219 = lean_ctor_get(x_213, 0); +lean_inc(x_219); +x_220 = lean_ctor_get(x_213, 1); +lean_inc(x_220); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_221 = x_213; +} else { + lean_dec_ref(x_213); + x_221 = lean_box(0); +} +x_222 = lean_ctor_get(x_219, 1); +lean_inc(x_222); +x_223 = lean_nat_dec_eq(x_209, x_222); +lean_dec(x_209); +if (x_223 == 0) +{ +lean_object* x_224; +lean_dec(x_222); +if (lean_is_scalar(x_221)) { + x_224 = lean_alloc_ctor(1, 2, 0); +} else { + x_224 = x_221; +} +lean_ctor_set(x_224, 0, x_219); +lean_ctor_set(x_224, 1, x_220); +return x_224; +} +else +{ +lean_object* x_225; lean_object* x_226; +lean_dec(x_221); +lean_dec(x_220); +x_225 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_226 = l_Std_Internal_Parsec_String_pstring(x_225, x_219); +if (lean_obj_tag(x_226) == 0) +{ +lean_object* x_227; lean_object* x_228; uint8_t x_229; lean_object* x_230; lean_object* x_231; +lean_dec(x_222); +x_227 = lean_ctor_get(x_226, 0); +lean_inc(x_227); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_228 = x_226; +} else { + lean_dec_ref(x_226); + x_228 = lean_box(0); +} +x_229 = 4; +x_230 = lean_box(x_229); +if (lean_is_scalar(x_228)) { + x_231 = lean_alloc_ctor(0, 2, 0); +} else { + x_231 = x_228; +} +lean_ctor_set(x_231, 0, x_227); +lean_ctor_set(x_231, 1, x_230); +return x_231; +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; uint8_t x_236; +x_232 = lean_ctor_get(x_226, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_226, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_234 = x_226; +} else { + lean_dec_ref(x_226); + x_234 = lean_box(0); +} +x_235 = lean_ctor_get(x_232, 1); +lean_inc(x_235); +x_236 = lean_nat_dec_eq(x_222, x_235); +lean_dec(x_235); +lean_dec(x_222); +if (x_236 == 0) +{ +lean_object* x_237; +if (lean_is_scalar(x_234)) { + x_237 = lean_alloc_ctor(1, 2, 0); +} else { + x_237 = x_234; +} +lean_ctor_set(x_237, 0, x_232); +lean_ctor_set(x_237, 1, x_233); +return x_237; +} +else +{ +lean_object* x_238; lean_object* x_239; +lean_dec(x_234); +lean_dec(x_233); +x_238 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_239 = l_Std_Internal_Parsec_String_pstring(x_238, x_232); +if (lean_obj_tag(x_239) == 0) +{ +lean_object* x_240; lean_object* x_241; uint8_t x_242; lean_object* x_243; lean_object* x_244; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_241 = x_239; +} else { + lean_dec_ref(x_239); + x_241 = lean_box(0); +} +x_242 = 5; +x_243 = lean_box(x_242); +if (lean_is_scalar(x_241)) { + x_244 = lean_alloc_ctor(0, 2, 0); +} else { + x_244 = x_241; +} +lean_ctor_set(x_244, 0, x_240); +lean_ctor_set(x_244, 1, x_243); +return x_244; +} +else +{ +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_245 = lean_ctor_get(x_239, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_239, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_247 = x_239; +} else { + lean_dec_ref(x_239); + x_247 = lean_box(0); +} +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); +} else { + x_248 = x_247; +} +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; +x_249 = lean_ctor_get(x_19, 0); +x_250 = lean_ctor_get(x_19, 1); +lean_inc(x_250); +lean_inc(x_249); +lean_dec(x_19); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +x_252 = lean_nat_dec_eq(x_16, x_251); +lean_dec(x_16); +if (x_252 == 0) +{ +lean_object* x_253; +lean_dec(x_251); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_249); +lean_ctor_set(x_253, 1, x_250); +return x_253; +} +else +{ +lean_object* x_254; lean_object* x_255; +lean_dec(x_250); +x_254 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2; +x_255 = l_Std_Internal_Parsec_String_pstring(x_254, x_249); +if (lean_obj_tag(x_255) == 0) +{ +lean_object* x_256; lean_object* x_257; uint8_t x_258; lean_object* x_259; lean_object* x_260; +lean_dec(x_251); +x_256 = lean_ctor_get(x_255, 0); +lean_inc(x_256); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_257 = x_255; +} else { + lean_dec_ref(x_255); + x_257 = lean_box(0); +} +x_258 = 1; +x_259 = lean_box(x_258); +if (lean_is_scalar(x_257)) { + x_260 = lean_alloc_ctor(0, 2, 0); +} else { + x_260 = x_257; +} +lean_ctor_set(x_260, 0, x_256); +lean_ctor_set(x_260, 1, x_259); +return x_260; +} +else +{ +lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; uint8_t x_265; +x_261 = lean_ctor_get(x_255, 0); +lean_inc(x_261); +x_262 = lean_ctor_get(x_255, 1); +lean_inc(x_262); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_263 = x_255; +} else { + lean_dec_ref(x_255); + x_263 = lean_box(0); +} +x_264 = lean_ctor_get(x_261, 1); +lean_inc(x_264); +x_265 = lean_nat_dec_eq(x_251, x_264); +lean_dec(x_251); +if (x_265 == 0) +{ +lean_object* x_266; +lean_dec(x_264); +if (lean_is_scalar(x_263)) { + x_266 = lean_alloc_ctor(1, 2, 0); +} else { + x_266 = x_263; +} +lean_ctor_set(x_266, 0, x_261); +lean_ctor_set(x_266, 1, x_262); +return x_266; +} +else +{ +lean_object* x_267; lean_object* x_268; +lean_dec(x_263); +lean_dec(x_262); +x_267 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +x_268 = l_Std_Internal_Parsec_String_pstring(x_267, x_261); +if (lean_obj_tag(x_268) == 0) +{ +lean_object* x_269; lean_object* x_270; uint8_t x_271; lean_object* x_272; lean_object* x_273; +lean_dec(x_264); +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + x_270 = x_268; +} else { + lean_dec_ref(x_268); + x_270 = lean_box(0); +} +x_271 = 2; +x_272 = lean_box(x_271); +if (lean_is_scalar(x_270)) { + x_273 = lean_alloc_ctor(0, 2, 0); +} else { + x_273 = x_270; +} +lean_ctor_set(x_273, 0, x_269); +lean_ctor_set(x_273, 1, x_272); +return x_273; +} +else +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; uint8_t x_278; +x_274 = lean_ctor_get(x_268, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_268, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_268)) { + lean_ctor_release(x_268, 0); + lean_ctor_release(x_268, 1); + x_276 = x_268; +} else { + lean_dec_ref(x_268); + x_276 = lean_box(0); +} +x_277 = lean_ctor_get(x_274, 1); +lean_inc(x_277); +x_278 = lean_nat_dec_eq(x_264, x_277); +lean_dec(x_264); +if (x_278 == 0) +{ +lean_object* x_279; +lean_dec(x_277); +if (lean_is_scalar(x_276)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_276; +} +lean_ctor_set(x_279, 0, x_274); +lean_ctor_set(x_279, 1, x_275); +return x_279; +} +else +{ +lean_object* x_280; lean_object* x_281; +lean_dec(x_276); +lean_dec(x_275); +x_280 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +x_281 = l_Std_Internal_Parsec_String_pstring(x_280, x_274); +if (lean_obj_tag(x_281) == 0) +{ +lean_object* x_282; lean_object* x_283; uint8_t x_284; lean_object* x_285; lean_object* x_286; +lean_dec(x_277); +x_282 = lean_ctor_get(x_281, 0); +lean_inc(x_282); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_283 = x_281; +} else { + lean_dec_ref(x_281); + x_283 = lean_box(0); +} +x_284 = 3; +x_285 = lean_box(x_284); +if (lean_is_scalar(x_283)) { + x_286 = lean_alloc_ctor(0, 2, 0); +} else { + x_286 = x_283; +} +lean_ctor_set(x_286, 0, x_282); +lean_ctor_set(x_286, 1, x_285); +return x_286; +} +else +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; uint8_t x_291; +x_287 = lean_ctor_get(x_281, 0); +lean_inc(x_287); +x_288 = lean_ctor_get(x_281, 1); +lean_inc(x_288); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_289 = x_281; +} else { + lean_dec_ref(x_281); + x_289 = lean_box(0); +} +x_290 = lean_ctor_get(x_287, 1); +lean_inc(x_290); +x_291 = lean_nat_dec_eq(x_277, x_290); +lean_dec(x_277); +if (x_291 == 0) +{ +lean_object* x_292; +lean_dec(x_290); +if (lean_is_scalar(x_289)) { + x_292 = lean_alloc_ctor(1, 2, 0); +} else { + x_292 = x_289; +} +lean_ctor_set(x_292, 0, x_287); +lean_ctor_set(x_292, 1, x_288); +return x_292; +} +else +{ +lean_object* x_293; lean_object* x_294; +lean_dec(x_289); +lean_dec(x_288); +x_293 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_294 = l_Std_Internal_Parsec_String_pstring(x_293, x_287); +if (lean_obj_tag(x_294) == 0) +{ +lean_object* x_295; lean_object* x_296; uint8_t x_297; lean_object* x_298; lean_object* x_299; +lean_dec(x_290); +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_296 = x_294; +} else { + lean_dec_ref(x_294); + x_296 = lean_box(0); +} +x_297 = 4; +x_298 = lean_box(x_297); +if (lean_is_scalar(x_296)) { + x_299 = lean_alloc_ctor(0, 2, 0); +} else { + x_299 = x_296; +} +lean_ctor_set(x_299, 0, x_295); +lean_ctor_set(x_299, 1, x_298); +return x_299; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; uint8_t x_304; +x_300 = lean_ctor_get(x_294, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_294, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_302 = x_294; +} else { + lean_dec_ref(x_294); + x_302 = lean_box(0); +} +x_303 = lean_ctor_get(x_300, 1); +lean_inc(x_303); +x_304 = lean_nat_dec_eq(x_290, x_303); +lean_dec(x_303); +lean_dec(x_290); +if (x_304 == 0) +{ +lean_object* x_305; +if (lean_is_scalar(x_302)) { + x_305 = lean_alloc_ctor(1, 2, 0); +} else { + x_305 = x_302; +} +lean_ctor_set(x_305, 0, x_300); +lean_ctor_set(x_305, 1, x_301); +return x_305; +} +else +{ +lean_object* x_306; lean_object* x_307; +lean_dec(x_302); +lean_dec(x_301); +x_306 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_307 = l_Std_Internal_Parsec_String_pstring(x_306, x_300); +if (lean_obj_tag(x_307) == 0) +{ +lean_object* x_308; lean_object* x_309; uint8_t x_310; lean_object* x_311; lean_object* x_312; +x_308 = lean_ctor_get(x_307, 0); +lean_inc(x_308); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_309 = x_307; +} else { + lean_dec_ref(x_307); + x_309 = lean_box(0); +} +x_310 = 5; +x_311 = lean_box(x_310); +if (lean_is_scalar(x_309)) { + x_312 = lean_alloc_ctor(0, 2, 0); +} else { + x_312 = x_309; +} +lean_ctor_set(x_312, 0, x_308); +lean_ctor_set(x_312, 1, x_311); +return x_312; +} +else +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +x_313 = lean_ctor_get(x_307, 0); +lean_inc(x_313); +x_314 = lean_ctor_get(x_307, 1); +lean_inc(x_314); +if (lean_is_exclusive(x_307)) { + lean_ctor_release(x_307, 0); + lean_ctor_release(x_307, 1); + x_315 = x_307; +} else { + lean_dec_ref(x_307); + x_315 = lean_box(0); +} +if (lean_is_scalar(x_315)) { + x_316 = lean_alloc_ctor(1, 2, 0); +} else { + x_316 = x_315; +} +lean_ctor_set(x_316, 0, x_313); +lean_ctor_set(x_316, 1, x_314); +return x_316; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; uint8_t x_321; +x_317 = lean_ctor_get(x_3, 0); +x_318 = lean_ctor_get(x_3, 1); +lean_inc(x_318); +lean_inc(x_317); +lean_dec(x_3); +x_319 = lean_ctor_get(x_1, 1); +lean_inc(x_319); +lean_dec(x_1); +x_320 = lean_ctor_get(x_317, 1); +lean_inc(x_320); +x_321 = lean_nat_dec_eq(x_319, x_320); +lean_dec(x_319); +if (x_321 == 0) +{ +lean_object* x_322; +lean_dec(x_320); +x_322 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_322, 0, x_317); +lean_ctor_set(x_322, 1, x_318); +return x_322; +} +else +{ +lean_object* x_323; lean_object* x_324; +lean_dec(x_318); +x_323 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1; +x_324 = l_Std_Internal_Parsec_String_pstring(x_323, x_317); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; uint8_t x_327; lean_object* x_328; lean_object* x_329; +lean_dec(x_320); +x_325 = lean_ctor_get(x_324, 0); +lean_inc(x_325); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_326 = x_324; +} else { + lean_dec_ref(x_324); + x_326 = lean_box(0); +} +x_327 = 0; +x_328 = lean_box(x_327); +if (lean_is_scalar(x_326)) { + x_329 = lean_alloc_ctor(0, 2, 0); +} else { + x_329 = x_326; +} +lean_ctor_set(x_329, 0, x_325); +lean_ctor_set(x_329, 1, x_328); +return x_329; +} +else +{ +lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; uint8_t x_334; +x_330 = lean_ctor_get(x_324, 0); +lean_inc(x_330); +x_331 = lean_ctor_get(x_324, 1); +lean_inc(x_331); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_332 = x_324; +} else { + lean_dec_ref(x_324); + x_332 = lean_box(0); +} +x_333 = lean_ctor_get(x_330, 1); +lean_inc(x_333); +x_334 = lean_nat_dec_eq(x_320, x_333); +lean_dec(x_320); +if (x_334 == 0) +{ +lean_object* x_335; +lean_dec(x_333); +if (lean_is_scalar(x_332)) { + x_335 = lean_alloc_ctor(1, 2, 0); +} else { + x_335 = x_332; +} +lean_ctor_set(x_335, 0, x_330); +lean_ctor_set(x_335, 1, x_331); +return x_335; +} +else +{ +lean_object* x_336; lean_object* x_337; +lean_dec(x_332); +lean_dec(x_331); +x_336 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2; +x_337 = l_Std_Internal_Parsec_String_pstring(x_336, x_330); +if (lean_obj_tag(x_337) == 0) +{ +lean_object* x_338; lean_object* x_339; uint8_t x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_333); +x_338 = lean_ctor_get(x_337, 0); +lean_inc(x_338); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_339 = x_337; +} else { + lean_dec_ref(x_337); + x_339 = lean_box(0); +} +x_340 = 1; +x_341 = lean_box(x_340); +if (lean_is_scalar(x_339)) { + x_342 = lean_alloc_ctor(0, 2, 0); +} else { + x_342 = x_339; +} +lean_ctor_set(x_342, 0, x_338); +lean_ctor_set(x_342, 1, x_341); +return x_342; +} +else +{ +lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; uint8_t x_347; +x_343 = lean_ctor_get(x_337, 0); +lean_inc(x_343); +x_344 = lean_ctor_get(x_337, 1); +lean_inc(x_344); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_345 = x_337; +} else { + lean_dec_ref(x_337); + x_345 = lean_box(0); +} +x_346 = lean_ctor_get(x_343, 1); +lean_inc(x_346); +x_347 = lean_nat_dec_eq(x_333, x_346); +lean_dec(x_333); +if (x_347 == 0) +{ +lean_object* x_348; +lean_dec(x_346); +if (lean_is_scalar(x_345)) { + x_348 = lean_alloc_ctor(1, 2, 0); +} else { + x_348 = x_345; +} +lean_ctor_set(x_348, 0, x_343); +lean_ctor_set(x_348, 1, x_344); +return x_348; +} +else +{ +lean_object* x_349; lean_object* x_350; +lean_dec(x_345); +lean_dec(x_344); +x_349 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3; +x_350 = l_Std_Internal_Parsec_String_pstring(x_349, x_343); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; uint8_t x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_346); +x_351 = lean_ctor_get(x_350, 0); +lean_inc(x_351); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + x_352 = x_350; +} else { + lean_dec_ref(x_350); + x_352 = lean_box(0); +} +x_353 = 2; +x_354 = lean_box(x_353); +if (lean_is_scalar(x_352)) { + x_355 = lean_alloc_ctor(0, 2, 0); +} else { + x_355 = x_352; +} +lean_ctor_set(x_355, 0, x_351); +lean_ctor_set(x_355, 1, x_354); +return x_355; +} +else +{ +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; +x_356 = lean_ctor_get(x_350, 0); +lean_inc(x_356); +x_357 = lean_ctor_get(x_350, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + lean_ctor_release(x_350, 1); + x_358 = x_350; +} else { + lean_dec_ref(x_350); + x_358 = lean_box(0); +} +x_359 = lean_ctor_get(x_356, 1); +lean_inc(x_359); +x_360 = lean_nat_dec_eq(x_346, x_359); +lean_dec(x_346); +if (x_360 == 0) +{ +lean_object* x_361; +lean_dec(x_359); +if (lean_is_scalar(x_358)) { + x_361 = lean_alloc_ctor(1, 2, 0); +} else { + x_361 = x_358; +} +lean_ctor_set(x_361, 0, x_356); +lean_ctor_set(x_361, 1, x_357); +return x_361; +} +else +{ +lean_object* x_362; lean_object* x_363; +lean_dec(x_358); +lean_dec(x_357); +x_362 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4; +x_363 = l_Std_Internal_Parsec_String_pstring(x_362, x_356); +if (lean_obj_tag(x_363) == 0) +{ +lean_object* x_364; lean_object* x_365; uint8_t x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_359); +x_364 = lean_ctor_get(x_363, 0); +lean_inc(x_364); +if (lean_is_exclusive(x_363)) { + lean_ctor_release(x_363, 0); + lean_ctor_release(x_363, 1); + x_365 = x_363; +} else { + lean_dec_ref(x_363); + x_365 = lean_box(0); +} +x_366 = 3; +x_367 = lean_box(x_366); +if (lean_is_scalar(x_365)) { + x_368 = lean_alloc_ctor(0, 2, 0); +} else { + x_368 = x_365; +} +lean_ctor_set(x_368, 0, x_364); +lean_ctor_set(x_368, 1, x_367); +return x_368; +} +else +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; uint8_t x_373; +x_369 = lean_ctor_get(x_363, 0); +lean_inc(x_369); +x_370 = lean_ctor_get(x_363, 1); +lean_inc(x_370); +if (lean_is_exclusive(x_363)) { + lean_ctor_release(x_363, 0); + lean_ctor_release(x_363, 1); + x_371 = x_363; +} else { + lean_dec_ref(x_363); + x_371 = lean_box(0); +} +x_372 = lean_ctor_get(x_369, 1); +lean_inc(x_372); +x_373 = lean_nat_dec_eq(x_359, x_372); +lean_dec(x_359); +if (x_373 == 0) +{ +lean_object* x_374; +lean_dec(x_372); +if (lean_is_scalar(x_371)) { + x_374 = lean_alloc_ctor(1, 2, 0); +} else { + x_374 = x_371; +} +lean_ctor_set(x_374, 0, x_369); +lean_ctor_set(x_374, 1, x_370); +return x_374; +} +else +{ +lean_object* x_375; lean_object* x_376; +lean_dec(x_371); +lean_dec(x_370); +x_375 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5; +x_376 = l_Std_Internal_Parsec_String_pstring(x_375, x_369); +if (lean_obj_tag(x_376) == 0) +{ +lean_object* x_377; lean_object* x_378; uint8_t x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_372); +x_377 = lean_ctor_get(x_376, 0); +lean_inc(x_377); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_378 = x_376; +} else { + lean_dec_ref(x_376); + x_378 = lean_box(0); +} +x_379 = 4; +x_380 = lean_box(x_379); +if (lean_is_scalar(x_378)) { + x_381 = lean_alloc_ctor(0, 2, 0); +} else { + x_381 = x_378; +} +lean_ctor_set(x_381, 0, x_377); +lean_ctor_set(x_381, 1, x_380); +return x_381; +} +else +{ +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; uint8_t x_386; +x_382 = lean_ctor_get(x_376, 0); +lean_inc(x_382); +x_383 = lean_ctor_get(x_376, 1); +lean_inc(x_383); +if (lean_is_exclusive(x_376)) { + lean_ctor_release(x_376, 0); + lean_ctor_release(x_376, 1); + x_384 = x_376; +} else { + lean_dec_ref(x_376); + x_384 = lean_box(0); +} +x_385 = lean_ctor_get(x_382, 1); +lean_inc(x_385); +x_386 = lean_nat_dec_eq(x_372, x_385); +lean_dec(x_385); +lean_dec(x_372); +if (x_386 == 0) +{ +lean_object* x_387; +if (lean_is_scalar(x_384)) { + x_387 = lean_alloc_ctor(1, 2, 0); +} else { + x_387 = x_384; +} +lean_ctor_set(x_387, 0, x_382); +lean_ctor_set(x_387, 1, x_383); +return x_387; +} +else +{ +lean_object* x_388; lean_object* x_389; +lean_dec(x_384); +lean_dec(x_383); +x_388 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6; +x_389 = l_Std_Internal_Parsec_String_pstring(x_388, x_382); +if (lean_obj_tag(x_389) == 0) +{ +lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; lean_object* x_394; +x_390 = lean_ctor_get(x_389, 0); +lean_inc(x_390); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_391 = x_389; +} else { + lean_dec_ref(x_389); + x_391 = lean_box(0); +} +x_392 = 5; +x_393 = lean_box(x_392); +if (lean_is_scalar(x_391)) { + x_394 = lean_alloc_ctor(0, 2, 0); +} else { + x_394 = x_391; +} +lean_ctor_set(x_394, 0, x_390); +lean_ctor_set(x_394, 1, x_393); +return x_394; +} +else +{ +lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; +x_395 = lean_ctor_get(x_389, 0); +lean_inc(x_395); +x_396 = lean_ctor_get(x_389, 1); +lean_inc(x_396); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_397 = x_389; +} else { + lean_dec_ref(x_389); + x_397 = lean_box(0); +} +if (lean_is_scalar(x_397)) { + x_398 = lean_alloc_ctor(1, 2, 0); +} else { + x_398 = x_397; +} +lean_ctor_set(x_398, 0, x_395); +lean_ctor_set(x_398, 1, x_396); +return x_398; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayNarrow(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 6; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 6; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_dec(x_16); +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 0; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +x_32 = lean_nat_dec_eq(x_16, x_31); +lean_dec(x_16); +if (x_32 == 0) +{ +lean_dec(x_31); +return x_19; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_free_object(x_19); +lean_dec(x_30); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1; +x_34 = l_Std_Internal_Parsec_String_pstring(x_33, x_29); +if (lean_obj_tag(x_34) == 0) +{ +uint8_t x_35; +lean_dec(x_31); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) +{ +lean_object* x_36; uint8_t x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_34, 1); +lean_dec(x_36); +x_37 = 1; +x_38 = lean_box(x_37); +lean_ctor_set(x_34, 1, x_38); +return x_34; +} +else +{ +lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_34, 0); +lean_inc(x_39); +lean_dec(x_34); +x_40 = 1; +x_41 = lean_box(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +} +else +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_34); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_34, 0); +x_45 = lean_ctor_get(x_34, 1); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +x_47 = lean_nat_dec_eq(x_31, x_46); +lean_dec(x_31); +if (x_47 == 0) +{ +lean_dec(x_46); +return x_34; +} +else +{ +lean_object* x_48; lean_object* x_49; +lean_free_object(x_34); +lean_dec(x_45); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +x_49 = l_Std_Internal_Parsec_String_pstring(x_48, x_44); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_50; +lean_dec(x_46); +x_50 = !lean_is_exclusive(x_49); +if (x_50 == 0) +{ +lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_49, 1); +lean_dec(x_51); +x_52 = 2; +x_53 = lean_box(x_52); +lean_ctor_set(x_49, 1, x_53); +return x_49; +} +else +{ +lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_49, 0); +lean_inc(x_54); +lean_dec(x_49); +x_55 = 2; +x_56 = lean_box(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +else +{ +uint8_t x_58; +x_58 = !lean_is_exclusive(x_49); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_59 = lean_ctor_get(x_49, 0); +x_60 = lean_ctor_get(x_49, 1); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +x_62 = lean_nat_dec_eq(x_46, x_61); +lean_dec(x_46); +if (x_62 == 0) +{ +lean_dec(x_61); +return x_49; +} +else +{ +lean_object* x_63; +lean_free_object(x_49); +lean_dec(x_60); +x_63 = l_Std_Internal_Parsec_String_pstring(x_33, x_59); +if (lean_obj_tag(x_63) == 0) +{ +uint8_t x_64; +lean_dec(x_61); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +lean_object* x_65; uint8_t x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_63, 1); +lean_dec(x_65); +x_66 = 3; +x_67 = lean_box(x_66); +lean_ctor_set(x_63, 1, x_67); +return x_63; +} +else +{ +lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; +x_68 = lean_ctor_get(x_63, 0); +lean_inc(x_68); +lean_dec(x_63); +x_69 = 3; +x_70 = lean_box(x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_68); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +else +{ +uint8_t x_72; +x_72 = !lean_is_exclusive(x_63); +if (x_72 == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_73 = lean_ctor_get(x_63, 0); +x_74 = lean_ctor_get(x_63, 1); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +x_76 = lean_nat_dec_eq(x_61, x_75); +lean_dec(x_61); +if (x_76 == 0) +{ +lean_dec(x_75); +return x_63; +} +else +{ +lean_object* x_77; lean_object* x_78; +lean_free_object(x_63); +lean_dec(x_74); +x_77 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_78 = l_Std_Internal_Parsec_String_pstring(x_77, x_73); +if (lean_obj_tag(x_78) == 0) +{ +uint8_t x_79; +lean_dec(x_75); +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) +{ +lean_object* x_80; uint8_t x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_78, 1); +lean_dec(x_80); +x_81 = 4; +x_82 = lean_box(x_81); +lean_ctor_set(x_78, 1, x_82); +return x_78; +} +else +{ +lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; +x_83 = lean_ctor_get(x_78, 0); +lean_inc(x_83); +lean_dec(x_78); +x_84 = 4; +x_85 = lean_box(x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_85); +return x_86; +} +} +else +{ +uint8_t x_87; +x_87 = !lean_is_exclusive(x_78); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_88 = lean_ctor_get(x_78, 0); +x_89 = lean_ctor_get(x_78, 1); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +x_91 = lean_nat_dec_eq(x_75, x_90); +lean_dec(x_90); +lean_dec(x_75); +if (x_91 == 0) +{ +return x_78; +} +else +{ +lean_object* x_92; +lean_free_object(x_78); +lean_dec(x_89); +x_92 = l_Std_Internal_Parsec_String_pstring(x_2, x_88); +if (lean_obj_tag(x_92) == 0) +{ +uint8_t x_93; +x_93 = !lean_is_exclusive(x_92); +if (x_93 == 0) +{ +lean_object* x_94; uint8_t x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_92, 1); +lean_dec(x_94); +x_95 = 5; +x_96 = lean_box(x_95); +lean_ctor_set(x_92, 1, x_96); +return x_92; +} +else +{ +lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_92, 0); +lean_inc(x_97); +lean_dec(x_92); +x_98 = 5; +x_99 = lean_box(x_98); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_99); +return x_100; +} +} +else +{ +uint8_t x_101; +x_101 = !lean_is_exclusive(x_92); +if (x_101 == 0) +{ +return x_92; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_92, 0); +x_103 = lean_ctor_get(x_92, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_92); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +return x_104; +} +} +} +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +x_105 = lean_ctor_get(x_78, 0); +x_106 = lean_ctor_get(x_78, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_78); +x_107 = lean_ctor_get(x_105, 1); +lean_inc(x_107); +x_108 = lean_nat_dec_eq(x_75, x_107); +lean_dec(x_107); +lean_dec(x_75); +if (x_108 == 0) +{ +lean_object* x_109; +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_105); +lean_ctor_set(x_109, 1, x_106); +return x_109; +} +else +{ +lean_object* x_110; +lean_dec(x_106); +x_110 = l_Std_Internal_Parsec_String_pstring(x_2, x_105); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_112 = x_110; +} else { + lean_dec_ref(x_110); + x_112 = lean_box(0); +} +x_113 = 5; +x_114 = lean_box(x_113); +if (lean_is_scalar(x_112)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_112; +} +lean_ctor_set(x_115, 0, x_111); +lean_ctor_set(x_115, 1, x_114); +return x_115; +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_116 = lean_ctor_get(x_110, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_110, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_118 = x_110; +} else { + lean_dec_ref(x_110); + x_118 = lean_box(0); +} +if (lean_is_scalar(x_118)) { + x_119 = lean_alloc_ctor(1, 2, 0); +} else { + x_119 = x_118; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_117); +return x_119; +} +} +} +} +} +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_120 = lean_ctor_get(x_63, 0); +x_121 = lean_ctor_get(x_63, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_63); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +x_123 = lean_nat_dec_eq(x_61, x_122); +lean_dec(x_61); +if (x_123 == 0) +{ +lean_object* x_124; +lean_dec(x_122); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_120); +lean_ctor_set(x_124, 1, x_121); +return x_124; +} +else +{ +lean_object* x_125; lean_object* x_126; +lean_dec(x_121); +x_125 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_126 = l_Std_Internal_Parsec_String_pstring(x_125, x_120); +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +lean_dec(x_122); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_128 = x_126; +} else { + lean_dec_ref(x_126); + x_128 = lean_box(0); +} +x_129 = 4; +x_130 = lean_box(x_129); +if (lean_is_scalar(x_128)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_128; +} +lean_ctor_set(x_131, 0, x_127); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_132 = lean_ctor_get(x_126, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_126, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_134 = x_126; +} else { + lean_dec_ref(x_126); + x_134 = lean_box(0); +} +x_135 = lean_ctor_get(x_132, 1); +lean_inc(x_135); +x_136 = lean_nat_dec_eq(x_122, x_135); +lean_dec(x_135); +lean_dec(x_122); +if (x_136 == 0) +{ +lean_object* x_137; +if (lean_is_scalar(x_134)) { + x_137 = lean_alloc_ctor(1, 2, 0); +} else { + x_137 = x_134; +} +lean_ctor_set(x_137, 0, x_132); +lean_ctor_set(x_137, 1, x_133); +return x_137; +} +else +{ +lean_object* x_138; +lean_dec(x_134); +lean_dec(x_133); +x_138 = l_Std_Internal_Parsec_String_pstring(x_2, x_132); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; +x_139 = lean_ctor_get(x_138, 0); +lean_inc(x_139); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_140 = x_138; +} else { + lean_dec_ref(x_138); + x_140 = lean_box(0); +} +x_141 = 5; +x_142 = lean_box(x_141); +if (lean_is_scalar(x_140)) { + x_143 = lean_alloc_ctor(0, 2, 0); +} else { + x_143 = x_140; +} +lean_ctor_set(x_143, 0, x_139); +lean_ctor_set(x_143, 1, x_142); +return x_143; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_144 = lean_ctor_get(x_138, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_138, 1); +lean_inc(x_145); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + lean_ctor_release(x_138, 1); + x_146 = x_138; +} else { + lean_dec_ref(x_138); + x_146 = lean_box(0); +} +if (lean_is_scalar(x_146)) { + x_147 = lean_alloc_ctor(1, 2, 0); +} else { + x_147 = x_146; +} +lean_ctor_set(x_147, 0, x_144); +lean_ctor_set(x_147, 1, x_145); +return x_147; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +x_148 = lean_ctor_get(x_49, 0); +x_149 = lean_ctor_get(x_49, 1); +lean_inc(x_149); +lean_inc(x_148); +lean_dec(x_49); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +x_151 = lean_nat_dec_eq(x_46, x_150); +lean_dec(x_46); +if (x_151 == 0) +{ +lean_object* x_152; +lean_dec(x_150); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_148); +lean_ctor_set(x_152, 1, x_149); +return x_152; +} +else +{ +lean_object* x_153; +lean_dec(x_149); +x_153 = l_Std_Internal_Parsec_String_pstring(x_33, x_148); +if (lean_obj_tag(x_153) == 0) +{ +lean_object* x_154; lean_object* x_155; uint8_t x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_150); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_155 = x_153; +} else { + lean_dec_ref(x_153); + x_155 = lean_box(0); +} +x_156 = 3; +x_157 = lean_box(x_156); +if (lean_is_scalar(x_155)) { + x_158 = lean_alloc_ctor(0, 2, 0); +} else { + x_158 = x_155; +} +lean_ctor_set(x_158, 0, x_154); +lean_ctor_set(x_158, 1, x_157); +return x_158; +} +else +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; uint8_t x_163; +x_159 = lean_ctor_get(x_153, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_153, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + x_161 = x_153; +} else { + lean_dec_ref(x_153); + x_161 = lean_box(0); +} +x_162 = lean_ctor_get(x_159, 1); +lean_inc(x_162); +x_163 = lean_nat_dec_eq(x_150, x_162); +lean_dec(x_150); +if (x_163 == 0) +{ +lean_object* x_164; +lean_dec(x_162); +if (lean_is_scalar(x_161)) { + x_164 = lean_alloc_ctor(1, 2, 0); +} else { + x_164 = x_161; +} +lean_ctor_set(x_164, 0, x_159); +lean_ctor_set(x_164, 1, x_160); +return x_164; +} +else +{ +lean_object* x_165; lean_object* x_166; +lean_dec(x_161); +lean_dec(x_160); +x_165 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_166 = l_Std_Internal_Parsec_String_pstring(x_165, x_159); +if (lean_obj_tag(x_166) == 0) +{ +lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_162); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_168 = x_166; +} else { + lean_dec_ref(x_166); + x_168 = lean_box(0); +} +x_169 = 4; +x_170 = lean_box(x_169); +if (lean_is_scalar(x_168)) { + x_171 = lean_alloc_ctor(0, 2, 0); +} else { + x_171 = x_168; +} +lean_ctor_set(x_171, 0, x_167); +lean_ctor_set(x_171, 1, x_170); +return x_171; +} +else +{ +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; +x_172 = lean_ctor_get(x_166, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_166, 1); +lean_inc(x_173); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_174 = x_166; +} else { + lean_dec_ref(x_166); + x_174 = lean_box(0); +} +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +x_176 = lean_nat_dec_eq(x_162, x_175); +lean_dec(x_175); +lean_dec(x_162); +if (x_176 == 0) +{ +lean_object* x_177; +if (lean_is_scalar(x_174)) { + x_177 = lean_alloc_ctor(1, 2, 0); +} else { + x_177 = x_174; +} +lean_ctor_set(x_177, 0, x_172); +lean_ctor_set(x_177, 1, x_173); +return x_177; +} +else +{ +lean_object* x_178; +lean_dec(x_174); +lean_dec(x_173); +x_178 = l_Std_Internal_Parsec_String_pstring(x_2, x_172); +if (lean_obj_tag(x_178) == 0) +{ +lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; lean_object* x_183; +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_180 = x_178; +} else { + lean_dec_ref(x_178); + x_180 = lean_box(0); +} +x_181 = 5; +x_182 = lean_box(x_181); +if (lean_is_scalar(x_180)) { + x_183 = lean_alloc_ctor(0, 2, 0); +} else { + x_183 = x_180; +} +lean_ctor_set(x_183, 0, x_179); +lean_ctor_set(x_183, 1, x_182); +return x_183; +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_184 = lean_ctor_get(x_178, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_178, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_178)) { + lean_ctor_release(x_178, 0); + lean_ctor_release(x_178, 1); + x_186 = x_178; +} else { + lean_dec_ref(x_178); + x_186 = lean_box(0); +} +if (lean_is_scalar(x_186)) { + x_187 = lean_alloc_ctor(1, 2, 0); +} else { + x_187 = x_186; +} +lean_ctor_set(x_187, 0, x_184); +lean_ctor_set(x_187, 1, x_185); +return x_187; +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +x_188 = lean_ctor_get(x_34, 0); +x_189 = lean_ctor_get(x_34, 1); +lean_inc(x_189); +lean_inc(x_188); +lean_dec(x_34); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +x_191 = lean_nat_dec_eq(x_31, x_190); +lean_dec(x_31); +if (x_191 == 0) +{ +lean_object* x_192; +lean_dec(x_190); +x_192 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_192, 0, x_188); +lean_ctor_set(x_192, 1, x_189); +return x_192; +} +else +{ +lean_object* x_193; lean_object* x_194; +lean_dec(x_189); +x_193 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +x_194 = l_Std_Internal_Parsec_String_pstring(x_193, x_188); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_190); +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_196 = x_194; +} else { + lean_dec_ref(x_194); + x_196 = lean_box(0); +} +x_197 = 2; +x_198 = lean_box(x_197); +if (lean_is_scalar(x_196)) { + x_199 = lean_alloc_ctor(0, 2, 0); +} else { + x_199 = x_196; +} +lean_ctor_set(x_199, 0, x_195); +lean_ctor_set(x_199, 1, x_198); +return x_199; +} +else +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; +x_200 = lean_ctor_get(x_194, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_194, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_202 = x_194; +} else { + lean_dec_ref(x_194); + x_202 = lean_box(0); +} +x_203 = lean_ctor_get(x_200, 1); +lean_inc(x_203); +x_204 = lean_nat_dec_eq(x_190, x_203); +lean_dec(x_190); +if (x_204 == 0) +{ +lean_object* x_205; +lean_dec(x_203); +if (lean_is_scalar(x_202)) { + x_205 = lean_alloc_ctor(1, 2, 0); +} else { + x_205 = x_202; +} +lean_ctor_set(x_205, 0, x_200); +lean_ctor_set(x_205, 1, x_201); +return x_205; +} +else +{ +lean_object* x_206; +lean_dec(x_202); +lean_dec(x_201); +x_206 = l_Std_Internal_Parsec_String_pstring(x_33, x_200); +if (lean_obj_tag(x_206) == 0) +{ +lean_object* x_207; lean_object* x_208; uint8_t x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_203); +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_208 = x_206; +} else { + lean_dec_ref(x_206); + x_208 = lean_box(0); +} +x_209 = 3; +x_210 = lean_box(x_209); +if (lean_is_scalar(x_208)) { + x_211 = lean_alloc_ctor(0, 2, 0); +} else { + x_211 = x_208; +} +lean_ctor_set(x_211, 0, x_207); +lean_ctor_set(x_211, 1, x_210); +return x_211; +} +else +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; uint8_t x_216; +x_212 = lean_ctor_get(x_206, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_206, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_214 = x_206; +} else { + lean_dec_ref(x_206); + x_214 = lean_box(0); +} +x_215 = lean_ctor_get(x_212, 1); +lean_inc(x_215); +x_216 = lean_nat_dec_eq(x_203, x_215); +lean_dec(x_203); +if (x_216 == 0) +{ +lean_object* x_217; +lean_dec(x_215); +if (lean_is_scalar(x_214)) { + x_217 = lean_alloc_ctor(1, 2, 0); +} else { + x_217 = x_214; +} +lean_ctor_set(x_217, 0, x_212); +lean_ctor_set(x_217, 1, x_213); +return x_217; +} +else +{ +lean_object* x_218; lean_object* x_219; +lean_dec(x_214); +lean_dec(x_213); +x_218 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_219 = l_Std_Internal_Parsec_String_pstring(x_218, x_212); +if (lean_obj_tag(x_219) == 0) +{ +lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_215); +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_221 = x_219; +} else { + lean_dec_ref(x_219); + x_221 = lean_box(0); +} +x_222 = 4; +x_223 = lean_box(x_222); +if (lean_is_scalar(x_221)) { + x_224 = lean_alloc_ctor(0, 2, 0); +} else { + x_224 = x_221; +} +lean_ctor_set(x_224, 0, x_220); +lean_ctor_set(x_224, 1, x_223); +return x_224; +} +else +{ +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; uint8_t x_229; +x_225 = lean_ctor_get(x_219, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_219, 1); +lean_inc(x_226); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_227 = x_219; +} else { + lean_dec_ref(x_219); + x_227 = lean_box(0); +} +x_228 = lean_ctor_get(x_225, 1); +lean_inc(x_228); +x_229 = lean_nat_dec_eq(x_215, x_228); +lean_dec(x_228); +lean_dec(x_215); +if (x_229 == 0) +{ +lean_object* x_230; +if (lean_is_scalar(x_227)) { + x_230 = lean_alloc_ctor(1, 2, 0); +} else { + x_230 = x_227; +} +lean_ctor_set(x_230, 0, x_225); +lean_ctor_set(x_230, 1, x_226); +return x_230; +} +else +{ +lean_object* x_231; +lean_dec(x_227); +lean_dec(x_226); +x_231 = l_Std_Internal_Parsec_String_pstring(x_2, x_225); +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_232; lean_object* x_233; uint8_t x_234; lean_object* x_235; lean_object* x_236; +x_232 = lean_ctor_get(x_231, 0); +lean_inc(x_232); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_233 = x_231; +} else { + lean_dec_ref(x_231); + x_233 = lean_box(0); +} +x_234 = 5; +x_235 = lean_box(x_234); +if (lean_is_scalar(x_233)) { + x_236 = lean_alloc_ctor(0, 2, 0); +} else { + x_236 = x_233; +} +lean_ctor_set(x_236, 0, x_232); +lean_ctor_set(x_236, 1, x_235); +return x_236; +} +else +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +x_237 = lean_ctor_get(x_231, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_231, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_239 = x_231; +} else { + lean_dec_ref(x_231); + x_239 = lean_box(0); +} +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); +} else { + x_240 = x_239; +} +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +return x_240; +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; uint8_t x_244; +x_241 = lean_ctor_get(x_19, 0); +x_242 = lean_ctor_get(x_19, 1); +lean_inc(x_242); +lean_inc(x_241); +lean_dec(x_19); +x_243 = lean_ctor_get(x_241, 1); +lean_inc(x_243); +x_244 = lean_nat_dec_eq(x_16, x_243); +lean_dec(x_16); +if (x_244 == 0) +{ +lean_object* x_245; +lean_dec(x_243); +x_245 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_245, 0, x_241); +lean_ctor_set(x_245, 1, x_242); +return x_245; +} +else +{ +lean_object* x_246; lean_object* x_247; +lean_dec(x_242); +x_246 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1; +x_247 = l_Std_Internal_Parsec_String_pstring(x_246, x_241); +if (lean_obj_tag(x_247) == 0) +{ +lean_object* x_248; lean_object* x_249; uint8_t x_250; lean_object* x_251; lean_object* x_252; +lean_dec(x_243); +x_248 = lean_ctor_get(x_247, 0); +lean_inc(x_248); +if (lean_is_exclusive(x_247)) { + lean_ctor_release(x_247, 0); + lean_ctor_release(x_247, 1); + x_249 = x_247; +} else { + lean_dec_ref(x_247); + x_249 = lean_box(0); +} +x_250 = 1; +x_251 = lean_box(x_250); +if (lean_is_scalar(x_249)) { + x_252 = lean_alloc_ctor(0, 2, 0); +} else { + x_252 = x_249; +} +lean_ctor_set(x_252, 0, x_248); +lean_ctor_set(x_252, 1, x_251); +return x_252; +} +else +{ +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; +x_253 = lean_ctor_get(x_247, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_247, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_247)) { + lean_ctor_release(x_247, 0); + lean_ctor_release(x_247, 1); + x_255 = x_247; +} else { + lean_dec_ref(x_247); + x_255 = lean_box(0); +} +x_256 = lean_ctor_get(x_253, 1); +lean_inc(x_256); +x_257 = lean_nat_dec_eq(x_243, x_256); +lean_dec(x_243); +if (x_257 == 0) +{ +lean_object* x_258; +lean_dec(x_256); +if (lean_is_scalar(x_255)) { + x_258 = lean_alloc_ctor(1, 2, 0); +} else { + x_258 = x_255; +} +lean_ctor_set(x_258, 0, x_253); +lean_ctor_set(x_258, 1, x_254); +return x_258; +} +else +{ +lean_object* x_259; lean_object* x_260; +lean_dec(x_255); +lean_dec(x_254); +x_259 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +x_260 = l_Std_Internal_Parsec_String_pstring(x_259, x_253); +if (lean_obj_tag(x_260) == 0) +{ +lean_object* x_261; lean_object* x_262; uint8_t x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_256); +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +if (lean_is_exclusive(x_260)) { + lean_ctor_release(x_260, 0); + lean_ctor_release(x_260, 1); + x_262 = x_260; +} else { + lean_dec_ref(x_260); + x_262 = lean_box(0); +} +x_263 = 2; +x_264 = lean_box(x_263); +if (lean_is_scalar(x_262)) { + x_265 = lean_alloc_ctor(0, 2, 0); +} else { + x_265 = x_262; +} +lean_ctor_set(x_265, 0, x_261); +lean_ctor_set(x_265, 1, x_264); +return x_265; +} +else +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; +x_266 = lean_ctor_get(x_260, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_260, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_260)) { + lean_ctor_release(x_260, 0); + lean_ctor_release(x_260, 1); + x_268 = x_260; +} else { + lean_dec_ref(x_260); + x_268 = lean_box(0); +} +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +x_270 = lean_nat_dec_eq(x_256, x_269); +lean_dec(x_256); +if (x_270 == 0) +{ +lean_object* x_271; +lean_dec(x_269); +if (lean_is_scalar(x_268)) { + x_271 = lean_alloc_ctor(1, 2, 0); +} else { + x_271 = x_268; +} +lean_ctor_set(x_271, 0, x_266); +lean_ctor_set(x_271, 1, x_267); +return x_271; +} +else +{ +lean_object* x_272; +lean_dec(x_268); +lean_dec(x_267); +x_272 = l_Std_Internal_Parsec_String_pstring(x_246, x_266); +if (lean_obj_tag(x_272) == 0) +{ +lean_object* x_273; lean_object* x_274; uint8_t x_275; lean_object* x_276; lean_object* x_277; +lean_dec(x_269); +x_273 = lean_ctor_get(x_272, 0); +lean_inc(x_273); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + lean_ctor_release(x_272, 1); + x_274 = x_272; +} else { + lean_dec_ref(x_272); + x_274 = lean_box(0); +} +x_275 = 3; +x_276 = lean_box(x_275); +if (lean_is_scalar(x_274)) { + x_277 = lean_alloc_ctor(0, 2, 0); +} else { + x_277 = x_274; +} +lean_ctor_set(x_277, 0, x_273); +lean_ctor_set(x_277, 1, x_276); +return x_277; +} +else +{ +lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; uint8_t x_282; +x_278 = lean_ctor_get(x_272, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_272, 1); +lean_inc(x_279); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + lean_ctor_release(x_272, 1); + x_280 = x_272; +} else { + lean_dec_ref(x_272); + x_280 = lean_box(0); +} +x_281 = lean_ctor_get(x_278, 1); +lean_inc(x_281); +x_282 = lean_nat_dec_eq(x_269, x_281); +lean_dec(x_269); +if (x_282 == 0) +{ +lean_object* x_283; +lean_dec(x_281); +if (lean_is_scalar(x_280)) { + x_283 = lean_alloc_ctor(1, 2, 0); +} else { + x_283 = x_280; +} +lean_ctor_set(x_283, 0, x_278); +lean_ctor_set(x_283, 1, x_279); +return x_283; +} +else +{ +lean_object* x_284; lean_object* x_285; +lean_dec(x_280); +lean_dec(x_279); +x_284 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_285 = l_Std_Internal_Parsec_String_pstring(x_284, x_278); +if (lean_obj_tag(x_285) == 0) +{ +lean_object* x_286; lean_object* x_287; uint8_t x_288; lean_object* x_289; lean_object* x_290; +lean_dec(x_281); +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_287 = x_285; +} else { + lean_dec_ref(x_285); + x_287 = lean_box(0); +} +x_288 = 4; +x_289 = lean_box(x_288); +if (lean_is_scalar(x_287)) { + x_290 = lean_alloc_ctor(0, 2, 0); +} else { + x_290 = x_287; +} +lean_ctor_set(x_290, 0, x_286); +lean_ctor_set(x_290, 1, x_289); +return x_290; +} +else +{ +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; uint8_t x_295; +x_291 = lean_ctor_get(x_285, 0); +lean_inc(x_291); +x_292 = lean_ctor_get(x_285, 1); +lean_inc(x_292); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_293 = x_285; +} else { + lean_dec_ref(x_285); + x_293 = lean_box(0); +} +x_294 = lean_ctor_get(x_291, 1); +lean_inc(x_294); +x_295 = lean_nat_dec_eq(x_281, x_294); +lean_dec(x_294); +lean_dec(x_281); +if (x_295 == 0) +{ +lean_object* x_296; +if (lean_is_scalar(x_293)) { + x_296 = lean_alloc_ctor(1, 2, 0); +} else { + x_296 = x_293; +} +lean_ctor_set(x_296, 0, x_291); +lean_ctor_set(x_296, 1, x_292); +return x_296; +} +else +{ +lean_object* x_297; +lean_dec(x_293); +lean_dec(x_292); +x_297 = l_Std_Internal_Parsec_String_pstring(x_2, x_291); +if (lean_obj_tag(x_297) == 0) +{ +lean_object* x_298; lean_object* x_299; uint8_t x_300; lean_object* x_301; lean_object* x_302; +x_298 = lean_ctor_get(x_297, 0); +lean_inc(x_298); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + lean_ctor_release(x_297, 1); + x_299 = x_297; +} else { + lean_dec_ref(x_297); + x_299 = lean_box(0); +} +x_300 = 5; +x_301 = lean_box(x_300); +if (lean_is_scalar(x_299)) { + x_302 = lean_alloc_ctor(0, 2, 0); +} else { + x_302 = x_299; +} +lean_ctor_set(x_302, 0, x_298); +lean_ctor_set(x_302, 1, x_301); +return x_302; +} +else +{ +lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; +x_303 = lean_ctor_get(x_297, 0); +lean_inc(x_303); +x_304 = lean_ctor_get(x_297, 1); +lean_inc(x_304); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + lean_ctor_release(x_297, 1); + x_305 = x_297; +} else { + lean_dec_ref(x_297); + x_305 = lean_box(0); +} +if (lean_is_scalar(x_305)) { + x_306 = lean_alloc_ctor(1, 2, 0); +} else { + x_306 = x_305; +} +lean_ctor_set(x_306, 0, x_303); +lean_ctor_set(x_306, 1, x_304); +return x_306; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +else +{ +lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; uint8_t x_311; +x_307 = lean_ctor_get(x_3, 0); +x_308 = lean_ctor_get(x_3, 1); +lean_inc(x_308); +lean_inc(x_307); +lean_dec(x_3); +x_309 = lean_ctor_get(x_1, 1); +lean_inc(x_309); +lean_dec(x_1); +x_310 = lean_ctor_get(x_307, 1); +lean_inc(x_310); +x_311 = lean_nat_dec_eq(x_309, x_310); +lean_dec(x_309); +if (x_311 == 0) +{ +lean_object* x_312; +lean_dec(x_310); +x_312 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_312, 0, x_307); +lean_ctor_set(x_312, 1, x_308); +return x_312; +} +else +{ +lean_object* x_313; lean_object* x_314; +lean_dec(x_308); +x_313 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7; +x_314 = l_Std_Internal_Parsec_String_pstring(x_313, x_307); +if (lean_obj_tag(x_314) == 0) +{ +lean_object* x_315; lean_object* x_316; uint8_t x_317; lean_object* x_318; lean_object* x_319; +lean_dec(x_310); +x_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_316 = x_314; +} else { + lean_dec_ref(x_314); + x_316 = lean_box(0); +} +x_317 = 0; +x_318 = lean_box(x_317); +if (lean_is_scalar(x_316)) { + x_319 = lean_alloc_ctor(0, 2, 0); +} else { + x_319 = x_316; +} +lean_ctor_set(x_319, 0, x_315); +lean_ctor_set(x_319, 1, x_318); +return x_319; +} +else +{ +lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; uint8_t x_324; +x_320 = lean_ctor_get(x_314, 0); +lean_inc(x_320); +x_321 = lean_ctor_get(x_314, 1); +lean_inc(x_321); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_322 = x_314; +} else { + lean_dec_ref(x_314); + x_322 = lean_box(0); +} +x_323 = lean_ctor_get(x_320, 1); +lean_inc(x_323); +x_324 = lean_nat_dec_eq(x_310, x_323); +lean_dec(x_310); +if (x_324 == 0) +{ +lean_object* x_325; +lean_dec(x_323); +if (lean_is_scalar(x_322)) { + x_325 = lean_alloc_ctor(1, 2, 0); +} else { + x_325 = x_322; +} +lean_ctor_set(x_325, 0, x_320); +lean_ctor_set(x_325, 1, x_321); +return x_325; +} +else +{ +lean_object* x_326; lean_object* x_327; +lean_dec(x_322); +lean_dec(x_321); +x_326 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1; +x_327 = l_Std_Internal_Parsec_String_pstring(x_326, x_320); +if (lean_obj_tag(x_327) == 0) +{ +lean_object* x_328; lean_object* x_329; uint8_t x_330; lean_object* x_331; lean_object* x_332; +lean_dec(x_323); +x_328 = lean_ctor_get(x_327, 0); +lean_inc(x_328); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + lean_ctor_release(x_327, 1); + x_329 = x_327; +} else { + lean_dec_ref(x_327); + x_329 = lean_box(0); +} +x_330 = 1; +x_331 = lean_box(x_330); +if (lean_is_scalar(x_329)) { + x_332 = lean_alloc_ctor(0, 2, 0); +} else { + x_332 = x_329; +} +lean_ctor_set(x_332, 0, x_328); +lean_ctor_set(x_332, 1, x_331); +return x_332; +} +else +{ +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; uint8_t x_337; +x_333 = lean_ctor_get(x_327, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_327, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_327)) { + lean_ctor_release(x_327, 0); + lean_ctor_release(x_327, 1); + x_335 = x_327; +} else { + lean_dec_ref(x_327); + x_335 = lean_box(0); +} +x_336 = lean_ctor_get(x_333, 1); +lean_inc(x_336); +x_337 = lean_nat_dec_eq(x_323, x_336); +lean_dec(x_323); +if (x_337 == 0) +{ +lean_object* x_338; +lean_dec(x_336); +if (lean_is_scalar(x_335)) { + x_338 = lean_alloc_ctor(1, 2, 0); +} else { + x_338 = x_335; +} +lean_ctor_set(x_338, 0, x_333); +lean_ctor_set(x_338, 1, x_334); +return x_338; +} +else +{ +lean_object* x_339; lean_object* x_340; +lean_dec(x_335); +lean_dec(x_334); +x_339 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2; +x_340 = l_Std_Internal_Parsec_String_pstring(x_339, x_333); +if (lean_obj_tag(x_340) == 0) +{ +lean_object* x_341; lean_object* x_342; uint8_t x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_336); +x_341 = lean_ctor_get(x_340, 0); +lean_inc(x_341); +if (lean_is_exclusive(x_340)) { + lean_ctor_release(x_340, 0); + lean_ctor_release(x_340, 1); + x_342 = x_340; +} else { + lean_dec_ref(x_340); + x_342 = lean_box(0); +} +x_343 = 2; +x_344 = lean_box(x_343); +if (lean_is_scalar(x_342)) { + x_345 = lean_alloc_ctor(0, 2, 0); +} else { + x_345 = x_342; +} +lean_ctor_set(x_345, 0, x_341); +lean_ctor_set(x_345, 1, x_344); +return x_345; +} +else +{ +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; uint8_t x_350; +x_346 = lean_ctor_get(x_340, 0); +lean_inc(x_346); +x_347 = lean_ctor_get(x_340, 1); +lean_inc(x_347); +if (lean_is_exclusive(x_340)) { + lean_ctor_release(x_340, 0); + lean_ctor_release(x_340, 1); + x_348 = x_340; +} else { + lean_dec_ref(x_340); + x_348 = lean_box(0); +} +x_349 = lean_ctor_get(x_346, 1); +lean_inc(x_349); +x_350 = lean_nat_dec_eq(x_336, x_349); +lean_dec(x_336); +if (x_350 == 0) +{ +lean_object* x_351; +lean_dec(x_349); +if (lean_is_scalar(x_348)) { + x_351 = lean_alloc_ctor(1, 2, 0); +} else { + x_351 = x_348; +} +lean_ctor_set(x_351, 0, x_346); +lean_ctor_set(x_351, 1, x_347); +return x_351; +} +else +{ +lean_object* x_352; +lean_dec(x_348); +lean_dec(x_347); +x_352 = l_Std_Internal_Parsec_String_pstring(x_326, x_346); +if (lean_obj_tag(x_352) == 0) +{ +lean_object* x_353; lean_object* x_354; uint8_t x_355; lean_object* x_356; lean_object* x_357; +lean_dec(x_349); +x_353 = lean_ctor_get(x_352, 0); +lean_inc(x_353); +if (lean_is_exclusive(x_352)) { + lean_ctor_release(x_352, 0); + lean_ctor_release(x_352, 1); + x_354 = x_352; +} else { + lean_dec_ref(x_352); + x_354 = lean_box(0); +} +x_355 = 3; +x_356 = lean_box(x_355); +if (lean_is_scalar(x_354)) { + x_357 = lean_alloc_ctor(0, 2, 0); +} else { + x_357 = x_354; +} +lean_ctor_set(x_357, 0, x_353); +lean_ctor_set(x_357, 1, x_356); +return x_357; +} +else +{ +lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; +x_358 = lean_ctor_get(x_352, 0); +lean_inc(x_358); +x_359 = lean_ctor_get(x_352, 1); +lean_inc(x_359); +if (lean_is_exclusive(x_352)) { + lean_ctor_release(x_352, 0); + lean_ctor_release(x_352, 1); + x_360 = x_352; +} else { + lean_dec_ref(x_352); + x_360 = lean_box(0); +} +x_361 = lean_ctor_get(x_358, 1); +lean_inc(x_361); +x_362 = lean_nat_dec_eq(x_349, x_361); +lean_dec(x_349); +if (x_362 == 0) +{ +lean_object* x_363; +lean_dec(x_361); +if (lean_is_scalar(x_360)) { + x_363 = lean_alloc_ctor(1, 2, 0); +} else { + x_363 = x_360; +} +lean_ctor_set(x_363, 0, x_358); +lean_ctor_set(x_363, 1, x_359); +return x_363; +} +else +{ +lean_object* x_364; lean_object* x_365; +lean_dec(x_360); +lean_dec(x_359); +x_364 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8; +x_365 = l_Std_Internal_Parsec_String_pstring(x_364, x_358); +if (lean_obj_tag(x_365) == 0) +{ +lean_object* x_366; lean_object* x_367; uint8_t x_368; lean_object* x_369; lean_object* x_370; +lean_dec(x_361); +x_366 = lean_ctor_get(x_365, 0); +lean_inc(x_366); +if (lean_is_exclusive(x_365)) { + lean_ctor_release(x_365, 0); + lean_ctor_release(x_365, 1); + x_367 = x_365; +} else { + lean_dec_ref(x_365); + x_367 = lean_box(0); +} +x_368 = 4; +x_369 = lean_box(x_368); +if (lean_is_scalar(x_367)) { + x_370 = lean_alloc_ctor(0, 2, 0); +} else { + x_370 = x_367; +} +lean_ctor_set(x_370, 0, x_366); +lean_ctor_set(x_370, 1, x_369); +return x_370; +} +else +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; uint8_t x_375; +x_371 = lean_ctor_get(x_365, 0); +lean_inc(x_371); +x_372 = lean_ctor_get(x_365, 1); +lean_inc(x_372); +if (lean_is_exclusive(x_365)) { + lean_ctor_release(x_365, 0); + lean_ctor_release(x_365, 1); + x_373 = x_365; +} else { + lean_dec_ref(x_365); + x_373 = lean_box(0); +} +x_374 = lean_ctor_get(x_371, 1); +lean_inc(x_374); +x_375 = lean_nat_dec_eq(x_361, x_374); +lean_dec(x_374); +lean_dec(x_361); +if (x_375 == 0) +{ +lean_object* x_376; +if (lean_is_scalar(x_373)) { + x_376 = lean_alloc_ctor(1, 2, 0); +} else { + x_376 = x_373; +} +lean_ctor_set(x_376, 0, x_371); +lean_ctor_set(x_376, 1, x_372); +return x_376; +} +else +{ +lean_object* x_377; +lean_dec(x_373); +lean_dec(x_372); +x_377 = l_Std_Internal_Parsec_String_pstring(x_2, x_371); +if (lean_obj_tag(x_377) == 0) +{ +lean_object* x_378; lean_object* x_379; uint8_t x_380; lean_object* x_381; lean_object* x_382; +x_378 = lean_ctor_get(x_377, 0); +lean_inc(x_378); +if (lean_is_exclusive(x_377)) { + lean_ctor_release(x_377, 0); + lean_ctor_release(x_377, 1); + x_379 = x_377; +} else { + lean_dec_ref(x_377); + x_379 = lean_box(0); +} +x_380 = 5; +x_381 = lean_box(x_380); +if (lean_is_scalar(x_379)) { + x_382 = lean_alloc_ctor(0, 2, 0); +} else { + x_382 = x_379; +} +lean_ctor_set(x_382, 0, x_378); +lean_ctor_set(x_382, 1, x_381); +return x_382; +} +else +{ +lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; +x_383 = lean_ctor_get(x_377, 0); +lean_inc(x_383); +x_384 = lean_ctor_get(x_377, 1); +lean_inc(x_384); +if (lean_is_exclusive(x_377)) { + lean_ctor_release(x_377, 0); + lean_ctor_release(x_377, 1); + x_385 = x_377; +} else { + lean_dec_ref(x_377); + x_385 = lean_box(0); +} +if (lean_is_scalar(x_385)) { + x_386 = lean_alloc_ctor(1, 2, 0); +} else { + x_386 = x_385; +} +lean_ctor_set(x_386, 0, x_383); +lean_ctor_set(x_386, 1, x_384); +return x_386; +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseEraNarrow(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterNumber(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +return x_43; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_43); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_54 = lean_ctor_get(x_30, 0); +x_55 = lean_ctor_get(x_30, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_30); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +x_57 = lean_nat_dec_eq(x_27, x_56); +lean_dec(x_56); +lean_dec(x_27); +if (x_57 == 0) +{ +lean_object* x_58; +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_54); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_55); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +x_60 = l_Std_Internal_Parsec_String_pstring(x_59, x_54); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_62 = x_60; +} else { + lean_dec_ref(x_60); + x_62 = lean_box(0); +} +x_63 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; +} +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_60, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_67 = x_60; +} else { + lean_dec_ref(x_60); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +} +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_17, 0); +x_70 = lean_ctor_get(x_17, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_17); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +x_72 = lean_nat_dec_eq(x_14, x_71); +lean_dec(x_14); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_71); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_69); +lean_ctor_set(x_73, 1, x_70); +return x_73; +} +else +{ +lean_object* x_74; lean_object* x_75; +lean_dec(x_70); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2; +x_75 = l_Std_Internal_Parsec_String_pstring(x_74, x_69); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_71); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_77 = x_75; +} else { + lean_dec_ref(x_75); + x_77 = lean_box(0); +} +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_77; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_75, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_75, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_82 = x_75; +} else { + lean_dec_ref(x_75); + x_82 = lean_box(0); +} +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +x_84 = lean_nat_dec_eq(x_71, x_83); +lean_dec(x_83); +lean_dec(x_71); +if (x_84 == 0) +{ +lean_object* x_85; +if (lean_is_scalar(x_82)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_82; +} +lean_ctor_set(x_85, 0, x_80); +lean_ctor_set(x_85, 1, x_81); +return x_85; +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_82); +lean_dec(x_81); +x_86 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +x_87 = l_Std_Internal_Parsec_String_pstring(x_86, x_80); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_89 = x_87; +} else { + lean_dec_ref(x_87); + x_89 = lean_box(0); +} +x_90 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_89; +} +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_87, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_94 = x_87; +} else { + lean_dec_ref(x_87); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_96 = lean_ctor_get(x_3, 0); +x_97 = lean_ctor_get(x_3, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_3); +x_98 = lean_ctor_get(x_1, 1); +lean_inc(x_98); +lean_dec(x_1); +x_99 = lean_ctor_get(x_96, 1); +lean_inc(x_99); +x_100 = lean_nat_dec_eq(x_98, x_99); +lean_dec(x_98); +if (x_100 == 0) +{ +lean_object* x_101; +lean_dec(x_99); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_96); +lean_ctor_set(x_101, 1, x_97); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_97); +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3; +x_103 = l_Std_Internal_Parsec_String_pstring(x_102, x_96); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_99); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_106); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_108 = lean_ctor_get(x_103, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_103, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_110 = x_103; +} else { + lean_dec_ref(x_103); + x_110 = lean_box(0); +} +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +x_112 = lean_nat_dec_eq(x_99, x_111); +lean_dec(x_99); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_111); +if (lean_is_scalar(x_110)) { + x_113 = lean_alloc_ctor(1, 2, 0); +} else { + x_113 = x_110; +} +lean_ctor_set(x_113, 0, x_108); +lean_ctor_set(x_113, 1, x_109); +return x_113; +} +else +{ +lean_object* x_114; lean_object* x_115; +lean_dec(x_110); +lean_dec(x_109); +x_114 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2; +x_115 = l_Std_Internal_Parsec_String_pstring(x_114, x_108); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_111); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_117 = x_115; +} else { + lean_dec_ref(x_115); + x_117 = lean_box(0); +} +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_117)) { + x_119 = lean_alloc_ctor(0, 2, 0); +} else { + x_119 = x_117; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_120 = lean_ctor_get(x_115, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_115, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_122 = x_115; +} else { + lean_dec_ref(x_115); + x_122 = lean_box(0); +} +x_123 = lean_ctor_get(x_120, 1); +lean_inc(x_123); +x_124 = lean_nat_dec_eq(x_111, x_123); +lean_dec(x_123); +lean_dec(x_111); +if (x_124 == 0) +{ +lean_object* x_125; +if (lean_is_scalar(x_122)) { + x_125 = lean_alloc_ctor(1, 2, 0); +} else { + x_125 = x_122; +} +lean_ctor_set(x_125, 0, x_120); +lean_ctor_set(x_125, 1, x_121); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; +lean_dec(x_122); +lean_dec(x_121); +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1; +x_127 = l_Std_Internal_Parsec_String_pstring(x_126, x_120); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_129 = x_127; +} else { + lean_dec_ref(x_127); + x_129 = lean_box(0); +} +x_130 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_129)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_129; +} +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_132 = lean_ctor_get(x_127, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_134 = x_127; +} else { + lean_dec_ref(x_127); + x_134 = lean_box(0); +} +if (lean_is_scalar(x_134)) { + x_135 = lean_alloc_ctor(1, 2, 0); +} else { + x_135 = x_134; +} +lean_ctor_set(x_135, 0, x_132); +lean_ctor_set(x_135, 1, x_133); +return x_135; +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +return x_43; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_43); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_54 = lean_ctor_get(x_30, 0); +x_55 = lean_ctor_get(x_30, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_30); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +x_57 = lean_nat_dec_eq(x_27, x_56); +lean_dec(x_56); +lean_dec(x_27); +if (x_57 == 0) +{ +lean_object* x_58; +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_54); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_55); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +x_60 = l_Std_Internal_Parsec_String_pstring(x_59, x_54); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_62 = x_60; +} else { + lean_dec_ref(x_60); + x_62 = lean_box(0); +} +x_63 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; +} +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_60, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_67 = x_60; +} else { + lean_dec_ref(x_60); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +} +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_17, 0); +x_70 = lean_ctor_get(x_17, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_17); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +x_72 = lean_nat_dec_eq(x_14, x_71); +lean_dec(x_14); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_71); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_69); +lean_ctor_set(x_73, 1, x_70); +return x_73; +} +else +{ +lean_object* x_74; lean_object* x_75; +lean_dec(x_70); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2; +x_75 = l_Std_Internal_Parsec_String_pstring(x_74, x_69); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_71); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_77 = x_75; +} else { + lean_dec_ref(x_75); + x_77 = lean_box(0); +} +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_77; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_75, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_75, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_82 = x_75; +} else { + lean_dec_ref(x_75); + x_82 = lean_box(0); +} +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +x_84 = lean_nat_dec_eq(x_71, x_83); +lean_dec(x_83); +lean_dec(x_71); +if (x_84 == 0) +{ +lean_object* x_85; +if (lean_is_scalar(x_82)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_82; +} +lean_ctor_set(x_85, 0, x_80); +lean_ctor_set(x_85, 1, x_81); +return x_85; +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_82); +lean_dec(x_81); +x_86 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +x_87 = l_Std_Internal_Parsec_String_pstring(x_86, x_80); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_89 = x_87; +} else { + lean_dec_ref(x_87); + x_89 = lean_box(0); +} +x_90 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_89; +} +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_87, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_94 = x_87; +} else { + lean_dec_ref(x_87); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_96 = lean_ctor_get(x_3, 0); +x_97 = lean_ctor_get(x_3, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_3); +x_98 = lean_ctor_get(x_1, 1); +lean_inc(x_98); +lean_dec(x_1); +x_99 = lean_ctor_get(x_96, 1); +lean_inc(x_99); +x_100 = lean_nat_dec_eq(x_98, x_99); +lean_dec(x_98); +if (x_100 == 0) +{ +lean_object* x_101; +lean_dec(x_99); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_96); +lean_ctor_set(x_101, 1, x_97); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_97); +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3; +x_103 = l_Std_Internal_Parsec_String_pstring(x_102, x_96); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_99); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_106); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_108 = lean_ctor_get(x_103, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_103, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_110 = x_103; +} else { + lean_dec_ref(x_103); + x_110 = lean_box(0); +} +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +x_112 = lean_nat_dec_eq(x_99, x_111); +lean_dec(x_99); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_111); +if (lean_is_scalar(x_110)) { + x_113 = lean_alloc_ctor(1, 2, 0); +} else { + x_113 = x_110; +} +lean_ctor_set(x_113, 0, x_108); +lean_ctor_set(x_113, 1, x_109); +return x_113; +} +else +{ +lean_object* x_114; lean_object* x_115; +lean_dec(x_110); +lean_dec(x_109); +x_114 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2; +x_115 = l_Std_Internal_Parsec_String_pstring(x_114, x_108); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_111); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_117 = x_115; +} else { + lean_dec_ref(x_115); + x_117 = lean_box(0); +} +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_117)) { + x_119 = lean_alloc_ctor(0, 2, 0); +} else { + x_119 = x_117; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_120 = lean_ctor_get(x_115, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_115, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_122 = x_115; +} else { + lean_dec_ref(x_115); + x_122 = lean_box(0); +} +x_123 = lean_ctor_get(x_120, 1); +lean_inc(x_123); +x_124 = lean_nat_dec_eq(x_111, x_123); +lean_dec(x_123); +lean_dec(x_111); +if (x_124 == 0) +{ +lean_object* x_125; +if (lean_is_scalar(x_122)) { + x_125 = lean_alloc_ctor(1, 2, 0); +} else { + x_125 = x_122; +} +lean_ctor_set(x_125, 0, x_120); +lean_ctor_set(x_125, 1, x_121); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; +lean_dec(x_122); +lean_dec(x_121); +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1; +x_127 = l_Std_Internal_Parsec_String_pstring(x_126, x_120); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_129 = x_127; +} else { + lean_dec_ref(x_127); + x_129 = lean_box(0); +} +x_130 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_129)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_129; +} +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_132 = lean_ctor_get(x_127, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_134 = x_127; +} else { + lean_dec_ref(x_127); + x_134 = lean_box(0); +} +if (lean_is_scalar(x_134)) { + x_135 = lean_alloc_ctor(1, 2, 0); +} else { + x_135 = x_134; +} +lean_ctor_set(x_135, 0, x_132); +lean_ctor_set(x_135, 1, x_133); +return x_135; +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_3); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_14); +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_3); +lean_dec(x_12); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3; +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_11); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_14); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_dec(x_19); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +lean_ctor_set(x_17, 1, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_nat_dec_eq(x_14, x_27); +lean_dec(x_14); +if (x_28 == 0) +{ +lean_dec(x_27); +return x_17; +} +else +{ +lean_object* x_29; lean_object* x_30; +lean_free_object(x_17); +lean_dec(x_26); +x_29 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2; +x_30 = l_Std_Internal_Parsec_String_pstring(x_29, x_25); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +lean_dec(x_27); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +lean_dec(x_32); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +lean_inc(x_34); +lean_dec(x_30); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_30); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_30, 0); +x_39 = lean_ctor_get(x_30, 1); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +x_41 = lean_nat_dec_eq(x_27, x_40); +lean_dec(x_40); +lean_dec(x_27); +if (x_41 == 0) +{ +return x_30; +} +else +{ +lean_object* x_42; lean_object* x_43; +lean_free_object(x_30); +lean_dec(x_39); +x_42 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +x_43 = l_Std_Internal_Parsec_String_pstring(x_42, x_38); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 1); +lean_dec(x_45); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +lean_ctor_set(x_43, 1, x_46); +return x_43; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_43, 0); +lean_inc(x_47); +lean_dec(x_43); +x_48 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +else +{ +uint8_t x_50; +x_50 = !lean_is_exclusive(x_43); +if (x_50 == 0) +{ +return x_43; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_43, 0); +x_52 = lean_ctor_get(x_43, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_43); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_54 = lean_ctor_get(x_30, 0); +x_55 = lean_ctor_get(x_30, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_30); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +x_57 = lean_nat_dec_eq(x_27, x_56); +lean_dec(x_56); +lean_dec(x_27); +if (x_57 == 0) +{ +lean_object* x_58; +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_54); +lean_ctor_set(x_58, 1, x_55); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; +lean_dec(x_55); +x_59 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +x_60 = l_Std_Internal_Parsec_String_pstring(x_59, x_54); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_62 = x_60; +} else { + lean_dec_ref(x_60); + x_62 = lean_box(0); +} +x_63 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_62)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_62; +} +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_65 = lean_ctor_get(x_60, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_67 = x_60; +} else { + lean_dec_ref(x_60); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +} +} +} +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = lean_ctor_get(x_17, 0); +x_70 = lean_ctor_get(x_17, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_17); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +x_72 = lean_nat_dec_eq(x_14, x_71); +lean_dec(x_14); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_71); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_69); +lean_ctor_set(x_73, 1, x_70); +return x_73; +} +else +{ +lean_object* x_74; lean_object* x_75; +lean_dec(x_70); +x_74 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2; +x_75 = l_Std_Internal_Parsec_String_pstring(x_74, x_69); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_71); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_77 = x_75; +} else { + lean_dec_ref(x_75); + x_77 = lean_box(0); +} +x_78 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_77)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_77; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_80 = lean_ctor_get(x_75, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_75, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_82 = x_75; +} else { + lean_dec_ref(x_75); + x_82 = lean_box(0); +} +x_83 = lean_ctor_get(x_80, 1); +lean_inc(x_83); +x_84 = lean_nat_dec_eq(x_71, x_83); +lean_dec(x_83); +lean_dec(x_71); +if (x_84 == 0) +{ +lean_object* x_85; +if (lean_is_scalar(x_82)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_82; +} +lean_ctor_set(x_85, 0, x_80); +lean_ctor_set(x_85, 1, x_81); +return x_85; +} +else +{ +lean_object* x_86; lean_object* x_87; +lean_dec(x_82); +lean_dec(x_81); +x_86 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +x_87 = l_Std_Internal_Parsec_String_pstring(x_86, x_80); +if (lean_obj_tag(x_87) == 0) +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_89 = x_87; +} else { + lean_dec_ref(x_87); + x_89 = lean_box(0); +} +x_90 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 2, 0); +} else { + x_91 = x_89; +} +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_87, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_87, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_94 = x_87; +} else { + lean_dec_ref(x_87); + x_94 = lean_box(0); +} +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); +} else { + x_95 = x_94; +} +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; +x_96 = lean_ctor_get(x_3, 0); +x_97 = lean_ctor_get(x_3, 1); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_3); +x_98 = lean_ctor_get(x_1, 1); +lean_inc(x_98); +lean_dec(x_1); +x_99 = lean_ctor_get(x_96, 1); +lean_inc(x_99); +x_100 = lean_nat_dec_eq(x_98, x_99); +lean_dec(x_98); +if (x_100 == 0) +{ +lean_object* x_101; +lean_dec(x_99); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_96); +lean_ctor_set(x_101, 1, x_97); +return x_101; +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_97); +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3; +x_103 = l_Std_Internal_Parsec_String_pstring(x_102, x_96); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_99); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_105 = x_103; +} else { + lean_dec_ref(x_103); + x_105 = lean_box(0); +} +x_106 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3; +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); +} else { + x_107 = x_105; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_106); +return x_107; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_108 = lean_ctor_get(x_103, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_103, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_110 = x_103; +} else { + lean_dec_ref(x_103); + x_110 = lean_box(0); +} +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +x_112 = lean_nat_dec_eq(x_99, x_111); +lean_dec(x_99); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_111); +if (lean_is_scalar(x_110)) { + x_113 = lean_alloc_ctor(1, 2, 0); +} else { + x_113 = x_110; +} +lean_ctor_set(x_113, 0, x_108); +lean_ctor_set(x_113, 1, x_109); +return x_113; +} +else +{ +lean_object* x_114; lean_object* x_115; +lean_dec(x_110); +lean_dec(x_109); +x_114 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2; +x_115 = l_Std_Internal_Parsec_String_pstring(x_114, x_108); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +lean_dec(x_111); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_117 = x_115; +} else { + lean_dec_ref(x_115); + x_117 = lean_box(0); +} +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1; +if (lean_is_scalar(x_117)) { + x_119 = lean_alloc_ctor(0, 2, 0); +} else { + x_119 = x_117; +} +lean_ctor_set(x_119, 0, x_116); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_120 = lean_ctor_get(x_115, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_115, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_122 = x_115; +} else { + lean_dec_ref(x_115); + x_122 = lean_box(0); +} +x_123 = lean_ctor_get(x_120, 1); +lean_inc(x_123); +x_124 = lean_nat_dec_eq(x_111, x_123); +lean_dec(x_123); +lean_dec(x_111); +if (x_124 == 0) +{ +lean_object* x_125; +if (lean_is_scalar(x_122)) { + x_125 = lean_alloc_ctor(1, 2, 0); +} else { + x_125 = x_122; +} +lean_ctor_set(x_125, 0, x_120); +lean_ctor_set(x_125, 1, x_121); +return x_125; +} +else +{ +lean_object* x_126; lean_object* x_127; +lean_dec(x_122); +lean_dec(x_121); +x_126 = l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1; +x_127 = l_Std_Internal_Parsec_String_pstring(x_126, x_120); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_129 = x_127; +} else { + lean_dec_ref(x_127); + x_129 = lean_box(0); +} +x_130 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +if (lean_is_scalar(x_129)) { + x_131 = lean_alloc_ctor(0, 2, 0); +} else { + x_131 = x_129; +} +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_130); +return x_131; +} +else +{ +lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_132 = lean_ctor_get(x_127, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_127, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_134 = x_127; +} else { + lean_dec_ref(x_127); + x_134 = lean_box(0); +} +if (lean_is_scalar(x_134)) { + x_135 = lean_alloc_ctor(1, 2, 0); +} else { + x_135 = x_134; +} +lean_ctor_set(x_135, 0, x_132); +lean_ctor_set(x_135, 1, x_133); +return x_135; +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerShort(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerLong(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerNarrow(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5; +lean_inc(x_1); +x_3 = l_Std_Internal_Parsec_String_pstring(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +lean_dec(x_1); +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 1); +lean_dec(x_5); +x_6 = 0; +x_7 = lean_box(x_6); +lean_ctor_set(x_3, 1, x_7); +return x_3; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = 0; +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_3); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_nat_dec_eq(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (x_17 == 0) +{ +return x_3; +} +else +{ +lean_object* x_18; lean_object* x_19; +lean_free_object(x_3); +lean_dec(x_14); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1; +x_19 = l_Std_Internal_Parsec_String_pstring(x_18, x_13); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_19, 1); +lean_dec(x_21); +x_22 = 1; +x_23 = lean_box(x_22); +lean_ctor_set(x_19, 1, x_23); +return x_19; +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_24); +lean_dec(x_19); +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_19); +if (x_28 == 0) +{ +return x_19; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_19, 0); +x_30 = lean_ctor_get(x_19, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_19); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_32 = lean_ctor_get(x_3, 0); +x_33 = lean_ctor_get(x_3, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_3); +x_34 = lean_ctor_get(x_1, 1); +lean_inc(x_34); +lean_dec(x_1); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +x_36 = lean_nat_dec_eq(x_34, x_35); +lean_dec(x_35); +lean_dec(x_34); +if (x_36 == 0) +{ +lean_object* x_37; +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_32); +lean_ctor_set(x_37, 1, x_33); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1; +x_39 = l_Std_Internal_Parsec_String_pstring(x_38, x_32); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; lean_object* x_44; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_41 = x_39; +} else { + lean_dec_ref(x_39); + x_41 = lean_box(0); +} +x_42 = 1; +x_43 = lean_box(x_42); +if (lean_is_scalar(x_41)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_41; +} +lean_ctor_set(x_44, 0, x_40); +lean_ctor_set(x_44, 1, x_43); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_39, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_39, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_47 = x_39; +} else { + lean_dec_ref(x_39); + x_47 = lean_box(0); +} +if (lean_is_scalar(x_47)) { + x_48 = lean_alloc_ctor(1, 2, 0); +} else { + x_48 = x_47; +} +lean_ctor_set(x_48, 0, x_45); +lean_ctor_set(x_48, 1, x_46); +return x_48; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_2, x_4); +if (x_6 == 0) +{ +lean_object* x_7; +lean_inc(x_1); +x_7 = lean_apply_1(x_1, x_5); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_array_push(x_3, x_9); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = x_10; +x_4 = x_12; +x_5 = x_8; +goto _start; +} +else +{ +uint8_t x_14; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_7); +if (x_14 == 0) +{ +return x_7; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_7, 0); +x_16 = lean_ctor_get(x_7, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_7); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +else +{ +lean_object* x_18; +lean_dec(x_4); +lean_dec(x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_5); +lean_ctor_set(x_18, 1, x_3); +return x_18; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1; +x_5 = lean_unsigned_to_nat(12u); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactly_go___rarg(x_1, x_2, x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_exactly___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; +x_6 = lean_nat_dec_le(x_2, x_4); +if (x_6 == 0) +{ +lean_object* x_7; +lean_inc(x_1); +x_7 = lean_apply_1(x_1, x_5); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; uint32_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_unbox_uint32(x_9); +lean_dec(x_9); +x_11 = lean_string_push(x_3, x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_4, x_12); +lean_dec(x_4); +x_3 = x_11; +x_4 = x_13; +x_5 = x_8; +goto _start; +} +else +{ +uint8_t x_15; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_7); +if (x_15 == 0) +{ +return x_7; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 0); +x_17 = lean_ctor_get(x_7, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_7); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +else +{ +lean_object* x_19; +lean_dec(x_4); +lean_dec(x_1); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_5); +lean_ctor_set(x_19, 1, x_3); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(x_1, x_2, x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseSigned(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1; +lean_inc(x_2); +x_4 = l_Std_Internal_Parsec_String_pstring(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_apply_1(x_1, x_5); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_6, 1); +x_9 = lean_nat_to_int(x_8); +x_10 = lean_int_neg(x_9); +lean_dec(x_9); +lean_ctor_set(x_6, 1, x_10); +return x_6; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_6, 0); +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_6); +x_13 = lean_nat_to_int(x_12); +x_14 = lean_int_neg(x_13); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_11); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_6); +if (x_16 == 0) +{ +return x_6; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_6); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_4); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +x_23 = lean_ctor_get(x_2, 1); +lean_inc(x_23); +lean_dec(x_2); +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +x_25 = lean_nat_dec_eq(x_23, x_24); +lean_dec(x_24); +lean_dec(x_23); +if (x_25 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_object* x_26; +lean_free_object(x_4); +lean_dec(x_22); +x_26 = lean_apply_1(x_1, x_21); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 1); +x_29 = lean_nat_to_int(x_28); +lean_ctor_set(x_26, 1, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_26, 0); +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_26); +x_32 = lean_nat_to_int(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_30); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_26); +if (x_34 == 0) +{ +return x_26; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_26, 0); +x_36 = lean_ctor_get(x_26, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_26); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_38 = lean_ctor_get(x_4, 0); +x_39 = lean_ctor_get(x_4, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_4); +x_40 = lean_ctor_get(x_2, 1); +lean_inc(x_40); +lean_dec(x_2); +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +x_42 = lean_nat_dec_eq(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_1); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_38); +lean_ctor_set(x_43, 1, x_39); +return x_43; +} +else +{ +lean_object* x_44; +lean_dec(x_39); +x_44 = lean_apply_1(x_1, x_38); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_47 = x_44; +} else { + lean_dec_ref(x_44); + x_47 = lean_box(0); +} +x_48 = lean_nat_to_int(x_46); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_47; +} +lean_ctor_set(x_49, 0, x_45); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_44, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_44, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_52 = x_44; +} else { + lean_dec_ref(x_44); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(1, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +} +} +} +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Char_isDigit___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1; +x_2 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(x_3, x_1, x_4, x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = l_String_toNat_x21(x_8); +lean_dec(x_8); +lean_ctor_set(x_6, 1, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +x_12 = l_String_toNat_x21(x_11); +lean_dec(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_10); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +else +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_6); +if (x_14 == 0) +{ +return x_6; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_6, 0); +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_6); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1; +lean_inc(x_2); +x_4 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; +lean_dec(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_unbox_uint32(x_6); +lean_dec(x_6); +x_8 = lean_string_push(x_1, x_7); +x_1 = x_8; +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_4, 0); +x_12 = lean_ctor_get(x_4, 1); +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +lean_dec(x_2); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_dec(x_12); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 1, x_1); +return x_4; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_4, 0); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_4); +x_18 = lean_ctor_get(x_2, 1); +lean_inc(x_18); +lean_dec(x_2); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +x_20 = lean_nat_dec_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_1); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_17); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_17); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_1); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(x_3, x_1, x_4, x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___spec__1(x_4, x_7); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 1); +x_12 = lean_string_append(x_8, x_11); +lean_dec(x_11); +x_13 = l_String_toNat_x21(x_12); +lean_dec(x_12); +lean_ctor_set(x_9, 1, x_13); +return x_9; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_9, 0); +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_9); +x_16 = lean_string_append(x_8, x_15); +lean_dec(x_15); +x_17 = l_String_toNat_x21(x_16); +lean_dec(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_14); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_8); +x_19 = !lean_is_exclusive(x_9); +if (x_19 == 0) +{ +return x_9; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_9, 0); +x_21 = lean_ctor_get(x_9, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_9); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_6); +if (x_23 == 0) +{ +return x_6; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_6, 0); +x_25 = lean_ctor_get(x_6, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_6); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = 48; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_exactlyChars_go(x_5, x_1, x_6, x_7, x_3); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 1); +x_11 = l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(x_2, x_4, x_10); +x_12 = l_String_toNat_x21(x_11); +lean_dec(x_11); +lean_ctor_set(x_8, 1, x_12); +return x_8; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_8, 0); +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_8); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_rightPad(x_2, x_4, x_14); +x_16 = l_String_toNat_x21(x_15); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_8); +if (x_18 == 0) +{ +return x_8; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_8, 0); +x_20 = lean_ctor_get(x_8, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_8); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1(uint32_t x_1) { +_start: +{ +lean_object* x_2; uint32_t x_25; uint8_t x_26; +x_25 = 65; +x_26 = lean_uint32_dec_le(x_25, x_1); +if (x_26 == 0) +{ +uint32_t x_27; uint8_t x_28; +x_27 = 97; +x_28 = lean_uint32_dec_le(x_27, x_1); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_box(0); +x_2 = x_29; +goto block_24; +} +else +{ +uint32_t x_30; uint8_t x_31; +x_30 = 122; +x_31 = lean_uint32_dec_le(x_1, x_30); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_box(0); +x_2 = x_32; +goto block_24; +} +else +{ +uint8_t x_33; +x_33 = 1; +return x_33; +} +} +} +else +{ +uint32_t x_34; uint8_t x_35; +x_34 = 90; +x_35 = lean_uint32_dec_le(x_1, x_34); +if (x_35 == 0) +{ +uint32_t x_36; uint8_t x_37; +x_36 = 97; +x_37 = lean_uint32_dec_le(x_36, x_1); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = lean_box(0); +x_2 = x_38; +goto block_24; +} +else +{ +uint32_t x_39; uint8_t x_40; +x_39 = 122; +x_40 = lean_uint32_dec_le(x_1, x_39); +if (x_40 == 0) +{ +lean_object* x_41; +x_41 = lean_box(0); +x_2 = x_41; +goto block_24; +} +else +{ +uint8_t x_42; +x_42 = 1; +return x_42; +} +} +} +else +{ +uint8_t x_43; +x_43 = 1; +return x_43; +} +} +block_24: +{ +uint32_t x_3; uint8_t x_4; +lean_dec(x_2); +x_3 = 48; +x_4 = lean_uint32_dec_le(x_3, x_1); +if (x_4 == 0) +{ +uint32_t x_5; uint8_t x_6; +x_5 = 95; +x_6 = lean_uint32_dec_eq(x_1, x_5); +if (x_6 == 0) +{ +uint32_t x_7; uint8_t x_8; +x_7 = 45; +x_8 = lean_uint32_dec_eq(x_1, x_7); +if (x_8 == 0) +{ +uint32_t x_9; uint8_t x_10; +x_9 = 47; +x_10 = lean_uint32_dec_eq(x_1, x_9); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = 1; +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = 1; +return x_12; +} +} +else +{ +uint32_t x_13; uint8_t x_14; +x_13 = 57; +x_14 = lean_uint32_dec_le(x_1, x_13); +if (x_14 == 0) +{ +uint32_t x_15; uint8_t x_16; +x_15 = 95; +x_16 = lean_uint32_dec_eq(x_1, x_15); +if (x_16 == 0) +{ +uint32_t x_17; uint8_t x_18; +x_17 = 45; +x_18 = lean_uint32_dec_eq(x_1, x_17); +if (x_18 == 0) +{ +uint32_t x_19; uint8_t x_20; +x_19 = 47; +x_20 = lean_uint32_dec_eq(x_1, x_19); +return x_20; +} +else +{ +uint8_t x_21; +x_21 = 1; +return x_21; +} +} +else +{ +uint8_t x_22; +x_22 = 1; +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = 1; +return x_23; +} +} +} +} +} +static lean_object* _init_l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1; +lean_inc(x_2); +x_4 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; +lean_dec(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_unbox_uint32(x_6); +lean_dec(x_6); +x_8 = lean_string_push(x_1, x_7); +x_1 = x_8; +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_4, 0); +x_12 = lean_ctor_get(x_4, 1); +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +lean_dec(x_2); +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (x_15 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_dec(x_12); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 1, x_1); +return x_4; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_4, 0); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_4); +x_18 = lean_ctor_get(x_2, 1); +lean_inc(x_18); +lean_dec(x_2); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +x_20 = lean_nat_dec_eq(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_1); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_17); +return x_21; +} +else +{ +lean_object* x_22; +lean_dec(x_17); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_1); +return x_22; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_string_utf8_byte_size(x_14); +x_17 = lean_nat_dec_lt(x_15, x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_15); +lean_dec(x_14); +x_18 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_18); +x_2 = x_19; +goto block_13; +} +else +{ +uint32_t x_20; uint32_t x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_35; +x_20 = 65; +x_21 = lean_string_utf8_get_fast(x_14, x_15); +x_22 = lean_uint32_dec_le(x_20, x_21); +x_23 = lean_string_utf8_next_fast(x_14, x_15); +lean_dec(x_15); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_14); +lean_ctor_set(x_24, 1, x_23); +if (x_22 == 0) +{ +uint32_t x_62; uint8_t x_63; +x_62 = 97; +x_63 = lean_uint32_dec_le(x_62, x_21); +if (x_63 == 0) +{ +lean_object* x_64; +x_64 = lean_box(0); +x_35 = x_64; +goto block_61; +} +else +{ +uint32_t x_65; uint8_t x_66; +x_65 = 122; +x_66 = lean_uint32_dec_le(x_21, x_65); +if (x_66 == 0) +{ +lean_object* x_67; +x_67 = lean_box(0); +x_35 = x_67; +goto block_61; +} +else +{ +uint8_t x_68; +x_68 = 1; +x_25 = x_68; +goto block_34; +} +} +} +else +{ +uint32_t x_69; uint8_t x_70; +x_69 = 90; +x_70 = lean_uint32_dec_le(x_21, x_69); +if (x_70 == 0) +{ +uint32_t x_71; uint8_t x_72; +x_71 = 97; +x_72 = lean_uint32_dec_le(x_71, x_21); +if (x_72 == 0) +{ +lean_object* x_73; +x_73 = lean_box(0); +x_35 = x_73; +goto block_61; +} +else +{ +uint32_t x_74; uint8_t x_75; +x_74 = 122; +x_75 = lean_uint32_dec_le(x_21, x_74); +if (x_75 == 0) +{ +lean_object* x_76; +x_76 = lean_box(0); +x_35 = x_76; +goto block_61; +} +else +{ +uint8_t x_77; +x_77 = 1; +x_25 = x_77; +goto block_34; +} +} +} +else +{ +uint8_t x_78; +x_78 = 1; +x_25 = x_78; +goto block_34; +} +} +block_34: +{ +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); +x_26 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1; +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_26); +x_2 = x_27; +goto block_13; +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_1); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_1, 1); +lean_dec(x_29); +x_30 = lean_ctor_get(x_1, 0); +lean_dec(x_30); +x_31 = lean_box_uint32(x_21); +lean_ctor_set(x_1, 1, x_31); +lean_ctor_set(x_1, 0, x_24); +x_2 = x_1; +goto block_13; +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_1); +x_32 = lean_box_uint32(x_21); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_24); +lean_ctor_set(x_33, 1, x_32); +x_2 = x_33; +goto block_13; +} +} +} +block_61: +{ +uint32_t x_36; uint8_t x_37; +lean_dec(x_35); +x_36 = 48; +x_37 = lean_uint32_dec_le(x_36, x_21); +if (x_37 == 0) +{ +uint32_t x_38; uint8_t x_39; +x_38 = 95; +x_39 = lean_uint32_dec_eq(x_21, x_38); +if (x_39 == 0) +{ +uint32_t x_40; uint8_t x_41; +x_40 = 45; +x_41 = lean_uint32_dec_eq(x_21, x_40); +if (x_41 == 0) +{ +uint32_t x_42; uint8_t x_43; +x_42 = 47; +x_43 = lean_uint32_dec_eq(x_21, x_42); +if (x_43 == 0) +{ +uint8_t x_44; +x_44 = 0; +x_25 = x_44; +goto block_34; +} +else +{ +uint8_t x_45; +x_45 = 1; +x_25 = x_45; +goto block_34; +} +} +else +{ +uint8_t x_46; +x_46 = 1; +x_25 = x_46; +goto block_34; +} +} +else +{ +uint8_t x_47; +x_47 = 1; +x_25 = x_47; +goto block_34; +} +} +else +{ +uint32_t x_48; uint8_t x_49; +x_48 = 57; +x_49 = lean_uint32_dec_le(x_21, x_48); +if (x_49 == 0) +{ +uint32_t x_50; uint8_t x_51; +x_50 = 95; +x_51 = lean_uint32_dec_eq(x_21, x_50); +if (x_51 == 0) +{ +uint32_t x_52; uint8_t x_53; +x_52 = 45; +x_53 = lean_uint32_dec_eq(x_21, x_52); +if (x_53 == 0) +{ +uint32_t x_54; uint8_t x_55; +x_54 = 47; +x_55 = lean_uint32_dec_eq(x_21, x_54); +if (x_55 == 0) +{ +uint8_t x_56; +x_56 = 0; +x_25 = x_56; +goto block_34; +} +else +{ +uint8_t x_57; +x_57 = 1; +x_25 = x_57; +goto block_34; +} +} +else +{ +uint8_t x_58; +x_58 = 1; +x_25 = x_58; +goto block_34; +} +} +else +{ +uint8_t x_59; +x_59 = 1; +x_25 = x_59; +goto block_34; +} +} +else +{ +uint8_t x_60; +x_60 = 1; +x_25 = x_60; +goto block_34; +} +} +} +} +block_13: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint32_t x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_6 = lean_unbox_uint32(x_4); +lean_dec(x_4); +x_7 = lean_string_push(x_5, x_6); +x_8 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1(x_7, x_3); +return x_8; +} +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_2); +if (x_9 == 0) +{ +return x_2; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_2, 0); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_2); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint32_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_3 = l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___lambda__1(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("need a natural number in the interval of ", 41, 41); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" to ", 4, 4); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_apply_1(x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = lean_nat_dec_le(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_7); +x_9 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_10 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1; +x_11 = lean_string_append(x_10, x_9); +lean_dec(x_9); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2; +x_13 = lean_string_append(x_11, x_12); +x_14 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); +x_15 = lean_string_append(x_13, x_14); +lean_dec(x_14); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_17 = lean_string_append(x_15, x_16); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_17); +return x_5; +} +else +{ +uint8_t x_18; +x_18 = lean_nat_dec_le(x_7, x_2); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_7); +x_19 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1; +x_21 = lean_string_append(x_20, x_19); +lean_dec(x_19); +x_22 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2; +x_23 = lean_string_append(x_21, x_22); +x_24 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); +x_25 = lean_string_append(x_23, x_24); +lean_dec(x_24); +x_26 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_27 = lean_string_append(x_25, x_26); +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 1, x_27); +return x_5; +} +else +{ +lean_object* x_28; +lean_dec(x_2); +lean_dec(x_1); +x_28 = lean_nat_to_int(x_7); +lean_ctor_set(x_5, 1, x_28); +return x_5; +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_5, 0); +x_30 = lean_ctor_get(x_5, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_5); +x_31 = lean_nat_dec_le(x_1, x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_30); +x_32 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_33 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1; +x_34 = lean_string_append(x_33, x_32); +lean_dec(x_32); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2; +x_36 = lean_string_append(x_34, x_35); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); +x_38 = lean_string_append(x_36, x_37); +lean_dec(x_37); +x_39 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_40 = lean_string_append(x_38, x_39); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_29); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +else +{ +uint8_t x_42; +x_42 = lean_nat_dec_le(x_30, x_2); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_30); +x_43 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_44 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1; +x_45 = lean_string_append(x_44, x_43); +lean_dec(x_43); +x_46 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2; +x_47 = lean_string_append(x_45, x_46); +x_48 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); +x_49 = lean_string_append(x_47, x_48); +lean_dec(x_48); +x_50 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_51 = lean_string_append(x_49, x_50); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_29); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +else +{ +lean_object* x_53; lean_object* x_54; +lean_dec(x_2); +lean_dec(x_1); +x_53 = lean_nat_to_int(x_30); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_29); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +} +} +else +{ +uint8_t x_55; +lean_dec(x_2); +lean_dec(x_1); +x_55 = !lean_is_exclusive(x_5); +if (x_55 == 0) +{ +return x_5; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_5, 0); +x_57 = lean_ctor_get(x_5, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_5); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_Reason_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Text_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_Reason_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2() { +_start: +{ +uint32_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 58; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = lean_string_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5() { +_start: +{ +uint32_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 45; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = lean_string_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9() { +_start: +{ +uint32_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 43; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = lean_string_push(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(uint8_t x_1, uint8_t x_2, uint8_t x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_182; lean_object* x_183; uint32_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; +x_203 = 43; +x_204 = lean_ctor_get(x_4, 0); +lean_inc(x_204); +x_205 = lean_ctor_get(x_4, 1); +lean_inc(x_205); +x_206 = lean_string_utf8_byte_size(x_204); +x_207 = lean_nat_dec_lt(x_205, x_206); +lean_dec(x_206); +if (x_207 == 0) +{ +lean_object* x_208; +lean_dec(x_205); +lean_dec(x_204); +x_208 = l_Std_Internal_Parsec_unexpectedEndOfInput; +lean_inc(x_4); +x_182 = x_4; +x_183 = x_208; +goto block_202; +} +else +{ +uint32_t x_209; lean_object* x_210; lean_object* x_211; uint8_t x_212; +x_209 = lean_string_utf8_get_fast(x_204, x_205); +x_210 = lean_string_utf8_next_fast(x_204, x_205); +lean_dec(x_205); +x_211 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_211, 0, x_204); +lean_ctor_set(x_211, 1, x_210); +x_212 = lean_uint32_dec_eq(x_209, x_203); +if (x_212 == 0) +{ +lean_object* x_213; +lean_dec(x_211); +x_213 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11; +lean_inc(x_4); +x_182 = x_4; +x_183 = x_213; +goto block_202; +} +else +{ +lean_object* x_214; +lean_dec(x_4); +x_214 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_5 = x_211; +x_6 = x_214; +goto block_181; +} +} +block_181: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_unsigned_to_nat(2u); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + x_11 = x_8; +} else { + lean_dec_ref(x_8); + x_11 = lean_box(0); +} +x_12 = lean_nat_to_int(x_10); +if (x_3 == 0) +{ +uint8_t x_175; +x_175 = 0; +x_13 = x_175; +goto block_174; +} +else +{ +uint8_t x_176; +x_176 = 1; +x_13 = x_176; +goto block_174; +} +block_174: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_109; lean_object* x_120; lean_object* x_141; lean_object* x_142; +switch (x_1) { +case 0: +{ +if (x_13 == 0) +{ +x_109 = x_9; +goto block_119; +} +else +{ +uint32_t x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; +x_149 = 58; +x_150 = lean_ctor_get(x_9, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_9, 1); +lean_inc(x_151); +x_152 = lean_string_utf8_byte_size(x_150); +x_153 = lean_nat_dec_lt(x_151, x_152); +lean_dec(x_152); +if (x_153 == 0) +{ +lean_object* x_154; lean_object* x_155; +lean_dec(x_151); +lean_dec(x_150); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_154 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_155 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_155, 0, x_9); +lean_ctor_set(x_155, 1, x_154); +return x_155; +} +else +{ +uint32_t x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; +x_156 = lean_string_utf8_get_fast(x_150, x_151); +x_157 = lean_string_utf8_next_fast(x_150, x_151); +lean_dec(x_151); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_150); +lean_ctor_set(x_158, 1, x_157); +x_159 = lean_uint32_dec_eq(x_156, x_149); +if (x_159 == 0) +{ +lean_object* x_160; lean_object* x_161; +lean_dec(x_158); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_160 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4; +x_161 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_161, 0, x_9); +lean_ctor_set(x_161, 1, x_160); +return x_161; +} +else +{ +lean_dec(x_9); +x_109 = x_158; +goto block_119; +} +} +} +} +case 1: +{ +lean_object* x_162; +x_162 = lean_box(0); +x_14 = x_9; +x_15 = x_162; +goto block_108; +} +default: +{ +if (x_13 == 0) +{ +lean_inc(x_9); +x_120 = x_9; +goto block_140; +} +else +{ +uint32_t x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; +x_163 = 58; +x_164 = lean_ctor_get(x_9, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_9, 1); +lean_inc(x_165); +x_166 = lean_string_utf8_byte_size(x_164); +x_167 = lean_nat_dec_lt(x_165, x_166); +lean_dec(x_166); +if (x_167 == 0) +{ +lean_object* x_168; +lean_dec(x_165); +lean_dec(x_164); +x_168 = l_Std_Internal_Parsec_unexpectedEndOfInput; +lean_inc(x_9); +x_141 = x_9; +x_142 = x_168; +goto block_148; +} +else +{ +uint32_t x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; +x_169 = lean_string_utf8_get_fast(x_164, x_165); +x_170 = lean_string_utf8_next_fast(x_164, x_165); +lean_dec(x_165); +x_171 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_171, 0, x_164); +lean_ctor_set(x_171, 1, x_170); +x_172 = lean_uint32_dec_eq(x_169, x_163); +if (x_172 == 0) +{ +lean_object* x_173; +lean_dec(x_171); +x_173 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4; +lean_inc(x_9); +x_141 = x_9; +x_142 = x_173; +goto block_148; +} +else +{ +x_120 = x_171; +goto block_140; +} +} +} +} +} +block_108: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_43; lean_object* x_54; lean_object* x_75; lean_object* x_76; +switch (x_2) { +case 0: +{ +if (x_13 == 0) +{ +x_43 = x_14; +goto block_53; +} +else +{ +uint32_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_83 = 58; +x_84 = lean_ctor_get(x_14, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_14, 1); +lean_inc(x_85); +x_86 = lean_string_utf8_byte_size(x_84); +x_87 = lean_nat_dec_lt(x_85, x_86); +lean_dec(x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +lean_dec(x_85); +lean_dec(x_84); +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_88 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_14); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +else +{ +uint32_t x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; +x_90 = lean_string_utf8_get_fast(x_84, x_85); +x_91 = lean_string_utf8_next_fast(x_84, x_85); +lean_dec(x_85); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_uint32_dec_eq(x_90, x_83); +if (x_93 == 0) +{ +lean_object* x_94; lean_object* x_95; +lean_dec(x_92); +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_94 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4; +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_14); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +else +{ +lean_dec(x_14); +x_43 = x_92; +goto block_53; +} +} +} +} +case 1: +{ +lean_object* x_96; +x_96 = lean_box(0); +x_16 = x_14; +x_17 = x_96; +goto block_42; +} +default: +{ +if (x_13 == 0) +{ +lean_inc(x_14); +x_54 = x_14; +goto block_74; +} +else +{ +uint32_t x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_97 = 58; +x_98 = lean_ctor_get(x_14, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_14, 1); +lean_inc(x_99); +x_100 = lean_string_utf8_byte_size(x_98); +x_101 = lean_nat_dec_lt(x_99, x_100); +lean_dec(x_100); +if (x_101 == 0) +{ +lean_object* x_102; +lean_dec(x_99); +lean_dec(x_98); +x_102 = l_Std_Internal_Parsec_unexpectedEndOfInput; +lean_inc(x_14); +x_75 = x_14; +x_76 = x_102; +goto block_82; +} +else +{ +uint32_t x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_103 = lean_string_utf8_get_fast(x_98, x_99); +x_104 = lean_string_utf8_next_fast(x_98, x_99); +lean_dec(x_99); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_98); +lean_ctor_set(x_105, 1, x_104); +x_106 = lean_uint32_dec_eq(x_103, x_97); +if (x_106 == 0) +{ +lean_object* x_107; +lean_dec(x_105); +x_107 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4; +lean_inc(x_14); +x_75 = x_14; +x_76 = x_107; +goto block_82; +} +else +{ +x_54 = x_105; +goto block_74; +} +} +} +} +} +block_42: +{ +lean_object* x_18; lean_object* x_19; +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1; +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_23 = lean_int_add(x_21, x_22); +lean_dec(x_21); +x_24 = lean_int_mul(x_23, x_6); +lean_dec(x_6); +lean_dec(x_23); +if (lean_is_scalar(x_11)) { + x_25 = lean_alloc_ctor(0, 2, 0); +} else { + x_25 = x_11; +} +lean_ctor_set(x_25, 0, x_16); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_17, 0); +lean_inc(x_26); +lean_dec(x_17); +x_27 = lean_int_add(x_21, x_26); +lean_dec(x_26); +lean_dec(x_21); +x_28 = lean_int_mul(x_27, x_6); +lean_dec(x_6); +lean_dec(x_27); +if (lean_is_scalar(x_11)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_11; +} +lean_ctor_set(x_29, 0, x_16); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_15, 0); +lean_inc(x_30); +lean_dec(x_15); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6; +x_32 = lean_int_mul(x_30, x_31); +lean_dec(x_30); +x_33 = lean_int_add(x_19, x_32); +lean_dec(x_32); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_35 = lean_int_add(x_33, x_34); +lean_dec(x_33); +x_36 = lean_int_mul(x_35, x_6); +lean_dec(x_6); +lean_dec(x_35); +if (lean_is_scalar(x_11)) { + x_37 = lean_alloc_ctor(0, 2, 0); +} else { + x_37 = x_11; +} +lean_ctor_set(x_37, 0, x_16); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_17, 0); +lean_inc(x_38); +lean_dec(x_17); +x_39 = lean_int_add(x_33, x_38); +lean_dec(x_38); +lean_dec(x_33); +x_40 = lean_int_mul(x_39, x_6); +lean_dec(x_6); +lean_dec(x_39); +if (lean_is_scalar(x_11)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_11; +} +lean_ctor_set(x_41, 0, x_16); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +block_53: +{ +lean_object* x_44; +x_44 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_7, x_43); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_nat_to_int(x_46); +x_48 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_16 = x_45; +x_17 = x_48; +goto block_42; +} +else +{ +uint8_t x_49; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_49 = !lean_is_exclusive(x_44); +if (x_49 == 0) +{ +return x_44; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_44, 0); +x_51 = lean_ctor_get(x_44, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_44); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +block_74: +{ +lean_object* x_55; +x_55 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_7, x_54); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_14); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); +x_58 = lean_nat_to_int(x_57); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_16 = x_56; +x_17 = x_59; +goto block_42; +} +else +{ +uint8_t x_60; +x_60 = !lean_is_exclusive(x_55); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_61 = lean_ctor_get(x_55, 0); +x_62 = lean_ctor_get(x_55, 1); +x_63 = lean_ctor_get(x_14, 1); +lean_inc(x_63); +lean_dec(x_14); +x_64 = lean_ctor_get(x_61, 1); +lean_inc(x_64); +x_65 = lean_nat_dec_eq(x_63, x_64); +lean_dec(x_64); +lean_dec(x_63); +if (x_65 == 0) +{ +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +return x_55; +} +else +{ +lean_object* x_66; +lean_free_object(x_55); +lean_dec(x_62); +x_66 = lean_box(0); +x_16 = x_61; +x_17 = x_66; +goto block_42; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_67 = lean_ctor_get(x_55, 0); +x_68 = lean_ctor_get(x_55, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_55); +x_69 = lean_ctor_get(x_14, 1); +lean_inc(x_69); +lean_dec(x_14); +x_70 = lean_ctor_get(x_67, 1); +lean_inc(x_70); +x_71 = lean_nat_dec_eq(x_69, x_70); +lean_dec(x_70); +lean_dec(x_69); +if (x_71 == 0) +{ +lean_object* x_72; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_67); +lean_ctor_set(x_72, 1, x_68); +return x_72; +} +else +{ +lean_object* x_73; +lean_dec(x_68); +x_73 = lean_box(0); +x_16 = x_67; +x_17 = x_73; +goto block_42; +} +} +} +} +block_82: +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_77 = lean_ctor_get(x_14, 1); +lean_inc(x_77); +lean_dec(x_14); +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +x_79 = lean_nat_dec_eq(x_77, x_78); +lean_dec(x_78); +lean_dec(x_77); +if (x_79 == 0) +{ +lean_object* x_80; +lean_dec(x_15); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_75); +lean_ctor_set(x_80, 1, x_76); +return x_80; +} +else +{ +lean_object* x_81; +lean_dec(x_76); +x_81 = lean_box(0); +x_16 = x_75; +x_17 = x_81; +goto block_42; +} +} +} +block_119: +{ +lean_object* x_110; +x_110 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_7, x_109); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_110, 1); +lean_inc(x_112); +lean_dec(x_110); +x_113 = lean_nat_to_int(x_112); +x_114 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_114, 0, x_113); +x_14 = x_111; +x_15 = x_114; +goto block_108; +} +else +{ +uint8_t x_115; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_115 = !lean_is_exclusive(x_110); +if (x_115 == 0) +{ +return x_110; +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_110, 0); +x_117 = lean_ctor_get(x_110, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_110); +x_118 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +return x_118; +} +} +} +block_140: +{ +lean_object* x_121; +x_121 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_7, x_120); +if (lean_obj_tag(x_121) == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_9); +x_122 = lean_ctor_get(x_121, 0); +lean_inc(x_122); +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_124 = lean_nat_to_int(x_123); +x_125 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_125, 0, x_124); +x_14 = x_122; +x_15 = x_125; +goto block_108; +} +else +{ +uint8_t x_126; +x_126 = !lean_is_exclusive(x_121); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_127 = lean_ctor_get(x_121, 0); +x_128 = lean_ctor_get(x_121, 1); +x_129 = lean_ctor_get(x_9, 1); +lean_inc(x_129); +lean_dec(x_9); +x_130 = lean_ctor_get(x_127, 1); +lean_inc(x_130); +x_131 = lean_nat_dec_eq(x_129, x_130); +lean_dec(x_130); +lean_dec(x_129); +if (x_131 == 0) +{ +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +return x_121; +} +else +{ +lean_object* x_132; +lean_free_object(x_121); +lean_dec(x_128); +x_132 = lean_box(0); +x_14 = x_127; +x_15 = x_132; +goto block_108; +} +} +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_133 = lean_ctor_get(x_121, 0); +x_134 = lean_ctor_get(x_121, 1); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_121); +x_135 = lean_ctor_get(x_9, 1); +lean_inc(x_135); +lean_dec(x_9); +x_136 = lean_ctor_get(x_133, 1); +lean_inc(x_136); +x_137 = lean_nat_dec_eq(x_135, x_136); +lean_dec(x_136); +lean_dec(x_135); +if (x_137 == 0) +{ +lean_object* x_138; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_138 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_138, 0, x_133); +lean_ctor_set(x_138, 1, x_134); +return x_138; +} +else +{ +lean_object* x_139; +lean_dec(x_134); +x_139 = lean_box(0); +x_14 = x_133; +x_15 = x_139; +goto block_108; +} +} +} +} +block_148: +{ +lean_object* x_143; lean_object* x_144; uint8_t x_145; +x_143 = lean_ctor_get(x_9, 1); +lean_inc(x_143); +lean_dec(x_9); +x_144 = lean_ctor_get(x_141, 1); +lean_inc(x_144); +x_145 = lean_nat_dec_eq(x_143, x_144); +lean_dec(x_144); +lean_dec(x_143); +if (x_145 == 0) +{ +lean_object* x_146; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_6); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_141); +lean_ctor_set(x_146, 1, x_142); +return x_146; +} +else +{ +lean_object* x_147; +lean_dec(x_142); +x_147 = lean_box(0); +x_14 = x_141; +x_15 = x_147; +goto block_108; +} +} +} +} +else +{ +uint8_t x_177; +lean_dec(x_6); +x_177 = !lean_is_exclusive(x_8); +if (x_177 == 0) +{ +return x_8; +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_8, 0); +x_179 = lean_ctor_get(x_8, 1); +lean_inc(x_179); +lean_inc(x_178); +lean_dec(x_8); +x_180 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set(x_180, 1, x_179); +return x_180; +} +} +} +block_202: +{ +lean_object* x_184; lean_object* x_185; uint8_t x_186; +x_184 = lean_ctor_get(x_4, 1); +lean_inc(x_184); +lean_dec(x_4); +x_185 = lean_ctor_get(x_182, 1); +lean_inc(x_185); +x_186 = lean_nat_dec_eq(x_184, x_185); +lean_dec(x_185); +lean_dec(x_184); +if (x_186 == 0) +{ +lean_object* x_187; +x_187 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_187, 0, x_182); +lean_ctor_set(x_187, 1, x_183); +return x_187; +} +else +{ +uint32_t x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; +lean_dec(x_183); +x_188 = 45; +x_189 = lean_ctor_get(x_182, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_182, 1); +lean_inc(x_190); +x_191 = lean_string_utf8_byte_size(x_189); +x_192 = lean_nat_dec_lt(x_190, x_191); +lean_dec(x_191); +if (x_192 == 0) +{ +lean_object* x_193; lean_object* x_194; +lean_dec(x_190); +lean_dec(x_189); +x_193 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_194 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_194, 0, x_182); +lean_ctor_set(x_194, 1, x_193); +return x_194; +} +else +{ +uint32_t x_195; lean_object* x_196; lean_object* x_197; uint8_t x_198; +x_195 = lean_string_utf8_get_fast(x_189, x_190); +x_196 = lean_string_utf8_next_fast(x_189, x_190); +lean_dec(x_190); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_189); +lean_ctor_set(x_197, 1, x_196); +x_198 = lean_uint32_dec_eq(x_195, x_188); +if (x_198 == 0) +{ +lean_object* x_199; lean_object* x_200; +lean_dec(x_197); +x_199 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7; +x_200 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_200, 0, x_182); +lean_ctor_set(x_200, 1, x_199); +return x_200; +} +else +{ +lean_object* x_201; +lean_dec(x_182); +x_201 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8; +x_5 = x_197; +x_6 = x_201; +goto block_181; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox(x_1); +lean_dec(x_1); +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_5, x_6, x_7, x_4); +return x_8; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_parseWith(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_3; +x_3 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_3) { +case 0: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_parseEraShort(x_2); +return x_4; +} +case 1: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_parseEraLong(x_2); +return x_5; +} +default: +{ +lean_object* x_6; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_parseEraNarrow(x_2); +return x_6; +} +} +} +case 1: +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +switch (lean_obj_tag(x_7)) { +case 0: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_unsigned_to_nat(2u); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_8, x_2); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 1); +x_12 = lean_nat_to_int(x_11); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1; +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_12); +lean_ctor_set(x_9, 1, x_14); +return x_9; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_9, 0); +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_9); +x_17 = lean_nat_to_int(x_16); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1; +x_19 = lean_int_add(x_18, x_17); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_15); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_9); +if (x_21 == 0) +{ +return x_9; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_9, 0); +x_23 = lean_ctor_get(x_9, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_9); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +case 1: +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_unsigned_to_nat(4u); +x_26 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNum(x_25, x_2); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 1); +x_29 = lean_nat_to_int(x_28); +lean_ctor_set(x_26, 1, x_29); +return x_26; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_26, 0); +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_26); +x_32 = lean_nat_to_int(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_30); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} +else +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_26); +if (x_34 == 0) +{ +return x_26; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_26, 0); +x_36 = lean_ctor_get(x_26, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_26); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +default: +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_7, 0); +lean_inc(x_38); +lean_dec(x_7); +x_39 = l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum(x_38, x_2); +lean_dec(x_38); +if (lean_obj_tag(x_39) == 0) +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_39, 1); +x_42 = lean_nat_to_int(x_41); +lean_ctor_set(x_39, 1, x_42); +return x_39; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_39, 0); +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_39); +x_45 = lean_nat_to_int(x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_43); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +uint8_t x_47; +x_47 = !lean_is_exclusive(x_39); +if (x_47 == 0) +{ +return x_39; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_39, 0); +x_49 = lean_ctor_get(x_39, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_39); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; +} +} +} +} +} +case 2: +{ +lean_object* x_51; +x_51 = lean_ctor_get(x_1, 0); +lean_inc(x_51); +lean_dec(x_1); +switch (lean_obj_tag(x_51)) { +case 0: +{ +lean_object* x_52; lean_object* x_53; +x_52 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2; +x_53 = l___private_Std_Time_Format_Basic_0__Std_Time_parseSigned(x_52, x_2); +if (lean_obj_tag(x_53) == 0) +{ +uint8_t x_54; +x_54 = !lean_is_exclusive(x_53); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_53, 1); +x_56 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1; +x_57 = lean_int_add(x_56, x_55); +lean_dec(x_55); +lean_ctor_set(x_53, 1, x_57); +return x_53; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_58 = lean_ctor_get(x_53, 0); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_53); +x_60 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1; +x_61 = lean_int_add(x_60, x_59); +lean_dec(x_59); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_58); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} +} +else +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_53); +if (x_63 == 0) +{ +return x_53; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_53, 0); +x_65 = lean_ctor_get(x_53, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_53); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +case 1: +{ +lean_object* x_67; lean_object* x_68; +x_67 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3; +x_68 = l___private_Std_Time_Format_Basic_0__Std_Time_parseSigned(x_67, x_2); +return x_68; +} +default: +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_51, 0); +lean_inc(x_69); +lean_dec(x_51); +x_70 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_70, 0, x_69); +x_71 = l___private_Std_Time_Format_Basic_0__Std_Time_parseSigned(x_70, x_2); +return x_71; +} +} +} +case 3: +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_72 = lean_ctor_get(x_1, 0); +lean_inc(x_72); +lean_dec(x_1); +x_73 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_73, 0, x_72); +x_74 = lean_unsigned_to_nat(1u); +x_75 = lean_unsigned_to_nat(366u); +x_76 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_74, x_75, x_73, x_2); +if (lean_obj_tag(x_76) == 0) +{ +uint8_t x_77; +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) +{ +lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_76, 1); +x_79 = 1; +x_80 = lean_box(x_79); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_78); +lean_ctor_set(x_76, 1, x_81); +return x_76; +} +else +{ +lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_76, 0); +x_83 = lean_ctor_get(x_76, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_76); +x_84 = 1; +x_85 = lean_box(x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_83); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 1, x_86); +return x_87; +} +} +else +{ +uint8_t x_88; +x_88 = !lean_is_exclusive(x_76); +if (x_88 == 0) +{ +return x_76; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_76, 0); +x_90 = lean_ctor_get(x_76, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_76); +x_91 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +return x_91; +} +} +} +case 4: +{ +lean_object* x_92; +x_92 = lean_ctor_get(x_1, 0); +lean_inc(x_92); +lean_dec(x_1); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +lean_dec(x_92); +x_94 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_94, 0, x_93); +x_95 = lean_unsigned_to_nat(1u); +x_96 = lean_unsigned_to_nat(12u); +x_97 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_95, x_96, x_94, x_2); +return x_97; +} +else +{ +lean_object* x_98; uint8_t x_99; +x_98 = lean_ctor_get(x_92, 0); +lean_inc(x_98); +lean_dec(x_92); +x_99 = lean_unbox(x_98); +lean_dec(x_98); +switch (x_99) { +case 0: +{ +lean_object* x_100; +x_100 = l_Std_Time_parseMonthShort(x_2); +return x_100; +} +case 1: +{ +lean_object* x_101; +x_101 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong(x_2); +return x_101; +} +default: +{ +lean_object* x_102; +x_102 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthNarrow(x_2); +return x_102; +} +} +} +} +case 5: +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_103 = lean_ctor_get(x_1, 0); +lean_inc(x_103); +lean_dec(x_1); +x_104 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_104, 0, x_103); +x_105 = lean_unsigned_to_nat(1u); +x_106 = lean_unsigned_to_nat(31u); +x_107 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_105, x_106, x_104, x_2); +return x_107; +} +case 6: +{ +lean_object* x_108; +x_108 = lean_ctor_get(x_1, 0); +lean_inc(x_108); +lean_dec(x_1); +if (lean_obj_tag(x_108) == 0) +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +lean_dec(x_108); +x_110 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_110, 0, x_109); +x_111 = lean_unsigned_to_nat(1u); +x_112 = lean_unsigned_to_nat(4u); +x_113 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_111, x_112, x_110, x_2); +return x_113; +} +else +{ +lean_object* x_114; uint8_t x_115; +x_114 = lean_ctor_get(x_108, 0); +lean_inc(x_114); +lean_dec(x_108); +x_115 = lean_unbox(x_114); +lean_dec(x_114); +switch (x_115) { +case 0: +{ +lean_object* x_116; +x_116 = l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterShort(x_2); +return x_116; +} +case 1: +{ +lean_object* x_117; +x_117 = l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterLong(x_2); +return x_117; +} +default: +{ +lean_object* x_118; +x_118 = l___private_Std_Time_Format_Basic_0__Std_Time_parseQuarterNumber(x_2); +return x_118; +} +} +} +} +case 7: +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_119 = lean_ctor_get(x_1, 0); +lean_inc(x_119); +lean_dec(x_1); +x_120 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_120, 0, x_119); +x_121 = lean_unsigned_to_nat(1u); +x_122 = lean_unsigned_to_nat(53u); +x_123 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_121, x_122, x_120, x_2); +return x_123; +} +case 8: +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_124 = lean_ctor_get(x_1, 0); +lean_inc(x_124); +lean_dec(x_1); +x_125 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_125, 0, x_124); +x_126 = lean_unsigned_to_nat(1u); +x_127 = lean_unsigned_to_nat(6u); +x_128 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_126, x_127, x_125, x_2); +return x_128; +} +case 9: +{ +uint8_t x_129; +x_129 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_129) { +case 0: +{ +lean_object* x_130; +x_130 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayShort(x_2); +return x_130; +} +case 1: +{ +lean_object* x_131; +x_131 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayLong(x_2); +return x_131; +} +default: +{ +lean_object* x_132; +x_132 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayNarrow(x_2); +return x_132; +} +} +} +case 10: +{ +lean_object* x_133; +x_133 = lean_ctor_get(x_1, 0); +lean_inc(x_133); +lean_dec(x_1); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_134 = lean_ctor_get(x_133, 0); +lean_inc(x_134); +lean_dec(x_133); +x_135 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_135, 0, x_134); +x_136 = lean_unsigned_to_nat(1u); +x_137 = lean_unsigned_to_nat(7u); +x_138 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_136, x_137, x_135, x_2); +if (lean_obj_tag(x_138) == 0) +{ +uint8_t x_139; +x_139 = !lean_is_exclusive(x_138); +if (x_139 == 0) +{ +lean_object* x_140; uint8_t x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_138, 1); +x_141 = l_Std_Time_Weekday_ofOrdinal(x_140); +lean_dec(x_140); +x_142 = lean_box(x_141); +lean_ctor_set(x_138, 1, x_142); +return x_138; +} +else +{ +lean_object* x_143; lean_object* x_144; uint8_t x_145; lean_object* x_146; lean_object* x_147; +x_143 = lean_ctor_get(x_138, 0); +x_144 = lean_ctor_get(x_138, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_138); +x_145 = l_Std_Time_Weekday_ofOrdinal(x_144); +lean_dec(x_144); +x_146 = lean_box(x_145); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_143); +lean_ctor_set(x_147, 1, x_146); +return x_147; +} +} +else +{ +uint8_t x_148; +x_148 = !lean_is_exclusive(x_138); +if (x_148 == 0) +{ +return x_138; +} +else +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_149 = lean_ctor_get(x_138, 0); +x_150 = lean_ctor_get(x_138, 1); +lean_inc(x_150); +lean_inc(x_149); +lean_dec(x_138); +x_151 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_151, 0, x_149); +lean_ctor_set(x_151, 1, x_150); +return x_151; +} +} +} +else +{ +lean_object* x_152; uint8_t x_153; +x_152 = lean_ctor_get(x_133, 0); +lean_inc(x_152); +lean_dec(x_133); +x_153 = lean_unbox(x_152); +lean_dec(x_152); +switch (x_153) { +case 0: +{ +lean_object* x_154; +x_154 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayShort(x_2); +return x_154; +} +case 1: +{ +lean_object* x_155; +x_155 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayLong(x_2); +return x_155; +} +default: +{ +lean_object* x_156; +x_156 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWeekdayNarrow(x_2); +return x_156; +} +} +} +} +case 11: +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_157 = lean_ctor_get(x_1, 0); +lean_inc(x_157); +lean_dec(x_1); +x_158 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_158, 0, x_157); +x_159 = lean_unsigned_to_nat(1u); +x_160 = lean_unsigned_to_nat(5u); +x_161 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_159, x_160, x_158, x_2); +return x_161; +} +case 12: +{ +uint8_t x_162; +x_162 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_162) { +case 0: +{ +lean_object* x_163; +x_163 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerShort(x_2); +return x_163; +} +case 1: +{ +lean_object* x_164; +x_164 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerLong(x_2); +return x_164; +} +default: +{ +lean_object* x_165; +x_165 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMarkerNarrow(x_2); +return x_165; +} +} +} +case 13: +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_166 = lean_ctor_get(x_1, 0); +lean_inc(x_166); +lean_dec(x_1); +x_167 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_167, 0, x_166); +x_168 = lean_unsigned_to_nat(1u); +x_169 = lean_unsigned_to_nat(12u); +x_170 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_168, x_169, x_167, x_2); +return x_170; +} +case 14: +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_171 = lean_ctor_get(x_1, 0); +lean_inc(x_171); +lean_dec(x_1); +x_172 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_172, 0, x_171); +x_173 = lean_unsigned_to_nat(0u); +x_174 = lean_unsigned_to_nat(11u); +x_175 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_173, x_174, x_172, x_2); +return x_175; +} +case 15: +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_176 = lean_ctor_get(x_1, 0); +lean_inc(x_176); +lean_dec(x_1); +x_177 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_177, 0, x_176); +x_178 = lean_unsigned_to_nat(1u); +x_179 = lean_unsigned_to_nat(24u); +x_180 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_178, x_179, x_177, x_2); +return x_180; +} +case 16: +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_181 = lean_ctor_get(x_1, 0); +lean_inc(x_181); +lean_dec(x_1); +x_182 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_182, 0, x_181); +x_183 = lean_unsigned_to_nat(0u); +x_184 = lean_unsigned_to_nat(23u); +x_185 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_183, x_184, x_182, x_2); +return x_185; +} +case 17: +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_186 = lean_ctor_get(x_1, 0); +lean_inc(x_186); +lean_dec(x_1); +x_187 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_187, 0, x_186); +x_188 = lean_unsigned_to_nat(0u); +x_189 = lean_unsigned_to_nat(59u); +x_190 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_188, x_189, x_187, x_2); +return x_190; +} +case 18: +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +x_191 = lean_ctor_get(x_1, 0); +lean_inc(x_191); +lean_dec(x_1); +x_192 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_192, 0, x_191); +x_193 = lean_unsigned_to_nat(0u); +x_194 = lean_unsigned_to_nat(60u); +x_195 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_193, x_194, x_192, x_2); +if (lean_obj_tag(x_195) == 0) +{ +uint8_t x_196; +x_196 = !lean_is_exclusive(x_195); +if (x_196 == 0) +{ +lean_object* x_197; uint8_t x_198; lean_object* x_199; lean_object* x_200; +x_197 = lean_ctor_get(x_195, 1); +x_198 = 1; +x_199 = lean_box(x_198); +x_200 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_197); +lean_ctor_set(x_195, 1, x_200); +return x_195; +} +else +{ +lean_object* x_201; lean_object* x_202; uint8_t x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_201 = lean_ctor_get(x_195, 0); +x_202 = lean_ctor_get(x_195, 1); +lean_inc(x_202); +lean_inc(x_201); +lean_dec(x_195); +x_203 = 1; +x_204 = lean_box(x_203); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_204); +lean_ctor_set(x_205, 1, x_202); +x_206 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_206, 0, x_201); +lean_ctor_set(x_206, 1, x_205); +return x_206; +} +} +else +{ +uint8_t x_207; +x_207 = !lean_is_exclusive(x_195); +if (x_207 == 0) +{ +return x_195; +} +else +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; +x_208 = lean_ctor_get(x_195, 0); +x_209 = lean_ctor_get(x_195, 1); +lean_inc(x_209); +lean_inc(x_208); +lean_dec(x_195); +x_210 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_210, 0, x_208); +lean_ctor_set(x_210, 1, x_209); +return x_210; +} +} +} +case 19: +{ +lean_object* x_211; +x_211 = lean_ctor_get(x_1, 0); +lean_inc(x_211); +lean_dec(x_1); +if (lean_obj_tag(x_211) == 0) +{ +lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_212 = lean_unsigned_to_nat(0u); +x_213 = lean_unsigned_to_nat(999999999u); +x_214 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4; +x_215 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_212, x_213, x_214, x_2); +return x_215; +} +else +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_216 = lean_ctor_get(x_211, 0); +lean_inc(x_216); +lean_dec(x_211); +x_217 = lean_unsigned_to_nat(9u); +x_218 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseFractionNum___boxed), 3, 2); +lean_closure_set(x_218, 0, x_216); +lean_closure_set(x_218, 1, x_217); +x_219 = lean_unsigned_to_nat(0u); +x_220 = lean_unsigned_to_nat(999999999u); +x_221 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_219, x_220, x_218, x_2); +return x_221; +} +} +case 21: +{ +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_222 = lean_ctor_get(x_1, 0); +lean_inc(x_222); +lean_dec(x_1); +x_223 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum___boxed), 2, 1); +lean_closure_set(x_223, 0, x_222); +x_224 = lean_unsigned_to_nat(0u); +x_225 = lean_unsigned_to_nat(999999999u); +x_226 = l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded(x_224, x_225, x_223, x_2); +return x_226; +} +case 23: +{ +lean_object* x_227; +x_227 = l___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier(x_2); +return x_227; +} +case 24: +{ +lean_object* x_228; +lean_dec(x_1); +x_228 = l___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier(x_2); +return x_228; +} +case 25: +{ +uint8_t x_229; +x_229 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +if (x_229 == 0) +{ +lean_object* x_230; lean_object* x_231; +x_230 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_231 = l_Std_Internal_Parsec_String_pstring(x_230, x_2); +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_232; uint8_t x_233; uint8_t x_234; lean_object* x_235; +x_232 = lean_ctor_get(x_231, 0); +lean_inc(x_232); +lean_dec(x_231); +x_233 = 1; +x_234 = 0; +x_235 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_233, x_233, x_234, x_232); +return x_235; +} +else +{ +uint8_t x_236; +x_236 = !lean_is_exclusive(x_231); +if (x_236 == 0) +{ +return x_231; +} +else +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_237 = lean_ctor_get(x_231, 0); +x_238 = lean_ctor_get(x_231, 1); +lean_inc(x_238); +lean_inc(x_237); +lean_dec(x_231); +x_239 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +return x_239; +} +} +} +else +{ +lean_object* x_240; lean_object* x_241; +x_240 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_241 = l_Std_Internal_Parsec_String_pstring(x_240, x_2); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; uint8_t x_243; uint8_t x_244; uint8_t x_245; lean_object* x_246; +x_242 = lean_ctor_get(x_241, 0); +lean_inc(x_242); +lean_dec(x_241); +x_243 = 0; +x_244 = 2; +x_245 = 0; +x_246 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_243, x_244, x_245, x_242); +return x_246; +} +else +{ +uint8_t x_247; +x_247 = !lean_is_exclusive(x_241); +if (x_247 == 0) +{ +return x_241; +} +else +{ +lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_248 = lean_ctor_get(x_241, 0); +x_249 = lean_ctor_get(x_241, 1); +lean_inc(x_249); +lean_inc(x_248); +lean_dec(x_241); +x_250 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +return x_250; +} +} +} +} +case 26: +{ +uint8_t x_251; lean_object* x_252; lean_object* x_253; +x_251 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_251) { +case 0: +{ +uint8_t x_271; uint8_t x_272; lean_object* x_273; +x_271 = 1; +x_272 = 0; +lean_inc(x_2); +x_273 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_271, x_271, x_272, x_2); +if (lean_obj_tag(x_273) == 0) +{ +uint8_t x_274; +lean_dec(x_2); +x_274 = !lean_is_exclusive(x_273); +if (x_274 == 0) +{ +return x_273; +} +else +{ +lean_object* x_275; lean_object* x_276; lean_object* x_277; +x_275 = lean_ctor_get(x_273, 0); +x_276 = lean_ctor_get(x_273, 1); +lean_inc(x_276); +lean_inc(x_275); +lean_dec(x_273); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_275); +lean_ctor_set(x_277, 1, x_276); +return x_277; +} +} +else +{ +lean_object* x_278; lean_object* x_279; +x_278 = lean_ctor_get(x_273, 0); +lean_inc(x_278); +x_279 = lean_ctor_get(x_273, 1); +lean_inc(x_279); +lean_dec(x_273); +x_252 = x_278; +x_253 = x_279; +goto block_270; +} +} +case 1: +{ +uint8_t x_280; uint8_t x_281; uint8_t x_282; lean_object* x_283; +x_280 = 0; +x_281 = 1; +x_282 = 0; +lean_inc(x_2); +x_283 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_280, x_281, x_282, x_2); +if (lean_obj_tag(x_283) == 0) +{ +uint8_t x_284; +lean_dec(x_2); +x_284 = !lean_is_exclusive(x_283); +if (x_284 == 0) +{ +return x_283; +} +else +{ +lean_object* x_285; lean_object* x_286; lean_object* x_287; +x_285 = lean_ctor_get(x_283, 0); +x_286 = lean_ctor_get(x_283, 1); +lean_inc(x_286); +lean_inc(x_285); +lean_dec(x_283); +x_287 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_287, 0, x_285); +lean_ctor_set(x_287, 1, x_286); +return x_287; +} +} +else +{ +lean_object* x_288; lean_object* x_289; +x_288 = lean_ctor_get(x_283, 0); +lean_inc(x_288); +x_289 = lean_ctor_get(x_283, 1); +lean_inc(x_289); +lean_dec(x_283); +x_252 = x_288; +x_253 = x_289; +goto block_270; +} +} +case 2: +{ +uint8_t x_290; uint8_t x_291; uint8_t x_292; lean_object* x_293; +x_290 = 0; +x_291 = 1; +x_292 = 1; +lean_inc(x_2); +x_293 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_290, x_291, x_292, x_2); +if (lean_obj_tag(x_293) == 0) +{ +uint8_t x_294; +lean_dec(x_2); +x_294 = !lean_is_exclusive(x_293); +if (x_294 == 0) +{ +return x_293; +} +else +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_295 = lean_ctor_get(x_293, 0); +x_296 = lean_ctor_get(x_293, 1); +lean_inc(x_296); +lean_inc(x_295); +lean_dec(x_293); +x_297 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_297, 0, x_295); +lean_ctor_set(x_297, 1, x_296); +return x_297; +} +} +else +{ +lean_object* x_298; lean_object* x_299; +x_298 = lean_ctor_get(x_293, 0); +lean_inc(x_298); +x_299 = lean_ctor_get(x_293, 1); +lean_inc(x_299); +lean_dec(x_293); +x_252 = x_298; +x_253 = x_299; +goto block_270; +} +} +case 3: +{ +uint8_t x_300; uint8_t x_301; lean_object* x_302; +x_300 = 0; +x_301 = 0; +lean_inc(x_2); +x_302 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_300, x_300, x_301, x_2); +if (lean_obj_tag(x_302) == 0) +{ +uint8_t x_303; +lean_dec(x_2); +x_303 = !lean_is_exclusive(x_302); +if (x_303 == 0) +{ +return x_302; +} +else +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; +x_304 = lean_ctor_get(x_302, 0); +x_305 = lean_ctor_get(x_302, 1); +lean_inc(x_305); +lean_inc(x_304); +lean_dec(x_302); +x_306 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_306, 0, x_304); +lean_ctor_set(x_306, 1, x_305); +return x_306; +} +} +else +{ +lean_object* x_307; lean_object* x_308; +x_307 = lean_ctor_get(x_302, 0); +lean_inc(x_307); +x_308 = lean_ctor_get(x_302, 1); +lean_inc(x_308); +lean_dec(x_302); +x_252 = x_307; +x_253 = x_308; +goto block_270; +} +} +default: +{ +uint8_t x_309; uint8_t x_310; lean_object* x_311; +x_309 = 0; +x_310 = 1; +lean_inc(x_2); +x_311 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_309, x_309, x_310, x_2); +if (lean_obj_tag(x_311) == 0) +{ +uint8_t x_312; +lean_dec(x_2); +x_312 = !lean_is_exclusive(x_311); +if (x_312 == 0) +{ +return x_311; +} +else +{ +lean_object* x_313; lean_object* x_314; lean_object* x_315; +x_313 = lean_ctor_get(x_311, 0); +x_314 = lean_ctor_get(x_311, 1); +lean_inc(x_314); +lean_inc(x_313); +lean_dec(x_311); +x_315 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_315, 0, x_313); +lean_ctor_set(x_315, 1, x_314); +return x_315; +} +} +else +{ +lean_object* x_316; lean_object* x_317; +x_316 = lean_ctor_get(x_311, 0); +lean_inc(x_316); +x_317 = lean_ctor_get(x_311, 1); +lean_inc(x_317); +lean_dec(x_311); +x_252 = x_316; +x_253 = x_317; +goto block_270; +} +} +} +block_270: +{ +lean_object* x_254; lean_object* x_255; uint8_t x_256; +x_254 = lean_ctor_get(x_2, 1); +lean_inc(x_254); +lean_dec(x_2); +x_255 = lean_ctor_get(x_252, 1); +lean_inc(x_255); +x_256 = lean_nat_dec_eq(x_254, x_255); +lean_dec(x_255); +lean_dec(x_254); +if (x_256 == 0) +{ +lean_object* x_257; +x_257 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_257, 0, x_252); +lean_ctor_set(x_257, 1, x_253); +return x_257; +} +else +{ +lean_object* x_258; lean_object* x_259; +lean_dec(x_253); +x_258 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5; +x_259 = l_Std_Internal_Parsec_String_pstring(x_258, x_252); +if (lean_obj_tag(x_259) == 0) +{ +uint8_t x_260; +x_260 = !lean_is_exclusive(x_259); +if (x_260 == 0) +{ +lean_object* x_261; lean_object* x_262; +x_261 = lean_ctor_get(x_259, 1); +lean_dec(x_261); +x_262 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +lean_ctor_set(x_259, 1, x_262); +return x_259; +} +else +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_259, 0); +lean_inc(x_263); +lean_dec(x_259); +x_264 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_265 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_265, 0, x_263); +lean_ctor_set(x_265, 1, x_264); +return x_265; +} +} +else +{ +uint8_t x_266; +x_266 = !lean_is_exclusive(x_259); +if (x_266 == 0) +{ +return x_259; +} +else +{ +lean_object* x_267; lean_object* x_268; lean_object* x_269; +x_267 = lean_ctor_get(x_259, 0); +x_268 = lean_ctor_get(x_259, 1); +lean_inc(x_268); +lean_inc(x_267); +lean_dec(x_259); +x_269 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_269, 0, x_267); +lean_ctor_set(x_269, 1, x_268); +return x_269; +} +} +} +} +} +case 27: +{ +uint8_t x_318; +x_318 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_318) { +case 0: +{ +uint8_t x_319; uint8_t x_320; uint8_t x_321; lean_object* x_322; +x_319 = 2; +x_320 = 1; +x_321 = 0; +x_322 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_319, x_320, x_321, x_2); +return x_322; +} +case 1: +{ +uint8_t x_323; uint8_t x_324; uint8_t x_325; lean_object* x_326; +x_323 = 0; +x_324 = 1; +x_325 = 0; +x_326 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_323, x_324, x_325, x_2); +return x_326; +} +case 2: +{ +uint8_t x_327; uint8_t x_328; uint8_t x_329; lean_object* x_330; +x_327 = 0; +x_328 = 2; +x_329 = 1; +x_330 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_327, x_328, x_329, x_2); +return x_330; +} +case 3: +{ +uint8_t x_331; uint8_t x_332; uint8_t x_333; lean_object* x_334; +x_331 = 0; +x_332 = 2; +x_333 = 0; +x_334 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_331, x_332, x_333, x_2); +return x_334; +} +default: +{ +uint8_t x_335; uint8_t x_336; lean_object* x_337; +x_335 = 0; +x_336 = 1; +x_337 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_335, x_335, x_336, x_2); +return x_337; +} +} +} +case 28: +{ +uint8_t x_338; +x_338 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +switch (x_338) { +case 0: +{ +uint8_t x_339; uint8_t x_340; uint8_t x_341; lean_object* x_342; +x_339 = 0; +x_340 = 1; +x_341 = 0; +x_342 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_339, x_340, x_341, x_2); +return x_342; +} +case 1: +{ +lean_object* x_343; lean_object* x_344; +x_343 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4; +x_344 = l_Std_Internal_Parsec_String_pstring(x_343, x_2); +if (lean_obj_tag(x_344) == 0) +{ +lean_object* x_345; uint8_t x_346; uint8_t x_347; uint8_t x_348; lean_object* x_349; +x_345 = lean_ctor_get(x_344, 0); +lean_inc(x_345); +lean_dec(x_344); +x_346 = 0; +x_347 = 1; +x_348 = 1; +lean_inc(x_345); +x_349 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_346, x_347, x_348, x_345); +if (lean_obj_tag(x_349) == 0) +{ +uint8_t x_350; +lean_dec(x_345); +x_350 = !lean_is_exclusive(x_349); +if (x_350 == 0) +{ +return x_349; +} +else +{ +lean_object* x_351; lean_object* x_352; lean_object* x_353; +x_351 = lean_ctor_get(x_349, 0); +x_352 = lean_ctor_get(x_349, 1); +lean_inc(x_352); +lean_inc(x_351); +lean_dec(x_349); +x_353 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_353, 0, x_351); +lean_ctor_set(x_353, 1, x_352); +return x_353; +} +} +else +{ +uint8_t x_354; +x_354 = !lean_is_exclusive(x_349); +if (x_354 == 0) +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; uint8_t x_359; +x_355 = lean_ctor_get(x_349, 0); +x_356 = lean_ctor_get(x_349, 1); +x_357 = lean_ctor_get(x_345, 1); +lean_inc(x_357); +lean_dec(x_345); +x_358 = lean_ctor_get(x_355, 1); +lean_inc(x_358); +x_359 = lean_nat_dec_eq(x_357, x_358); +lean_dec(x_358); +lean_dec(x_357); +if (x_359 == 0) +{ +return x_349; +} +else +{ +lean_object* x_360; +lean_dec(x_356); +x_360 = l_Std_Time_TimeZone_Offset_zero; +lean_ctor_set_tag(x_349, 0); +lean_ctor_set(x_349, 1, x_360); +return x_349; +} +} +else +{ +lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; uint8_t x_365; +x_361 = lean_ctor_get(x_349, 0); +x_362 = lean_ctor_get(x_349, 1); +lean_inc(x_362); +lean_inc(x_361); +lean_dec(x_349); +x_363 = lean_ctor_get(x_345, 1); +lean_inc(x_363); +lean_dec(x_345); +x_364 = lean_ctor_get(x_361, 1); +lean_inc(x_364); +x_365 = lean_nat_dec_eq(x_363, x_364); +lean_dec(x_364); +lean_dec(x_363); +if (x_365 == 0) +{ +lean_object* x_366; +x_366 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_366, 0, x_361); +lean_ctor_set(x_366, 1, x_362); +return x_366; +} +else +{ +lean_object* x_367; lean_object* x_368; +lean_dec(x_362); +x_367 = l_Std_Time_TimeZone_Offset_zero; +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_361); +lean_ctor_set(x_368, 1, x_367); +return x_368; +} +} +} +} +else +{ +uint8_t x_369; +x_369 = !lean_is_exclusive(x_344); +if (x_369 == 0) +{ +return x_344; +} +else +{ +lean_object* x_370; lean_object* x_371; lean_object* x_372; +x_370 = lean_ctor_get(x_344, 0); +x_371 = lean_ctor_get(x_344, 1); +lean_inc(x_371); +lean_inc(x_370); +lean_dec(x_344); +x_372 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_372, 0, x_370); +lean_ctor_set(x_372, 1, x_371); +return x_372; +} +} +} +default: +{ +lean_object* x_373; lean_object* x_374; +x_373 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5; +lean_inc(x_2); +x_374 = l_Std_Internal_Parsec_String_pstring(x_373, x_2); +if (lean_obj_tag(x_374) == 0) +{ +uint8_t x_375; +lean_dec(x_2); +x_375 = !lean_is_exclusive(x_374); +if (x_375 == 0) +{ +lean_object* x_376; lean_object* x_377; +x_376 = lean_ctor_get(x_374, 1); +lean_dec(x_376); +x_377 = l_Std_Time_TimeZone_Offset_zero; +lean_ctor_set(x_374, 1, x_377); +return x_374; +} +else +{ +lean_object* x_378; lean_object* x_379; lean_object* x_380; +x_378 = lean_ctor_get(x_374, 0); +lean_inc(x_378); +lean_dec(x_374); +x_379 = l_Std_Time_TimeZone_Offset_zero; +x_380 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_380, 0, x_378); +lean_ctor_set(x_380, 1, x_379); +return x_380; +} +} +else +{ +uint8_t x_381; +x_381 = !lean_is_exclusive(x_374); +if (x_381 == 0) +{ +lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; uint8_t x_386; +x_382 = lean_ctor_get(x_374, 0); +x_383 = lean_ctor_get(x_374, 1); +x_384 = lean_ctor_get(x_2, 1); +lean_inc(x_384); +lean_dec(x_2); +x_385 = lean_ctor_get(x_382, 1); +lean_inc(x_385); +x_386 = lean_nat_dec_eq(x_384, x_385); +lean_dec(x_385); +lean_dec(x_384); +if (x_386 == 0) +{ +return x_374; +} +else +{ +uint8_t x_387; uint8_t x_388; uint8_t x_389; lean_object* x_390; +lean_free_object(x_374); +lean_dec(x_383); +x_387 = 0; +x_388 = 2; +x_389 = 1; +x_390 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_387, x_388, x_389, x_382); +return x_390; +} +} +else +{ +lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; uint8_t x_395; +x_391 = lean_ctor_get(x_374, 0); +x_392 = lean_ctor_get(x_374, 1); +lean_inc(x_392); +lean_inc(x_391); +lean_dec(x_374); +x_393 = lean_ctor_get(x_2, 1); +lean_inc(x_393); +lean_dec(x_2); +x_394 = lean_ctor_get(x_391, 1); +lean_inc(x_394); +x_395 = lean_nat_dec_eq(x_393, x_394); +lean_dec(x_394); +lean_dec(x_393); +if (x_395 == 0) +{ +lean_object* x_396; +x_396 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_396, 0, x_391); +lean_ctor_set(x_396, 1, x_392); +return x_396; +} +else +{ +uint8_t x_397; uint8_t x_398; uint8_t x_399; lean_object* x_400; +lean_dec(x_392); +x_397 = 0; +x_398 = 2; +x_399 = 1; +x_400 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset(x_397, x_398, x_399, x_391); +return x_400; +} +} +} +} +} +} +default: +{ +lean_object* x_401; lean_object* x_402; +x_401 = lean_ctor_get(x_1, 0); +lean_inc(x_401); +lean_dec(x_1); +x_402 = l___private_Std_Time_Format_Basic_0__Std_Time_parseAtLeastNum(x_401, x_2); +lean_dec(x_401); +if (lean_obj_tag(x_402) == 0) +{ +uint8_t x_403; +x_403 = !lean_is_exclusive(x_402); +if (x_403 == 0) +{ +lean_object* x_404; lean_object* x_405; +x_404 = lean_ctor_get(x_402, 1); +x_405 = lean_nat_to_int(x_404); +lean_ctor_set(x_402, 1, x_405); +return x_402; +} +else +{ +lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; +x_406 = lean_ctor_get(x_402, 0); +x_407 = lean_ctor_get(x_402, 1); +lean_inc(x_407); +lean_inc(x_406); +lean_dec(x_402); +x_408 = lean_nat_to_int(x_407); +x_409 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_409, 0, x_406); +lean_ctor_set(x_409, 1, x_408); +return x_409; +} +} +else +{ +uint8_t x_410; +x_410 = !lean_is_exclusive(x_402); +if (x_410 == 0) +{ +return x_402; +} +else +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_411 = lean_ctor_get(x_402, 0); +x_412 = lean_ctor_get(x_402, 1); +lean_inc(x_412); +lean_inc(x_411); +lean_dec(x_402); +x_413 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_413, 0, x_411); +lean_ctor_set(x_413, 1, x_412); +return x_413; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier(x_1, x_5, x_2); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith(x_5, x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_dec(x_3); +lean_dec(x_2); +lean_inc(x_4); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_2); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_apply_2(x_3, x_7, x_6); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_3); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_5, 0); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_apply_2(x_2, x_10, x_9); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0____private_Std_Time_Format_Basic_0__Std_Time_FormatType_match__1_splitter___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_insert(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_2)) { +case 0: +{ +lean_object* x_4; uint8_t x_5; +lean_dec(x_2); +x_4 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_4, 0, x_3); +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_1, 0); +lean_dec(x_6); +lean_ctor_set(x_1, 0, x_4); +return x_1; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_1, 2); +x_9 = lean_ctor_get(x_1, 3); +x_10 = lean_ctor_get(x_1, 4); +x_11 = lean_ctor_get(x_1, 5); +x_12 = lean_ctor_get(x_1, 6); +x_13 = lean_ctor_get(x_1, 7); +x_14 = lean_ctor_get(x_1, 8); +x_15 = lean_ctor_get(x_1, 9); +x_16 = lean_ctor_get(x_1, 10); +x_17 = lean_ctor_get(x_1, 11); +x_18 = lean_ctor_get(x_1, 12); +x_19 = lean_ctor_get(x_1, 13); +x_20 = lean_ctor_get(x_1, 14); +x_21 = lean_ctor_get(x_1, 15); +x_22 = lean_ctor_get(x_1, 16); +x_23 = lean_ctor_get(x_1, 17); +x_24 = lean_ctor_get(x_1, 18); +x_25 = lean_ctor_get(x_1, 19); +x_26 = lean_ctor_get(x_1, 20); +x_27 = lean_ctor_get(x_1, 21); +x_28 = lean_ctor_get(x_1, 22); +x_29 = lean_ctor_get(x_1, 23); +x_30 = lean_ctor_get(x_1, 24); +x_31 = lean_ctor_get(x_1, 25); +x_32 = lean_ctor_get(x_1, 26); +x_33 = lean_ctor_get(x_1, 27); +x_34 = lean_ctor_get(x_1, 28); +x_35 = lean_ctor_get(x_1, 29); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_1); +x_36 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_36, 0, x_4); +lean_ctor_set(x_36, 1, x_7); +lean_ctor_set(x_36, 2, x_8); +lean_ctor_set(x_36, 3, x_9); +lean_ctor_set(x_36, 4, x_10); +lean_ctor_set(x_36, 5, x_11); +lean_ctor_set(x_36, 6, x_12); +lean_ctor_set(x_36, 7, x_13); +lean_ctor_set(x_36, 8, x_14); +lean_ctor_set(x_36, 9, x_15); +lean_ctor_set(x_36, 10, x_16); +lean_ctor_set(x_36, 11, x_17); +lean_ctor_set(x_36, 12, x_18); +lean_ctor_set(x_36, 13, x_19); +lean_ctor_set(x_36, 14, x_20); +lean_ctor_set(x_36, 15, x_21); +lean_ctor_set(x_36, 16, x_22); +lean_ctor_set(x_36, 17, x_23); +lean_ctor_set(x_36, 18, x_24); +lean_ctor_set(x_36, 19, x_25); +lean_ctor_set(x_36, 20, x_26); +lean_ctor_set(x_36, 21, x_27); +lean_ctor_set(x_36, 22, x_28); +lean_ctor_set(x_36, 23, x_29); +lean_ctor_set(x_36, 24, x_30); +lean_ctor_set(x_36, 25, x_31); +lean_ctor_set(x_36, 26, x_32); +lean_ctor_set(x_36, 27, x_33); +lean_ctor_set(x_36, 28, x_34); +lean_ctor_set(x_36, 29, x_35); +return x_36; +} +} +case 1: +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_2); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_2, 0); +lean_dec(x_38); +x_39 = !lean_is_exclusive(x_1); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_1, 1); +lean_dec(x_40); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_41 = lean_ctor_get(x_1, 0); +x_42 = lean_ctor_get(x_1, 2); +x_43 = lean_ctor_get(x_1, 3); +x_44 = lean_ctor_get(x_1, 4); +x_45 = lean_ctor_get(x_1, 5); +x_46 = lean_ctor_get(x_1, 6); +x_47 = lean_ctor_get(x_1, 7); +x_48 = lean_ctor_get(x_1, 8); +x_49 = lean_ctor_get(x_1, 9); +x_50 = lean_ctor_get(x_1, 10); +x_51 = lean_ctor_get(x_1, 11); +x_52 = lean_ctor_get(x_1, 12); +x_53 = lean_ctor_get(x_1, 13); +x_54 = lean_ctor_get(x_1, 14); +x_55 = lean_ctor_get(x_1, 15); +x_56 = lean_ctor_get(x_1, 16); +x_57 = lean_ctor_get(x_1, 17); +x_58 = lean_ctor_get(x_1, 18); +x_59 = lean_ctor_get(x_1, 19); +x_60 = lean_ctor_get(x_1, 20); +x_61 = lean_ctor_get(x_1, 21); +x_62 = lean_ctor_get(x_1, 22); +x_63 = lean_ctor_get(x_1, 23); +x_64 = lean_ctor_get(x_1, 24); +x_65 = lean_ctor_get(x_1, 25); +x_66 = lean_ctor_get(x_1, 26); +x_67 = lean_ctor_get(x_1, 27); +x_68 = lean_ctor_get(x_1, 28); +x_69 = lean_ctor_get(x_1, 29); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_3); +x_70 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_70, 0, x_41); +lean_ctor_set(x_70, 1, x_2); +lean_ctor_set(x_70, 2, x_42); +lean_ctor_set(x_70, 3, x_43); +lean_ctor_set(x_70, 4, x_44); +lean_ctor_set(x_70, 5, x_45); +lean_ctor_set(x_70, 6, x_46); +lean_ctor_set(x_70, 7, x_47); +lean_ctor_set(x_70, 8, x_48); +lean_ctor_set(x_70, 9, x_49); +lean_ctor_set(x_70, 10, x_50); +lean_ctor_set(x_70, 11, x_51); +lean_ctor_set(x_70, 12, x_52); +lean_ctor_set(x_70, 13, x_53); +lean_ctor_set(x_70, 14, x_54); +lean_ctor_set(x_70, 15, x_55); +lean_ctor_set(x_70, 16, x_56); +lean_ctor_set(x_70, 17, x_57); +lean_ctor_set(x_70, 18, x_58); +lean_ctor_set(x_70, 19, x_59); +lean_ctor_set(x_70, 20, x_60); +lean_ctor_set(x_70, 21, x_61); +lean_ctor_set(x_70, 22, x_62); +lean_ctor_set(x_70, 23, x_63); +lean_ctor_set(x_70, 24, x_64); +lean_ctor_set(x_70, 25, x_65); +lean_ctor_set(x_70, 26, x_66); +lean_ctor_set(x_70, 27, x_67); +lean_ctor_set(x_70, 28, x_68); +lean_ctor_set(x_70, 29, x_69); +return x_70; +} +} +else +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_2); +x_71 = lean_ctor_get(x_1, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_1, 2); +lean_inc(x_72); +x_73 = lean_ctor_get(x_1, 3); +lean_inc(x_73); +x_74 = lean_ctor_get(x_1, 4); +lean_inc(x_74); +x_75 = lean_ctor_get(x_1, 5); +lean_inc(x_75); +x_76 = lean_ctor_get(x_1, 6); +lean_inc(x_76); +x_77 = lean_ctor_get(x_1, 7); +lean_inc(x_77); +x_78 = lean_ctor_get(x_1, 8); +lean_inc(x_78); +x_79 = lean_ctor_get(x_1, 9); +lean_inc(x_79); +x_80 = lean_ctor_get(x_1, 10); +lean_inc(x_80); +x_81 = lean_ctor_get(x_1, 11); +lean_inc(x_81); +x_82 = lean_ctor_get(x_1, 12); +lean_inc(x_82); +x_83 = lean_ctor_get(x_1, 13); +lean_inc(x_83); +x_84 = lean_ctor_get(x_1, 14); +lean_inc(x_84); +x_85 = lean_ctor_get(x_1, 15); +lean_inc(x_85); +x_86 = lean_ctor_get(x_1, 16); +lean_inc(x_86); +x_87 = lean_ctor_get(x_1, 17); +lean_inc(x_87); +x_88 = lean_ctor_get(x_1, 18); +lean_inc(x_88); +x_89 = lean_ctor_get(x_1, 19); +lean_inc(x_89); +x_90 = lean_ctor_get(x_1, 20); +lean_inc(x_90); +x_91 = lean_ctor_get(x_1, 21); +lean_inc(x_91); +x_92 = lean_ctor_get(x_1, 22); +lean_inc(x_92); +x_93 = lean_ctor_get(x_1, 23); +lean_inc(x_93); +x_94 = lean_ctor_get(x_1, 24); +lean_inc(x_94); +x_95 = lean_ctor_get(x_1, 25); +lean_inc(x_95); +x_96 = lean_ctor_get(x_1, 26); +lean_inc(x_96); +x_97 = lean_ctor_get(x_1, 27); +lean_inc(x_97); +x_98 = lean_ctor_get(x_1, 28); +lean_inc(x_98); +x_99 = lean_ctor_get(x_1, 29); +lean_inc(x_99); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_100 = x_1; +} else { + lean_dec_ref(x_1); + x_100 = lean_box(0); +} +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_3); +if (lean_is_scalar(x_100)) { + x_102 = lean_alloc_ctor(0, 30, 0); +} else { + x_102 = x_100; +} +lean_ctor_set(x_102, 0, x_71); +lean_ctor_set(x_102, 1, x_101); +lean_ctor_set(x_102, 2, x_72); +lean_ctor_set(x_102, 3, x_73); +lean_ctor_set(x_102, 4, x_74); +lean_ctor_set(x_102, 5, x_75); +lean_ctor_set(x_102, 6, x_76); +lean_ctor_set(x_102, 7, x_77); +lean_ctor_set(x_102, 8, x_78); +lean_ctor_set(x_102, 9, x_79); +lean_ctor_set(x_102, 10, x_80); +lean_ctor_set(x_102, 11, x_81); +lean_ctor_set(x_102, 12, x_82); +lean_ctor_set(x_102, 13, x_83); +lean_ctor_set(x_102, 14, x_84); +lean_ctor_set(x_102, 15, x_85); +lean_ctor_set(x_102, 16, x_86); +lean_ctor_set(x_102, 17, x_87); +lean_ctor_set(x_102, 18, x_88); +lean_ctor_set(x_102, 19, x_89); +lean_ctor_set(x_102, 20, x_90); +lean_ctor_set(x_102, 21, x_91); +lean_ctor_set(x_102, 22, x_92); +lean_ctor_set(x_102, 23, x_93); +lean_ctor_set(x_102, 24, x_94); +lean_ctor_set(x_102, 25, x_95); +lean_ctor_set(x_102, 26, x_96); +lean_ctor_set(x_102, 27, x_97); +lean_ctor_set(x_102, 28, x_98); +lean_ctor_set(x_102, 29, x_99); +return x_102; +} +} +case 2: +{ +uint8_t x_103; +x_103 = !lean_is_exclusive(x_2); +if (x_103 == 0) +{ +lean_object* x_104; uint8_t x_105; +x_104 = lean_ctor_get(x_2, 0); +lean_dec(x_104); +x_105 = !lean_is_exclusive(x_1); +if (x_105 == 0) +{ +lean_object* x_106; +x_106 = lean_ctor_get(x_1, 2); +lean_dec(x_106); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_107 = lean_ctor_get(x_1, 0); +x_108 = lean_ctor_get(x_1, 1); +x_109 = lean_ctor_get(x_1, 3); +x_110 = lean_ctor_get(x_1, 4); +x_111 = lean_ctor_get(x_1, 5); +x_112 = lean_ctor_get(x_1, 6); +x_113 = lean_ctor_get(x_1, 7); +x_114 = lean_ctor_get(x_1, 8); +x_115 = lean_ctor_get(x_1, 9); +x_116 = lean_ctor_get(x_1, 10); +x_117 = lean_ctor_get(x_1, 11); +x_118 = lean_ctor_get(x_1, 12); +x_119 = lean_ctor_get(x_1, 13); +x_120 = lean_ctor_get(x_1, 14); +x_121 = lean_ctor_get(x_1, 15); +x_122 = lean_ctor_get(x_1, 16); +x_123 = lean_ctor_get(x_1, 17); +x_124 = lean_ctor_get(x_1, 18); +x_125 = lean_ctor_get(x_1, 19); +x_126 = lean_ctor_get(x_1, 20); +x_127 = lean_ctor_get(x_1, 21); +x_128 = lean_ctor_get(x_1, 22); +x_129 = lean_ctor_get(x_1, 23); +x_130 = lean_ctor_get(x_1, 24); +x_131 = lean_ctor_get(x_1, 25); +x_132 = lean_ctor_get(x_1, 26); +x_133 = lean_ctor_get(x_1, 27); +x_134 = lean_ctor_get(x_1, 28); +x_135 = lean_ctor_get(x_1, 29); +lean_inc(x_135); +lean_inc(x_134); +lean_inc(x_133); +lean_inc(x_132); +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); +lean_inc(x_125); +lean_inc(x_124); +lean_inc(x_123); +lean_inc(x_122); +lean_inc(x_121); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); +lean_inc(x_116); +lean_inc(x_115); +lean_inc(x_114); +lean_inc(x_113); +lean_inc(x_112); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_136 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_136, 0, x_107); +lean_ctor_set(x_136, 1, x_108); +lean_ctor_set(x_136, 2, x_2); +lean_ctor_set(x_136, 3, x_109); +lean_ctor_set(x_136, 4, x_110); +lean_ctor_set(x_136, 5, x_111); +lean_ctor_set(x_136, 6, x_112); +lean_ctor_set(x_136, 7, x_113); +lean_ctor_set(x_136, 8, x_114); +lean_ctor_set(x_136, 9, x_115); +lean_ctor_set(x_136, 10, x_116); +lean_ctor_set(x_136, 11, x_117); +lean_ctor_set(x_136, 12, x_118); +lean_ctor_set(x_136, 13, x_119); +lean_ctor_set(x_136, 14, x_120); +lean_ctor_set(x_136, 15, x_121); +lean_ctor_set(x_136, 16, x_122); +lean_ctor_set(x_136, 17, x_123); +lean_ctor_set(x_136, 18, x_124); +lean_ctor_set(x_136, 19, x_125); +lean_ctor_set(x_136, 20, x_126); +lean_ctor_set(x_136, 21, x_127); +lean_ctor_set(x_136, 22, x_128); +lean_ctor_set(x_136, 23, x_129); +lean_ctor_set(x_136, 24, x_130); +lean_ctor_set(x_136, 25, x_131); +lean_ctor_set(x_136, 26, x_132); +lean_ctor_set(x_136, 27, x_133); +lean_ctor_set(x_136, 28, x_134); +lean_ctor_set(x_136, 29, x_135); +return x_136; +} +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +lean_dec(x_2); +x_137 = lean_ctor_get(x_1, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_1, 1); +lean_inc(x_138); +x_139 = lean_ctor_get(x_1, 3); +lean_inc(x_139); +x_140 = lean_ctor_get(x_1, 4); +lean_inc(x_140); +x_141 = lean_ctor_get(x_1, 5); +lean_inc(x_141); +x_142 = lean_ctor_get(x_1, 6); +lean_inc(x_142); +x_143 = lean_ctor_get(x_1, 7); +lean_inc(x_143); +x_144 = lean_ctor_get(x_1, 8); +lean_inc(x_144); +x_145 = lean_ctor_get(x_1, 9); +lean_inc(x_145); +x_146 = lean_ctor_get(x_1, 10); +lean_inc(x_146); +x_147 = lean_ctor_get(x_1, 11); +lean_inc(x_147); +x_148 = lean_ctor_get(x_1, 12); +lean_inc(x_148); +x_149 = lean_ctor_get(x_1, 13); +lean_inc(x_149); +x_150 = lean_ctor_get(x_1, 14); +lean_inc(x_150); +x_151 = lean_ctor_get(x_1, 15); +lean_inc(x_151); +x_152 = lean_ctor_get(x_1, 16); +lean_inc(x_152); +x_153 = lean_ctor_get(x_1, 17); +lean_inc(x_153); +x_154 = lean_ctor_get(x_1, 18); +lean_inc(x_154); +x_155 = lean_ctor_get(x_1, 19); +lean_inc(x_155); +x_156 = lean_ctor_get(x_1, 20); +lean_inc(x_156); +x_157 = lean_ctor_get(x_1, 21); +lean_inc(x_157); +x_158 = lean_ctor_get(x_1, 22); +lean_inc(x_158); +x_159 = lean_ctor_get(x_1, 23); +lean_inc(x_159); +x_160 = lean_ctor_get(x_1, 24); +lean_inc(x_160); +x_161 = lean_ctor_get(x_1, 25); +lean_inc(x_161); +x_162 = lean_ctor_get(x_1, 26); +lean_inc(x_162); +x_163 = lean_ctor_get(x_1, 27); +lean_inc(x_163); +x_164 = lean_ctor_get(x_1, 28); +lean_inc(x_164); +x_165 = lean_ctor_get(x_1, 29); +lean_inc(x_165); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_166 = x_1; +} else { + lean_dec_ref(x_1); + x_166 = lean_box(0); +} +x_167 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_167, 0, x_3); +if (lean_is_scalar(x_166)) { + x_168 = lean_alloc_ctor(0, 30, 0); +} else { + x_168 = x_166; +} +lean_ctor_set(x_168, 0, x_137); +lean_ctor_set(x_168, 1, x_138); +lean_ctor_set(x_168, 2, x_167); +lean_ctor_set(x_168, 3, x_139); +lean_ctor_set(x_168, 4, x_140); +lean_ctor_set(x_168, 5, x_141); +lean_ctor_set(x_168, 6, x_142); +lean_ctor_set(x_168, 7, x_143); +lean_ctor_set(x_168, 8, x_144); +lean_ctor_set(x_168, 9, x_145); +lean_ctor_set(x_168, 10, x_146); +lean_ctor_set(x_168, 11, x_147); +lean_ctor_set(x_168, 12, x_148); +lean_ctor_set(x_168, 13, x_149); +lean_ctor_set(x_168, 14, x_150); +lean_ctor_set(x_168, 15, x_151); +lean_ctor_set(x_168, 16, x_152); +lean_ctor_set(x_168, 17, x_153); +lean_ctor_set(x_168, 18, x_154); +lean_ctor_set(x_168, 19, x_155); +lean_ctor_set(x_168, 20, x_156); +lean_ctor_set(x_168, 21, x_157); +lean_ctor_set(x_168, 22, x_158); +lean_ctor_set(x_168, 23, x_159); +lean_ctor_set(x_168, 24, x_160); +lean_ctor_set(x_168, 25, x_161); +lean_ctor_set(x_168, 26, x_162); +lean_ctor_set(x_168, 27, x_163); +lean_ctor_set(x_168, 28, x_164); +lean_ctor_set(x_168, 29, x_165); +return x_168; +} +} +case 3: +{ +uint8_t x_169; +x_169 = !lean_is_exclusive(x_2); +if (x_169 == 0) +{ +lean_object* x_170; uint8_t x_171; +x_170 = lean_ctor_get(x_2, 0); +lean_dec(x_170); +x_171 = !lean_is_exclusive(x_1); +if (x_171 == 0) +{ +lean_object* x_172; +x_172 = lean_ctor_get(x_1, 3); +lean_dec(x_172); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 3, x_2); +return x_1; +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +x_173 = lean_ctor_get(x_1, 0); +x_174 = lean_ctor_get(x_1, 1); +x_175 = lean_ctor_get(x_1, 2); +x_176 = lean_ctor_get(x_1, 4); +x_177 = lean_ctor_get(x_1, 5); +x_178 = lean_ctor_get(x_1, 6); +x_179 = lean_ctor_get(x_1, 7); +x_180 = lean_ctor_get(x_1, 8); +x_181 = lean_ctor_get(x_1, 9); +x_182 = lean_ctor_get(x_1, 10); +x_183 = lean_ctor_get(x_1, 11); +x_184 = lean_ctor_get(x_1, 12); +x_185 = lean_ctor_get(x_1, 13); +x_186 = lean_ctor_get(x_1, 14); +x_187 = lean_ctor_get(x_1, 15); +x_188 = lean_ctor_get(x_1, 16); +x_189 = lean_ctor_get(x_1, 17); +x_190 = lean_ctor_get(x_1, 18); +x_191 = lean_ctor_get(x_1, 19); +x_192 = lean_ctor_get(x_1, 20); +x_193 = lean_ctor_get(x_1, 21); +x_194 = lean_ctor_get(x_1, 22); +x_195 = lean_ctor_get(x_1, 23); +x_196 = lean_ctor_get(x_1, 24); +x_197 = lean_ctor_get(x_1, 25); +x_198 = lean_ctor_get(x_1, 26); +x_199 = lean_ctor_get(x_1, 27); +x_200 = lean_ctor_get(x_1, 28); +x_201 = lean_ctor_get(x_1, 29); +lean_inc(x_201); +lean_inc(x_200); +lean_inc(x_199); +lean_inc(x_198); +lean_inc(x_197); +lean_inc(x_196); +lean_inc(x_195); +lean_inc(x_194); +lean_inc(x_193); +lean_inc(x_192); +lean_inc(x_191); +lean_inc(x_190); +lean_inc(x_189); +lean_inc(x_188); +lean_inc(x_187); +lean_inc(x_186); +lean_inc(x_185); +lean_inc(x_184); +lean_inc(x_183); +lean_inc(x_182); +lean_inc(x_181); +lean_inc(x_180); +lean_inc(x_179); +lean_inc(x_178); +lean_inc(x_177); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_202 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_202, 0, x_173); +lean_ctor_set(x_202, 1, x_174); +lean_ctor_set(x_202, 2, x_175); +lean_ctor_set(x_202, 3, x_2); +lean_ctor_set(x_202, 4, x_176); +lean_ctor_set(x_202, 5, x_177); +lean_ctor_set(x_202, 6, x_178); +lean_ctor_set(x_202, 7, x_179); +lean_ctor_set(x_202, 8, x_180); +lean_ctor_set(x_202, 9, x_181); +lean_ctor_set(x_202, 10, x_182); +lean_ctor_set(x_202, 11, x_183); +lean_ctor_set(x_202, 12, x_184); +lean_ctor_set(x_202, 13, x_185); +lean_ctor_set(x_202, 14, x_186); +lean_ctor_set(x_202, 15, x_187); +lean_ctor_set(x_202, 16, x_188); +lean_ctor_set(x_202, 17, x_189); +lean_ctor_set(x_202, 18, x_190); +lean_ctor_set(x_202, 19, x_191); +lean_ctor_set(x_202, 20, x_192); +lean_ctor_set(x_202, 21, x_193); +lean_ctor_set(x_202, 22, x_194); +lean_ctor_set(x_202, 23, x_195); +lean_ctor_set(x_202, 24, x_196); +lean_ctor_set(x_202, 25, x_197); +lean_ctor_set(x_202, 26, x_198); +lean_ctor_set(x_202, 27, x_199); +lean_ctor_set(x_202, 28, x_200); +lean_ctor_set(x_202, 29, x_201); +return x_202; +} +} +else +{ +lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +lean_dec(x_2); +x_203 = lean_ctor_get(x_1, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_1, 1); +lean_inc(x_204); +x_205 = lean_ctor_get(x_1, 2); +lean_inc(x_205); +x_206 = lean_ctor_get(x_1, 4); +lean_inc(x_206); +x_207 = lean_ctor_get(x_1, 5); +lean_inc(x_207); +x_208 = lean_ctor_get(x_1, 6); +lean_inc(x_208); +x_209 = lean_ctor_get(x_1, 7); +lean_inc(x_209); +x_210 = lean_ctor_get(x_1, 8); +lean_inc(x_210); +x_211 = lean_ctor_get(x_1, 9); +lean_inc(x_211); +x_212 = lean_ctor_get(x_1, 10); +lean_inc(x_212); +x_213 = lean_ctor_get(x_1, 11); +lean_inc(x_213); +x_214 = lean_ctor_get(x_1, 12); +lean_inc(x_214); +x_215 = lean_ctor_get(x_1, 13); +lean_inc(x_215); +x_216 = lean_ctor_get(x_1, 14); +lean_inc(x_216); +x_217 = lean_ctor_get(x_1, 15); +lean_inc(x_217); +x_218 = lean_ctor_get(x_1, 16); +lean_inc(x_218); +x_219 = lean_ctor_get(x_1, 17); +lean_inc(x_219); +x_220 = lean_ctor_get(x_1, 18); +lean_inc(x_220); +x_221 = lean_ctor_get(x_1, 19); +lean_inc(x_221); +x_222 = lean_ctor_get(x_1, 20); +lean_inc(x_222); +x_223 = lean_ctor_get(x_1, 21); +lean_inc(x_223); +x_224 = lean_ctor_get(x_1, 22); +lean_inc(x_224); +x_225 = lean_ctor_get(x_1, 23); +lean_inc(x_225); +x_226 = lean_ctor_get(x_1, 24); +lean_inc(x_226); +x_227 = lean_ctor_get(x_1, 25); +lean_inc(x_227); +x_228 = lean_ctor_get(x_1, 26); +lean_inc(x_228); +x_229 = lean_ctor_get(x_1, 27); +lean_inc(x_229); +x_230 = lean_ctor_get(x_1, 28); +lean_inc(x_230); +x_231 = lean_ctor_get(x_1, 29); +lean_inc(x_231); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_232 = x_1; +} else { + lean_dec_ref(x_1); + x_232 = lean_box(0); +} +x_233 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_233, 0, x_3); +if (lean_is_scalar(x_232)) { + x_234 = lean_alloc_ctor(0, 30, 0); +} else { + x_234 = x_232; +} +lean_ctor_set(x_234, 0, x_203); +lean_ctor_set(x_234, 1, x_204); +lean_ctor_set(x_234, 2, x_205); +lean_ctor_set(x_234, 3, x_233); +lean_ctor_set(x_234, 4, x_206); +lean_ctor_set(x_234, 5, x_207); +lean_ctor_set(x_234, 6, x_208); +lean_ctor_set(x_234, 7, x_209); +lean_ctor_set(x_234, 8, x_210); +lean_ctor_set(x_234, 9, x_211); +lean_ctor_set(x_234, 10, x_212); +lean_ctor_set(x_234, 11, x_213); +lean_ctor_set(x_234, 12, x_214); +lean_ctor_set(x_234, 13, x_215); +lean_ctor_set(x_234, 14, x_216); +lean_ctor_set(x_234, 15, x_217); +lean_ctor_set(x_234, 16, x_218); +lean_ctor_set(x_234, 17, x_219); +lean_ctor_set(x_234, 18, x_220); +lean_ctor_set(x_234, 19, x_221); +lean_ctor_set(x_234, 20, x_222); +lean_ctor_set(x_234, 21, x_223); +lean_ctor_set(x_234, 22, x_224); +lean_ctor_set(x_234, 23, x_225); +lean_ctor_set(x_234, 24, x_226); +lean_ctor_set(x_234, 25, x_227); +lean_ctor_set(x_234, 26, x_228); +lean_ctor_set(x_234, 27, x_229); +lean_ctor_set(x_234, 28, x_230); +lean_ctor_set(x_234, 29, x_231); +return x_234; +} +} +case 4: +{ +uint8_t x_235; +x_235 = !lean_is_exclusive(x_2); +if (x_235 == 0) +{ +lean_object* x_236; uint8_t x_237; +x_236 = lean_ctor_get(x_2, 0); +lean_dec(x_236); +x_237 = !lean_is_exclusive(x_1); +if (x_237 == 0) +{ +lean_object* x_238; +x_238 = lean_ctor_get(x_1, 4); +lean_dec(x_238); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 4, x_2); +return x_1; +} +else +{ +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_239 = lean_ctor_get(x_1, 0); +x_240 = lean_ctor_get(x_1, 1); +x_241 = lean_ctor_get(x_1, 2); +x_242 = lean_ctor_get(x_1, 3); +x_243 = lean_ctor_get(x_1, 5); +x_244 = lean_ctor_get(x_1, 6); +x_245 = lean_ctor_get(x_1, 7); +x_246 = lean_ctor_get(x_1, 8); +x_247 = lean_ctor_get(x_1, 9); +x_248 = lean_ctor_get(x_1, 10); +x_249 = lean_ctor_get(x_1, 11); +x_250 = lean_ctor_get(x_1, 12); +x_251 = lean_ctor_get(x_1, 13); +x_252 = lean_ctor_get(x_1, 14); +x_253 = lean_ctor_get(x_1, 15); +x_254 = lean_ctor_get(x_1, 16); +x_255 = lean_ctor_get(x_1, 17); +x_256 = lean_ctor_get(x_1, 18); +x_257 = lean_ctor_get(x_1, 19); +x_258 = lean_ctor_get(x_1, 20); +x_259 = lean_ctor_get(x_1, 21); +x_260 = lean_ctor_get(x_1, 22); +x_261 = lean_ctor_get(x_1, 23); +x_262 = lean_ctor_get(x_1, 24); +x_263 = lean_ctor_get(x_1, 25); +x_264 = lean_ctor_get(x_1, 26); +x_265 = lean_ctor_get(x_1, 27); +x_266 = lean_ctor_get(x_1, 28); +x_267 = lean_ctor_get(x_1, 29); +lean_inc(x_267); +lean_inc(x_266); +lean_inc(x_265); +lean_inc(x_264); +lean_inc(x_263); +lean_inc(x_262); +lean_inc(x_261); +lean_inc(x_260); +lean_inc(x_259); +lean_inc(x_258); +lean_inc(x_257); +lean_inc(x_256); +lean_inc(x_255); +lean_inc(x_254); +lean_inc(x_253); +lean_inc(x_252); +lean_inc(x_251); +lean_inc(x_250); +lean_inc(x_249); +lean_inc(x_248); +lean_inc(x_247); +lean_inc(x_246); +lean_inc(x_245); +lean_inc(x_244); +lean_inc(x_243); +lean_inc(x_242); +lean_inc(x_241); +lean_inc(x_240); +lean_inc(x_239); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_268 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_268, 0, x_239); +lean_ctor_set(x_268, 1, x_240); +lean_ctor_set(x_268, 2, x_241); +lean_ctor_set(x_268, 3, x_242); +lean_ctor_set(x_268, 4, x_2); +lean_ctor_set(x_268, 5, x_243); +lean_ctor_set(x_268, 6, x_244); +lean_ctor_set(x_268, 7, x_245); +lean_ctor_set(x_268, 8, x_246); +lean_ctor_set(x_268, 9, x_247); +lean_ctor_set(x_268, 10, x_248); +lean_ctor_set(x_268, 11, x_249); +lean_ctor_set(x_268, 12, x_250); +lean_ctor_set(x_268, 13, x_251); +lean_ctor_set(x_268, 14, x_252); +lean_ctor_set(x_268, 15, x_253); +lean_ctor_set(x_268, 16, x_254); +lean_ctor_set(x_268, 17, x_255); +lean_ctor_set(x_268, 18, x_256); +lean_ctor_set(x_268, 19, x_257); +lean_ctor_set(x_268, 20, x_258); +lean_ctor_set(x_268, 21, x_259); +lean_ctor_set(x_268, 22, x_260); +lean_ctor_set(x_268, 23, x_261); +lean_ctor_set(x_268, 24, x_262); +lean_ctor_set(x_268, 25, x_263); +lean_ctor_set(x_268, 26, x_264); +lean_ctor_set(x_268, 27, x_265); +lean_ctor_set(x_268, 28, x_266); +lean_ctor_set(x_268, 29, x_267); +return x_268; +} +} +else +{ +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_2); +x_269 = lean_ctor_get(x_1, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_1, 1); +lean_inc(x_270); +x_271 = lean_ctor_get(x_1, 2); +lean_inc(x_271); +x_272 = lean_ctor_get(x_1, 3); +lean_inc(x_272); +x_273 = lean_ctor_get(x_1, 5); +lean_inc(x_273); +x_274 = lean_ctor_get(x_1, 6); +lean_inc(x_274); +x_275 = lean_ctor_get(x_1, 7); +lean_inc(x_275); +x_276 = lean_ctor_get(x_1, 8); +lean_inc(x_276); +x_277 = lean_ctor_get(x_1, 9); +lean_inc(x_277); +x_278 = lean_ctor_get(x_1, 10); +lean_inc(x_278); +x_279 = lean_ctor_get(x_1, 11); +lean_inc(x_279); +x_280 = lean_ctor_get(x_1, 12); +lean_inc(x_280); +x_281 = lean_ctor_get(x_1, 13); +lean_inc(x_281); +x_282 = lean_ctor_get(x_1, 14); +lean_inc(x_282); +x_283 = lean_ctor_get(x_1, 15); +lean_inc(x_283); +x_284 = lean_ctor_get(x_1, 16); +lean_inc(x_284); +x_285 = lean_ctor_get(x_1, 17); +lean_inc(x_285); +x_286 = lean_ctor_get(x_1, 18); +lean_inc(x_286); +x_287 = lean_ctor_get(x_1, 19); +lean_inc(x_287); +x_288 = lean_ctor_get(x_1, 20); +lean_inc(x_288); +x_289 = lean_ctor_get(x_1, 21); +lean_inc(x_289); +x_290 = lean_ctor_get(x_1, 22); +lean_inc(x_290); +x_291 = lean_ctor_get(x_1, 23); +lean_inc(x_291); +x_292 = lean_ctor_get(x_1, 24); +lean_inc(x_292); +x_293 = lean_ctor_get(x_1, 25); +lean_inc(x_293); +x_294 = lean_ctor_get(x_1, 26); +lean_inc(x_294); +x_295 = lean_ctor_get(x_1, 27); +lean_inc(x_295); +x_296 = lean_ctor_get(x_1, 28); +lean_inc(x_296); +x_297 = lean_ctor_get(x_1, 29); +lean_inc(x_297); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_298 = x_1; +} else { + lean_dec_ref(x_1); + x_298 = lean_box(0); +} +x_299 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_299, 0, x_3); +if (lean_is_scalar(x_298)) { + x_300 = lean_alloc_ctor(0, 30, 0); +} else { + x_300 = x_298; +} +lean_ctor_set(x_300, 0, x_269); +lean_ctor_set(x_300, 1, x_270); +lean_ctor_set(x_300, 2, x_271); +lean_ctor_set(x_300, 3, x_272); +lean_ctor_set(x_300, 4, x_299); +lean_ctor_set(x_300, 5, x_273); +lean_ctor_set(x_300, 6, x_274); +lean_ctor_set(x_300, 7, x_275); +lean_ctor_set(x_300, 8, x_276); +lean_ctor_set(x_300, 9, x_277); +lean_ctor_set(x_300, 10, x_278); +lean_ctor_set(x_300, 11, x_279); +lean_ctor_set(x_300, 12, x_280); +lean_ctor_set(x_300, 13, x_281); +lean_ctor_set(x_300, 14, x_282); +lean_ctor_set(x_300, 15, x_283); +lean_ctor_set(x_300, 16, x_284); +lean_ctor_set(x_300, 17, x_285); +lean_ctor_set(x_300, 18, x_286); +lean_ctor_set(x_300, 19, x_287); +lean_ctor_set(x_300, 20, x_288); +lean_ctor_set(x_300, 21, x_289); +lean_ctor_set(x_300, 22, x_290); +lean_ctor_set(x_300, 23, x_291); +lean_ctor_set(x_300, 24, x_292); +lean_ctor_set(x_300, 25, x_293); +lean_ctor_set(x_300, 26, x_294); +lean_ctor_set(x_300, 27, x_295); +lean_ctor_set(x_300, 28, x_296); +lean_ctor_set(x_300, 29, x_297); +return x_300; +} +} +case 5: +{ +uint8_t x_301; +x_301 = !lean_is_exclusive(x_2); +if (x_301 == 0) +{ +lean_object* x_302; uint8_t x_303; +x_302 = lean_ctor_get(x_2, 0); +lean_dec(x_302); +x_303 = !lean_is_exclusive(x_1); +if (x_303 == 0) +{ +lean_object* x_304; +x_304 = lean_ctor_get(x_1, 5); +lean_dec(x_304); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 5, x_2); +return x_1; +} +else +{ +lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_305 = lean_ctor_get(x_1, 0); +x_306 = lean_ctor_get(x_1, 1); +x_307 = lean_ctor_get(x_1, 2); +x_308 = lean_ctor_get(x_1, 3); +x_309 = lean_ctor_get(x_1, 4); +x_310 = lean_ctor_get(x_1, 6); +x_311 = lean_ctor_get(x_1, 7); +x_312 = lean_ctor_get(x_1, 8); +x_313 = lean_ctor_get(x_1, 9); +x_314 = lean_ctor_get(x_1, 10); +x_315 = lean_ctor_get(x_1, 11); +x_316 = lean_ctor_get(x_1, 12); +x_317 = lean_ctor_get(x_1, 13); +x_318 = lean_ctor_get(x_1, 14); +x_319 = lean_ctor_get(x_1, 15); +x_320 = lean_ctor_get(x_1, 16); +x_321 = lean_ctor_get(x_1, 17); +x_322 = lean_ctor_get(x_1, 18); +x_323 = lean_ctor_get(x_1, 19); +x_324 = lean_ctor_get(x_1, 20); +x_325 = lean_ctor_get(x_1, 21); +x_326 = lean_ctor_get(x_1, 22); +x_327 = lean_ctor_get(x_1, 23); +x_328 = lean_ctor_get(x_1, 24); +x_329 = lean_ctor_get(x_1, 25); +x_330 = lean_ctor_get(x_1, 26); +x_331 = lean_ctor_get(x_1, 27); +x_332 = lean_ctor_get(x_1, 28); +x_333 = lean_ctor_get(x_1, 29); +lean_inc(x_333); +lean_inc(x_332); +lean_inc(x_331); +lean_inc(x_330); +lean_inc(x_329); +lean_inc(x_328); +lean_inc(x_327); +lean_inc(x_326); +lean_inc(x_325); +lean_inc(x_324); +lean_inc(x_323); +lean_inc(x_322); +lean_inc(x_321); +lean_inc(x_320); +lean_inc(x_319); +lean_inc(x_318); +lean_inc(x_317); +lean_inc(x_316); +lean_inc(x_315); +lean_inc(x_314); +lean_inc(x_313); +lean_inc(x_312); +lean_inc(x_311); +lean_inc(x_310); +lean_inc(x_309); +lean_inc(x_308); +lean_inc(x_307); +lean_inc(x_306); +lean_inc(x_305); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_334 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_334, 0, x_305); +lean_ctor_set(x_334, 1, x_306); +lean_ctor_set(x_334, 2, x_307); +lean_ctor_set(x_334, 3, x_308); +lean_ctor_set(x_334, 4, x_309); +lean_ctor_set(x_334, 5, x_2); +lean_ctor_set(x_334, 6, x_310); +lean_ctor_set(x_334, 7, x_311); +lean_ctor_set(x_334, 8, x_312); +lean_ctor_set(x_334, 9, x_313); +lean_ctor_set(x_334, 10, x_314); +lean_ctor_set(x_334, 11, x_315); +lean_ctor_set(x_334, 12, x_316); +lean_ctor_set(x_334, 13, x_317); +lean_ctor_set(x_334, 14, x_318); +lean_ctor_set(x_334, 15, x_319); +lean_ctor_set(x_334, 16, x_320); +lean_ctor_set(x_334, 17, x_321); +lean_ctor_set(x_334, 18, x_322); +lean_ctor_set(x_334, 19, x_323); +lean_ctor_set(x_334, 20, x_324); +lean_ctor_set(x_334, 21, x_325); +lean_ctor_set(x_334, 22, x_326); +lean_ctor_set(x_334, 23, x_327); +lean_ctor_set(x_334, 24, x_328); +lean_ctor_set(x_334, 25, x_329); +lean_ctor_set(x_334, 26, x_330); +lean_ctor_set(x_334, 27, x_331); +lean_ctor_set(x_334, 28, x_332); +lean_ctor_set(x_334, 29, x_333); +return x_334; +} +} +else +{ +lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; +lean_dec(x_2); +x_335 = lean_ctor_get(x_1, 0); +lean_inc(x_335); +x_336 = lean_ctor_get(x_1, 1); +lean_inc(x_336); +x_337 = lean_ctor_get(x_1, 2); +lean_inc(x_337); +x_338 = lean_ctor_get(x_1, 3); +lean_inc(x_338); +x_339 = lean_ctor_get(x_1, 4); +lean_inc(x_339); +x_340 = lean_ctor_get(x_1, 6); +lean_inc(x_340); +x_341 = lean_ctor_get(x_1, 7); +lean_inc(x_341); +x_342 = lean_ctor_get(x_1, 8); +lean_inc(x_342); +x_343 = lean_ctor_get(x_1, 9); +lean_inc(x_343); +x_344 = lean_ctor_get(x_1, 10); +lean_inc(x_344); +x_345 = lean_ctor_get(x_1, 11); +lean_inc(x_345); +x_346 = lean_ctor_get(x_1, 12); +lean_inc(x_346); +x_347 = lean_ctor_get(x_1, 13); +lean_inc(x_347); +x_348 = lean_ctor_get(x_1, 14); +lean_inc(x_348); +x_349 = lean_ctor_get(x_1, 15); +lean_inc(x_349); +x_350 = lean_ctor_get(x_1, 16); +lean_inc(x_350); +x_351 = lean_ctor_get(x_1, 17); +lean_inc(x_351); +x_352 = lean_ctor_get(x_1, 18); +lean_inc(x_352); +x_353 = lean_ctor_get(x_1, 19); +lean_inc(x_353); +x_354 = lean_ctor_get(x_1, 20); +lean_inc(x_354); +x_355 = lean_ctor_get(x_1, 21); +lean_inc(x_355); +x_356 = lean_ctor_get(x_1, 22); +lean_inc(x_356); +x_357 = lean_ctor_get(x_1, 23); +lean_inc(x_357); +x_358 = lean_ctor_get(x_1, 24); +lean_inc(x_358); +x_359 = lean_ctor_get(x_1, 25); +lean_inc(x_359); +x_360 = lean_ctor_get(x_1, 26); +lean_inc(x_360); +x_361 = lean_ctor_get(x_1, 27); +lean_inc(x_361); +x_362 = lean_ctor_get(x_1, 28); +lean_inc(x_362); +x_363 = lean_ctor_get(x_1, 29); +lean_inc(x_363); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_364 = x_1; +} else { + lean_dec_ref(x_1); + x_364 = lean_box(0); +} +x_365 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_365, 0, x_3); +if (lean_is_scalar(x_364)) { + x_366 = lean_alloc_ctor(0, 30, 0); +} else { + x_366 = x_364; +} +lean_ctor_set(x_366, 0, x_335); +lean_ctor_set(x_366, 1, x_336); +lean_ctor_set(x_366, 2, x_337); +lean_ctor_set(x_366, 3, x_338); +lean_ctor_set(x_366, 4, x_339); +lean_ctor_set(x_366, 5, x_365); +lean_ctor_set(x_366, 6, x_340); +lean_ctor_set(x_366, 7, x_341); +lean_ctor_set(x_366, 8, x_342); +lean_ctor_set(x_366, 9, x_343); +lean_ctor_set(x_366, 10, x_344); +lean_ctor_set(x_366, 11, x_345); +lean_ctor_set(x_366, 12, x_346); +lean_ctor_set(x_366, 13, x_347); +lean_ctor_set(x_366, 14, x_348); +lean_ctor_set(x_366, 15, x_349); +lean_ctor_set(x_366, 16, x_350); +lean_ctor_set(x_366, 17, x_351); +lean_ctor_set(x_366, 18, x_352); +lean_ctor_set(x_366, 19, x_353); +lean_ctor_set(x_366, 20, x_354); +lean_ctor_set(x_366, 21, x_355); +lean_ctor_set(x_366, 22, x_356); +lean_ctor_set(x_366, 23, x_357); +lean_ctor_set(x_366, 24, x_358); +lean_ctor_set(x_366, 25, x_359); +lean_ctor_set(x_366, 26, x_360); +lean_ctor_set(x_366, 27, x_361); +lean_ctor_set(x_366, 28, x_362); +lean_ctor_set(x_366, 29, x_363); +return x_366; +} +} +case 6: +{ +uint8_t x_367; +x_367 = !lean_is_exclusive(x_2); +if (x_367 == 0) +{ +lean_object* x_368; uint8_t x_369; +x_368 = lean_ctor_get(x_2, 0); +lean_dec(x_368); +x_369 = !lean_is_exclusive(x_1); +if (x_369 == 0) +{ +lean_object* x_370; +x_370 = lean_ctor_get(x_1, 6); +lean_dec(x_370); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 6, x_2); +return x_1; +} +else +{ +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; +x_371 = lean_ctor_get(x_1, 0); +x_372 = lean_ctor_get(x_1, 1); +x_373 = lean_ctor_get(x_1, 2); +x_374 = lean_ctor_get(x_1, 3); +x_375 = lean_ctor_get(x_1, 4); +x_376 = lean_ctor_get(x_1, 5); +x_377 = lean_ctor_get(x_1, 7); +x_378 = lean_ctor_get(x_1, 8); +x_379 = lean_ctor_get(x_1, 9); +x_380 = lean_ctor_get(x_1, 10); +x_381 = lean_ctor_get(x_1, 11); +x_382 = lean_ctor_get(x_1, 12); +x_383 = lean_ctor_get(x_1, 13); +x_384 = lean_ctor_get(x_1, 14); +x_385 = lean_ctor_get(x_1, 15); +x_386 = lean_ctor_get(x_1, 16); +x_387 = lean_ctor_get(x_1, 17); +x_388 = lean_ctor_get(x_1, 18); +x_389 = lean_ctor_get(x_1, 19); +x_390 = lean_ctor_get(x_1, 20); +x_391 = lean_ctor_get(x_1, 21); +x_392 = lean_ctor_get(x_1, 22); +x_393 = lean_ctor_get(x_1, 23); +x_394 = lean_ctor_get(x_1, 24); +x_395 = lean_ctor_get(x_1, 25); +x_396 = lean_ctor_get(x_1, 26); +x_397 = lean_ctor_get(x_1, 27); +x_398 = lean_ctor_get(x_1, 28); +x_399 = lean_ctor_get(x_1, 29); +lean_inc(x_399); +lean_inc(x_398); +lean_inc(x_397); +lean_inc(x_396); +lean_inc(x_395); +lean_inc(x_394); +lean_inc(x_393); +lean_inc(x_392); +lean_inc(x_391); +lean_inc(x_390); +lean_inc(x_389); +lean_inc(x_388); +lean_inc(x_387); +lean_inc(x_386); +lean_inc(x_385); +lean_inc(x_384); +lean_inc(x_383); +lean_inc(x_382); +lean_inc(x_381); +lean_inc(x_380); +lean_inc(x_379); +lean_inc(x_378); +lean_inc(x_377); +lean_inc(x_376); +lean_inc(x_375); +lean_inc(x_374); +lean_inc(x_373); +lean_inc(x_372); +lean_inc(x_371); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_400 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_400, 0, x_371); +lean_ctor_set(x_400, 1, x_372); +lean_ctor_set(x_400, 2, x_373); +lean_ctor_set(x_400, 3, x_374); +lean_ctor_set(x_400, 4, x_375); +lean_ctor_set(x_400, 5, x_376); +lean_ctor_set(x_400, 6, x_2); +lean_ctor_set(x_400, 7, x_377); +lean_ctor_set(x_400, 8, x_378); +lean_ctor_set(x_400, 9, x_379); +lean_ctor_set(x_400, 10, x_380); +lean_ctor_set(x_400, 11, x_381); +lean_ctor_set(x_400, 12, x_382); +lean_ctor_set(x_400, 13, x_383); +lean_ctor_set(x_400, 14, x_384); +lean_ctor_set(x_400, 15, x_385); +lean_ctor_set(x_400, 16, x_386); +lean_ctor_set(x_400, 17, x_387); +lean_ctor_set(x_400, 18, x_388); +lean_ctor_set(x_400, 19, x_389); +lean_ctor_set(x_400, 20, x_390); +lean_ctor_set(x_400, 21, x_391); +lean_ctor_set(x_400, 22, x_392); +lean_ctor_set(x_400, 23, x_393); +lean_ctor_set(x_400, 24, x_394); +lean_ctor_set(x_400, 25, x_395); +lean_ctor_set(x_400, 26, x_396); +lean_ctor_set(x_400, 27, x_397); +lean_ctor_set(x_400, 28, x_398); +lean_ctor_set(x_400, 29, x_399); +return x_400; +} +} +else +{ +lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; +lean_dec(x_2); +x_401 = lean_ctor_get(x_1, 0); +lean_inc(x_401); +x_402 = lean_ctor_get(x_1, 1); +lean_inc(x_402); +x_403 = lean_ctor_get(x_1, 2); +lean_inc(x_403); +x_404 = lean_ctor_get(x_1, 3); +lean_inc(x_404); +x_405 = lean_ctor_get(x_1, 4); +lean_inc(x_405); +x_406 = lean_ctor_get(x_1, 5); +lean_inc(x_406); +x_407 = lean_ctor_get(x_1, 7); +lean_inc(x_407); +x_408 = lean_ctor_get(x_1, 8); +lean_inc(x_408); +x_409 = lean_ctor_get(x_1, 9); +lean_inc(x_409); +x_410 = lean_ctor_get(x_1, 10); +lean_inc(x_410); +x_411 = lean_ctor_get(x_1, 11); +lean_inc(x_411); +x_412 = lean_ctor_get(x_1, 12); +lean_inc(x_412); +x_413 = lean_ctor_get(x_1, 13); +lean_inc(x_413); +x_414 = lean_ctor_get(x_1, 14); +lean_inc(x_414); +x_415 = lean_ctor_get(x_1, 15); +lean_inc(x_415); +x_416 = lean_ctor_get(x_1, 16); +lean_inc(x_416); +x_417 = lean_ctor_get(x_1, 17); +lean_inc(x_417); +x_418 = lean_ctor_get(x_1, 18); +lean_inc(x_418); +x_419 = lean_ctor_get(x_1, 19); +lean_inc(x_419); +x_420 = lean_ctor_get(x_1, 20); +lean_inc(x_420); +x_421 = lean_ctor_get(x_1, 21); +lean_inc(x_421); +x_422 = lean_ctor_get(x_1, 22); +lean_inc(x_422); +x_423 = lean_ctor_get(x_1, 23); +lean_inc(x_423); +x_424 = lean_ctor_get(x_1, 24); +lean_inc(x_424); +x_425 = lean_ctor_get(x_1, 25); +lean_inc(x_425); +x_426 = lean_ctor_get(x_1, 26); +lean_inc(x_426); +x_427 = lean_ctor_get(x_1, 27); +lean_inc(x_427); +x_428 = lean_ctor_get(x_1, 28); +lean_inc(x_428); +x_429 = lean_ctor_get(x_1, 29); +lean_inc(x_429); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_430 = x_1; +} else { + lean_dec_ref(x_1); + x_430 = lean_box(0); +} +x_431 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_431, 0, x_3); +if (lean_is_scalar(x_430)) { + x_432 = lean_alloc_ctor(0, 30, 0); +} else { + x_432 = x_430; +} +lean_ctor_set(x_432, 0, x_401); +lean_ctor_set(x_432, 1, x_402); +lean_ctor_set(x_432, 2, x_403); +lean_ctor_set(x_432, 3, x_404); +lean_ctor_set(x_432, 4, x_405); +lean_ctor_set(x_432, 5, x_406); +lean_ctor_set(x_432, 6, x_431); +lean_ctor_set(x_432, 7, x_407); +lean_ctor_set(x_432, 8, x_408); +lean_ctor_set(x_432, 9, x_409); +lean_ctor_set(x_432, 10, x_410); +lean_ctor_set(x_432, 11, x_411); +lean_ctor_set(x_432, 12, x_412); +lean_ctor_set(x_432, 13, x_413); +lean_ctor_set(x_432, 14, x_414); +lean_ctor_set(x_432, 15, x_415); +lean_ctor_set(x_432, 16, x_416); +lean_ctor_set(x_432, 17, x_417); +lean_ctor_set(x_432, 18, x_418); +lean_ctor_set(x_432, 19, x_419); +lean_ctor_set(x_432, 20, x_420); +lean_ctor_set(x_432, 21, x_421); +lean_ctor_set(x_432, 22, x_422); +lean_ctor_set(x_432, 23, x_423); +lean_ctor_set(x_432, 24, x_424); +lean_ctor_set(x_432, 25, x_425); +lean_ctor_set(x_432, 26, x_426); +lean_ctor_set(x_432, 27, x_427); +lean_ctor_set(x_432, 28, x_428); +lean_ctor_set(x_432, 29, x_429); +return x_432; +} +} +case 7: +{ +uint8_t x_433; +x_433 = !lean_is_exclusive(x_2); +if (x_433 == 0) +{ +lean_object* x_434; uint8_t x_435; +x_434 = lean_ctor_get(x_2, 0); +lean_dec(x_434); +x_435 = !lean_is_exclusive(x_1); +if (x_435 == 0) +{ +lean_object* x_436; +x_436 = lean_ctor_get(x_1, 7); +lean_dec(x_436); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 7, x_2); +return x_1; +} +else +{ +lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; +x_437 = lean_ctor_get(x_1, 0); +x_438 = lean_ctor_get(x_1, 1); +x_439 = lean_ctor_get(x_1, 2); +x_440 = lean_ctor_get(x_1, 3); +x_441 = lean_ctor_get(x_1, 4); +x_442 = lean_ctor_get(x_1, 5); +x_443 = lean_ctor_get(x_1, 6); +x_444 = lean_ctor_get(x_1, 8); +x_445 = lean_ctor_get(x_1, 9); +x_446 = lean_ctor_get(x_1, 10); +x_447 = lean_ctor_get(x_1, 11); +x_448 = lean_ctor_get(x_1, 12); +x_449 = lean_ctor_get(x_1, 13); +x_450 = lean_ctor_get(x_1, 14); +x_451 = lean_ctor_get(x_1, 15); +x_452 = lean_ctor_get(x_1, 16); +x_453 = lean_ctor_get(x_1, 17); +x_454 = lean_ctor_get(x_1, 18); +x_455 = lean_ctor_get(x_1, 19); +x_456 = lean_ctor_get(x_1, 20); +x_457 = lean_ctor_get(x_1, 21); +x_458 = lean_ctor_get(x_1, 22); +x_459 = lean_ctor_get(x_1, 23); +x_460 = lean_ctor_get(x_1, 24); +x_461 = lean_ctor_get(x_1, 25); +x_462 = lean_ctor_get(x_1, 26); +x_463 = lean_ctor_get(x_1, 27); +x_464 = lean_ctor_get(x_1, 28); +x_465 = lean_ctor_get(x_1, 29); +lean_inc(x_465); +lean_inc(x_464); +lean_inc(x_463); +lean_inc(x_462); +lean_inc(x_461); +lean_inc(x_460); +lean_inc(x_459); +lean_inc(x_458); +lean_inc(x_457); +lean_inc(x_456); +lean_inc(x_455); +lean_inc(x_454); +lean_inc(x_453); +lean_inc(x_452); +lean_inc(x_451); +lean_inc(x_450); +lean_inc(x_449); +lean_inc(x_448); +lean_inc(x_447); +lean_inc(x_446); +lean_inc(x_445); +lean_inc(x_444); +lean_inc(x_443); +lean_inc(x_442); +lean_inc(x_441); +lean_inc(x_440); +lean_inc(x_439); +lean_inc(x_438); +lean_inc(x_437); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_466 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_466, 0, x_437); +lean_ctor_set(x_466, 1, x_438); +lean_ctor_set(x_466, 2, x_439); +lean_ctor_set(x_466, 3, x_440); +lean_ctor_set(x_466, 4, x_441); +lean_ctor_set(x_466, 5, x_442); +lean_ctor_set(x_466, 6, x_443); +lean_ctor_set(x_466, 7, x_2); +lean_ctor_set(x_466, 8, x_444); +lean_ctor_set(x_466, 9, x_445); +lean_ctor_set(x_466, 10, x_446); +lean_ctor_set(x_466, 11, x_447); +lean_ctor_set(x_466, 12, x_448); +lean_ctor_set(x_466, 13, x_449); +lean_ctor_set(x_466, 14, x_450); +lean_ctor_set(x_466, 15, x_451); +lean_ctor_set(x_466, 16, x_452); +lean_ctor_set(x_466, 17, x_453); +lean_ctor_set(x_466, 18, x_454); +lean_ctor_set(x_466, 19, x_455); +lean_ctor_set(x_466, 20, x_456); +lean_ctor_set(x_466, 21, x_457); +lean_ctor_set(x_466, 22, x_458); +lean_ctor_set(x_466, 23, x_459); +lean_ctor_set(x_466, 24, x_460); +lean_ctor_set(x_466, 25, x_461); +lean_ctor_set(x_466, 26, x_462); +lean_ctor_set(x_466, 27, x_463); +lean_ctor_set(x_466, 28, x_464); +lean_ctor_set(x_466, 29, x_465); +return x_466; +} +} +else +{ +lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; +lean_dec(x_2); +x_467 = lean_ctor_get(x_1, 0); +lean_inc(x_467); +x_468 = lean_ctor_get(x_1, 1); +lean_inc(x_468); +x_469 = lean_ctor_get(x_1, 2); +lean_inc(x_469); +x_470 = lean_ctor_get(x_1, 3); +lean_inc(x_470); +x_471 = lean_ctor_get(x_1, 4); +lean_inc(x_471); +x_472 = lean_ctor_get(x_1, 5); +lean_inc(x_472); +x_473 = lean_ctor_get(x_1, 6); +lean_inc(x_473); +x_474 = lean_ctor_get(x_1, 8); +lean_inc(x_474); +x_475 = lean_ctor_get(x_1, 9); +lean_inc(x_475); +x_476 = lean_ctor_get(x_1, 10); +lean_inc(x_476); +x_477 = lean_ctor_get(x_1, 11); +lean_inc(x_477); +x_478 = lean_ctor_get(x_1, 12); +lean_inc(x_478); +x_479 = lean_ctor_get(x_1, 13); +lean_inc(x_479); +x_480 = lean_ctor_get(x_1, 14); +lean_inc(x_480); +x_481 = lean_ctor_get(x_1, 15); +lean_inc(x_481); +x_482 = lean_ctor_get(x_1, 16); +lean_inc(x_482); +x_483 = lean_ctor_get(x_1, 17); +lean_inc(x_483); +x_484 = lean_ctor_get(x_1, 18); +lean_inc(x_484); +x_485 = lean_ctor_get(x_1, 19); +lean_inc(x_485); +x_486 = lean_ctor_get(x_1, 20); +lean_inc(x_486); +x_487 = lean_ctor_get(x_1, 21); +lean_inc(x_487); +x_488 = lean_ctor_get(x_1, 22); +lean_inc(x_488); +x_489 = lean_ctor_get(x_1, 23); +lean_inc(x_489); +x_490 = lean_ctor_get(x_1, 24); +lean_inc(x_490); +x_491 = lean_ctor_get(x_1, 25); +lean_inc(x_491); +x_492 = lean_ctor_get(x_1, 26); +lean_inc(x_492); +x_493 = lean_ctor_get(x_1, 27); +lean_inc(x_493); +x_494 = lean_ctor_get(x_1, 28); +lean_inc(x_494); +x_495 = lean_ctor_get(x_1, 29); +lean_inc(x_495); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_496 = x_1; +} else { + lean_dec_ref(x_1); + x_496 = lean_box(0); +} +x_497 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_497, 0, x_3); +if (lean_is_scalar(x_496)) { + x_498 = lean_alloc_ctor(0, 30, 0); +} else { + x_498 = x_496; +} +lean_ctor_set(x_498, 0, x_467); +lean_ctor_set(x_498, 1, x_468); +lean_ctor_set(x_498, 2, x_469); +lean_ctor_set(x_498, 3, x_470); +lean_ctor_set(x_498, 4, x_471); +lean_ctor_set(x_498, 5, x_472); +lean_ctor_set(x_498, 6, x_473); +lean_ctor_set(x_498, 7, x_497); +lean_ctor_set(x_498, 8, x_474); +lean_ctor_set(x_498, 9, x_475); +lean_ctor_set(x_498, 10, x_476); +lean_ctor_set(x_498, 11, x_477); +lean_ctor_set(x_498, 12, x_478); +lean_ctor_set(x_498, 13, x_479); +lean_ctor_set(x_498, 14, x_480); +lean_ctor_set(x_498, 15, x_481); +lean_ctor_set(x_498, 16, x_482); +lean_ctor_set(x_498, 17, x_483); +lean_ctor_set(x_498, 18, x_484); +lean_ctor_set(x_498, 19, x_485); +lean_ctor_set(x_498, 20, x_486); +lean_ctor_set(x_498, 21, x_487); +lean_ctor_set(x_498, 22, x_488); +lean_ctor_set(x_498, 23, x_489); +lean_ctor_set(x_498, 24, x_490); +lean_ctor_set(x_498, 25, x_491); +lean_ctor_set(x_498, 26, x_492); +lean_ctor_set(x_498, 27, x_493); +lean_ctor_set(x_498, 28, x_494); +lean_ctor_set(x_498, 29, x_495); +return x_498; +} +} +case 8: +{ +uint8_t x_499; +x_499 = !lean_is_exclusive(x_2); +if (x_499 == 0) +{ +lean_object* x_500; uint8_t x_501; +x_500 = lean_ctor_get(x_2, 0); +lean_dec(x_500); +x_501 = !lean_is_exclusive(x_1); +if (x_501 == 0) +{ +lean_object* x_502; +x_502 = lean_ctor_get(x_1, 8); +lean_dec(x_502); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 8, x_2); +return x_1; +} +else +{ +lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; +x_503 = lean_ctor_get(x_1, 0); +x_504 = lean_ctor_get(x_1, 1); +x_505 = lean_ctor_get(x_1, 2); +x_506 = lean_ctor_get(x_1, 3); +x_507 = lean_ctor_get(x_1, 4); +x_508 = lean_ctor_get(x_1, 5); +x_509 = lean_ctor_get(x_1, 6); +x_510 = lean_ctor_get(x_1, 7); +x_511 = lean_ctor_get(x_1, 9); +x_512 = lean_ctor_get(x_1, 10); +x_513 = lean_ctor_get(x_1, 11); +x_514 = lean_ctor_get(x_1, 12); +x_515 = lean_ctor_get(x_1, 13); +x_516 = lean_ctor_get(x_1, 14); +x_517 = lean_ctor_get(x_1, 15); +x_518 = lean_ctor_get(x_1, 16); +x_519 = lean_ctor_get(x_1, 17); +x_520 = lean_ctor_get(x_1, 18); +x_521 = lean_ctor_get(x_1, 19); +x_522 = lean_ctor_get(x_1, 20); +x_523 = lean_ctor_get(x_1, 21); +x_524 = lean_ctor_get(x_1, 22); +x_525 = lean_ctor_get(x_1, 23); +x_526 = lean_ctor_get(x_1, 24); +x_527 = lean_ctor_get(x_1, 25); +x_528 = lean_ctor_get(x_1, 26); +x_529 = lean_ctor_get(x_1, 27); +x_530 = lean_ctor_get(x_1, 28); +x_531 = lean_ctor_get(x_1, 29); +lean_inc(x_531); +lean_inc(x_530); +lean_inc(x_529); +lean_inc(x_528); +lean_inc(x_527); +lean_inc(x_526); +lean_inc(x_525); +lean_inc(x_524); +lean_inc(x_523); +lean_inc(x_522); +lean_inc(x_521); +lean_inc(x_520); +lean_inc(x_519); +lean_inc(x_518); +lean_inc(x_517); +lean_inc(x_516); +lean_inc(x_515); +lean_inc(x_514); +lean_inc(x_513); +lean_inc(x_512); +lean_inc(x_511); +lean_inc(x_510); +lean_inc(x_509); +lean_inc(x_508); +lean_inc(x_507); +lean_inc(x_506); +lean_inc(x_505); +lean_inc(x_504); +lean_inc(x_503); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_532 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_532, 0, x_503); +lean_ctor_set(x_532, 1, x_504); +lean_ctor_set(x_532, 2, x_505); +lean_ctor_set(x_532, 3, x_506); +lean_ctor_set(x_532, 4, x_507); +lean_ctor_set(x_532, 5, x_508); +lean_ctor_set(x_532, 6, x_509); +lean_ctor_set(x_532, 7, x_510); +lean_ctor_set(x_532, 8, x_2); +lean_ctor_set(x_532, 9, x_511); +lean_ctor_set(x_532, 10, x_512); +lean_ctor_set(x_532, 11, x_513); +lean_ctor_set(x_532, 12, x_514); +lean_ctor_set(x_532, 13, x_515); +lean_ctor_set(x_532, 14, x_516); +lean_ctor_set(x_532, 15, x_517); +lean_ctor_set(x_532, 16, x_518); +lean_ctor_set(x_532, 17, x_519); +lean_ctor_set(x_532, 18, x_520); +lean_ctor_set(x_532, 19, x_521); +lean_ctor_set(x_532, 20, x_522); +lean_ctor_set(x_532, 21, x_523); +lean_ctor_set(x_532, 22, x_524); +lean_ctor_set(x_532, 23, x_525); +lean_ctor_set(x_532, 24, x_526); +lean_ctor_set(x_532, 25, x_527); +lean_ctor_set(x_532, 26, x_528); +lean_ctor_set(x_532, 27, x_529); +lean_ctor_set(x_532, 28, x_530); +lean_ctor_set(x_532, 29, x_531); +return x_532; +} +} +else +{ +lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; lean_object* x_564; +lean_dec(x_2); +x_533 = lean_ctor_get(x_1, 0); +lean_inc(x_533); +x_534 = lean_ctor_get(x_1, 1); +lean_inc(x_534); +x_535 = lean_ctor_get(x_1, 2); +lean_inc(x_535); +x_536 = lean_ctor_get(x_1, 3); +lean_inc(x_536); +x_537 = lean_ctor_get(x_1, 4); +lean_inc(x_537); +x_538 = lean_ctor_get(x_1, 5); +lean_inc(x_538); +x_539 = lean_ctor_get(x_1, 6); +lean_inc(x_539); +x_540 = lean_ctor_get(x_1, 7); +lean_inc(x_540); +x_541 = lean_ctor_get(x_1, 9); +lean_inc(x_541); +x_542 = lean_ctor_get(x_1, 10); +lean_inc(x_542); +x_543 = lean_ctor_get(x_1, 11); +lean_inc(x_543); +x_544 = lean_ctor_get(x_1, 12); +lean_inc(x_544); +x_545 = lean_ctor_get(x_1, 13); +lean_inc(x_545); +x_546 = lean_ctor_get(x_1, 14); +lean_inc(x_546); +x_547 = lean_ctor_get(x_1, 15); +lean_inc(x_547); +x_548 = lean_ctor_get(x_1, 16); +lean_inc(x_548); +x_549 = lean_ctor_get(x_1, 17); +lean_inc(x_549); +x_550 = lean_ctor_get(x_1, 18); +lean_inc(x_550); +x_551 = lean_ctor_get(x_1, 19); +lean_inc(x_551); +x_552 = lean_ctor_get(x_1, 20); +lean_inc(x_552); +x_553 = lean_ctor_get(x_1, 21); +lean_inc(x_553); +x_554 = lean_ctor_get(x_1, 22); +lean_inc(x_554); +x_555 = lean_ctor_get(x_1, 23); +lean_inc(x_555); +x_556 = lean_ctor_get(x_1, 24); +lean_inc(x_556); +x_557 = lean_ctor_get(x_1, 25); +lean_inc(x_557); +x_558 = lean_ctor_get(x_1, 26); +lean_inc(x_558); +x_559 = lean_ctor_get(x_1, 27); +lean_inc(x_559); +x_560 = lean_ctor_get(x_1, 28); +lean_inc(x_560); +x_561 = lean_ctor_get(x_1, 29); +lean_inc(x_561); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_562 = x_1; +} else { + lean_dec_ref(x_1); + x_562 = lean_box(0); +} +x_563 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_563, 0, x_3); +if (lean_is_scalar(x_562)) { + x_564 = lean_alloc_ctor(0, 30, 0); +} else { + x_564 = x_562; +} +lean_ctor_set(x_564, 0, x_533); +lean_ctor_set(x_564, 1, x_534); +lean_ctor_set(x_564, 2, x_535); +lean_ctor_set(x_564, 3, x_536); +lean_ctor_set(x_564, 4, x_537); +lean_ctor_set(x_564, 5, x_538); +lean_ctor_set(x_564, 6, x_539); +lean_ctor_set(x_564, 7, x_540); +lean_ctor_set(x_564, 8, x_563); +lean_ctor_set(x_564, 9, x_541); +lean_ctor_set(x_564, 10, x_542); +lean_ctor_set(x_564, 11, x_543); +lean_ctor_set(x_564, 12, x_544); +lean_ctor_set(x_564, 13, x_545); +lean_ctor_set(x_564, 14, x_546); +lean_ctor_set(x_564, 15, x_547); +lean_ctor_set(x_564, 16, x_548); +lean_ctor_set(x_564, 17, x_549); +lean_ctor_set(x_564, 18, x_550); +lean_ctor_set(x_564, 19, x_551); +lean_ctor_set(x_564, 20, x_552); +lean_ctor_set(x_564, 21, x_553); +lean_ctor_set(x_564, 22, x_554); +lean_ctor_set(x_564, 23, x_555); +lean_ctor_set(x_564, 24, x_556); +lean_ctor_set(x_564, 25, x_557); +lean_ctor_set(x_564, 26, x_558); +lean_ctor_set(x_564, 27, x_559); +lean_ctor_set(x_564, 28, x_560); +lean_ctor_set(x_564, 29, x_561); +return x_564; +} +} +case 9: +{ +uint8_t x_565; +lean_dec(x_2); +x_565 = !lean_is_exclusive(x_1); +if (x_565 == 0) +{ +lean_object* x_566; lean_object* x_567; +x_566 = lean_ctor_get(x_1, 9); +lean_dec(x_566); +x_567 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_567, 0, x_3); +lean_ctor_set(x_1, 9, x_567); +return x_1; +} +else +{ +lean_object* x_568; lean_object* x_569; lean_object* x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; +x_568 = lean_ctor_get(x_1, 0); +x_569 = lean_ctor_get(x_1, 1); +x_570 = lean_ctor_get(x_1, 2); +x_571 = lean_ctor_get(x_1, 3); +x_572 = lean_ctor_get(x_1, 4); +x_573 = lean_ctor_get(x_1, 5); +x_574 = lean_ctor_get(x_1, 6); +x_575 = lean_ctor_get(x_1, 7); +x_576 = lean_ctor_get(x_1, 8); +x_577 = lean_ctor_get(x_1, 10); +x_578 = lean_ctor_get(x_1, 11); +x_579 = lean_ctor_get(x_1, 12); +x_580 = lean_ctor_get(x_1, 13); +x_581 = lean_ctor_get(x_1, 14); +x_582 = lean_ctor_get(x_1, 15); +x_583 = lean_ctor_get(x_1, 16); +x_584 = lean_ctor_get(x_1, 17); +x_585 = lean_ctor_get(x_1, 18); +x_586 = lean_ctor_get(x_1, 19); +x_587 = lean_ctor_get(x_1, 20); +x_588 = lean_ctor_get(x_1, 21); +x_589 = lean_ctor_get(x_1, 22); +x_590 = lean_ctor_get(x_1, 23); +x_591 = lean_ctor_get(x_1, 24); +x_592 = lean_ctor_get(x_1, 25); +x_593 = lean_ctor_get(x_1, 26); +x_594 = lean_ctor_get(x_1, 27); +x_595 = lean_ctor_get(x_1, 28); +x_596 = lean_ctor_get(x_1, 29); +lean_inc(x_596); +lean_inc(x_595); +lean_inc(x_594); +lean_inc(x_593); +lean_inc(x_592); +lean_inc(x_591); +lean_inc(x_590); +lean_inc(x_589); +lean_inc(x_588); +lean_inc(x_587); +lean_inc(x_586); +lean_inc(x_585); +lean_inc(x_584); +lean_inc(x_583); +lean_inc(x_582); +lean_inc(x_581); +lean_inc(x_580); +lean_inc(x_579); +lean_inc(x_578); +lean_inc(x_577); +lean_inc(x_576); +lean_inc(x_575); +lean_inc(x_574); +lean_inc(x_573); +lean_inc(x_572); +lean_inc(x_571); +lean_inc(x_570); +lean_inc(x_569); +lean_inc(x_568); +lean_dec(x_1); +x_597 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_597, 0, x_3); +x_598 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_598, 0, x_568); +lean_ctor_set(x_598, 1, x_569); +lean_ctor_set(x_598, 2, x_570); +lean_ctor_set(x_598, 3, x_571); +lean_ctor_set(x_598, 4, x_572); +lean_ctor_set(x_598, 5, x_573); +lean_ctor_set(x_598, 6, x_574); +lean_ctor_set(x_598, 7, x_575); +lean_ctor_set(x_598, 8, x_576); +lean_ctor_set(x_598, 9, x_597); +lean_ctor_set(x_598, 10, x_577); +lean_ctor_set(x_598, 11, x_578); +lean_ctor_set(x_598, 12, x_579); +lean_ctor_set(x_598, 13, x_580); +lean_ctor_set(x_598, 14, x_581); +lean_ctor_set(x_598, 15, x_582); +lean_ctor_set(x_598, 16, x_583); +lean_ctor_set(x_598, 17, x_584); +lean_ctor_set(x_598, 18, x_585); +lean_ctor_set(x_598, 19, x_586); +lean_ctor_set(x_598, 20, x_587); +lean_ctor_set(x_598, 21, x_588); +lean_ctor_set(x_598, 22, x_589); +lean_ctor_set(x_598, 23, x_590); +lean_ctor_set(x_598, 24, x_591); +lean_ctor_set(x_598, 25, x_592); +lean_ctor_set(x_598, 26, x_593); +lean_ctor_set(x_598, 27, x_594); +lean_ctor_set(x_598, 28, x_595); +lean_ctor_set(x_598, 29, x_596); +return x_598; +} +} +case 10: +{ +uint8_t x_599; +x_599 = !lean_is_exclusive(x_2); +if (x_599 == 0) +{ +lean_object* x_600; uint8_t x_601; +x_600 = lean_ctor_get(x_2, 0); +lean_dec(x_600); +x_601 = !lean_is_exclusive(x_1); +if (x_601 == 0) +{ +lean_object* x_602; +x_602 = lean_ctor_get(x_1, 10); +lean_dec(x_602); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 10, x_2); +return x_1; +} +else +{ +lean_object* x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; +x_603 = lean_ctor_get(x_1, 0); +x_604 = lean_ctor_get(x_1, 1); +x_605 = lean_ctor_get(x_1, 2); +x_606 = lean_ctor_get(x_1, 3); +x_607 = lean_ctor_get(x_1, 4); +x_608 = lean_ctor_get(x_1, 5); +x_609 = lean_ctor_get(x_1, 6); +x_610 = lean_ctor_get(x_1, 7); +x_611 = lean_ctor_get(x_1, 8); +x_612 = lean_ctor_get(x_1, 9); +x_613 = lean_ctor_get(x_1, 11); +x_614 = lean_ctor_get(x_1, 12); +x_615 = lean_ctor_get(x_1, 13); +x_616 = lean_ctor_get(x_1, 14); +x_617 = lean_ctor_get(x_1, 15); +x_618 = lean_ctor_get(x_1, 16); +x_619 = lean_ctor_get(x_1, 17); +x_620 = lean_ctor_get(x_1, 18); +x_621 = lean_ctor_get(x_1, 19); +x_622 = lean_ctor_get(x_1, 20); +x_623 = lean_ctor_get(x_1, 21); +x_624 = lean_ctor_get(x_1, 22); +x_625 = lean_ctor_get(x_1, 23); +x_626 = lean_ctor_get(x_1, 24); +x_627 = lean_ctor_get(x_1, 25); +x_628 = lean_ctor_get(x_1, 26); +x_629 = lean_ctor_get(x_1, 27); +x_630 = lean_ctor_get(x_1, 28); +x_631 = lean_ctor_get(x_1, 29); +lean_inc(x_631); +lean_inc(x_630); +lean_inc(x_629); +lean_inc(x_628); +lean_inc(x_627); +lean_inc(x_626); +lean_inc(x_625); +lean_inc(x_624); +lean_inc(x_623); +lean_inc(x_622); +lean_inc(x_621); +lean_inc(x_620); +lean_inc(x_619); +lean_inc(x_618); +lean_inc(x_617); +lean_inc(x_616); +lean_inc(x_615); +lean_inc(x_614); +lean_inc(x_613); +lean_inc(x_612); +lean_inc(x_611); +lean_inc(x_610); +lean_inc(x_609); +lean_inc(x_608); +lean_inc(x_607); +lean_inc(x_606); +lean_inc(x_605); +lean_inc(x_604); +lean_inc(x_603); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_632 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_632, 0, x_603); +lean_ctor_set(x_632, 1, x_604); +lean_ctor_set(x_632, 2, x_605); +lean_ctor_set(x_632, 3, x_606); +lean_ctor_set(x_632, 4, x_607); +lean_ctor_set(x_632, 5, x_608); +lean_ctor_set(x_632, 6, x_609); +lean_ctor_set(x_632, 7, x_610); +lean_ctor_set(x_632, 8, x_611); +lean_ctor_set(x_632, 9, x_612); +lean_ctor_set(x_632, 10, x_2); +lean_ctor_set(x_632, 11, x_613); +lean_ctor_set(x_632, 12, x_614); +lean_ctor_set(x_632, 13, x_615); +lean_ctor_set(x_632, 14, x_616); +lean_ctor_set(x_632, 15, x_617); +lean_ctor_set(x_632, 16, x_618); +lean_ctor_set(x_632, 17, x_619); +lean_ctor_set(x_632, 18, x_620); +lean_ctor_set(x_632, 19, x_621); +lean_ctor_set(x_632, 20, x_622); +lean_ctor_set(x_632, 21, x_623); +lean_ctor_set(x_632, 22, x_624); +lean_ctor_set(x_632, 23, x_625); +lean_ctor_set(x_632, 24, x_626); +lean_ctor_set(x_632, 25, x_627); +lean_ctor_set(x_632, 26, x_628); +lean_ctor_set(x_632, 27, x_629); +lean_ctor_set(x_632, 28, x_630); +lean_ctor_set(x_632, 29, x_631); +return x_632; +} +} +else +{ +lean_object* x_633; lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; +lean_dec(x_2); +x_633 = lean_ctor_get(x_1, 0); +lean_inc(x_633); +x_634 = lean_ctor_get(x_1, 1); +lean_inc(x_634); +x_635 = lean_ctor_get(x_1, 2); +lean_inc(x_635); +x_636 = lean_ctor_get(x_1, 3); +lean_inc(x_636); +x_637 = lean_ctor_get(x_1, 4); +lean_inc(x_637); +x_638 = lean_ctor_get(x_1, 5); +lean_inc(x_638); +x_639 = lean_ctor_get(x_1, 6); +lean_inc(x_639); +x_640 = lean_ctor_get(x_1, 7); +lean_inc(x_640); +x_641 = lean_ctor_get(x_1, 8); +lean_inc(x_641); +x_642 = lean_ctor_get(x_1, 9); +lean_inc(x_642); +x_643 = lean_ctor_get(x_1, 11); +lean_inc(x_643); +x_644 = lean_ctor_get(x_1, 12); +lean_inc(x_644); +x_645 = lean_ctor_get(x_1, 13); +lean_inc(x_645); +x_646 = lean_ctor_get(x_1, 14); +lean_inc(x_646); +x_647 = lean_ctor_get(x_1, 15); +lean_inc(x_647); +x_648 = lean_ctor_get(x_1, 16); +lean_inc(x_648); +x_649 = lean_ctor_get(x_1, 17); +lean_inc(x_649); +x_650 = lean_ctor_get(x_1, 18); +lean_inc(x_650); +x_651 = lean_ctor_get(x_1, 19); +lean_inc(x_651); +x_652 = lean_ctor_get(x_1, 20); +lean_inc(x_652); +x_653 = lean_ctor_get(x_1, 21); +lean_inc(x_653); +x_654 = lean_ctor_get(x_1, 22); +lean_inc(x_654); +x_655 = lean_ctor_get(x_1, 23); +lean_inc(x_655); +x_656 = lean_ctor_get(x_1, 24); +lean_inc(x_656); +x_657 = lean_ctor_get(x_1, 25); +lean_inc(x_657); +x_658 = lean_ctor_get(x_1, 26); +lean_inc(x_658); +x_659 = lean_ctor_get(x_1, 27); +lean_inc(x_659); +x_660 = lean_ctor_get(x_1, 28); +lean_inc(x_660); +x_661 = lean_ctor_get(x_1, 29); +lean_inc(x_661); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_662 = x_1; +} else { + lean_dec_ref(x_1); + x_662 = lean_box(0); +} +x_663 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_663, 0, x_3); +if (lean_is_scalar(x_662)) { + x_664 = lean_alloc_ctor(0, 30, 0); +} else { + x_664 = x_662; +} +lean_ctor_set(x_664, 0, x_633); +lean_ctor_set(x_664, 1, x_634); +lean_ctor_set(x_664, 2, x_635); +lean_ctor_set(x_664, 3, x_636); +lean_ctor_set(x_664, 4, x_637); +lean_ctor_set(x_664, 5, x_638); +lean_ctor_set(x_664, 6, x_639); +lean_ctor_set(x_664, 7, x_640); +lean_ctor_set(x_664, 8, x_641); +lean_ctor_set(x_664, 9, x_642); +lean_ctor_set(x_664, 10, x_663); +lean_ctor_set(x_664, 11, x_643); +lean_ctor_set(x_664, 12, x_644); +lean_ctor_set(x_664, 13, x_645); +lean_ctor_set(x_664, 14, x_646); +lean_ctor_set(x_664, 15, x_647); +lean_ctor_set(x_664, 16, x_648); +lean_ctor_set(x_664, 17, x_649); +lean_ctor_set(x_664, 18, x_650); +lean_ctor_set(x_664, 19, x_651); +lean_ctor_set(x_664, 20, x_652); +lean_ctor_set(x_664, 21, x_653); +lean_ctor_set(x_664, 22, x_654); +lean_ctor_set(x_664, 23, x_655); +lean_ctor_set(x_664, 24, x_656); +lean_ctor_set(x_664, 25, x_657); +lean_ctor_set(x_664, 26, x_658); +lean_ctor_set(x_664, 27, x_659); +lean_ctor_set(x_664, 28, x_660); +lean_ctor_set(x_664, 29, x_661); +return x_664; +} +} +case 11: +{ +uint8_t x_665; +x_665 = !lean_is_exclusive(x_2); +if (x_665 == 0) +{ +lean_object* x_666; uint8_t x_667; +x_666 = lean_ctor_get(x_2, 0); +lean_dec(x_666); +x_667 = !lean_is_exclusive(x_1); +if (x_667 == 0) +{ +lean_object* x_668; +x_668 = lean_ctor_get(x_1, 11); +lean_dec(x_668); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 11, x_2); +return x_1; +} +else +{ +lean_object* x_669; lean_object* x_670; lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; +x_669 = lean_ctor_get(x_1, 0); +x_670 = lean_ctor_get(x_1, 1); +x_671 = lean_ctor_get(x_1, 2); +x_672 = lean_ctor_get(x_1, 3); +x_673 = lean_ctor_get(x_1, 4); +x_674 = lean_ctor_get(x_1, 5); +x_675 = lean_ctor_get(x_1, 6); +x_676 = lean_ctor_get(x_1, 7); +x_677 = lean_ctor_get(x_1, 8); +x_678 = lean_ctor_get(x_1, 9); +x_679 = lean_ctor_get(x_1, 10); +x_680 = lean_ctor_get(x_1, 12); +x_681 = lean_ctor_get(x_1, 13); +x_682 = lean_ctor_get(x_1, 14); +x_683 = lean_ctor_get(x_1, 15); +x_684 = lean_ctor_get(x_1, 16); +x_685 = lean_ctor_get(x_1, 17); +x_686 = lean_ctor_get(x_1, 18); +x_687 = lean_ctor_get(x_1, 19); +x_688 = lean_ctor_get(x_1, 20); +x_689 = lean_ctor_get(x_1, 21); +x_690 = lean_ctor_get(x_1, 22); +x_691 = lean_ctor_get(x_1, 23); +x_692 = lean_ctor_get(x_1, 24); +x_693 = lean_ctor_get(x_1, 25); +x_694 = lean_ctor_get(x_1, 26); +x_695 = lean_ctor_get(x_1, 27); +x_696 = lean_ctor_get(x_1, 28); +x_697 = lean_ctor_get(x_1, 29); +lean_inc(x_697); +lean_inc(x_696); +lean_inc(x_695); +lean_inc(x_694); +lean_inc(x_693); +lean_inc(x_692); +lean_inc(x_691); +lean_inc(x_690); +lean_inc(x_689); +lean_inc(x_688); +lean_inc(x_687); +lean_inc(x_686); +lean_inc(x_685); +lean_inc(x_684); +lean_inc(x_683); +lean_inc(x_682); +lean_inc(x_681); +lean_inc(x_680); +lean_inc(x_679); +lean_inc(x_678); +lean_inc(x_677); +lean_inc(x_676); +lean_inc(x_675); +lean_inc(x_674); +lean_inc(x_673); +lean_inc(x_672); +lean_inc(x_671); +lean_inc(x_670); +lean_inc(x_669); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_698 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_698, 0, x_669); +lean_ctor_set(x_698, 1, x_670); +lean_ctor_set(x_698, 2, x_671); +lean_ctor_set(x_698, 3, x_672); +lean_ctor_set(x_698, 4, x_673); +lean_ctor_set(x_698, 5, x_674); +lean_ctor_set(x_698, 6, x_675); +lean_ctor_set(x_698, 7, x_676); +lean_ctor_set(x_698, 8, x_677); +lean_ctor_set(x_698, 9, x_678); +lean_ctor_set(x_698, 10, x_679); +lean_ctor_set(x_698, 11, x_2); +lean_ctor_set(x_698, 12, x_680); +lean_ctor_set(x_698, 13, x_681); +lean_ctor_set(x_698, 14, x_682); +lean_ctor_set(x_698, 15, x_683); +lean_ctor_set(x_698, 16, x_684); +lean_ctor_set(x_698, 17, x_685); +lean_ctor_set(x_698, 18, x_686); +lean_ctor_set(x_698, 19, x_687); +lean_ctor_set(x_698, 20, x_688); +lean_ctor_set(x_698, 21, x_689); +lean_ctor_set(x_698, 22, x_690); +lean_ctor_set(x_698, 23, x_691); +lean_ctor_set(x_698, 24, x_692); +lean_ctor_set(x_698, 25, x_693); +lean_ctor_set(x_698, 26, x_694); +lean_ctor_set(x_698, 27, x_695); +lean_ctor_set(x_698, 28, x_696); +lean_ctor_set(x_698, 29, x_697); +return x_698; +} +} +else +{ +lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; lean_object* x_705; lean_object* x_706; lean_object* x_707; lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; lean_object* x_722; lean_object* x_723; lean_object* x_724; lean_object* x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; +lean_dec(x_2); +x_699 = lean_ctor_get(x_1, 0); +lean_inc(x_699); +x_700 = lean_ctor_get(x_1, 1); +lean_inc(x_700); +x_701 = lean_ctor_get(x_1, 2); +lean_inc(x_701); +x_702 = lean_ctor_get(x_1, 3); +lean_inc(x_702); +x_703 = lean_ctor_get(x_1, 4); +lean_inc(x_703); +x_704 = lean_ctor_get(x_1, 5); +lean_inc(x_704); +x_705 = lean_ctor_get(x_1, 6); +lean_inc(x_705); +x_706 = lean_ctor_get(x_1, 7); +lean_inc(x_706); +x_707 = lean_ctor_get(x_1, 8); +lean_inc(x_707); +x_708 = lean_ctor_get(x_1, 9); +lean_inc(x_708); +x_709 = lean_ctor_get(x_1, 10); +lean_inc(x_709); +x_710 = lean_ctor_get(x_1, 12); +lean_inc(x_710); +x_711 = lean_ctor_get(x_1, 13); +lean_inc(x_711); +x_712 = lean_ctor_get(x_1, 14); +lean_inc(x_712); +x_713 = lean_ctor_get(x_1, 15); +lean_inc(x_713); +x_714 = lean_ctor_get(x_1, 16); +lean_inc(x_714); +x_715 = lean_ctor_get(x_1, 17); +lean_inc(x_715); +x_716 = lean_ctor_get(x_1, 18); +lean_inc(x_716); +x_717 = lean_ctor_get(x_1, 19); +lean_inc(x_717); +x_718 = lean_ctor_get(x_1, 20); +lean_inc(x_718); +x_719 = lean_ctor_get(x_1, 21); +lean_inc(x_719); +x_720 = lean_ctor_get(x_1, 22); +lean_inc(x_720); +x_721 = lean_ctor_get(x_1, 23); +lean_inc(x_721); +x_722 = lean_ctor_get(x_1, 24); +lean_inc(x_722); +x_723 = lean_ctor_get(x_1, 25); +lean_inc(x_723); +x_724 = lean_ctor_get(x_1, 26); +lean_inc(x_724); +x_725 = lean_ctor_get(x_1, 27); +lean_inc(x_725); +x_726 = lean_ctor_get(x_1, 28); +lean_inc(x_726); +x_727 = lean_ctor_get(x_1, 29); +lean_inc(x_727); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_728 = x_1; +} else { + lean_dec_ref(x_1); + x_728 = lean_box(0); +} +x_729 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_729, 0, x_3); +if (lean_is_scalar(x_728)) { + x_730 = lean_alloc_ctor(0, 30, 0); +} else { + x_730 = x_728; +} +lean_ctor_set(x_730, 0, x_699); +lean_ctor_set(x_730, 1, x_700); +lean_ctor_set(x_730, 2, x_701); +lean_ctor_set(x_730, 3, x_702); +lean_ctor_set(x_730, 4, x_703); +lean_ctor_set(x_730, 5, x_704); +lean_ctor_set(x_730, 6, x_705); +lean_ctor_set(x_730, 7, x_706); +lean_ctor_set(x_730, 8, x_707); +lean_ctor_set(x_730, 9, x_708); +lean_ctor_set(x_730, 10, x_709); +lean_ctor_set(x_730, 11, x_729); +lean_ctor_set(x_730, 12, x_710); +lean_ctor_set(x_730, 13, x_711); +lean_ctor_set(x_730, 14, x_712); +lean_ctor_set(x_730, 15, x_713); +lean_ctor_set(x_730, 16, x_714); +lean_ctor_set(x_730, 17, x_715); +lean_ctor_set(x_730, 18, x_716); +lean_ctor_set(x_730, 19, x_717); +lean_ctor_set(x_730, 20, x_718); +lean_ctor_set(x_730, 21, x_719); +lean_ctor_set(x_730, 22, x_720); +lean_ctor_set(x_730, 23, x_721); +lean_ctor_set(x_730, 24, x_722); +lean_ctor_set(x_730, 25, x_723); +lean_ctor_set(x_730, 26, x_724); +lean_ctor_set(x_730, 27, x_725); +lean_ctor_set(x_730, 28, x_726); +lean_ctor_set(x_730, 29, x_727); +return x_730; +} +} +case 12: +{ +uint8_t x_731; +lean_dec(x_2); +x_731 = !lean_is_exclusive(x_1); +if (x_731 == 0) +{ +lean_object* x_732; lean_object* x_733; +x_732 = lean_ctor_get(x_1, 12); +lean_dec(x_732); +x_733 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_733, 0, x_3); +lean_ctor_set(x_1, 12, x_733); +return x_1; +} +else +{ +lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; lean_object* x_739; lean_object* x_740; lean_object* x_741; lean_object* x_742; lean_object* x_743; lean_object* x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; lean_object* x_757; lean_object* x_758; lean_object* x_759; lean_object* x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; +x_734 = lean_ctor_get(x_1, 0); +x_735 = lean_ctor_get(x_1, 1); +x_736 = lean_ctor_get(x_1, 2); +x_737 = lean_ctor_get(x_1, 3); +x_738 = lean_ctor_get(x_1, 4); +x_739 = lean_ctor_get(x_1, 5); +x_740 = lean_ctor_get(x_1, 6); +x_741 = lean_ctor_get(x_1, 7); +x_742 = lean_ctor_get(x_1, 8); +x_743 = lean_ctor_get(x_1, 9); +x_744 = lean_ctor_get(x_1, 10); +x_745 = lean_ctor_get(x_1, 11); +x_746 = lean_ctor_get(x_1, 13); +x_747 = lean_ctor_get(x_1, 14); +x_748 = lean_ctor_get(x_1, 15); +x_749 = lean_ctor_get(x_1, 16); +x_750 = lean_ctor_get(x_1, 17); +x_751 = lean_ctor_get(x_1, 18); +x_752 = lean_ctor_get(x_1, 19); +x_753 = lean_ctor_get(x_1, 20); +x_754 = lean_ctor_get(x_1, 21); +x_755 = lean_ctor_get(x_1, 22); +x_756 = lean_ctor_get(x_1, 23); +x_757 = lean_ctor_get(x_1, 24); +x_758 = lean_ctor_get(x_1, 25); +x_759 = lean_ctor_get(x_1, 26); +x_760 = lean_ctor_get(x_1, 27); +x_761 = lean_ctor_get(x_1, 28); +x_762 = lean_ctor_get(x_1, 29); +lean_inc(x_762); +lean_inc(x_761); +lean_inc(x_760); +lean_inc(x_759); +lean_inc(x_758); +lean_inc(x_757); +lean_inc(x_756); +lean_inc(x_755); +lean_inc(x_754); +lean_inc(x_753); +lean_inc(x_752); +lean_inc(x_751); +lean_inc(x_750); +lean_inc(x_749); +lean_inc(x_748); +lean_inc(x_747); +lean_inc(x_746); +lean_inc(x_745); +lean_inc(x_744); +lean_inc(x_743); +lean_inc(x_742); +lean_inc(x_741); +lean_inc(x_740); +lean_inc(x_739); +lean_inc(x_738); +lean_inc(x_737); +lean_inc(x_736); +lean_inc(x_735); +lean_inc(x_734); +lean_dec(x_1); +x_763 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_763, 0, x_3); +x_764 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_764, 0, x_734); +lean_ctor_set(x_764, 1, x_735); +lean_ctor_set(x_764, 2, x_736); +lean_ctor_set(x_764, 3, x_737); +lean_ctor_set(x_764, 4, x_738); +lean_ctor_set(x_764, 5, x_739); +lean_ctor_set(x_764, 6, x_740); +lean_ctor_set(x_764, 7, x_741); +lean_ctor_set(x_764, 8, x_742); +lean_ctor_set(x_764, 9, x_743); +lean_ctor_set(x_764, 10, x_744); +lean_ctor_set(x_764, 11, x_745); +lean_ctor_set(x_764, 12, x_763); +lean_ctor_set(x_764, 13, x_746); +lean_ctor_set(x_764, 14, x_747); +lean_ctor_set(x_764, 15, x_748); +lean_ctor_set(x_764, 16, x_749); +lean_ctor_set(x_764, 17, x_750); +lean_ctor_set(x_764, 18, x_751); +lean_ctor_set(x_764, 19, x_752); +lean_ctor_set(x_764, 20, x_753); +lean_ctor_set(x_764, 21, x_754); +lean_ctor_set(x_764, 22, x_755); +lean_ctor_set(x_764, 23, x_756); +lean_ctor_set(x_764, 24, x_757); +lean_ctor_set(x_764, 25, x_758); +lean_ctor_set(x_764, 26, x_759); +lean_ctor_set(x_764, 27, x_760); +lean_ctor_set(x_764, 28, x_761); +lean_ctor_set(x_764, 29, x_762); +return x_764; +} +} +case 13: +{ +uint8_t x_765; +x_765 = !lean_is_exclusive(x_2); +if (x_765 == 0) +{ +lean_object* x_766; uint8_t x_767; +x_766 = lean_ctor_get(x_2, 0); +lean_dec(x_766); +x_767 = !lean_is_exclusive(x_1); +if (x_767 == 0) +{ +lean_object* x_768; +x_768 = lean_ctor_get(x_1, 13); +lean_dec(x_768); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 13, x_2); +return x_1; +} +else +{ +lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; lean_object* x_774; lean_object* x_775; lean_object* x_776; lean_object* x_777; lean_object* x_778; lean_object* x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; lean_object* x_792; lean_object* x_793; lean_object* x_794; lean_object* x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; +x_769 = lean_ctor_get(x_1, 0); +x_770 = lean_ctor_get(x_1, 1); +x_771 = lean_ctor_get(x_1, 2); +x_772 = lean_ctor_get(x_1, 3); +x_773 = lean_ctor_get(x_1, 4); +x_774 = lean_ctor_get(x_1, 5); +x_775 = lean_ctor_get(x_1, 6); +x_776 = lean_ctor_get(x_1, 7); +x_777 = lean_ctor_get(x_1, 8); +x_778 = lean_ctor_get(x_1, 9); +x_779 = lean_ctor_get(x_1, 10); +x_780 = lean_ctor_get(x_1, 11); +x_781 = lean_ctor_get(x_1, 12); +x_782 = lean_ctor_get(x_1, 14); +x_783 = lean_ctor_get(x_1, 15); +x_784 = lean_ctor_get(x_1, 16); +x_785 = lean_ctor_get(x_1, 17); +x_786 = lean_ctor_get(x_1, 18); +x_787 = lean_ctor_get(x_1, 19); +x_788 = lean_ctor_get(x_1, 20); +x_789 = lean_ctor_get(x_1, 21); +x_790 = lean_ctor_get(x_1, 22); +x_791 = lean_ctor_get(x_1, 23); +x_792 = lean_ctor_get(x_1, 24); +x_793 = lean_ctor_get(x_1, 25); +x_794 = lean_ctor_get(x_1, 26); +x_795 = lean_ctor_get(x_1, 27); +x_796 = lean_ctor_get(x_1, 28); +x_797 = lean_ctor_get(x_1, 29); +lean_inc(x_797); +lean_inc(x_796); +lean_inc(x_795); +lean_inc(x_794); +lean_inc(x_793); +lean_inc(x_792); +lean_inc(x_791); +lean_inc(x_790); +lean_inc(x_789); +lean_inc(x_788); +lean_inc(x_787); +lean_inc(x_786); +lean_inc(x_785); +lean_inc(x_784); +lean_inc(x_783); +lean_inc(x_782); +lean_inc(x_781); +lean_inc(x_780); +lean_inc(x_779); +lean_inc(x_778); +lean_inc(x_777); +lean_inc(x_776); +lean_inc(x_775); +lean_inc(x_774); +lean_inc(x_773); +lean_inc(x_772); +lean_inc(x_771); +lean_inc(x_770); +lean_inc(x_769); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_798 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_798, 0, x_769); +lean_ctor_set(x_798, 1, x_770); +lean_ctor_set(x_798, 2, x_771); +lean_ctor_set(x_798, 3, x_772); +lean_ctor_set(x_798, 4, x_773); +lean_ctor_set(x_798, 5, x_774); +lean_ctor_set(x_798, 6, x_775); +lean_ctor_set(x_798, 7, x_776); +lean_ctor_set(x_798, 8, x_777); +lean_ctor_set(x_798, 9, x_778); +lean_ctor_set(x_798, 10, x_779); +lean_ctor_set(x_798, 11, x_780); +lean_ctor_set(x_798, 12, x_781); +lean_ctor_set(x_798, 13, x_2); +lean_ctor_set(x_798, 14, x_782); +lean_ctor_set(x_798, 15, x_783); +lean_ctor_set(x_798, 16, x_784); +lean_ctor_set(x_798, 17, x_785); +lean_ctor_set(x_798, 18, x_786); +lean_ctor_set(x_798, 19, x_787); +lean_ctor_set(x_798, 20, x_788); +lean_ctor_set(x_798, 21, x_789); +lean_ctor_set(x_798, 22, x_790); +lean_ctor_set(x_798, 23, x_791); +lean_ctor_set(x_798, 24, x_792); +lean_ctor_set(x_798, 25, x_793); +lean_ctor_set(x_798, 26, x_794); +lean_ctor_set(x_798, 27, x_795); +lean_ctor_set(x_798, 28, x_796); +lean_ctor_set(x_798, 29, x_797); +return x_798; +} +} +else +{ +lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; lean_object* x_809; lean_object* x_810; lean_object* x_811; lean_object* x_812; lean_object* x_813; lean_object* x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; lean_object* x_827; lean_object* x_828; lean_object* x_829; lean_object* x_830; +lean_dec(x_2); +x_799 = lean_ctor_get(x_1, 0); +lean_inc(x_799); +x_800 = lean_ctor_get(x_1, 1); +lean_inc(x_800); +x_801 = lean_ctor_get(x_1, 2); +lean_inc(x_801); +x_802 = lean_ctor_get(x_1, 3); +lean_inc(x_802); +x_803 = lean_ctor_get(x_1, 4); +lean_inc(x_803); +x_804 = lean_ctor_get(x_1, 5); +lean_inc(x_804); +x_805 = lean_ctor_get(x_1, 6); +lean_inc(x_805); +x_806 = lean_ctor_get(x_1, 7); +lean_inc(x_806); +x_807 = lean_ctor_get(x_1, 8); +lean_inc(x_807); +x_808 = lean_ctor_get(x_1, 9); +lean_inc(x_808); +x_809 = lean_ctor_get(x_1, 10); +lean_inc(x_809); +x_810 = lean_ctor_get(x_1, 11); +lean_inc(x_810); +x_811 = lean_ctor_get(x_1, 12); +lean_inc(x_811); +x_812 = lean_ctor_get(x_1, 14); +lean_inc(x_812); +x_813 = lean_ctor_get(x_1, 15); +lean_inc(x_813); +x_814 = lean_ctor_get(x_1, 16); +lean_inc(x_814); +x_815 = lean_ctor_get(x_1, 17); +lean_inc(x_815); +x_816 = lean_ctor_get(x_1, 18); +lean_inc(x_816); +x_817 = lean_ctor_get(x_1, 19); +lean_inc(x_817); +x_818 = lean_ctor_get(x_1, 20); +lean_inc(x_818); +x_819 = lean_ctor_get(x_1, 21); +lean_inc(x_819); +x_820 = lean_ctor_get(x_1, 22); +lean_inc(x_820); +x_821 = lean_ctor_get(x_1, 23); +lean_inc(x_821); +x_822 = lean_ctor_get(x_1, 24); +lean_inc(x_822); +x_823 = lean_ctor_get(x_1, 25); +lean_inc(x_823); +x_824 = lean_ctor_get(x_1, 26); +lean_inc(x_824); +x_825 = lean_ctor_get(x_1, 27); +lean_inc(x_825); +x_826 = lean_ctor_get(x_1, 28); +lean_inc(x_826); +x_827 = lean_ctor_get(x_1, 29); +lean_inc(x_827); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_828 = x_1; +} else { + lean_dec_ref(x_1); + x_828 = lean_box(0); +} +x_829 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_829, 0, x_3); +if (lean_is_scalar(x_828)) { + x_830 = lean_alloc_ctor(0, 30, 0); +} else { + x_830 = x_828; +} +lean_ctor_set(x_830, 0, x_799); +lean_ctor_set(x_830, 1, x_800); +lean_ctor_set(x_830, 2, x_801); +lean_ctor_set(x_830, 3, x_802); +lean_ctor_set(x_830, 4, x_803); +lean_ctor_set(x_830, 5, x_804); +lean_ctor_set(x_830, 6, x_805); +lean_ctor_set(x_830, 7, x_806); +lean_ctor_set(x_830, 8, x_807); +lean_ctor_set(x_830, 9, x_808); +lean_ctor_set(x_830, 10, x_809); +lean_ctor_set(x_830, 11, x_810); +lean_ctor_set(x_830, 12, x_811); +lean_ctor_set(x_830, 13, x_829); +lean_ctor_set(x_830, 14, x_812); +lean_ctor_set(x_830, 15, x_813); +lean_ctor_set(x_830, 16, x_814); +lean_ctor_set(x_830, 17, x_815); +lean_ctor_set(x_830, 18, x_816); +lean_ctor_set(x_830, 19, x_817); +lean_ctor_set(x_830, 20, x_818); +lean_ctor_set(x_830, 21, x_819); +lean_ctor_set(x_830, 22, x_820); +lean_ctor_set(x_830, 23, x_821); +lean_ctor_set(x_830, 24, x_822); +lean_ctor_set(x_830, 25, x_823); +lean_ctor_set(x_830, 26, x_824); +lean_ctor_set(x_830, 27, x_825); +lean_ctor_set(x_830, 28, x_826); +lean_ctor_set(x_830, 29, x_827); +return x_830; +} +} +case 14: +{ +uint8_t x_831; +x_831 = !lean_is_exclusive(x_2); +if (x_831 == 0) +{ +lean_object* x_832; uint8_t x_833; +x_832 = lean_ctor_get(x_2, 0); +lean_dec(x_832); +x_833 = !lean_is_exclusive(x_1); +if (x_833 == 0) +{ +lean_object* x_834; +x_834 = lean_ctor_get(x_1, 14); +lean_dec(x_834); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 14, x_2); +return x_1; +} +else +{ +lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; lean_object* x_844; lean_object* x_845; lean_object* x_846; lean_object* x_847; lean_object* x_848; lean_object* x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; lean_object* x_862; lean_object* x_863; lean_object* x_864; +x_835 = lean_ctor_get(x_1, 0); +x_836 = lean_ctor_get(x_1, 1); +x_837 = lean_ctor_get(x_1, 2); +x_838 = lean_ctor_get(x_1, 3); +x_839 = lean_ctor_get(x_1, 4); +x_840 = lean_ctor_get(x_1, 5); +x_841 = lean_ctor_get(x_1, 6); +x_842 = lean_ctor_get(x_1, 7); +x_843 = lean_ctor_get(x_1, 8); +x_844 = lean_ctor_get(x_1, 9); +x_845 = lean_ctor_get(x_1, 10); +x_846 = lean_ctor_get(x_1, 11); +x_847 = lean_ctor_get(x_1, 12); +x_848 = lean_ctor_get(x_1, 13); +x_849 = lean_ctor_get(x_1, 15); +x_850 = lean_ctor_get(x_1, 16); +x_851 = lean_ctor_get(x_1, 17); +x_852 = lean_ctor_get(x_1, 18); +x_853 = lean_ctor_get(x_1, 19); +x_854 = lean_ctor_get(x_1, 20); +x_855 = lean_ctor_get(x_1, 21); +x_856 = lean_ctor_get(x_1, 22); +x_857 = lean_ctor_get(x_1, 23); +x_858 = lean_ctor_get(x_1, 24); +x_859 = lean_ctor_get(x_1, 25); +x_860 = lean_ctor_get(x_1, 26); +x_861 = lean_ctor_get(x_1, 27); +x_862 = lean_ctor_get(x_1, 28); +x_863 = lean_ctor_get(x_1, 29); +lean_inc(x_863); +lean_inc(x_862); +lean_inc(x_861); +lean_inc(x_860); +lean_inc(x_859); +lean_inc(x_858); +lean_inc(x_857); +lean_inc(x_856); +lean_inc(x_855); +lean_inc(x_854); +lean_inc(x_853); +lean_inc(x_852); +lean_inc(x_851); +lean_inc(x_850); +lean_inc(x_849); +lean_inc(x_848); +lean_inc(x_847); +lean_inc(x_846); +lean_inc(x_845); +lean_inc(x_844); +lean_inc(x_843); +lean_inc(x_842); +lean_inc(x_841); +lean_inc(x_840); +lean_inc(x_839); +lean_inc(x_838); +lean_inc(x_837); +lean_inc(x_836); +lean_inc(x_835); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_864 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_864, 0, x_835); +lean_ctor_set(x_864, 1, x_836); +lean_ctor_set(x_864, 2, x_837); +lean_ctor_set(x_864, 3, x_838); +lean_ctor_set(x_864, 4, x_839); +lean_ctor_set(x_864, 5, x_840); +lean_ctor_set(x_864, 6, x_841); +lean_ctor_set(x_864, 7, x_842); +lean_ctor_set(x_864, 8, x_843); +lean_ctor_set(x_864, 9, x_844); +lean_ctor_set(x_864, 10, x_845); +lean_ctor_set(x_864, 11, x_846); +lean_ctor_set(x_864, 12, x_847); +lean_ctor_set(x_864, 13, x_848); +lean_ctor_set(x_864, 14, x_2); +lean_ctor_set(x_864, 15, x_849); +lean_ctor_set(x_864, 16, x_850); +lean_ctor_set(x_864, 17, x_851); +lean_ctor_set(x_864, 18, x_852); +lean_ctor_set(x_864, 19, x_853); +lean_ctor_set(x_864, 20, x_854); +lean_ctor_set(x_864, 21, x_855); +lean_ctor_set(x_864, 22, x_856); +lean_ctor_set(x_864, 23, x_857); +lean_ctor_set(x_864, 24, x_858); +lean_ctor_set(x_864, 25, x_859); +lean_ctor_set(x_864, 26, x_860); +lean_ctor_set(x_864, 27, x_861); +lean_ctor_set(x_864, 28, x_862); +lean_ctor_set(x_864, 29, x_863); +return x_864; +} +} +else +{ +lean_object* x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; lean_object* x_879; lean_object* x_880; lean_object* x_881; lean_object* x_882; lean_object* x_883; lean_object* x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; +lean_dec(x_2); +x_865 = lean_ctor_get(x_1, 0); +lean_inc(x_865); +x_866 = lean_ctor_get(x_1, 1); +lean_inc(x_866); +x_867 = lean_ctor_get(x_1, 2); +lean_inc(x_867); +x_868 = lean_ctor_get(x_1, 3); +lean_inc(x_868); +x_869 = lean_ctor_get(x_1, 4); +lean_inc(x_869); +x_870 = lean_ctor_get(x_1, 5); +lean_inc(x_870); +x_871 = lean_ctor_get(x_1, 6); +lean_inc(x_871); +x_872 = lean_ctor_get(x_1, 7); +lean_inc(x_872); +x_873 = lean_ctor_get(x_1, 8); +lean_inc(x_873); +x_874 = lean_ctor_get(x_1, 9); +lean_inc(x_874); +x_875 = lean_ctor_get(x_1, 10); +lean_inc(x_875); +x_876 = lean_ctor_get(x_1, 11); +lean_inc(x_876); +x_877 = lean_ctor_get(x_1, 12); +lean_inc(x_877); +x_878 = lean_ctor_get(x_1, 13); +lean_inc(x_878); +x_879 = lean_ctor_get(x_1, 15); +lean_inc(x_879); +x_880 = lean_ctor_get(x_1, 16); +lean_inc(x_880); +x_881 = lean_ctor_get(x_1, 17); +lean_inc(x_881); +x_882 = lean_ctor_get(x_1, 18); +lean_inc(x_882); +x_883 = lean_ctor_get(x_1, 19); +lean_inc(x_883); +x_884 = lean_ctor_get(x_1, 20); +lean_inc(x_884); +x_885 = lean_ctor_get(x_1, 21); +lean_inc(x_885); +x_886 = lean_ctor_get(x_1, 22); +lean_inc(x_886); +x_887 = lean_ctor_get(x_1, 23); +lean_inc(x_887); +x_888 = lean_ctor_get(x_1, 24); +lean_inc(x_888); +x_889 = lean_ctor_get(x_1, 25); +lean_inc(x_889); +x_890 = lean_ctor_get(x_1, 26); +lean_inc(x_890); +x_891 = lean_ctor_get(x_1, 27); +lean_inc(x_891); +x_892 = lean_ctor_get(x_1, 28); +lean_inc(x_892); +x_893 = lean_ctor_get(x_1, 29); +lean_inc(x_893); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_894 = x_1; +} else { + lean_dec_ref(x_1); + x_894 = lean_box(0); +} +x_895 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_895, 0, x_3); +if (lean_is_scalar(x_894)) { + x_896 = lean_alloc_ctor(0, 30, 0); +} else { + x_896 = x_894; +} +lean_ctor_set(x_896, 0, x_865); +lean_ctor_set(x_896, 1, x_866); +lean_ctor_set(x_896, 2, x_867); +lean_ctor_set(x_896, 3, x_868); +lean_ctor_set(x_896, 4, x_869); +lean_ctor_set(x_896, 5, x_870); +lean_ctor_set(x_896, 6, x_871); +lean_ctor_set(x_896, 7, x_872); +lean_ctor_set(x_896, 8, x_873); +lean_ctor_set(x_896, 9, x_874); +lean_ctor_set(x_896, 10, x_875); +lean_ctor_set(x_896, 11, x_876); +lean_ctor_set(x_896, 12, x_877); +lean_ctor_set(x_896, 13, x_878); +lean_ctor_set(x_896, 14, x_895); +lean_ctor_set(x_896, 15, x_879); +lean_ctor_set(x_896, 16, x_880); +lean_ctor_set(x_896, 17, x_881); +lean_ctor_set(x_896, 18, x_882); +lean_ctor_set(x_896, 19, x_883); +lean_ctor_set(x_896, 20, x_884); +lean_ctor_set(x_896, 21, x_885); +lean_ctor_set(x_896, 22, x_886); +lean_ctor_set(x_896, 23, x_887); +lean_ctor_set(x_896, 24, x_888); +lean_ctor_set(x_896, 25, x_889); +lean_ctor_set(x_896, 26, x_890); +lean_ctor_set(x_896, 27, x_891); +lean_ctor_set(x_896, 28, x_892); +lean_ctor_set(x_896, 29, x_893); +return x_896; +} +} +case 15: +{ +uint8_t x_897; +x_897 = !lean_is_exclusive(x_2); +if (x_897 == 0) +{ +lean_object* x_898; uint8_t x_899; +x_898 = lean_ctor_get(x_2, 0); +lean_dec(x_898); +x_899 = !lean_is_exclusive(x_1); +if (x_899 == 0) +{ +lean_object* x_900; +x_900 = lean_ctor_get(x_1, 15); +lean_dec(x_900); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 15, x_2); +return x_1; +} +else +{ +lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; lean_object* x_914; lean_object* x_915; lean_object* x_916; lean_object* x_917; lean_object* x_918; lean_object* x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; +x_901 = lean_ctor_get(x_1, 0); +x_902 = lean_ctor_get(x_1, 1); +x_903 = lean_ctor_get(x_1, 2); +x_904 = lean_ctor_get(x_1, 3); +x_905 = lean_ctor_get(x_1, 4); +x_906 = lean_ctor_get(x_1, 5); +x_907 = lean_ctor_get(x_1, 6); +x_908 = lean_ctor_get(x_1, 7); +x_909 = lean_ctor_get(x_1, 8); +x_910 = lean_ctor_get(x_1, 9); +x_911 = lean_ctor_get(x_1, 10); +x_912 = lean_ctor_get(x_1, 11); +x_913 = lean_ctor_get(x_1, 12); +x_914 = lean_ctor_get(x_1, 13); +x_915 = lean_ctor_get(x_1, 14); +x_916 = lean_ctor_get(x_1, 16); +x_917 = lean_ctor_get(x_1, 17); +x_918 = lean_ctor_get(x_1, 18); +x_919 = lean_ctor_get(x_1, 19); +x_920 = lean_ctor_get(x_1, 20); +x_921 = lean_ctor_get(x_1, 21); +x_922 = lean_ctor_get(x_1, 22); +x_923 = lean_ctor_get(x_1, 23); +x_924 = lean_ctor_get(x_1, 24); +x_925 = lean_ctor_get(x_1, 25); +x_926 = lean_ctor_get(x_1, 26); +x_927 = lean_ctor_get(x_1, 27); +x_928 = lean_ctor_get(x_1, 28); +x_929 = lean_ctor_get(x_1, 29); +lean_inc(x_929); +lean_inc(x_928); +lean_inc(x_927); +lean_inc(x_926); +lean_inc(x_925); +lean_inc(x_924); +lean_inc(x_923); +lean_inc(x_922); +lean_inc(x_921); +lean_inc(x_920); +lean_inc(x_919); +lean_inc(x_918); +lean_inc(x_917); +lean_inc(x_916); +lean_inc(x_915); +lean_inc(x_914); +lean_inc(x_913); +lean_inc(x_912); +lean_inc(x_911); +lean_inc(x_910); +lean_inc(x_909); +lean_inc(x_908); +lean_inc(x_907); +lean_inc(x_906); +lean_inc(x_905); +lean_inc(x_904); +lean_inc(x_903); +lean_inc(x_902); +lean_inc(x_901); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_930 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_930, 0, x_901); +lean_ctor_set(x_930, 1, x_902); +lean_ctor_set(x_930, 2, x_903); +lean_ctor_set(x_930, 3, x_904); +lean_ctor_set(x_930, 4, x_905); +lean_ctor_set(x_930, 5, x_906); +lean_ctor_set(x_930, 6, x_907); +lean_ctor_set(x_930, 7, x_908); +lean_ctor_set(x_930, 8, x_909); +lean_ctor_set(x_930, 9, x_910); +lean_ctor_set(x_930, 10, x_911); +lean_ctor_set(x_930, 11, x_912); +lean_ctor_set(x_930, 12, x_913); +lean_ctor_set(x_930, 13, x_914); +lean_ctor_set(x_930, 14, x_915); +lean_ctor_set(x_930, 15, x_2); +lean_ctor_set(x_930, 16, x_916); +lean_ctor_set(x_930, 17, x_917); +lean_ctor_set(x_930, 18, x_918); +lean_ctor_set(x_930, 19, x_919); +lean_ctor_set(x_930, 20, x_920); +lean_ctor_set(x_930, 21, x_921); +lean_ctor_set(x_930, 22, x_922); +lean_ctor_set(x_930, 23, x_923); +lean_ctor_set(x_930, 24, x_924); +lean_ctor_set(x_930, 25, x_925); +lean_ctor_set(x_930, 26, x_926); +lean_ctor_set(x_930, 27, x_927); +lean_ctor_set(x_930, 28, x_928); +lean_ctor_set(x_930, 29, x_929); +return x_930; +} +} +else +{ +lean_object* x_931; lean_object* x_932; lean_object* x_933; lean_object* x_934; lean_object* x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; lean_object* x_949; lean_object* x_950; lean_object* x_951; lean_object* x_952; lean_object* x_953; lean_object* x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; +lean_dec(x_2); +x_931 = lean_ctor_get(x_1, 0); +lean_inc(x_931); +x_932 = lean_ctor_get(x_1, 1); +lean_inc(x_932); +x_933 = lean_ctor_get(x_1, 2); +lean_inc(x_933); +x_934 = lean_ctor_get(x_1, 3); +lean_inc(x_934); +x_935 = lean_ctor_get(x_1, 4); +lean_inc(x_935); +x_936 = lean_ctor_get(x_1, 5); +lean_inc(x_936); +x_937 = lean_ctor_get(x_1, 6); +lean_inc(x_937); +x_938 = lean_ctor_get(x_1, 7); +lean_inc(x_938); +x_939 = lean_ctor_get(x_1, 8); +lean_inc(x_939); +x_940 = lean_ctor_get(x_1, 9); +lean_inc(x_940); +x_941 = lean_ctor_get(x_1, 10); +lean_inc(x_941); +x_942 = lean_ctor_get(x_1, 11); +lean_inc(x_942); +x_943 = lean_ctor_get(x_1, 12); +lean_inc(x_943); +x_944 = lean_ctor_get(x_1, 13); +lean_inc(x_944); +x_945 = lean_ctor_get(x_1, 14); +lean_inc(x_945); +x_946 = lean_ctor_get(x_1, 16); +lean_inc(x_946); +x_947 = lean_ctor_get(x_1, 17); +lean_inc(x_947); +x_948 = lean_ctor_get(x_1, 18); +lean_inc(x_948); +x_949 = lean_ctor_get(x_1, 19); +lean_inc(x_949); +x_950 = lean_ctor_get(x_1, 20); +lean_inc(x_950); +x_951 = lean_ctor_get(x_1, 21); +lean_inc(x_951); +x_952 = lean_ctor_get(x_1, 22); +lean_inc(x_952); +x_953 = lean_ctor_get(x_1, 23); +lean_inc(x_953); +x_954 = lean_ctor_get(x_1, 24); +lean_inc(x_954); +x_955 = lean_ctor_get(x_1, 25); +lean_inc(x_955); +x_956 = lean_ctor_get(x_1, 26); +lean_inc(x_956); +x_957 = lean_ctor_get(x_1, 27); +lean_inc(x_957); +x_958 = lean_ctor_get(x_1, 28); +lean_inc(x_958); +x_959 = lean_ctor_get(x_1, 29); +lean_inc(x_959); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_960 = x_1; +} else { + lean_dec_ref(x_1); + x_960 = lean_box(0); +} +x_961 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_961, 0, x_3); +if (lean_is_scalar(x_960)) { + x_962 = lean_alloc_ctor(0, 30, 0); +} else { + x_962 = x_960; +} +lean_ctor_set(x_962, 0, x_931); +lean_ctor_set(x_962, 1, x_932); +lean_ctor_set(x_962, 2, x_933); +lean_ctor_set(x_962, 3, x_934); +lean_ctor_set(x_962, 4, x_935); +lean_ctor_set(x_962, 5, x_936); +lean_ctor_set(x_962, 6, x_937); +lean_ctor_set(x_962, 7, x_938); +lean_ctor_set(x_962, 8, x_939); +lean_ctor_set(x_962, 9, x_940); +lean_ctor_set(x_962, 10, x_941); +lean_ctor_set(x_962, 11, x_942); +lean_ctor_set(x_962, 12, x_943); +lean_ctor_set(x_962, 13, x_944); +lean_ctor_set(x_962, 14, x_945); +lean_ctor_set(x_962, 15, x_961); +lean_ctor_set(x_962, 16, x_946); +lean_ctor_set(x_962, 17, x_947); +lean_ctor_set(x_962, 18, x_948); +lean_ctor_set(x_962, 19, x_949); +lean_ctor_set(x_962, 20, x_950); +lean_ctor_set(x_962, 21, x_951); +lean_ctor_set(x_962, 22, x_952); +lean_ctor_set(x_962, 23, x_953); +lean_ctor_set(x_962, 24, x_954); +lean_ctor_set(x_962, 25, x_955); +lean_ctor_set(x_962, 26, x_956); +lean_ctor_set(x_962, 27, x_957); +lean_ctor_set(x_962, 28, x_958); +lean_ctor_set(x_962, 29, x_959); +return x_962; +} +} +case 16: +{ +uint8_t x_963; +x_963 = !lean_is_exclusive(x_2); +if (x_963 == 0) +{ +lean_object* x_964; uint8_t x_965; +x_964 = lean_ctor_get(x_2, 0); +lean_dec(x_964); +x_965 = !lean_is_exclusive(x_1); +if (x_965 == 0) +{ +lean_object* x_966; +x_966 = lean_ctor_get(x_1, 16); +lean_dec(x_966); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 16, x_2); +return x_1; +} +else +{ +lean_object* x_967; lean_object* x_968; lean_object* x_969; lean_object* x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; lean_object* x_984; lean_object* x_985; lean_object* x_986; lean_object* x_987; lean_object* x_988; lean_object* x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; +x_967 = lean_ctor_get(x_1, 0); +x_968 = lean_ctor_get(x_1, 1); +x_969 = lean_ctor_get(x_1, 2); +x_970 = lean_ctor_get(x_1, 3); +x_971 = lean_ctor_get(x_1, 4); +x_972 = lean_ctor_get(x_1, 5); +x_973 = lean_ctor_get(x_1, 6); +x_974 = lean_ctor_get(x_1, 7); +x_975 = lean_ctor_get(x_1, 8); +x_976 = lean_ctor_get(x_1, 9); +x_977 = lean_ctor_get(x_1, 10); +x_978 = lean_ctor_get(x_1, 11); +x_979 = lean_ctor_get(x_1, 12); +x_980 = lean_ctor_get(x_1, 13); +x_981 = lean_ctor_get(x_1, 14); +x_982 = lean_ctor_get(x_1, 15); +x_983 = lean_ctor_get(x_1, 17); +x_984 = lean_ctor_get(x_1, 18); +x_985 = lean_ctor_get(x_1, 19); +x_986 = lean_ctor_get(x_1, 20); +x_987 = lean_ctor_get(x_1, 21); +x_988 = lean_ctor_get(x_1, 22); +x_989 = lean_ctor_get(x_1, 23); +x_990 = lean_ctor_get(x_1, 24); +x_991 = lean_ctor_get(x_1, 25); +x_992 = lean_ctor_get(x_1, 26); +x_993 = lean_ctor_get(x_1, 27); +x_994 = lean_ctor_get(x_1, 28); +x_995 = lean_ctor_get(x_1, 29); +lean_inc(x_995); +lean_inc(x_994); +lean_inc(x_993); +lean_inc(x_992); +lean_inc(x_991); +lean_inc(x_990); +lean_inc(x_989); +lean_inc(x_988); +lean_inc(x_987); +lean_inc(x_986); +lean_inc(x_985); +lean_inc(x_984); +lean_inc(x_983); +lean_inc(x_982); +lean_inc(x_981); +lean_inc(x_980); +lean_inc(x_979); +lean_inc(x_978); +lean_inc(x_977); +lean_inc(x_976); +lean_inc(x_975); +lean_inc(x_974); +lean_inc(x_973); +lean_inc(x_972); +lean_inc(x_971); +lean_inc(x_970); +lean_inc(x_969); +lean_inc(x_968); +lean_inc(x_967); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_996 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_996, 0, x_967); +lean_ctor_set(x_996, 1, x_968); +lean_ctor_set(x_996, 2, x_969); +lean_ctor_set(x_996, 3, x_970); +lean_ctor_set(x_996, 4, x_971); +lean_ctor_set(x_996, 5, x_972); +lean_ctor_set(x_996, 6, x_973); +lean_ctor_set(x_996, 7, x_974); +lean_ctor_set(x_996, 8, x_975); +lean_ctor_set(x_996, 9, x_976); +lean_ctor_set(x_996, 10, x_977); +lean_ctor_set(x_996, 11, x_978); +lean_ctor_set(x_996, 12, x_979); +lean_ctor_set(x_996, 13, x_980); +lean_ctor_set(x_996, 14, x_981); +lean_ctor_set(x_996, 15, x_982); +lean_ctor_set(x_996, 16, x_2); +lean_ctor_set(x_996, 17, x_983); +lean_ctor_set(x_996, 18, x_984); +lean_ctor_set(x_996, 19, x_985); +lean_ctor_set(x_996, 20, x_986); +lean_ctor_set(x_996, 21, x_987); +lean_ctor_set(x_996, 22, x_988); +lean_ctor_set(x_996, 23, x_989); +lean_ctor_set(x_996, 24, x_990); +lean_ctor_set(x_996, 25, x_991); +lean_ctor_set(x_996, 26, x_992); +lean_ctor_set(x_996, 27, x_993); +lean_ctor_set(x_996, 28, x_994); +lean_ctor_set(x_996, 29, x_995); +return x_996; +} +} +else +{ +lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; lean_object* x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; lean_object* x_1019; lean_object* x_1020; lean_object* x_1021; lean_object* x_1022; lean_object* x_1023; lean_object* x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; +lean_dec(x_2); +x_997 = lean_ctor_get(x_1, 0); +lean_inc(x_997); +x_998 = lean_ctor_get(x_1, 1); +lean_inc(x_998); +x_999 = lean_ctor_get(x_1, 2); +lean_inc(x_999); +x_1000 = lean_ctor_get(x_1, 3); +lean_inc(x_1000); +x_1001 = lean_ctor_get(x_1, 4); +lean_inc(x_1001); +x_1002 = lean_ctor_get(x_1, 5); +lean_inc(x_1002); +x_1003 = lean_ctor_get(x_1, 6); +lean_inc(x_1003); +x_1004 = lean_ctor_get(x_1, 7); +lean_inc(x_1004); +x_1005 = lean_ctor_get(x_1, 8); +lean_inc(x_1005); +x_1006 = lean_ctor_get(x_1, 9); +lean_inc(x_1006); +x_1007 = lean_ctor_get(x_1, 10); +lean_inc(x_1007); +x_1008 = lean_ctor_get(x_1, 11); +lean_inc(x_1008); +x_1009 = lean_ctor_get(x_1, 12); +lean_inc(x_1009); +x_1010 = lean_ctor_get(x_1, 13); +lean_inc(x_1010); +x_1011 = lean_ctor_get(x_1, 14); +lean_inc(x_1011); +x_1012 = lean_ctor_get(x_1, 15); +lean_inc(x_1012); +x_1013 = lean_ctor_get(x_1, 17); +lean_inc(x_1013); +x_1014 = lean_ctor_get(x_1, 18); +lean_inc(x_1014); +x_1015 = lean_ctor_get(x_1, 19); +lean_inc(x_1015); +x_1016 = lean_ctor_get(x_1, 20); +lean_inc(x_1016); +x_1017 = lean_ctor_get(x_1, 21); +lean_inc(x_1017); +x_1018 = lean_ctor_get(x_1, 22); +lean_inc(x_1018); +x_1019 = lean_ctor_get(x_1, 23); +lean_inc(x_1019); +x_1020 = lean_ctor_get(x_1, 24); +lean_inc(x_1020); +x_1021 = lean_ctor_get(x_1, 25); +lean_inc(x_1021); +x_1022 = lean_ctor_get(x_1, 26); +lean_inc(x_1022); +x_1023 = lean_ctor_get(x_1, 27); +lean_inc(x_1023); +x_1024 = lean_ctor_get(x_1, 28); +lean_inc(x_1024); +x_1025 = lean_ctor_get(x_1, 29); +lean_inc(x_1025); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1026 = x_1; +} else { + lean_dec_ref(x_1); + x_1026 = lean_box(0); +} +x_1027 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1027, 0, x_3); +if (lean_is_scalar(x_1026)) { + x_1028 = lean_alloc_ctor(0, 30, 0); +} else { + x_1028 = x_1026; +} +lean_ctor_set(x_1028, 0, x_997); +lean_ctor_set(x_1028, 1, x_998); +lean_ctor_set(x_1028, 2, x_999); +lean_ctor_set(x_1028, 3, x_1000); +lean_ctor_set(x_1028, 4, x_1001); +lean_ctor_set(x_1028, 5, x_1002); +lean_ctor_set(x_1028, 6, x_1003); +lean_ctor_set(x_1028, 7, x_1004); +lean_ctor_set(x_1028, 8, x_1005); +lean_ctor_set(x_1028, 9, x_1006); +lean_ctor_set(x_1028, 10, x_1007); +lean_ctor_set(x_1028, 11, x_1008); +lean_ctor_set(x_1028, 12, x_1009); +lean_ctor_set(x_1028, 13, x_1010); +lean_ctor_set(x_1028, 14, x_1011); +lean_ctor_set(x_1028, 15, x_1012); +lean_ctor_set(x_1028, 16, x_1027); +lean_ctor_set(x_1028, 17, x_1013); +lean_ctor_set(x_1028, 18, x_1014); +lean_ctor_set(x_1028, 19, x_1015); +lean_ctor_set(x_1028, 20, x_1016); +lean_ctor_set(x_1028, 21, x_1017); +lean_ctor_set(x_1028, 22, x_1018); +lean_ctor_set(x_1028, 23, x_1019); +lean_ctor_set(x_1028, 24, x_1020); +lean_ctor_set(x_1028, 25, x_1021); +lean_ctor_set(x_1028, 26, x_1022); +lean_ctor_set(x_1028, 27, x_1023); +lean_ctor_set(x_1028, 28, x_1024); +lean_ctor_set(x_1028, 29, x_1025); +return x_1028; +} +} +case 17: +{ +uint8_t x_1029; +x_1029 = !lean_is_exclusive(x_2); +if (x_1029 == 0) +{ +lean_object* x_1030; uint8_t x_1031; +x_1030 = lean_ctor_get(x_2, 0); +lean_dec(x_1030); +x_1031 = !lean_is_exclusive(x_1); +if (x_1031 == 0) +{ +lean_object* x_1032; +x_1032 = lean_ctor_get(x_1, 17); +lean_dec(x_1032); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 17, x_2); +return x_1; +} +else +{ +lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; lean_object* x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; lean_object* x_1054; lean_object* x_1055; lean_object* x_1056; lean_object* x_1057; lean_object* x_1058; lean_object* x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; +x_1033 = lean_ctor_get(x_1, 0); +x_1034 = lean_ctor_get(x_1, 1); +x_1035 = lean_ctor_get(x_1, 2); +x_1036 = lean_ctor_get(x_1, 3); +x_1037 = lean_ctor_get(x_1, 4); +x_1038 = lean_ctor_get(x_1, 5); +x_1039 = lean_ctor_get(x_1, 6); +x_1040 = lean_ctor_get(x_1, 7); +x_1041 = lean_ctor_get(x_1, 8); +x_1042 = lean_ctor_get(x_1, 9); +x_1043 = lean_ctor_get(x_1, 10); +x_1044 = lean_ctor_get(x_1, 11); +x_1045 = lean_ctor_get(x_1, 12); +x_1046 = lean_ctor_get(x_1, 13); +x_1047 = lean_ctor_get(x_1, 14); +x_1048 = lean_ctor_get(x_1, 15); +x_1049 = lean_ctor_get(x_1, 16); +x_1050 = lean_ctor_get(x_1, 18); +x_1051 = lean_ctor_get(x_1, 19); +x_1052 = lean_ctor_get(x_1, 20); +x_1053 = lean_ctor_get(x_1, 21); +x_1054 = lean_ctor_get(x_1, 22); +x_1055 = lean_ctor_get(x_1, 23); +x_1056 = lean_ctor_get(x_1, 24); +x_1057 = lean_ctor_get(x_1, 25); +x_1058 = lean_ctor_get(x_1, 26); +x_1059 = lean_ctor_get(x_1, 27); +x_1060 = lean_ctor_get(x_1, 28); +x_1061 = lean_ctor_get(x_1, 29); +lean_inc(x_1061); +lean_inc(x_1060); +lean_inc(x_1059); +lean_inc(x_1058); +lean_inc(x_1057); +lean_inc(x_1056); +lean_inc(x_1055); +lean_inc(x_1054); +lean_inc(x_1053); +lean_inc(x_1052); +lean_inc(x_1051); +lean_inc(x_1050); +lean_inc(x_1049); +lean_inc(x_1048); +lean_inc(x_1047); +lean_inc(x_1046); +lean_inc(x_1045); +lean_inc(x_1044); +lean_inc(x_1043); +lean_inc(x_1042); +lean_inc(x_1041); +lean_inc(x_1040); +lean_inc(x_1039); +lean_inc(x_1038); +lean_inc(x_1037); +lean_inc(x_1036); +lean_inc(x_1035); +lean_inc(x_1034); +lean_inc(x_1033); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1062 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1062, 0, x_1033); +lean_ctor_set(x_1062, 1, x_1034); +lean_ctor_set(x_1062, 2, x_1035); +lean_ctor_set(x_1062, 3, x_1036); +lean_ctor_set(x_1062, 4, x_1037); +lean_ctor_set(x_1062, 5, x_1038); +lean_ctor_set(x_1062, 6, x_1039); +lean_ctor_set(x_1062, 7, x_1040); +lean_ctor_set(x_1062, 8, x_1041); +lean_ctor_set(x_1062, 9, x_1042); +lean_ctor_set(x_1062, 10, x_1043); +lean_ctor_set(x_1062, 11, x_1044); +lean_ctor_set(x_1062, 12, x_1045); +lean_ctor_set(x_1062, 13, x_1046); +lean_ctor_set(x_1062, 14, x_1047); +lean_ctor_set(x_1062, 15, x_1048); +lean_ctor_set(x_1062, 16, x_1049); +lean_ctor_set(x_1062, 17, x_2); +lean_ctor_set(x_1062, 18, x_1050); +lean_ctor_set(x_1062, 19, x_1051); +lean_ctor_set(x_1062, 20, x_1052); +lean_ctor_set(x_1062, 21, x_1053); +lean_ctor_set(x_1062, 22, x_1054); +lean_ctor_set(x_1062, 23, x_1055); +lean_ctor_set(x_1062, 24, x_1056); +lean_ctor_set(x_1062, 25, x_1057); +lean_ctor_set(x_1062, 26, x_1058); +lean_ctor_set(x_1062, 27, x_1059); +lean_ctor_set(x_1062, 28, x_1060); +lean_ctor_set(x_1062, 29, x_1061); +return x_1062; +} +} +else +{ +lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; lean_object* x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; lean_object* x_1089; lean_object* x_1090; lean_object* x_1091; lean_object* x_1092; lean_object* x_1093; lean_object* x_1094; +lean_dec(x_2); +x_1063 = lean_ctor_get(x_1, 0); +lean_inc(x_1063); +x_1064 = lean_ctor_get(x_1, 1); +lean_inc(x_1064); +x_1065 = lean_ctor_get(x_1, 2); +lean_inc(x_1065); +x_1066 = lean_ctor_get(x_1, 3); +lean_inc(x_1066); +x_1067 = lean_ctor_get(x_1, 4); +lean_inc(x_1067); +x_1068 = lean_ctor_get(x_1, 5); +lean_inc(x_1068); +x_1069 = lean_ctor_get(x_1, 6); +lean_inc(x_1069); +x_1070 = lean_ctor_get(x_1, 7); +lean_inc(x_1070); +x_1071 = lean_ctor_get(x_1, 8); +lean_inc(x_1071); +x_1072 = lean_ctor_get(x_1, 9); +lean_inc(x_1072); +x_1073 = lean_ctor_get(x_1, 10); +lean_inc(x_1073); +x_1074 = lean_ctor_get(x_1, 11); +lean_inc(x_1074); +x_1075 = lean_ctor_get(x_1, 12); +lean_inc(x_1075); +x_1076 = lean_ctor_get(x_1, 13); +lean_inc(x_1076); +x_1077 = lean_ctor_get(x_1, 14); +lean_inc(x_1077); +x_1078 = lean_ctor_get(x_1, 15); +lean_inc(x_1078); +x_1079 = lean_ctor_get(x_1, 16); +lean_inc(x_1079); +x_1080 = lean_ctor_get(x_1, 18); +lean_inc(x_1080); +x_1081 = lean_ctor_get(x_1, 19); +lean_inc(x_1081); +x_1082 = lean_ctor_get(x_1, 20); +lean_inc(x_1082); +x_1083 = lean_ctor_get(x_1, 21); +lean_inc(x_1083); +x_1084 = lean_ctor_get(x_1, 22); +lean_inc(x_1084); +x_1085 = lean_ctor_get(x_1, 23); +lean_inc(x_1085); +x_1086 = lean_ctor_get(x_1, 24); +lean_inc(x_1086); +x_1087 = lean_ctor_get(x_1, 25); +lean_inc(x_1087); +x_1088 = lean_ctor_get(x_1, 26); +lean_inc(x_1088); +x_1089 = lean_ctor_get(x_1, 27); +lean_inc(x_1089); +x_1090 = lean_ctor_get(x_1, 28); +lean_inc(x_1090); +x_1091 = lean_ctor_get(x_1, 29); +lean_inc(x_1091); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1092 = x_1; +} else { + lean_dec_ref(x_1); + x_1092 = lean_box(0); +} +x_1093 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1093, 0, x_3); +if (lean_is_scalar(x_1092)) { + x_1094 = lean_alloc_ctor(0, 30, 0); +} else { + x_1094 = x_1092; +} +lean_ctor_set(x_1094, 0, x_1063); +lean_ctor_set(x_1094, 1, x_1064); +lean_ctor_set(x_1094, 2, x_1065); +lean_ctor_set(x_1094, 3, x_1066); +lean_ctor_set(x_1094, 4, x_1067); +lean_ctor_set(x_1094, 5, x_1068); +lean_ctor_set(x_1094, 6, x_1069); +lean_ctor_set(x_1094, 7, x_1070); +lean_ctor_set(x_1094, 8, x_1071); +lean_ctor_set(x_1094, 9, x_1072); +lean_ctor_set(x_1094, 10, x_1073); +lean_ctor_set(x_1094, 11, x_1074); +lean_ctor_set(x_1094, 12, x_1075); +lean_ctor_set(x_1094, 13, x_1076); +lean_ctor_set(x_1094, 14, x_1077); +lean_ctor_set(x_1094, 15, x_1078); +lean_ctor_set(x_1094, 16, x_1079); +lean_ctor_set(x_1094, 17, x_1093); +lean_ctor_set(x_1094, 18, x_1080); +lean_ctor_set(x_1094, 19, x_1081); +lean_ctor_set(x_1094, 20, x_1082); +lean_ctor_set(x_1094, 21, x_1083); +lean_ctor_set(x_1094, 22, x_1084); +lean_ctor_set(x_1094, 23, x_1085); +lean_ctor_set(x_1094, 24, x_1086); +lean_ctor_set(x_1094, 25, x_1087); +lean_ctor_set(x_1094, 26, x_1088); +lean_ctor_set(x_1094, 27, x_1089); +lean_ctor_set(x_1094, 28, x_1090); +lean_ctor_set(x_1094, 29, x_1091); +return x_1094; +} +} +case 18: +{ +uint8_t x_1095; +x_1095 = !lean_is_exclusive(x_2); +if (x_1095 == 0) +{ +lean_object* x_1096; uint8_t x_1097; +x_1096 = lean_ctor_get(x_2, 0); +lean_dec(x_1096); +x_1097 = !lean_is_exclusive(x_1); +if (x_1097 == 0) +{ +lean_object* x_1098; +x_1098 = lean_ctor_get(x_1, 18); +lean_dec(x_1098); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 18, x_2); +return x_1; +} +else +{ +lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; lean_object* x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; lean_object* x_1124; lean_object* x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; +x_1099 = lean_ctor_get(x_1, 0); +x_1100 = lean_ctor_get(x_1, 1); +x_1101 = lean_ctor_get(x_1, 2); +x_1102 = lean_ctor_get(x_1, 3); +x_1103 = lean_ctor_get(x_1, 4); +x_1104 = lean_ctor_get(x_1, 5); +x_1105 = lean_ctor_get(x_1, 6); +x_1106 = lean_ctor_get(x_1, 7); +x_1107 = lean_ctor_get(x_1, 8); +x_1108 = lean_ctor_get(x_1, 9); +x_1109 = lean_ctor_get(x_1, 10); +x_1110 = lean_ctor_get(x_1, 11); +x_1111 = lean_ctor_get(x_1, 12); +x_1112 = lean_ctor_get(x_1, 13); +x_1113 = lean_ctor_get(x_1, 14); +x_1114 = lean_ctor_get(x_1, 15); +x_1115 = lean_ctor_get(x_1, 16); +x_1116 = lean_ctor_get(x_1, 17); +x_1117 = lean_ctor_get(x_1, 19); +x_1118 = lean_ctor_get(x_1, 20); +x_1119 = lean_ctor_get(x_1, 21); +x_1120 = lean_ctor_get(x_1, 22); +x_1121 = lean_ctor_get(x_1, 23); +x_1122 = lean_ctor_get(x_1, 24); +x_1123 = lean_ctor_get(x_1, 25); +x_1124 = lean_ctor_get(x_1, 26); +x_1125 = lean_ctor_get(x_1, 27); +x_1126 = lean_ctor_get(x_1, 28); +x_1127 = lean_ctor_get(x_1, 29); +lean_inc(x_1127); +lean_inc(x_1126); +lean_inc(x_1125); +lean_inc(x_1124); +lean_inc(x_1123); +lean_inc(x_1122); +lean_inc(x_1121); +lean_inc(x_1120); +lean_inc(x_1119); +lean_inc(x_1118); +lean_inc(x_1117); +lean_inc(x_1116); +lean_inc(x_1115); +lean_inc(x_1114); +lean_inc(x_1113); +lean_inc(x_1112); +lean_inc(x_1111); +lean_inc(x_1110); +lean_inc(x_1109); +lean_inc(x_1108); +lean_inc(x_1107); +lean_inc(x_1106); +lean_inc(x_1105); +lean_inc(x_1104); +lean_inc(x_1103); +lean_inc(x_1102); +lean_inc(x_1101); +lean_inc(x_1100); +lean_inc(x_1099); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1128 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1128, 0, x_1099); +lean_ctor_set(x_1128, 1, x_1100); +lean_ctor_set(x_1128, 2, x_1101); +lean_ctor_set(x_1128, 3, x_1102); +lean_ctor_set(x_1128, 4, x_1103); +lean_ctor_set(x_1128, 5, x_1104); +lean_ctor_set(x_1128, 6, x_1105); +lean_ctor_set(x_1128, 7, x_1106); +lean_ctor_set(x_1128, 8, x_1107); +lean_ctor_set(x_1128, 9, x_1108); +lean_ctor_set(x_1128, 10, x_1109); +lean_ctor_set(x_1128, 11, x_1110); +lean_ctor_set(x_1128, 12, x_1111); +lean_ctor_set(x_1128, 13, x_1112); +lean_ctor_set(x_1128, 14, x_1113); +lean_ctor_set(x_1128, 15, x_1114); +lean_ctor_set(x_1128, 16, x_1115); +lean_ctor_set(x_1128, 17, x_1116); +lean_ctor_set(x_1128, 18, x_2); +lean_ctor_set(x_1128, 19, x_1117); +lean_ctor_set(x_1128, 20, x_1118); +lean_ctor_set(x_1128, 21, x_1119); +lean_ctor_set(x_1128, 22, x_1120); +lean_ctor_set(x_1128, 23, x_1121); +lean_ctor_set(x_1128, 24, x_1122); +lean_ctor_set(x_1128, 25, x_1123); +lean_ctor_set(x_1128, 26, x_1124); +lean_ctor_set(x_1128, 27, x_1125); +lean_ctor_set(x_1128, 28, x_1126); +lean_ctor_set(x_1128, 29, x_1127); +return x_1128; +} +} +else +{ +lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; lean_object* x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; lean_object* x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; +lean_dec(x_2); +x_1129 = lean_ctor_get(x_1, 0); +lean_inc(x_1129); +x_1130 = lean_ctor_get(x_1, 1); +lean_inc(x_1130); +x_1131 = lean_ctor_get(x_1, 2); +lean_inc(x_1131); +x_1132 = lean_ctor_get(x_1, 3); +lean_inc(x_1132); +x_1133 = lean_ctor_get(x_1, 4); +lean_inc(x_1133); +x_1134 = lean_ctor_get(x_1, 5); +lean_inc(x_1134); +x_1135 = lean_ctor_get(x_1, 6); +lean_inc(x_1135); +x_1136 = lean_ctor_get(x_1, 7); +lean_inc(x_1136); +x_1137 = lean_ctor_get(x_1, 8); +lean_inc(x_1137); +x_1138 = lean_ctor_get(x_1, 9); +lean_inc(x_1138); +x_1139 = lean_ctor_get(x_1, 10); +lean_inc(x_1139); +x_1140 = lean_ctor_get(x_1, 11); +lean_inc(x_1140); +x_1141 = lean_ctor_get(x_1, 12); +lean_inc(x_1141); +x_1142 = lean_ctor_get(x_1, 13); +lean_inc(x_1142); +x_1143 = lean_ctor_get(x_1, 14); +lean_inc(x_1143); +x_1144 = lean_ctor_get(x_1, 15); +lean_inc(x_1144); +x_1145 = lean_ctor_get(x_1, 16); +lean_inc(x_1145); +x_1146 = lean_ctor_get(x_1, 17); +lean_inc(x_1146); +x_1147 = lean_ctor_get(x_1, 19); +lean_inc(x_1147); +x_1148 = lean_ctor_get(x_1, 20); +lean_inc(x_1148); +x_1149 = lean_ctor_get(x_1, 21); +lean_inc(x_1149); +x_1150 = lean_ctor_get(x_1, 22); +lean_inc(x_1150); +x_1151 = lean_ctor_get(x_1, 23); +lean_inc(x_1151); +x_1152 = lean_ctor_get(x_1, 24); +lean_inc(x_1152); +x_1153 = lean_ctor_get(x_1, 25); +lean_inc(x_1153); +x_1154 = lean_ctor_get(x_1, 26); +lean_inc(x_1154); +x_1155 = lean_ctor_get(x_1, 27); +lean_inc(x_1155); +x_1156 = lean_ctor_get(x_1, 28); +lean_inc(x_1156); +x_1157 = lean_ctor_get(x_1, 29); +lean_inc(x_1157); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1158 = x_1; +} else { + lean_dec_ref(x_1); + x_1158 = lean_box(0); +} +x_1159 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1159, 0, x_3); +if (lean_is_scalar(x_1158)) { + x_1160 = lean_alloc_ctor(0, 30, 0); +} else { + x_1160 = x_1158; +} +lean_ctor_set(x_1160, 0, x_1129); +lean_ctor_set(x_1160, 1, x_1130); +lean_ctor_set(x_1160, 2, x_1131); +lean_ctor_set(x_1160, 3, x_1132); +lean_ctor_set(x_1160, 4, x_1133); +lean_ctor_set(x_1160, 5, x_1134); +lean_ctor_set(x_1160, 6, x_1135); +lean_ctor_set(x_1160, 7, x_1136); +lean_ctor_set(x_1160, 8, x_1137); +lean_ctor_set(x_1160, 9, x_1138); +lean_ctor_set(x_1160, 10, x_1139); +lean_ctor_set(x_1160, 11, x_1140); +lean_ctor_set(x_1160, 12, x_1141); +lean_ctor_set(x_1160, 13, x_1142); +lean_ctor_set(x_1160, 14, x_1143); +lean_ctor_set(x_1160, 15, x_1144); +lean_ctor_set(x_1160, 16, x_1145); +lean_ctor_set(x_1160, 17, x_1146); +lean_ctor_set(x_1160, 18, x_1159); +lean_ctor_set(x_1160, 19, x_1147); +lean_ctor_set(x_1160, 20, x_1148); +lean_ctor_set(x_1160, 21, x_1149); +lean_ctor_set(x_1160, 22, x_1150); +lean_ctor_set(x_1160, 23, x_1151); +lean_ctor_set(x_1160, 24, x_1152); +lean_ctor_set(x_1160, 25, x_1153); +lean_ctor_set(x_1160, 26, x_1154); +lean_ctor_set(x_1160, 27, x_1155); +lean_ctor_set(x_1160, 28, x_1156); +lean_ctor_set(x_1160, 29, x_1157); +return x_1160; +} +} +case 19: +{ +uint8_t x_1161; +x_1161 = !lean_is_exclusive(x_2); +if (x_1161 == 0) +{ +lean_object* x_1162; uint8_t x_1163; +x_1162 = lean_ctor_get(x_2, 0); +lean_dec(x_1162); +x_1163 = !lean_is_exclusive(x_1); +if (x_1163 == 0) +{ +lean_object* x_1164; +x_1164 = lean_ctor_get(x_1, 19); +lean_dec(x_1164); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 19, x_2); +return x_1; +} +else +{ +lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; lean_object* x_1170; lean_object* x_1171; lean_object* x_1172; lean_object* x_1173; lean_object* x_1174; lean_object* x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; lean_object* x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; +x_1165 = lean_ctor_get(x_1, 0); +x_1166 = lean_ctor_get(x_1, 1); +x_1167 = lean_ctor_get(x_1, 2); +x_1168 = lean_ctor_get(x_1, 3); +x_1169 = lean_ctor_get(x_1, 4); +x_1170 = lean_ctor_get(x_1, 5); +x_1171 = lean_ctor_get(x_1, 6); +x_1172 = lean_ctor_get(x_1, 7); +x_1173 = lean_ctor_get(x_1, 8); +x_1174 = lean_ctor_get(x_1, 9); +x_1175 = lean_ctor_get(x_1, 10); +x_1176 = lean_ctor_get(x_1, 11); +x_1177 = lean_ctor_get(x_1, 12); +x_1178 = lean_ctor_get(x_1, 13); +x_1179 = lean_ctor_get(x_1, 14); +x_1180 = lean_ctor_get(x_1, 15); +x_1181 = lean_ctor_get(x_1, 16); +x_1182 = lean_ctor_get(x_1, 17); +x_1183 = lean_ctor_get(x_1, 18); +x_1184 = lean_ctor_get(x_1, 20); +x_1185 = lean_ctor_get(x_1, 21); +x_1186 = lean_ctor_get(x_1, 22); +x_1187 = lean_ctor_get(x_1, 23); +x_1188 = lean_ctor_get(x_1, 24); +x_1189 = lean_ctor_get(x_1, 25); +x_1190 = lean_ctor_get(x_1, 26); +x_1191 = lean_ctor_get(x_1, 27); +x_1192 = lean_ctor_get(x_1, 28); +x_1193 = lean_ctor_get(x_1, 29); +lean_inc(x_1193); +lean_inc(x_1192); +lean_inc(x_1191); +lean_inc(x_1190); +lean_inc(x_1189); +lean_inc(x_1188); +lean_inc(x_1187); +lean_inc(x_1186); +lean_inc(x_1185); +lean_inc(x_1184); +lean_inc(x_1183); +lean_inc(x_1182); +lean_inc(x_1181); +lean_inc(x_1180); +lean_inc(x_1179); +lean_inc(x_1178); +lean_inc(x_1177); +lean_inc(x_1176); +lean_inc(x_1175); +lean_inc(x_1174); +lean_inc(x_1173); +lean_inc(x_1172); +lean_inc(x_1171); +lean_inc(x_1170); +lean_inc(x_1169); +lean_inc(x_1168); +lean_inc(x_1167); +lean_inc(x_1166); +lean_inc(x_1165); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1194 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1194, 0, x_1165); +lean_ctor_set(x_1194, 1, x_1166); +lean_ctor_set(x_1194, 2, x_1167); +lean_ctor_set(x_1194, 3, x_1168); +lean_ctor_set(x_1194, 4, x_1169); +lean_ctor_set(x_1194, 5, x_1170); +lean_ctor_set(x_1194, 6, x_1171); +lean_ctor_set(x_1194, 7, x_1172); +lean_ctor_set(x_1194, 8, x_1173); +lean_ctor_set(x_1194, 9, x_1174); +lean_ctor_set(x_1194, 10, x_1175); +lean_ctor_set(x_1194, 11, x_1176); +lean_ctor_set(x_1194, 12, x_1177); +lean_ctor_set(x_1194, 13, x_1178); +lean_ctor_set(x_1194, 14, x_1179); +lean_ctor_set(x_1194, 15, x_1180); +lean_ctor_set(x_1194, 16, x_1181); +lean_ctor_set(x_1194, 17, x_1182); +lean_ctor_set(x_1194, 18, x_1183); +lean_ctor_set(x_1194, 19, x_2); +lean_ctor_set(x_1194, 20, x_1184); +lean_ctor_set(x_1194, 21, x_1185); +lean_ctor_set(x_1194, 22, x_1186); +lean_ctor_set(x_1194, 23, x_1187); +lean_ctor_set(x_1194, 24, x_1188); +lean_ctor_set(x_1194, 25, x_1189); +lean_ctor_set(x_1194, 26, x_1190); +lean_ctor_set(x_1194, 27, x_1191); +lean_ctor_set(x_1194, 28, x_1192); +lean_ctor_set(x_1194, 29, x_1193); +return x_1194; +} +} +else +{ +lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; lean_object* x_1205; lean_object* x_1206; lean_object* x_1207; lean_object* x_1208; lean_object* x_1209; lean_object* x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; lean_object* x_1226; +lean_dec(x_2); +x_1195 = lean_ctor_get(x_1, 0); +lean_inc(x_1195); +x_1196 = lean_ctor_get(x_1, 1); +lean_inc(x_1196); +x_1197 = lean_ctor_get(x_1, 2); +lean_inc(x_1197); +x_1198 = lean_ctor_get(x_1, 3); +lean_inc(x_1198); +x_1199 = lean_ctor_get(x_1, 4); +lean_inc(x_1199); +x_1200 = lean_ctor_get(x_1, 5); +lean_inc(x_1200); +x_1201 = lean_ctor_get(x_1, 6); +lean_inc(x_1201); +x_1202 = lean_ctor_get(x_1, 7); +lean_inc(x_1202); +x_1203 = lean_ctor_get(x_1, 8); +lean_inc(x_1203); +x_1204 = lean_ctor_get(x_1, 9); +lean_inc(x_1204); +x_1205 = lean_ctor_get(x_1, 10); +lean_inc(x_1205); +x_1206 = lean_ctor_get(x_1, 11); +lean_inc(x_1206); +x_1207 = lean_ctor_get(x_1, 12); +lean_inc(x_1207); +x_1208 = lean_ctor_get(x_1, 13); +lean_inc(x_1208); +x_1209 = lean_ctor_get(x_1, 14); +lean_inc(x_1209); +x_1210 = lean_ctor_get(x_1, 15); +lean_inc(x_1210); +x_1211 = lean_ctor_get(x_1, 16); +lean_inc(x_1211); +x_1212 = lean_ctor_get(x_1, 17); +lean_inc(x_1212); +x_1213 = lean_ctor_get(x_1, 18); +lean_inc(x_1213); +x_1214 = lean_ctor_get(x_1, 20); +lean_inc(x_1214); +x_1215 = lean_ctor_get(x_1, 21); +lean_inc(x_1215); +x_1216 = lean_ctor_get(x_1, 22); +lean_inc(x_1216); +x_1217 = lean_ctor_get(x_1, 23); +lean_inc(x_1217); +x_1218 = lean_ctor_get(x_1, 24); +lean_inc(x_1218); +x_1219 = lean_ctor_get(x_1, 25); +lean_inc(x_1219); +x_1220 = lean_ctor_get(x_1, 26); +lean_inc(x_1220); +x_1221 = lean_ctor_get(x_1, 27); +lean_inc(x_1221); +x_1222 = lean_ctor_get(x_1, 28); +lean_inc(x_1222); +x_1223 = lean_ctor_get(x_1, 29); +lean_inc(x_1223); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1224 = x_1; +} else { + lean_dec_ref(x_1); + x_1224 = lean_box(0); +} +x_1225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1225, 0, x_3); +if (lean_is_scalar(x_1224)) { + x_1226 = lean_alloc_ctor(0, 30, 0); +} else { + x_1226 = x_1224; +} +lean_ctor_set(x_1226, 0, x_1195); +lean_ctor_set(x_1226, 1, x_1196); +lean_ctor_set(x_1226, 2, x_1197); +lean_ctor_set(x_1226, 3, x_1198); +lean_ctor_set(x_1226, 4, x_1199); +lean_ctor_set(x_1226, 5, x_1200); +lean_ctor_set(x_1226, 6, x_1201); +lean_ctor_set(x_1226, 7, x_1202); +lean_ctor_set(x_1226, 8, x_1203); +lean_ctor_set(x_1226, 9, x_1204); +lean_ctor_set(x_1226, 10, x_1205); +lean_ctor_set(x_1226, 11, x_1206); +lean_ctor_set(x_1226, 12, x_1207); +lean_ctor_set(x_1226, 13, x_1208); +lean_ctor_set(x_1226, 14, x_1209); +lean_ctor_set(x_1226, 15, x_1210); +lean_ctor_set(x_1226, 16, x_1211); +lean_ctor_set(x_1226, 17, x_1212); +lean_ctor_set(x_1226, 18, x_1213); +lean_ctor_set(x_1226, 19, x_1225); +lean_ctor_set(x_1226, 20, x_1214); +lean_ctor_set(x_1226, 21, x_1215); +lean_ctor_set(x_1226, 22, x_1216); +lean_ctor_set(x_1226, 23, x_1217); +lean_ctor_set(x_1226, 24, x_1218); +lean_ctor_set(x_1226, 25, x_1219); +lean_ctor_set(x_1226, 26, x_1220); +lean_ctor_set(x_1226, 27, x_1221); +lean_ctor_set(x_1226, 28, x_1222); +lean_ctor_set(x_1226, 29, x_1223); +return x_1226; +} +} +case 20: +{ +uint8_t x_1227; +x_1227 = !lean_is_exclusive(x_2); +if (x_1227 == 0) +{ +lean_object* x_1228; uint8_t x_1229; +x_1228 = lean_ctor_get(x_2, 0); +lean_dec(x_1228); +x_1229 = !lean_is_exclusive(x_1); +if (x_1229 == 0) +{ +lean_object* x_1230; +x_1230 = lean_ctor_get(x_1, 20); +lean_dec(x_1230); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 20, x_2); +return x_1; +} +else +{ +lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; lean_object* x_1240; lean_object* x_1241; lean_object* x_1242; lean_object* x_1243; lean_object* x_1244; lean_object* x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; +x_1231 = lean_ctor_get(x_1, 0); +x_1232 = lean_ctor_get(x_1, 1); +x_1233 = lean_ctor_get(x_1, 2); +x_1234 = lean_ctor_get(x_1, 3); +x_1235 = lean_ctor_get(x_1, 4); +x_1236 = lean_ctor_get(x_1, 5); +x_1237 = lean_ctor_get(x_1, 6); +x_1238 = lean_ctor_get(x_1, 7); +x_1239 = lean_ctor_get(x_1, 8); +x_1240 = lean_ctor_get(x_1, 9); +x_1241 = lean_ctor_get(x_1, 10); +x_1242 = lean_ctor_get(x_1, 11); +x_1243 = lean_ctor_get(x_1, 12); +x_1244 = lean_ctor_get(x_1, 13); +x_1245 = lean_ctor_get(x_1, 14); +x_1246 = lean_ctor_get(x_1, 15); +x_1247 = lean_ctor_get(x_1, 16); +x_1248 = lean_ctor_get(x_1, 17); +x_1249 = lean_ctor_get(x_1, 18); +x_1250 = lean_ctor_get(x_1, 19); +x_1251 = lean_ctor_get(x_1, 21); +x_1252 = lean_ctor_get(x_1, 22); +x_1253 = lean_ctor_get(x_1, 23); +x_1254 = lean_ctor_get(x_1, 24); +x_1255 = lean_ctor_get(x_1, 25); +x_1256 = lean_ctor_get(x_1, 26); +x_1257 = lean_ctor_get(x_1, 27); +x_1258 = lean_ctor_get(x_1, 28); +x_1259 = lean_ctor_get(x_1, 29); +lean_inc(x_1259); +lean_inc(x_1258); +lean_inc(x_1257); +lean_inc(x_1256); +lean_inc(x_1255); +lean_inc(x_1254); +lean_inc(x_1253); +lean_inc(x_1252); +lean_inc(x_1251); +lean_inc(x_1250); +lean_inc(x_1249); +lean_inc(x_1248); +lean_inc(x_1247); +lean_inc(x_1246); +lean_inc(x_1245); +lean_inc(x_1244); +lean_inc(x_1243); +lean_inc(x_1242); +lean_inc(x_1241); +lean_inc(x_1240); +lean_inc(x_1239); +lean_inc(x_1238); +lean_inc(x_1237); +lean_inc(x_1236); +lean_inc(x_1235); +lean_inc(x_1234); +lean_inc(x_1233); +lean_inc(x_1232); +lean_inc(x_1231); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1260 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1260, 0, x_1231); +lean_ctor_set(x_1260, 1, x_1232); +lean_ctor_set(x_1260, 2, x_1233); +lean_ctor_set(x_1260, 3, x_1234); +lean_ctor_set(x_1260, 4, x_1235); +lean_ctor_set(x_1260, 5, x_1236); +lean_ctor_set(x_1260, 6, x_1237); +lean_ctor_set(x_1260, 7, x_1238); +lean_ctor_set(x_1260, 8, x_1239); +lean_ctor_set(x_1260, 9, x_1240); +lean_ctor_set(x_1260, 10, x_1241); +lean_ctor_set(x_1260, 11, x_1242); +lean_ctor_set(x_1260, 12, x_1243); +lean_ctor_set(x_1260, 13, x_1244); +lean_ctor_set(x_1260, 14, x_1245); +lean_ctor_set(x_1260, 15, x_1246); +lean_ctor_set(x_1260, 16, x_1247); +lean_ctor_set(x_1260, 17, x_1248); +lean_ctor_set(x_1260, 18, x_1249); +lean_ctor_set(x_1260, 19, x_1250); +lean_ctor_set(x_1260, 20, x_2); +lean_ctor_set(x_1260, 21, x_1251); +lean_ctor_set(x_1260, 22, x_1252); +lean_ctor_set(x_1260, 23, x_1253); +lean_ctor_set(x_1260, 24, x_1254); +lean_ctor_set(x_1260, 25, x_1255); +lean_ctor_set(x_1260, 26, x_1256); +lean_ctor_set(x_1260, 27, x_1257); +lean_ctor_set(x_1260, 28, x_1258); +lean_ctor_set(x_1260, 29, x_1259); +return x_1260; +} +} +else +{ +lean_object* x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; lean_object* x_1275; lean_object* x_1276; lean_object* x_1277; lean_object* x_1278; lean_object* x_1279; lean_object* x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; +lean_dec(x_2); +x_1261 = lean_ctor_get(x_1, 0); +lean_inc(x_1261); +x_1262 = lean_ctor_get(x_1, 1); +lean_inc(x_1262); +x_1263 = lean_ctor_get(x_1, 2); +lean_inc(x_1263); +x_1264 = lean_ctor_get(x_1, 3); +lean_inc(x_1264); +x_1265 = lean_ctor_get(x_1, 4); +lean_inc(x_1265); +x_1266 = lean_ctor_get(x_1, 5); +lean_inc(x_1266); +x_1267 = lean_ctor_get(x_1, 6); +lean_inc(x_1267); +x_1268 = lean_ctor_get(x_1, 7); +lean_inc(x_1268); +x_1269 = lean_ctor_get(x_1, 8); +lean_inc(x_1269); +x_1270 = lean_ctor_get(x_1, 9); +lean_inc(x_1270); +x_1271 = lean_ctor_get(x_1, 10); +lean_inc(x_1271); +x_1272 = lean_ctor_get(x_1, 11); +lean_inc(x_1272); +x_1273 = lean_ctor_get(x_1, 12); +lean_inc(x_1273); +x_1274 = lean_ctor_get(x_1, 13); +lean_inc(x_1274); +x_1275 = lean_ctor_get(x_1, 14); +lean_inc(x_1275); +x_1276 = lean_ctor_get(x_1, 15); +lean_inc(x_1276); +x_1277 = lean_ctor_get(x_1, 16); +lean_inc(x_1277); +x_1278 = lean_ctor_get(x_1, 17); +lean_inc(x_1278); +x_1279 = lean_ctor_get(x_1, 18); +lean_inc(x_1279); +x_1280 = lean_ctor_get(x_1, 19); +lean_inc(x_1280); +x_1281 = lean_ctor_get(x_1, 21); +lean_inc(x_1281); +x_1282 = lean_ctor_get(x_1, 22); +lean_inc(x_1282); +x_1283 = lean_ctor_get(x_1, 23); +lean_inc(x_1283); +x_1284 = lean_ctor_get(x_1, 24); +lean_inc(x_1284); +x_1285 = lean_ctor_get(x_1, 25); +lean_inc(x_1285); +x_1286 = lean_ctor_get(x_1, 26); +lean_inc(x_1286); +x_1287 = lean_ctor_get(x_1, 27); +lean_inc(x_1287); +x_1288 = lean_ctor_get(x_1, 28); +lean_inc(x_1288); +x_1289 = lean_ctor_get(x_1, 29); +lean_inc(x_1289); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1290 = x_1; +} else { + lean_dec_ref(x_1); + x_1290 = lean_box(0); +} +x_1291 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1291, 0, x_3); +if (lean_is_scalar(x_1290)) { + x_1292 = lean_alloc_ctor(0, 30, 0); +} else { + x_1292 = x_1290; +} +lean_ctor_set(x_1292, 0, x_1261); +lean_ctor_set(x_1292, 1, x_1262); +lean_ctor_set(x_1292, 2, x_1263); +lean_ctor_set(x_1292, 3, x_1264); +lean_ctor_set(x_1292, 4, x_1265); +lean_ctor_set(x_1292, 5, x_1266); +lean_ctor_set(x_1292, 6, x_1267); +lean_ctor_set(x_1292, 7, x_1268); +lean_ctor_set(x_1292, 8, x_1269); +lean_ctor_set(x_1292, 9, x_1270); +lean_ctor_set(x_1292, 10, x_1271); +lean_ctor_set(x_1292, 11, x_1272); +lean_ctor_set(x_1292, 12, x_1273); +lean_ctor_set(x_1292, 13, x_1274); +lean_ctor_set(x_1292, 14, x_1275); +lean_ctor_set(x_1292, 15, x_1276); +lean_ctor_set(x_1292, 16, x_1277); +lean_ctor_set(x_1292, 17, x_1278); +lean_ctor_set(x_1292, 18, x_1279); +lean_ctor_set(x_1292, 19, x_1280); +lean_ctor_set(x_1292, 20, x_1291); +lean_ctor_set(x_1292, 21, x_1281); +lean_ctor_set(x_1292, 22, x_1282); +lean_ctor_set(x_1292, 23, x_1283); +lean_ctor_set(x_1292, 24, x_1284); +lean_ctor_set(x_1292, 25, x_1285); +lean_ctor_set(x_1292, 26, x_1286); +lean_ctor_set(x_1292, 27, x_1287); +lean_ctor_set(x_1292, 28, x_1288); +lean_ctor_set(x_1292, 29, x_1289); +return x_1292; +} +} +case 21: +{ +uint8_t x_1293; +x_1293 = !lean_is_exclusive(x_2); +if (x_1293 == 0) +{ +lean_object* x_1294; uint8_t x_1295; +x_1294 = lean_ctor_get(x_2, 0); +lean_dec(x_1294); +x_1295 = !lean_is_exclusive(x_1); +if (x_1295 == 0) +{ +lean_object* x_1296; +x_1296 = lean_ctor_get(x_1, 21); +lean_dec(x_1296); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 21, x_2); +return x_1; +} +else +{ +lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; lean_object* x_1310; lean_object* x_1311; lean_object* x_1312; lean_object* x_1313; lean_object* x_1314; lean_object* x_1315; lean_object* x_1316; lean_object* x_1317; lean_object* x_1318; lean_object* x_1319; lean_object* x_1320; lean_object* x_1321; lean_object* x_1322; lean_object* x_1323; lean_object* x_1324; lean_object* x_1325; lean_object* x_1326; +x_1297 = lean_ctor_get(x_1, 0); +x_1298 = lean_ctor_get(x_1, 1); +x_1299 = lean_ctor_get(x_1, 2); +x_1300 = lean_ctor_get(x_1, 3); +x_1301 = lean_ctor_get(x_1, 4); +x_1302 = lean_ctor_get(x_1, 5); +x_1303 = lean_ctor_get(x_1, 6); +x_1304 = lean_ctor_get(x_1, 7); +x_1305 = lean_ctor_get(x_1, 8); +x_1306 = lean_ctor_get(x_1, 9); +x_1307 = lean_ctor_get(x_1, 10); +x_1308 = lean_ctor_get(x_1, 11); +x_1309 = lean_ctor_get(x_1, 12); +x_1310 = lean_ctor_get(x_1, 13); +x_1311 = lean_ctor_get(x_1, 14); +x_1312 = lean_ctor_get(x_1, 15); +x_1313 = lean_ctor_get(x_1, 16); +x_1314 = lean_ctor_get(x_1, 17); +x_1315 = lean_ctor_get(x_1, 18); +x_1316 = lean_ctor_get(x_1, 19); +x_1317 = lean_ctor_get(x_1, 20); +x_1318 = lean_ctor_get(x_1, 22); +x_1319 = lean_ctor_get(x_1, 23); +x_1320 = lean_ctor_get(x_1, 24); +x_1321 = lean_ctor_get(x_1, 25); +x_1322 = lean_ctor_get(x_1, 26); +x_1323 = lean_ctor_get(x_1, 27); +x_1324 = lean_ctor_get(x_1, 28); +x_1325 = lean_ctor_get(x_1, 29); +lean_inc(x_1325); +lean_inc(x_1324); +lean_inc(x_1323); +lean_inc(x_1322); +lean_inc(x_1321); +lean_inc(x_1320); +lean_inc(x_1319); +lean_inc(x_1318); +lean_inc(x_1317); +lean_inc(x_1316); +lean_inc(x_1315); +lean_inc(x_1314); +lean_inc(x_1313); +lean_inc(x_1312); +lean_inc(x_1311); +lean_inc(x_1310); +lean_inc(x_1309); +lean_inc(x_1308); +lean_inc(x_1307); +lean_inc(x_1306); +lean_inc(x_1305); +lean_inc(x_1304); +lean_inc(x_1303); +lean_inc(x_1302); +lean_inc(x_1301); +lean_inc(x_1300); +lean_inc(x_1299); +lean_inc(x_1298); +lean_inc(x_1297); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1326 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1326, 0, x_1297); +lean_ctor_set(x_1326, 1, x_1298); +lean_ctor_set(x_1326, 2, x_1299); +lean_ctor_set(x_1326, 3, x_1300); +lean_ctor_set(x_1326, 4, x_1301); +lean_ctor_set(x_1326, 5, x_1302); +lean_ctor_set(x_1326, 6, x_1303); +lean_ctor_set(x_1326, 7, x_1304); +lean_ctor_set(x_1326, 8, x_1305); +lean_ctor_set(x_1326, 9, x_1306); +lean_ctor_set(x_1326, 10, x_1307); +lean_ctor_set(x_1326, 11, x_1308); +lean_ctor_set(x_1326, 12, x_1309); +lean_ctor_set(x_1326, 13, x_1310); +lean_ctor_set(x_1326, 14, x_1311); +lean_ctor_set(x_1326, 15, x_1312); +lean_ctor_set(x_1326, 16, x_1313); +lean_ctor_set(x_1326, 17, x_1314); +lean_ctor_set(x_1326, 18, x_1315); +lean_ctor_set(x_1326, 19, x_1316); +lean_ctor_set(x_1326, 20, x_1317); +lean_ctor_set(x_1326, 21, x_2); +lean_ctor_set(x_1326, 22, x_1318); +lean_ctor_set(x_1326, 23, x_1319); +lean_ctor_set(x_1326, 24, x_1320); +lean_ctor_set(x_1326, 25, x_1321); +lean_ctor_set(x_1326, 26, x_1322); +lean_ctor_set(x_1326, 27, x_1323); +lean_ctor_set(x_1326, 28, x_1324); +lean_ctor_set(x_1326, 29, x_1325); +return x_1326; +} +} +else +{ +lean_object* x_1327; lean_object* x_1328; lean_object* x_1329; lean_object* x_1330; lean_object* x_1331; lean_object* x_1332; lean_object* x_1333; lean_object* x_1334; lean_object* x_1335; lean_object* x_1336; lean_object* x_1337; lean_object* x_1338; lean_object* x_1339; lean_object* x_1340; lean_object* x_1341; lean_object* x_1342; lean_object* x_1343; lean_object* x_1344; lean_object* x_1345; lean_object* x_1346; lean_object* x_1347; lean_object* x_1348; lean_object* x_1349; lean_object* x_1350; lean_object* x_1351; lean_object* x_1352; lean_object* x_1353; lean_object* x_1354; lean_object* x_1355; lean_object* x_1356; lean_object* x_1357; lean_object* x_1358; +lean_dec(x_2); +x_1327 = lean_ctor_get(x_1, 0); +lean_inc(x_1327); +x_1328 = lean_ctor_get(x_1, 1); +lean_inc(x_1328); +x_1329 = lean_ctor_get(x_1, 2); +lean_inc(x_1329); +x_1330 = lean_ctor_get(x_1, 3); +lean_inc(x_1330); +x_1331 = lean_ctor_get(x_1, 4); +lean_inc(x_1331); +x_1332 = lean_ctor_get(x_1, 5); +lean_inc(x_1332); +x_1333 = lean_ctor_get(x_1, 6); +lean_inc(x_1333); +x_1334 = lean_ctor_get(x_1, 7); +lean_inc(x_1334); +x_1335 = lean_ctor_get(x_1, 8); +lean_inc(x_1335); +x_1336 = lean_ctor_get(x_1, 9); +lean_inc(x_1336); +x_1337 = lean_ctor_get(x_1, 10); +lean_inc(x_1337); +x_1338 = lean_ctor_get(x_1, 11); +lean_inc(x_1338); +x_1339 = lean_ctor_get(x_1, 12); +lean_inc(x_1339); +x_1340 = lean_ctor_get(x_1, 13); +lean_inc(x_1340); +x_1341 = lean_ctor_get(x_1, 14); +lean_inc(x_1341); +x_1342 = lean_ctor_get(x_1, 15); +lean_inc(x_1342); +x_1343 = lean_ctor_get(x_1, 16); +lean_inc(x_1343); +x_1344 = lean_ctor_get(x_1, 17); +lean_inc(x_1344); +x_1345 = lean_ctor_get(x_1, 18); +lean_inc(x_1345); +x_1346 = lean_ctor_get(x_1, 19); +lean_inc(x_1346); +x_1347 = lean_ctor_get(x_1, 20); +lean_inc(x_1347); +x_1348 = lean_ctor_get(x_1, 22); +lean_inc(x_1348); +x_1349 = lean_ctor_get(x_1, 23); +lean_inc(x_1349); +x_1350 = lean_ctor_get(x_1, 24); +lean_inc(x_1350); +x_1351 = lean_ctor_get(x_1, 25); +lean_inc(x_1351); +x_1352 = lean_ctor_get(x_1, 26); +lean_inc(x_1352); +x_1353 = lean_ctor_get(x_1, 27); +lean_inc(x_1353); +x_1354 = lean_ctor_get(x_1, 28); +lean_inc(x_1354); +x_1355 = lean_ctor_get(x_1, 29); +lean_inc(x_1355); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1356 = x_1; +} else { + lean_dec_ref(x_1); + x_1356 = lean_box(0); +} +x_1357 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1357, 0, x_3); +if (lean_is_scalar(x_1356)) { + x_1358 = lean_alloc_ctor(0, 30, 0); +} else { + x_1358 = x_1356; +} +lean_ctor_set(x_1358, 0, x_1327); +lean_ctor_set(x_1358, 1, x_1328); +lean_ctor_set(x_1358, 2, x_1329); +lean_ctor_set(x_1358, 3, x_1330); +lean_ctor_set(x_1358, 4, x_1331); +lean_ctor_set(x_1358, 5, x_1332); +lean_ctor_set(x_1358, 6, x_1333); +lean_ctor_set(x_1358, 7, x_1334); +lean_ctor_set(x_1358, 8, x_1335); +lean_ctor_set(x_1358, 9, x_1336); +lean_ctor_set(x_1358, 10, x_1337); +lean_ctor_set(x_1358, 11, x_1338); +lean_ctor_set(x_1358, 12, x_1339); +lean_ctor_set(x_1358, 13, x_1340); +lean_ctor_set(x_1358, 14, x_1341); +lean_ctor_set(x_1358, 15, x_1342); +lean_ctor_set(x_1358, 16, x_1343); +lean_ctor_set(x_1358, 17, x_1344); +lean_ctor_set(x_1358, 18, x_1345); +lean_ctor_set(x_1358, 19, x_1346); +lean_ctor_set(x_1358, 20, x_1347); +lean_ctor_set(x_1358, 21, x_1357); +lean_ctor_set(x_1358, 22, x_1348); +lean_ctor_set(x_1358, 23, x_1349); +lean_ctor_set(x_1358, 24, x_1350); +lean_ctor_set(x_1358, 25, x_1351); +lean_ctor_set(x_1358, 26, x_1352); +lean_ctor_set(x_1358, 27, x_1353); +lean_ctor_set(x_1358, 28, x_1354); +lean_ctor_set(x_1358, 29, x_1355); +return x_1358; +} +} +case 22: +{ +uint8_t x_1359; +x_1359 = !lean_is_exclusive(x_2); +if (x_1359 == 0) +{ +lean_object* x_1360; uint8_t x_1361; +x_1360 = lean_ctor_get(x_2, 0); +lean_dec(x_1360); +x_1361 = !lean_is_exclusive(x_1); +if (x_1361 == 0) +{ +lean_object* x_1362; +x_1362 = lean_ctor_get(x_1, 22); +lean_dec(x_1362); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +lean_ctor_set(x_1, 22, x_2); +return x_1; +} +else +{ +lean_object* x_1363; lean_object* x_1364; lean_object* x_1365; lean_object* x_1366; lean_object* x_1367; lean_object* x_1368; lean_object* x_1369; lean_object* x_1370; lean_object* x_1371; lean_object* x_1372; lean_object* x_1373; lean_object* x_1374; lean_object* x_1375; lean_object* x_1376; lean_object* x_1377; lean_object* x_1378; lean_object* x_1379; lean_object* x_1380; lean_object* x_1381; lean_object* x_1382; lean_object* x_1383; lean_object* x_1384; lean_object* x_1385; lean_object* x_1386; lean_object* x_1387; lean_object* x_1388; lean_object* x_1389; lean_object* x_1390; lean_object* x_1391; lean_object* x_1392; +x_1363 = lean_ctor_get(x_1, 0); +x_1364 = lean_ctor_get(x_1, 1); +x_1365 = lean_ctor_get(x_1, 2); +x_1366 = lean_ctor_get(x_1, 3); +x_1367 = lean_ctor_get(x_1, 4); +x_1368 = lean_ctor_get(x_1, 5); +x_1369 = lean_ctor_get(x_1, 6); +x_1370 = lean_ctor_get(x_1, 7); +x_1371 = lean_ctor_get(x_1, 8); +x_1372 = lean_ctor_get(x_1, 9); +x_1373 = lean_ctor_get(x_1, 10); +x_1374 = lean_ctor_get(x_1, 11); +x_1375 = lean_ctor_get(x_1, 12); +x_1376 = lean_ctor_get(x_1, 13); +x_1377 = lean_ctor_get(x_1, 14); +x_1378 = lean_ctor_get(x_1, 15); +x_1379 = lean_ctor_get(x_1, 16); +x_1380 = lean_ctor_get(x_1, 17); +x_1381 = lean_ctor_get(x_1, 18); +x_1382 = lean_ctor_get(x_1, 19); +x_1383 = lean_ctor_get(x_1, 20); +x_1384 = lean_ctor_get(x_1, 21); +x_1385 = lean_ctor_get(x_1, 23); +x_1386 = lean_ctor_get(x_1, 24); +x_1387 = lean_ctor_get(x_1, 25); +x_1388 = lean_ctor_get(x_1, 26); +x_1389 = lean_ctor_get(x_1, 27); +x_1390 = lean_ctor_get(x_1, 28); +x_1391 = lean_ctor_get(x_1, 29); +lean_inc(x_1391); +lean_inc(x_1390); +lean_inc(x_1389); +lean_inc(x_1388); +lean_inc(x_1387); +lean_inc(x_1386); +lean_inc(x_1385); +lean_inc(x_1384); +lean_inc(x_1383); +lean_inc(x_1382); +lean_inc(x_1381); +lean_inc(x_1380); +lean_inc(x_1379); +lean_inc(x_1378); +lean_inc(x_1377); +lean_inc(x_1376); +lean_inc(x_1375); +lean_inc(x_1374); +lean_inc(x_1373); +lean_inc(x_1372); +lean_inc(x_1371); +lean_inc(x_1370); +lean_inc(x_1369); +lean_inc(x_1368); +lean_inc(x_1367); +lean_inc(x_1366); +lean_inc(x_1365); +lean_inc(x_1364); +lean_inc(x_1363); +lean_dec(x_1); +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_3); +x_1392 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1392, 0, x_1363); +lean_ctor_set(x_1392, 1, x_1364); +lean_ctor_set(x_1392, 2, x_1365); +lean_ctor_set(x_1392, 3, x_1366); +lean_ctor_set(x_1392, 4, x_1367); +lean_ctor_set(x_1392, 5, x_1368); +lean_ctor_set(x_1392, 6, x_1369); +lean_ctor_set(x_1392, 7, x_1370); +lean_ctor_set(x_1392, 8, x_1371); +lean_ctor_set(x_1392, 9, x_1372); +lean_ctor_set(x_1392, 10, x_1373); +lean_ctor_set(x_1392, 11, x_1374); +lean_ctor_set(x_1392, 12, x_1375); +lean_ctor_set(x_1392, 13, x_1376); +lean_ctor_set(x_1392, 14, x_1377); +lean_ctor_set(x_1392, 15, x_1378); +lean_ctor_set(x_1392, 16, x_1379); +lean_ctor_set(x_1392, 17, x_1380); +lean_ctor_set(x_1392, 18, x_1381); +lean_ctor_set(x_1392, 19, x_1382); +lean_ctor_set(x_1392, 20, x_1383); +lean_ctor_set(x_1392, 21, x_1384); +lean_ctor_set(x_1392, 22, x_2); +lean_ctor_set(x_1392, 23, x_1385); +lean_ctor_set(x_1392, 24, x_1386); +lean_ctor_set(x_1392, 25, x_1387); +lean_ctor_set(x_1392, 26, x_1388); +lean_ctor_set(x_1392, 27, x_1389); +lean_ctor_set(x_1392, 28, x_1390); +lean_ctor_set(x_1392, 29, x_1391); +return x_1392; +} +} +else +{ +lean_object* x_1393; lean_object* x_1394; lean_object* x_1395; lean_object* x_1396; lean_object* x_1397; lean_object* x_1398; lean_object* x_1399; lean_object* x_1400; lean_object* x_1401; lean_object* x_1402; lean_object* x_1403; lean_object* x_1404; lean_object* x_1405; lean_object* x_1406; lean_object* x_1407; lean_object* x_1408; lean_object* x_1409; lean_object* x_1410; lean_object* x_1411; lean_object* x_1412; lean_object* x_1413; lean_object* x_1414; lean_object* x_1415; lean_object* x_1416; lean_object* x_1417; lean_object* x_1418; lean_object* x_1419; lean_object* x_1420; lean_object* x_1421; lean_object* x_1422; lean_object* x_1423; lean_object* x_1424; +lean_dec(x_2); +x_1393 = lean_ctor_get(x_1, 0); +lean_inc(x_1393); +x_1394 = lean_ctor_get(x_1, 1); +lean_inc(x_1394); +x_1395 = lean_ctor_get(x_1, 2); +lean_inc(x_1395); +x_1396 = lean_ctor_get(x_1, 3); +lean_inc(x_1396); +x_1397 = lean_ctor_get(x_1, 4); +lean_inc(x_1397); +x_1398 = lean_ctor_get(x_1, 5); +lean_inc(x_1398); +x_1399 = lean_ctor_get(x_1, 6); +lean_inc(x_1399); +x_1400 = lean_ctor_get(x_1, 7); +lean_inc(x_1400); +x_1401 = lean_ctor_get(x_1, 8); +lean_inc(x_1401); +x_1402 = lean_ctor_get(x_1, 9); +lean_inc(x_1402); +x_1403 = lean_ctor_get(x_1, 10); +lean_inc(x_1403); +x_1404 = lean_ctor_get(x_1, 11); +lean_inc(x_1404); +x_1405 = lean_ctor_get(x_1, 12); +lean_inc(x_1405); +x_1406 = lean_ctor_get(x_1, 13); +lean_inc(x_1406); +x_1407 = lean_ctor_get(x_1, 14); +lean_inc(x_1407); +x_1408 = lean_ctor_get(x_1, 15); +lean_inc(x_1408); +x_1409 = lean_ctor_get(x_1, 16); +lean_inc(x_1409); +x_1410 = lean_ctor_get(x_1, 17); +lean_inc(x_1410); +x_1411 = lean_ctor_get(x_1, 18); +lean_inc(x_1411); +x_1412 = lean_ctor_get(x_1, 19); +lean_inc(x_1412); +x_1413 = lean_ctor_get(x_1, 20); +lean_inc(x_1413); +x_1414 = lean_ctor_get(x_1, 21); +lean_inc(x_1414); +x_1415 = lean_ctor_get(x_1, 23); +lean_inc(x_1415); +x_1416 = lean_ctor_get(x_1, 24); +lean_inc(x_1416); +x_1417 = lean_ctor_get(x_1, 25); +lean_inc(x_1417); +x_1418 = lean_ctor_get(x_1, 26); +lean_inc(x_1418); +x_1419 = lean_ctor_get(x_1, 27); +lean_inc(x_1419); +x_1420 = lean_ctor_get(x_1, 28); +lean_inc(x_1420); +x_1421 = lean_ctor_get(x_1, 29); +lean_inc(x_1421); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + lean_ctor_release(x_1, 4); + lean_ctor_release(x_1, 5); + lean_ctor_release(x_1, 6); + lean_ctor_release(x_1, 7); + lean_ctor_release(x_1, 8); + lean_ctor_release(x_1, 9); + lean_ctor_release(x_1, 10); + lean_ctor_release(x_1, 11); + lean_ctor_release(x_1, 12); + lean_ctor_release(x_1, 13); + lean_ctor_release(x_1, 14); + lean_ctor_release(x_1, 15); + lean_ctor_release(x_1, 16); + lean_ctor_release(x_1, 17); + lean_ctor_release(x_1, 18); + lean_ctor_release(x_1, 19); + lean_ctor_release(x_1, 20); + lean_ctor_release(x_1, 21); + lean_ctor_release(x_1, 22); + lean_ctor_release(x_1, 23); + lean_ctor_release(x_1, 24); + lean_ctor_release(x_1, 25); + lean_ctor_release(x_1, 26); + lean_ctor_release(x_1, 27); + lean_ctor_release(x_1, 28); + lean_ctor_release(x_1, 29); + x_1422 = x_1; +} else { + lean_dec_ref(x_1); + x_1422 = lean_box(0); +} +x_1423 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1423, 0, x_3); +if (lean_is_scalar(x_1422)) { + x_1424 = lean_alloc_ctor(0, 30, 0); +} else { + x_1424 = x_1422; +} +lean_ctor_set(x_1424, 0, x_1393); +lean_ctor_set(x_1424, 1, x_1394); +lean_ctor_set(x_1424, 2, x_1395); +lean_ctor_set(x_1424, 3, x_1396); +lean_ctor_set(x_1424, 4, x_1397); +lean_ctor_set(x_1424, 5, x_1398); +lean_ctor_set(x_1424, 6, x_1399); +lean_ctor_set(x_1424, 7, x_1400); +lean_ctor_set(x_1424, 8, x_1401); +lean_ctor_set(x_1424, 9, x_1402); +lean_ctor_set(x_1424, 10, x_1403); +lean_ctor_set(x_1424, 11, x_1404); +lean_ctor_set(x_1424, 12, x_1405); +lean_ctor_set(x_1424, 13, x_1406); +lean_ctor_set(x_1424, 14, x_1407); +lean_ctor_set(x_1424, 15, x_1408); +lean_ctor_set(x_1424, 16, x_1409); +lean_ctor_set(x_1424, 17, x_1410); +lean_ctor_set(x_1424, 18, x_1411); +lean_ctor_set(x_1424, 19, x_1412); +lean_ctor_set(x_1424, 20, x_1413); +lean_ctor_set(x_1424, 21, x_1414); +lean_ctor_set(x_1424, 22, x_1423); +lean_ctor_set(x_1424, 23, x_1415); +lean_ctor_set(x_1424, 24, x_1416); +lean_ctor_set(x_1424, 25, x_1417); +lean_ctor_set(x_1424, 26, x_1418); +lean_ctor_set(x_1424, 27, x_1419); +lean_ctor_set(x_1424, 28, x_1420); +lean_ctor_set(x_1424, 29, x_1421); +return x_1424; +} +} +case 23: +{ +uint8_t x_1425; +x_1425 = !lean_is_exclusive(x_1); +if (x_1425 == 0) +{ +lean_object* x_1426; lean_object* x_1427; +x_1426 = lean_ctor_get(x_1, 23); +lean_dec(x_1426); +x_1427 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1427, 0, x_3); +lean_ctor_set(x_1, 23, x_1427); +return x_1; +} +else +{ +lean_object* x_1428; lean_object* x_1429; lean_object* x_1430; lean_object* x_1431; lean_object* x_1432; lean_object* x_1433; lean_object* x_1434; lean_object* x_1435; lean_object* x_1436; lean_object* x_1437; lean_object* x_1438; lean_object* x_1439; lean_object* x_1440; lean_object* x_1441; lean_object* x_1442; lean_object* x_1443; lean_object* x_1444; lean_object* x_1445; lean_object* x_1446; lean_object* x_1447; lean_object* x_1448; lean_object* x_1449; lean_object* x_1450; lean_object* x_1451; lean_object* x_1452; lean_object* x_1453; lean_object* x_1454; lean_object* x_1455; lean_object* x_1456; lean_object* x_1457; lean_object* x_1458; +x_1428 = lean_ctor_get(x_1, 0); +x_1429 = lean_ctor_get(x_1, 1); +x_1430 = lean_ctor_get(x_1, 2); +x_1431 = lean_ctor_get(x_1, 3); +x_1432 = lean_ctor_get(x_1, 4); +x_1433 = lean_ctor_get(x_1, 5); +x_1434 = lean_ctor_get(x_1, 6); +x_1435 = lean_ctor_get(x_1, 7); +x_1436 = lean_ctor_get(x_1, 8); +x_1437 = lean_ctor_get(x_1, 9); +x_1438 = lean_ctor_get(x_1, 10); +x_1439 = lean_ctor_get(x_1, 11); +x_1440 = lean_ctor_get(x_1, 12); +x_1441 = lean_ctor_get(x_1, 13); +x_1442 = lean_ctor_get(x_1, 14); +x_1443 = lean_ctor_get(x_1, 15); +x_1444 = lean_ctor_get(x_1, 16); +x_1445 = lean_ctor_get(x_1, 17); +x_1446 = lean_ctor_get(x_1, 18); +x_1447 = lean_ctor_get(x_1, 19); +x_1448 = lean_ctor_get(x_1, 20); +x_1449 = lean_ctor_get(x_1, 21); +x_1450 = lean_ctor_get(x_1, 22); +x_1451 = lean_ctor_get(x_1, 24); +x_1452 = lean_ctor_get(x_1, 25); +x_1453 = lean_ctor_get(x_1, 26); +x_1454 = lean_ctor_get(x_1, 27); +x_1455 = lean_ctor_get(x_1, 28); +x_1456 = lean_ctor_get(x_1, 29); +lean_inc(x_1456); +lean_inc(x_1455); +lean_inc(x_1454); +lean_inc(x_1453); +lean_inc(x_1452); +lean_inc(x_1451); +lean_inc(x_1450); +lean_inc(x_1449); +lean_inc(x_1448); +lean_inc(x_1447); +lean_inc(x_1446); +lean_inc(x_1445); +lean_inc(x_1444); +lean_inc(x_1443); +lean_inc(x_1442); +lean_inc(x_1441); +lean_inc(x_1440); +lean_inc(x_1439); +lean_inc(x_1438); +lean_inc(x_1437); +lean_inc(x_1436); +lean_inc(x_1435); +lean_inc(x_1434); +lean_inc(x_1433); +lean_inc(x_1432); +lean_inc(x_1431); +lean_inc(x_1430); +lean_inc(x_1429); +lean_inc(x_1428); +lean_dec(x_1); +x_1457 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1457, 0, x_3); +x_1458 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1458, 0, x_1428); +lean_ctor_set(x_1458, 1, x_1429); +lean_ctor_set(x_1458, 2, x_1430); +lean_ctor_set(x_1458, 3, x_1431); +lean_ctor_set(x_1458, 4, x_1432); +lean_ctor_set(x_1458, 5, x_1433); +lean_ctor_set(x_1458, 6, x_1434); +lean_ctor_set(x_1458, 7, x_1435); +lean_ctor_set(x_1458, 8, x_1436); +lean_ctor_set(x_1458, 9, x_1437); +lean_ctor_set(x_1458, 10, x_1438); +lean_ctor_set(x_1458, 11, x_1439); +lean_ctor_set(x_1458, 12, x_1440); +lean_ctor_set(x_1458, 13, x_1441); +lean_ctor_set(x_1458, 14, x_1442); +lean_ctor_set(x_1458, 15, x_1443); +lean_ctor_set(x_1458, 16, x_1444); +lean_ctor_set(x_1458, 17, x_1445); +lean_ctor_set(x_1458, 18, x_1446); +lean_ctor_set(x_1458, 19, x_1447); +lean_ctor_set(x_1458, 20, x_1448); +lean_ctor_set(x_1458, 21, x_1449); +lean_ctor_set(x_1458, 22, x_1450); +lean_ctor_set(x_1458, 23, x_1457); +lean_ctor_set(x_1458, 24, x_1451); +lean_ctor_set(x_1458, 25, x_1452); +lean_ctor_set(x_1458, 26, x_1453); +lean_ctor_set(x_1458, 27, x_1454); +lean_ctor_set(x_1458, 28, x_1455); +lean_ctor_set(x_1458, 29, x_1456); +return x_1458; +} +} +case 24: +{ +uint8_t x_1459; +x_1459 = lean_ctor_get_uint8(x_2, 0); +lean_dec(x_2); +if (x_1459 == 0) +{ +uint8_t x_1460; +x_1460 = !lean_is_exclusive(x_1); +if (x_1460 == 0) +{ +lean_object* x_1461; lean_object* x_1462; +x_1461 = lean_ctor_get(x_1, 25); +lean_dec(x_1461); +x_1462 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1462, 0, x_3); +lean_ctor_set(x_1, 25, x_1462); +return x_1; +} +else +{ +lean_object* x_1463; lean_object* x_1464; lean_object* x_1465; lean_object* x_1466; lean_object* x_1467; lean_object* x_1468; lean_object* x_1469; lean_object* x_1470; lean_object* x_1471; lean_object* x_1472; lean_object* x_1473; lean_object* x_1474; lean_object* x_1475; lean_object* x_1476; lean_object* x_1477; lean_object* x_1478; lean_object* x_1479; lean_object* x_1480; lean_object* x_1481; lean_object* x_1482; lean_object* x_1483; lean_object* x_1484; lean_object* x_1485; lean_object* x_1486; lean_object* x_1487; lean_object* x_1488; lean_object* x_1489; lean_object* x_1490; lean_object* x_1491; lean_object* x_1492; lean_object* x_1493; +x_1463 = lean_ctor_get(x_1, 0); +x_1464 = lean_ctor_get(x_1, 1); +x_1465 = lean_ctor_get(x_1, 2); +x_1466 = lean_ctor_get(x_1, 3); +x_1467 = lean_ctor_get(x_1, 4); +x_1468 = lean_ctor_get(x_1, 5); +x_1469 = lean_ctor_get(x_1, 6); +x_1470 = lean_ctor_get(x_1, 7); +x_1471 = lean_ctor_get(x_1, 8); +x_1472 = lean_ctor_get(x_1, 9); +x_1473 = lean_ctor_get(x_1, 10); +x_1474 = lean_ctor_get(x_1, 11); +x_1475 = lean_ctor_get(x_1, 12); +x_1476 = lean_ctor_get(x_1, 13); +x_1477 = lean_ctor_get(x_1, 14); +x_1478 = lean_ctor_get(x_1, 15); +x_1479 = lean_ctor_get(x_1, 16); +x_1480 = lean_ctor_get(x_1, 17); +x_1481 = lean_ctor_get(x_1, 18); +x_1482 = lean_ctor_get(x_1, 19); +x_1483 = lean_ctor_get(x_1, 20); +x_1484 = lean_ctor_get(x_1, 21); +x_1485 = lean_ctor_get(x_1, 22); +x_1486 = lean_ctor_get(x_1, 23); +x_1487 = lean_ctor_get(x_1, 24); +x_1488 = lean_ctor_get(x_1, 26); +x_1489 = lean_ctor_get(x_1, 27); +x_1490 = lean_ctor_get(x_1, 28); +x_1491 = lean_ctor_get(x_1, 29); +lean_inc(x_1491); +lean_inc(x_1490); +lean_inc(x_1489); +lean_inc(x_1488); +lean_inc(x_1487); +lean_inc(x_1486); +lean_inc(x_1485); +lean_inc(x_1484); +lean_inc(x_1483); +lean_inc(x_1482); +lean_inc(x_1481); +lean_inc(x_1480); +lean_inc(x_1479); +lean_inc(x_1478); +lean_inc(x_1477); +lean_inc(x_1476); +lean_inc(x_1475); +lean_inc(x_1474); +lean_inc(x_1473); +lean_inc(x_1472); +lean_inc(x_1471); +lean_inc(x_1470); +lean_inc(x_1469); +lean_inc(x_1468); +lean_inc(x_1467); +lean_inc(x_1466); +lean_inc(x_1465); +lean_inc(x_1464); +lean_inc(x_1463); +lean_dec(x_1); +x_1492 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1492, 0, x_3); +x_1493 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1493, 0, x_1463); +lean_ctor_set(x_1493, 1, x_1464); +lean_ctor_set(x_1493, 2, x_1465); +lean_ctor_set(x_1493, 3, x_1466); +lean_ctor_set(x_1493, 4, x_1467); +lean_ctor_set(x_1493, 5, x_1468); +lean_ctor_set(x_1493, 6, x_1469); +lean_ctor_set(x_1493, 7, x_1470); +lean_ctor_set(x_1493, 8, x_1471); +lean_ctor_set(x_1493, 9, x_1472); +lean_ctor_set(x_1493, 10, x_1473); +lean_ctor_set(x_1493, 11, x_1474); +lean_ctor_set(x_1493, 12, x_1475); +lean_ctor_set(x_1493, 13, x_1476); +lean_ctor_set(x_1493, 14, x_1477); +lean_ctor_set(x_1493, 15, x_1478); +lean_ctor_set(x_1493, 16, x_1479); +lean_ctor_set(x_1493, 17, x_1480); +lean_ctor_set(x_1493, 18, x_1481); +lean_ctor_set(x_1493, 19, x_1482); +lean_ctor_set(x_1493, 20, x_1483); +lean_ctor_set(x_1493, 21, x_1484); +lean_ctor_set(x_1493, 22, x_1485); +lean_ctor_set(x_1493, 23, x_1486); +lean_ctor_set(x_1493, 24, x_1487); +lean_ctor_set(x_1493, 25, x_1492); +lean_ctor_set(x_1493, 26, x_1488); +lean_ctor_set(x_1493, 27, x_1489); +lean_ctor_set(x_1493, 28, x_1490); +lean_ctor_set(x_1493, 29, x_1491); +return x_1493; +} +} +else +{ +uint8_t x_1494; +x_1494 = !lean_is_exclusive(x_1); +if (x_1494 == 0) +{ +lean_object* x_1495; lean_object* x_1496; +x_1495 = lean_ctor_get(x_1, 24); +lean_dec(x_1495); +x_1496 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1496, 0, x_3); +lean_ctor_set(x_1, 24, x_1496); +return x_1; +} +else +{ +lean_object* x_1497; lean_object* x_1498; lean_object* x_1499; lean_object* x_1500; lean_object* x_1501; lean_object* x_1502; lean_object* x_1503; lean_object* x_1504; lean_object* x_1505; lean_object* x_1506; lean_object* x_1507; lean_object* x_1508; lean_object* x_1509; lean_object* x_1510; lean_object* x_1511; lean_object* x_1512; lean_object* x_1513; lean_object* x_1514; lean_object* x_1515; lean_object* x_1516; lean_object* x_1517; lean_object* x_1518; lean_object* x_1519; lean_object* x_1520; lean_object* x_1521; lean_object* x_1522; lean_object* x_1523; lean_object* x_1524; lean_object* x_1525; lean_object* x_1526; lean_object* x_1527; +x_1497 = lean_ctor_get(x_1, 0); +x_1498 = lean_ctor_get(x_1, 1); +x_1499 = lean_ctor_get(x_1, 2); +x_1500 = lean_ctor_get(x_1, 3); +x_1501 = lean_ctor_get(x_1, 4); +x_1502 = lean_ctor_get(x_1, 5); +x_1503 = lean_ctor_get(x_1, 6); +x_1504 = lean_ctor_get(x_1, 7); +x_1505 = lean_ctor_get(x_1, 8); +x_1506 = lean_ctor_get(x_1, 9); +x_1507 = lean_ctor_get(x_1, 10); +x_1508 = lean_ctor_get(x_1, 11); +x_1509 = lean_ctor_get(x_1, 12); +x_1510 = lean_ctor_get(x_1, 13); +x_1511 = lean_ctor_get(x_1, 14); +x_1512 = lean_ctor_get(x_1, 15); +x_1513 = lean_ctor_get(x_1, 16); +x_1514 = lean_ctor_get(x_1, 17); +x_1515 = lean_ctor_get(x_1, 18); +x_1516 = lean_ctor_get(x_1, 19); +x_1517 = lean_ctor_get(x_1, 20); +x_1518 = lean_ctor_get(x_1, 21); +x_1519 = lean_ctor_get(x_1, 22); +x_1520 = lean_ctor_get(x_1, 23); +x_1521 = lean_ctor_get(x_1, 25); +x_1522 = lean_ctor_get(x_1, 26); +x_1523 = lean_ctor_get(x_1, 27); +x_1524 = lean_ctor_get(x_1, 28); +x_1525 = lean_ctor_get(x_1, 29); +lean_inc(x_1525); +lean_inc(x_1524); +lean_inc(x_1523); +lean_inc(x_1522); +lean_inc(x_1521); +lean_inc(x_1520); +lean_inc(x_1519); +lean_inc(x_1518); +lean_inc(x_1517); +lean_inc(x_1516); +lean_inc(x_1515); +lean_inc(x_1514); +lean_inc(x_1513); +lean_inc(x_1512); +lean_inc(x_1511); +lean_inc(x_1510); +lean_inc(x_1509); +lean_inc(x_1508); +lean_inc(x_1507); +lean_inc(x_1506); +lean_inc(x_1505); +lean_inc(x_1504); +lean_inc(x_1503); +lean_inc(x_1502); +lean_inc(x_1501); +lean_inc(x_1500); +lean_inc(x_1499); +lean_inc(x_1498); +lean_inc(x_1497); +lean_dec(x_1); +x_1526 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1526, 0, x_3); +x_1527 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1527, 0, x_1497); +lean_ctor_set(x_1527, 1, x_1498); +lean_ctor_set(x_1527, 2, x_1499); +lean_ctor_set(x_1527, 3, x_1500); +lean_ctor_set(x_1527, 4, x_1501); +lean_ctor_set(x_1527, 5, x_1502); +lean_ctor_set(x_1527, 6, x_1503); +lean_ctor_set(x_1527, 7, x_1504); +lean_ctor_set(x_1527, 8, x_1505); +lean_ctor_set(x_1527, 9, x_1506); +lean_ctor_set(x_1527, 10, x_1507); +lean_ctor_set(x_1527, 11, x_1508); +lean_ctor_set(x_1527, 12, x_1509); +lean_ctor_set(x_1527, 13, x_1510); +lean_ctor_set(x_1527, 14, x_1511); +lean_ctor_set(x_1527, 15, x_1512); +lean_ctor_set(x_1527, 16, x_1513); +lean_ctor_set(x_1527, 17, x_1514); +lean_ctor_set(x_1527, 18, x_1515); +lean_ctor_set(x_1527, 19, x_1516); +lean_ctor_set(x_1527, 20, x_1517); +lean_ctor_set(x_1527, 21, x_1518); +lean_ctor_set(x_1527, 22, x_1519); +lean_ctor_set(x_1527, 23, x_1520); +lean_ctor_set(x_1527, 24, x_1526); +lean_ctor_set(x_1527, 25, x_1521); +lean_ctor_set(x_1527, 26, x_1522); +lean_ctor_set(x_1527, 27, x_1523); +lean_ctor_set(x_1527, 28, x_1524); +lean_ctor_set(x_1527, 29, x_1525); +return x_1527; +} +} +} +case 25: +{ +uint8_t x_1528; +lean_dec(x_2); +x_1528 = !lean_is_exclusive(x_1); +if (x_1528 == 0) +{ +lean_object* x_1529; lean_object* x_1530; +x_1529 = lean_ctor_get(x_1, 26); +lean_dec(x_1529); +x_1530 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1530, 0, x_3); +lean_ctor_set(x_1, 26, x_1530); +return x_1; +} +else +{ +lean_object* x_1531; lean_object* x_1532; lean_object* x_1533; lean_object* x_1534; lean_object* x_1535; lean_object* x_1536; lean_object* x_1537; lean_object* x_1538; lean_object* x_1539; lean_object* x_1540; lean_object* x_1541; lean_object* x_1542; lean_object* x_1543; lean_object* x_1544; lean_object* x_1545; lean_object* x_1546; lean_object* x_1547; lean_object* x_1548; lean_object* x_1549; lean_object* x_1550; lean_object* x_1551; lean_object* x_1552; lean_object* x_1553; lean_object* x_1554; lean_object* x_1555; lean_object* x_1556; lean_object* x_1557; lean_object* x_1558; lean_object* x_1559; lean_object* x_1560; lean_object* x_1561; +x_1531 = lean_ctor_get(x_1, 0); +x_1532 = lean_ctor_get(x_1, 1); +x_1533 = lean_ctor_get(x_1, 2); +x_1534 = lean_ctor_get(x_1, 3); +x_1535 = lean_ctor_get(x_1, 4); +x_1536 = lean_ctor_get(x_1, 5); +x_1537 = lean_ctor_get(x_1, 6); +x_1538 = lean_ctor_get(x_1, 7); +x_1539 = lean_ctor_get(x_1, 8); +x_1540 = lean_ctor_get(x_1, 9); +x_1541 = lean_ctor_get(x_1, 10); +x_1542 = lean_ctor_get(x_1, 11); +x_1543 = lean_ctor_get(x_1, 12); +x_1544 = lean_ctor_get(x_1, 13); +x_1545 = lean_ctor_get(x_1, 14); +x_1546 = lean_ctor_get(x_1, 15); +x_1547 = lean_ctor_get(x_1, 16); +x_1548 = lean_ctor_get(x_1, 17); +x_1549 = lean_ctor_get(x_1, 18); +x_1550 = lean_ctor_get(x_1, 19); +x_1551 = lean_ctor_get(x_1, 20); +x_1552 = lean_ctor_get(x_1, 21); +x_1553 = lean_ctor_get(x_1, 22); +x_1554 = lean_ctor_get(x_1, 23); +x_1555 = lean_ctor_get(x_1, 24); +x_1556 = lean_ctor_get(x_1, 25); +x_1557 = lean_ctor_get(x_1, 27); +x_1558 = lean_ctor_get(x_1, 28); +x_1559 = lean_ctor_get(x_1, 29); +lean_inc(x_1559); +lean_inc(x_1558); +lean_inc(x_1557); +lean_inc(x_1556); +lean_inc(x_1555); +lean_inc(x_1554); +lean_inc(x_1553); +lean_inc(x_1552); +lean_inc(x_1551); +lean_inc(x_1550); +lean_inc(x_1549); +lean_inc(x_1548); +lean_inc(x_1547); +lean_inc(x_1546); +lean_inc(x_1545); +lean_inc(x_1544); +lean_inc(x_1543); +lean_inc(x_1542); +lean_inc(x_1541); +lean_inc(x_1540); +lean_inc(x_1539); +lean_inc(x_1538); +lean_inc(x_1537); +lean_inc(x_1536); +lean_inc(x_1535); +lean_inc(x_1534); +lean_inc(x_1533); +lean_inc(x_1532); +lean_inc(x_1531); +lean_dec(x_1); +x_1560 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1560, 0, x_3); +x_1561 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1561, 0, x_1531); +lean_ctor_set(x_1561, 1, x_1532); +lean_ctor_set(x_1561, 2, x_1533); +lean_ctor_set(x_1561, 3, x_1534); +lean_ctor_set(x_1561, 4, x_1535); +lean_ctor_set(x_1561, 5, x_1536); +lean_ctor_set(x_1561, 6, x_1537); +lean_ctor_set(x_1561, 7, x_1538); +lean_ctor_set(x_1561, 8, x_1539); +lean_ctor_set(x_1561, 9, x_1540); +lean_ctor_set(x_1561, 10, x_1541); +lean_ctor_set(x_1561, 11, x_1542); +lean_ctor_set(x_1561, 12, x_1543); +lean_ctor_set(x_1561, 13, x_1544); +lean_ctor_set(x_1561, 14, x_1545); +lean_ctor_set(x_1561, 15, x_1546); +lean_ctor_set(x_1561, 16, x_1547); +lean_ctor_set(x_1561, 17, x_1548); +lean_ctor_set(x_1561, 18, x_1549); +lean_ctor_set(x_1561, 19, x_1550); +lean_ctor_set(x_1561, 20, x_1551); +lean_ctor_set(x_1561, 21, x_1552); +lean_ctor_set(x_1561, 22, x_1553); +lean_ctor_set(x_1561, 23, x_1554); +lean_ctor_set(x_1561, 24, x_1555); +lean_ctor_set(x_1561, 25, x_1556); +lean_ctor_set(x_1561, 26, x_1560); +lean_ctor_set(x_1561, 27, x_1557); +lean_ctor_set(x_1561, 28, x_1558); +lean_ctor_set(x_1561, 29, x_1559); +return x_1561; +} +} +case 26: +{ +uint8_t x_1562; +lean_dec(x_2); +x_1562 = !lean_is_exclusive(x_1); +if (x_1562 == 0) +{ +lean_object* x_1563; lean_object* x_1564; +x_1563 = lean_ctor_get(x_1, 27); +lean_dec(x_1563); +x_1564 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1564, 0, x_3); +lean_ctor_set(x_1, 27, x_1564); +return x_1; +} +else +{ +lean_object* x_1565; lean_object* x_1566; lean_object* x_1567; lean_object* x_1568; lean_object* x_1569; lean_object* x_1570; lean_object* x_1571; lean_object* x_1572; lean_object* x_1573; lean_object* x_1574; lean_object* x_1575; lean_object* x_1576; lean_object* x_1577; lean_object* x_1578; lean_object* x_1579; lean_object* x_1580; lean_object* x_1581; lean_object* x_1582; lean_object* x_1583; lean_object* x_1584; lean_object* x_1585; lean_object* x_1586; lean_object* x_1587; lean_object* x_1588; lean_object* x_1589; lean_object* x_1590; lean_object* x_1591; lean_object* x_1592; lean_object* x_1593; lean_object* x_1594; lean_object* x_1595; +x_1565 = lean_ctor_get(x_1, 0); +x_1566 = lean_ctor_get(x_1, 1); +x_1567 = lean_ctor_get(x_1, 2); +x_1568 = lean_ctor_get(x_1, 3); +x_1569 = lean_ctor_get(x_1, 4); +x_1570 = lean_ctor_get(x_1, 5); +x_1571 = lean_ctor_get(x_1, 6); +x_1572 = lean_ctor_get(x_1, 7); +x_1573 = lean_ctor_get(x_1, 8); +x_1574 = lean_ctor_get(x_1, 9); +x_1575 = lean_ctor_get(x_1, 10); +x_1576 = lean_ctor_get(x_1, 11); +x_1577 = lean_ctor_get(x_1, 12); +x_1578 = lean_ctor_get(x_1, 13); +x_1579 = lean_ctor_get(x_1, 14); +x_1580 = lean_ctor_get(x_1, 15); +x_1581 = lean_ctor_get(x_1, 16); +x_1582 = lean_ctor_get(x_1, 17); +x_1583 = lean_ctor_get(x_1, 18); +x_1584 = lean_ctor_get(x_1, 19); +x_1585 = lean_ctor_get(x_1, 20); +x_1586 = lean_ctor_get(x_1, 21); +x_1587 = lean_ctor_get(x_1, 22); +x_1588 = lean_ctor_get(x_1, 23); +x_1589 = lean_ctor_get(x_1, 24); +x_1590 = lean_ctor_get(x_1, 25); +x_1591 = lean_ctor_get(x_1, 26); +x_1592 = lean_ctor_get(x_1, 28); +x_1593 = lean_ctor_get(x_1, 29); +lean_inc(x_1593); +lean_inc(x_1592); +lean_inc(x_1591); +lean_inc(x_1590); +lean_inc(x_1589); +lean_inc(x_1588); +lean_inc(x_1587); +lean_inc(x_1586); +lean_inc(x_1585); +lean_inc(x_1584); +lean_inc(x_1583); +lean_inc(x_1582); +lean_inc(x_1581); +lean_inc(x_1580); +lean_inc(x_1579); +lean_inc(x_1578); +lean_inc(x_1577); +lean_inc(x_1576); +lean_inc(x_1575); +lean_inc(x_1574); +lean_inc(x_1573); +lean_inc(x_1572); +lean_inc(x_1571); +lean_inc(x_1570); +lean_inc(x_1569); +lean_inc(x_1568); +lean_inc(x_1567); +lean_inc(x_1566); +lean_inc(x_1565); +lean_dec(x_1); +x_1594 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1594, 0, x_3); +x_1595 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1595, 0, x_1565); +lean_ctor_set(x_1595, 1, x_1566); +lean_ctor_set(x_1595, 2, x_1567); +lean_ctor_set(x_1595, 3, x_1568); +lean_ctor_set(x_1595, 4, x_1569); +lean_ctor_set(x_1595, 5, x_1570); +lean_ctor_set(x_1595, 6, x_1571); +lean_ctor_set(x_1595, 7, x_1572); +lean_ctor_set(x_1595, 8, x_1573); +lean_ctor_set(x_1595, 9, x_1574); +lean_ctor_set(x_1595, 10, x_1575); +lean_ctor_set(x_1595, 11, x_1576); +lean_ctor_set(x_1595, 12, x_1577); +lean_ctor_set(x_1595, 13, x_1578); +lean_ctor_set(x_1595, 14, x_1579); +lean_ctor_set(x_1595, 15, x_1580); +lean_ctor_set(x_1595, 16, x_1581); +lean_ctor_set(x_1595, 17, x_1582); +lean_ctor_set(x_1595, 18, x_1583); +lean_ctor_set(x_1595, 19, x_1584); +lean_ctor_set(x_1595, 20, x_1585); +lean_ctor_set(x_1595, 21, x_1586); +lean_ctor_set(x_1595, 22, x_1587); +lean_ctor_set(x_1595, 23, x_1588); +lean_ctor_set(x_1595, 24, x_1589); +lean_ctor_set(x_1595, 25, x_1590); +lean_ctor_set(x_1595, 26, x_1591); +lean_ctor_set(x_1595, 27, x_1594); +lean_ctor_set(x_1595, 28, x_1592); +lean_ctor_set(x_1595, 29, x_1593); +return x_1595; +} +} +case 27: +{ +uint8_t x_1596; +lean_dec(x_2); +x_1596 = !lean_is_exclusive(x_1); +if (x_1596 == 0) +{ +lean_object* x_1597; lean_object* x_1598; +x_1597 = lean_ctor_get(x_1, 28); +lean_dec(x_1597); +x_1598 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1598, 0, x_3); +lean_ctor_set(x_1, 28, x_1598); +return x_1; +} +else +{ +lean_object* x_1599; lean_object* x_1600; lean_object* x_1601; lean_object* x_1602; lean_object* x_1603; lean_object* x_1604; lean_object* x_1605; lean_object* x_1606; lean_object* x_1607; lean_object* x_1608; lean_object* x_1609; lean_object* x_1610; lean_object* x_1611; lean_object* x_1612; lean_object* x_1613; lean_object* x_1614; lean_object* x_1615; lean_object* x_1616; lean_object* x_1617; lean_object* x_1618; lean_object* x_1619; lean_object* x_1620; lean_object* x_1621; lean_object* x_1622; lean_object* x_1623; lean_object* x_1624; lean_object* x_1625; lean_object* x_1626; lean_object* x_1627; lean_object* x_1628; lean_object* x_1629; +x_1599 = lean_ctor_get(x_1, 0); +x_1600 = lean_ctor_get(x_1, 1); +x_1601 = lean_ctor_get(x_1, 2); +x_1602 = lean_ctor_get(x_1, 3); +x_1603 = lean_ctor_get(x_1, 4); +x_1604 = lean_ctor_get(x_1, 5); +x_1605 = lean_ctor_get(x_1, 6); +x_1606 = lean_ctor_get(x_1, 7); +x_1607 = lean_ctor_get(x_1, 8); +x_1608 = lean_ctor_get(x_1, 9); +x_1609 = lean_ctor_get(x_1, 10); +x_1610 = lean_ctor_get(x_1, 11); +x_1611 = lean_ctor_get(x_1, 12); +x_1612 = lean_ctor_get(x_1, 13); +x_1613 = lean_ctor_get(x_1, 14); +x_1614 = lean_ctor_get(x_1, 15); +x_1615 = lean_ctor_get(x_1, 16); +x_1616 = lean_ctor_get(x_1, 17); +x_1617 = lean_ctor_get(x_1, 18); +x_1618 = lean_ctor_get(x_1, 19); +x_1619 = lean_ctor_get(x_1, 20); +x_1620 = lean_ctor_get(x_1, 21); +x_1621 = lean_ctor_get(x_1, 22); +x_1622 = lean_ctor_get(x_1, 23); +x_1623 = lean_ctor_get(x_1, 24); +x_1624 = lean_ctor_get(x_1, 25); +x_1625 = lean_ctor_get(x_1, 26); +x_1626 = lean_ctor_get(x_1, 27); +x_1627 = lean_ctor_get(x_1, 29); +lean_inc(x_1627); +lean_inc(x_1626); +lean_inc(x_1625); +lean_inc(x_1624); +lean_inc(x_1623); +lean_inc(x_1622); +lean_inc(x_1621); +lean_inc(x_1620); +lean_inc(x_1619); +lean_inc(x_1618); +lean_inc(x_1617); +lean_inc(x_1616); +lean_inc(x_1615); +lean_inc(x_1614); +lean_inc(x_1613); +lean_inc(x_1612); +lean_inc(x_1611); +lean_inc(x_1610); +lean_inc(x_1609); +lean_inc(x_1608); +lean_inc(x_1607); +lean_inc(x_1606); +lean_inc(x_1605); +lean_inc(x_1604); +lean_inc(x_1603); +lean_inc(x_1602); +lean_inc(x_1601); +lean_inc(x_1600); +lean_inc(x_1599); +lean_dec(x_1); +x_1628 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1628, 0, x_3); +x_1629 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1629, 0, x_1599); +lean_ctor_set(x_1629, 1, x_1600); +lean_ctor_set(x_1629, 2, x_1601); +lean_ctor_set(x_1629, 3, x_1602); +lean_ctor_set(x_1629, 4, x_1603); +lean_ctor_set(x_1629, 5, x_1604); +lean_ctor_set(x_1629, 6, x_1605); +lean_ctor_set(x_1629, 7, x_1606); +lean_ctor_set(x_1629, 8, x_1607); +lean_ctor_set(x_1629, 9, x_1608); +lean_ctor_set(x_1629, 10, x_1609); +lean_ctor_set(x_1629, 11, x_1610); +lean_ctor_set(x_1629, 12, x_1611); +lean_ctor_set(x_1629, 13, x_1612); +lean_ctor_set(x_1629, 14, x_1613); +lean_ctor_set(x_1629, 15, x_1614); +lean_ctor_set(x_1629, 16, x_1615); +lean_ctor_set(x_1629, 17, x_1616); +lean_ctor_set(x_1629, 18, x_1617); +lean_ctor_set(x_1629, 19, x_1618); +lean_ctor_set(x_1629, 20, x_1619); +lean_ctor_set(x_1629, 21, x_1620); +lean_ctor_set(x_1629, 22, x_1621); +lean_ctor_set(x_1629, 23, x_1622); +lean_ctor_set(x_1629, 24, x_1623); +lean_ctor_set(x_1629, 25, x_1624); +lean_ctor_set(x_1629, 26, x_1625); +lean_ctor_set(x_1629, 27, x_1626); +lean_ctor_set(x_1629, 28, x_1628); +lean_ctor_set(x_1629, 29, x_1627); +return x_1629; +} +} +default: +{ +uint8_t x_1630; +lean_dec(x_2); +x_1630 = !lean_is_exclusive(x_1); +if (x_1630 == 0) +{ +lean_object* x_1631; lean_object* x_1632; +x_1631 = lean_ctor_get(x_1, 29); +lean_dec(x_1631); +x_1632 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1632, 0, x_3); +lean_ctor_set(x_1, 29, x_1632); +return x_1; +} +else +{ +lean_object* x_1633; lean_object* x_1634; lean_object* x_1635; lean_object* x_1636; lean_object* x_1637; lean_object* x_1638; lean_object* x_1639; lean_object* x_1640; lean_object* x_1641; lean_object* x_1642; lean_object* x_1643; lean_object* x_1644; lean_object* x_1645; lean_object* x_1646; lean_object* x_1647; lean_object* x_1648; lean_object* x_1649; lean_object* x_1650; lean_object* x_1651; lean_object* x_1652; lean_object* x_1653; lean_object* x_1654; lean_object* x_1655; lean_object* x_1656; lean_object* x_1657; lean_object* x_1658; lean_object* x_1659; lean_object* x_1660; lean_object* x_1661; lean_object* x_1662; lean_object* x_1663; +x_1633 = lean_ctor_get(x_1, 0); +x_1634 = lean_ctor_get(x_1, 1); +x_1635 = lean_ctor_get(x_1, 2); +x_1636 = lean_ctor_get(x_1, 3); +x_1637 = lean_ctor_get(x_1, 4); +x_1638 = lean_ctor_get(x_1, 5); +x_1639 = lean_ctor_get(x_1, 6); +x_1640 = lean_ctor_get(x_1, 7); +x_1641 = lean_ctor_get(x_1, 8); +x_1642 = lean_ctor_get(x_1, 9); +x_1643 = lean_ctor_get(x_1, 10); +x_1644 = lean_ctor_get(x_1, 11); +x_1645 = lean_ctor_get(x_1, 12); +x_1646 = lean_ctor_get(x_1, 13); +x_1647 = lean_ctor_get(x_1, 14); +x_1648 = lean_ctor_get(x_1, 15); +x_1649 = lean_ctor_get(x_1, 16); +x_1650 = lean_ctor_get(x_1, 17); +x_1651 = lean_ctor_get(x_1, 18); +x_1652 = lean_ctor_get(x_1, 19); +x_1653 = lean_ctor_get(x_1, 20); +x_1654 = lean_ctor_get(x_1, 21); +x_1655 = lean_ctor_get(x_1, 22); +x_1656 = lean_ctor_get(x_1, 23); +x_1657 = lean_ctor_get(x_1, 24); +x_1658 = lean_ctor_get(x_1, 25); +x_1659 = lean_ctor_get(x_1, 26); +x_1660 = lean_ctor_get(x_1, 27); +x_1661 = lean_ctor_get(x_1, 28); +lean_inc(x_1661); +lean_inc(x_1660); +lean_inc(x_1659); +lean_inc(x_1658); +lean_inc(x_1657); +lean_inc(x_1656); +lean_inc(x_1655); +lean_inc(x_1654); +lean_inc(x_1653); +lean_inc(x_1652); +lean_inc(x_1651); +lean_inc(x_1650); +lean_inc(x_1649); +lean_inc(x_1648); +lean_inc(x_1647); +lean_inc(x_1646); +lean_inc(x_1645); +lean_inc(x_1644); +lean_inc(x_1643); +lean_inc(x_1642); +lean_inc(x_1641); +lean_inc(x_1640); +lean_inc(x_1639); +lean_inc(x_1638); +lean_inc(x_1637); +lean_inc(x_1636); +lean_inc(x_1635); +lean_inc(x_1634); +lean_inc(x_1633); +lean_dec(x_1); +x_1662 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_1662, 0, x_3); +x_1663 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_1663, 0, x_1633); +lean_ctor_set(x_1663, 1, x_1634); +lean_ctor_set(x_1663, 2, x_1635); +lean_ctor_set(x_1663, 3, x_1636); +lean_ctor_set(x_1663, 4, x_1637); +lean_ctor_set(x_1663, 5, x_1638); +lean_ctor_set(x_1663, 6, x_1639); +lean_ctor_set(x_1663, 7, x_1640); +lean_ctor_set(x_1663, 8, x_1641); +lean_ctor_set(x_1663, 9, x_1642); +lean_ctor_set(x_1663, 10, x_1643); +lean_ctor_set(x_1663, 11, x_1644); +lean_ctor_set(x_1663, 12, x_1645); +lean_ctor_set(x_1663, 13, x_1646); +lean_ctor_set(x_1663, 14, x_1647); +lean_ctor_set(x_1663, 15, x_1648); +lean_ctor_set(x_1663, 16, x_1649); +lean_ctor_set(x_1663, 17, x_1650); +lean_ctor_set(x_1663, 18, x_1651); +lean_ctor_set(x_1663, 19, x_1652); +lean_ctor_set(x_1663, 20, x_1653); +lean_ctor_set(x_1663, 21, x_1654); +lean_ctor_set(x_1663, 22, x_1655); +lean_ctor_set(x_1663, 23, x_1656); +lean_ctor_set(x_1663, 24, x_1657); +lean_ctor_set(x_1663, 25, x_1658); +lean_ctor_set(x_1663, 26, x_1659); +lean_ctor_set(x_1663, 27, x_1660); +lean_ctor_set(x_1663, 28, x_1661); +lean_ctor_set(x_1663, 29, x_1662); +return x_1663; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra(lean_object* x_1, uint8_t x_2) { +_start: +{ +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_4 = lean_int_add(x_1, x_3); +x_5 = lean_int_neg(x_4); +lean_dec(x_4); +return x_5; +} +else +{ +lean_inc(x_1); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra(x_1, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_dec_le(x_1, x_3); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_2); +lean_dec(x_7); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = lean_int_mul(x_2, x_3); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +lean_dec(x_5); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_10); +lean_dec(x_10); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1; +x_2 = l_Array_back_x3f___rarg(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(999999999u); +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Fin_ofNat(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18; +x_2 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_53 = lean_ctor_get(x_1, 26); +lean_inc(x_53); +x_54 = lean_ctor_get(x_1, 23); +lean_inc(x_54); +x_55 = lean_ctor_get(x_1, 25); +lean_inc(x_55); +x_56 = lean_ctor_get(x_1, 4); +lean_inc(x_56); +x_57 = lean_ctor_get(x_1, 5); +lean_inc(x_57); +x_58 = lean_ctor_get(x_1, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_1, 2); +lean_inc(x_59); +x_60 = lean_ctor_get(x_1, 12); +lean_inc(x_60); +x_61 = lean_ctor_get(x_1, 17); +lean_inc(x_61); +x_62 = lean_ctor_get(x_1, 18); +lean_inc(x_62); +x_63 = lean_ctor_get(x_1, 21); +lean_inc(x_63); +x_64 = lean_ctor_get(x_1, 22); +lean_inc(x_64); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_247; +x_247 = lean_ctor_get(x_1, 27); +lean_inc(x_247); +if (lean_obj_tag(x_247) == 0) +{ +lean_object* x_248; +x_248 = lean_ctor_get(x_1, 28); +lean_inc(x_248); +if (lean_obj_tag(x_248) == 0) +{ +lean_object* x_249; +x_249 = lean_ctor_get(x_1, 29); +lean_inc(x_249); +if (lean_obj_tag(x_249) == 0) +{ +lean_object* x_250; +x_250 = l_Std_Time_TimeZone_Offset_zero; +x_65 = x_250; +goto block_246; +} +else +{ +lean_object* x_251; +x_251 = lean_ctor_get(x_249, 0); +lean_inc(x_251); +lean_dec(x_249); +x_65 = x_251; +goto block_246; +} +} +else +{ +lean_object* x_252; +x_252 = lean_ctor_get(x_248, 0); +lean_inc(x_252); +lean_dec(x_248); +x_65 = x_252; +goto block_246; +} +} +else +{ +lean_object* x_253; +x_253 = lean_ctor_get(x_247, 0); +lean_inc(x_253); +lean_dec(x_247); +x_65 = x_253; +goto block_246; +} +} +else +{ +lean_object* x_254; +x_254 = lean_ctor_get(x_53, 0); +lean_inc(x_254); +lean_dec(x_53); +x_65 = x_254; +goto block_246; +} +block_52: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_5; +x_5 = lean_box(0); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_int_neg(x_8); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +x_11 = l_Std_Time_PlainTime_toSeconds(x_10); +x_12 = lean_int_add(x_11, x_9); +lean_dec(x_11); +x_13 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1; +x_14 = lean_int_ediv(x_12, x_13); +lean_dec(x_12); +x_15 = l_Std_Time_PlainTime_toNanoseconds(x_10); +lean_dec(x_10); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +x_17 = lean_int_mul(x_9, x_16); +lean_dec(x_9); +x_18 = lean_int_add(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +x_19 = l_Std_Time_PlainTime_ofNanoseconds(x_18); +lean_dec(x_18); +x_20 = lean_ctor_get(x_7, 0); +lean_inc(x_20); +x_21 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_20); +x_22 = lean_int_add(x_21, x_14); +lean_dec(x_14); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +x_25 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_24); +x_26 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1___boxed), 2, 1); +lean_closure_set(x_26, 0, x_7); +x_27 = lean_mk_thunk(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_3, 0, x_28); +return x_3; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_29 = lean_ctor_get(x_3, 0); +lean_inc(x_29); +lean_dec(x_3); +x_30 = lean_ctor_get(x_4, 0); +x_31 = lean_int_neg(x_30); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +x_33 = l_Std_Time_PlainTime_toSeconds(x_32); +x_34 = lean_int_add(x_33, x_31); +lean_dec(x_33); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1; +x_36 = lean_int_ediv(x_34, x_35); +lean_dec(x_34); +x_37 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_38 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +x_39 = lean_int_mul(x_31, x_38); +lean_dec(x_31); +x_40 = lean_int_add(x_37, x_39); +lean_dec(x_39); +lean_dec(x_37); +x_41 = l_Std_Time_PlainTime_ofNanoseconds(x_40); +lean_dec(x_40); +x_42 = lean_ctor_get(x_29, 0); +lean_inc(x_42); +x_43 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_42); +x_44 = lean_int_add(x_43, x_36); +lean_dec(x_36); +lean_dec(x_43); +x_45 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_44); +lean_dec(x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_41); +x_47 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_46); +x_48 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1___boxed), 2, 1); +lean_closure_set(x_48, 0, x_29); +x_49 = lean_mk_thunk(x_48); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_47); +lean_ctor_set(x_50, 1, x_49); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +return x_51; +} +} +} +block_246: +{ +lean_object* x_66; +if (lean_obj_tag(x_54) == 0) +{ +lean_object* x_241; +x_241 = lean_ctor_get(x_1, 24); +lean_inc(x_241); +if (lean_obj_tag(x_241) == 0) +{ +uint8_t x_242; lean_object* x_243; +x_242 = 1; +lean_inc(x_65); +x_243 = l_Std_Time_TimeZone_Offset_toIsoString(x_65, x_242); +x_66 = x_243; +goto block_240; +} +else +{ +lean_object* x_244; +x_244 = lean_ctor_get(x_241, 0); +lean_inc(x_244); +lean_dec(x_241); +x_66 = x_244; +goto block_240; +} +} +else +{ +lean_object* x_245; +x_245 = lean_ctor_get(x_54, 0); +lean_inc(x_245); +lean_dec(x_54); +x_66 = x_245; +goto block_240; +} +block_240: +{ +lean_object* x_67; +if (lean_obj_tag(x_55) == 0) +{ +uint8_t x_237; lean_object* x_238; +x_237 = 1; +lean_inc(x_65); +x_238 = l_Std_Time_TimeZone_Offset_toIsoString(x_65, x_237); +x_67 = x_238; +goto block_236; +} +else +{ +lean_object* x_239; +x_239 = lean_ctor_get(x_55, 0); +lean_inc(x_239); +lean_dec(x_55); +x_67 = x_239; +goto block_236; +} +block_236: +{ +lean_object* x_68; lean_object* x_130; +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_234; +x_234 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20; +x_130 = x_234; +goto block_233; +} +else +{ +lean_object* x_235; +x_235 = lean_ctor_get(x_56, 0); +lean_inc(x_235); +lean_dec(x_56); +x_130 = x_235; +goto block_233; +} +block_129: +{ +if (lean_obj_tag(x_68) == 0) +{ +lean_object* x_69; +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +x_69 = lean_box(0); +return x_69; +} +else +{ +lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_70 = lean_ctor_get(x_68, 0); +lean_inc(x_70); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + x_71 = x_68; +} else { + lean_dec_ref(x_68); + x_71 = lean_box(0); +} +x_72 = 0; +x_73 = 0; +x_74 = 1; +x_75 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_75, 0, x_65); +lean_ctor_set(x_75, 1, x_67); +lean_ctor_set(x_75, 2, x_66); +lean_ctor_set_uint8(x_75, sizeof(void*)*3, x_72); +lean_ctor_set_uint8(x_75, sizeof(void*)*3 + 1, x_73); +lean_ctor_set_uint8(x_75, sizeof(void*)*3 + 2, x_74); +x_76 = l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1; +lean_inc(x_75); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +x_78 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_70); +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +lean_inc(x_79); +x_80 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2___boxed), 2, 1); +lean_closure_set(x_80, 0, x_79); +x_81 = lean_unsigned_to_nat(0u); +x_82 = l_Array_findIdx_x3f_loop___rarg(x_80, x_76, x_81); +x_83 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +x_84 = lean_int_mul(x_79, x_83); +x_85 = lean_ctor_get(x_78, 1); +lean_inc(x_85); +lean_dec(x_78); +x_86 = lean_int_add(x_84, x_85); +lean_dec(x_85); +lean_dec(x_84); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_101; +lean_dec(x_79); +x_101 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2; +if (lean_obj_tag(x_101) == 0) +{ +x_87 = x_75; +goto block_100; +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_75); +x_102 = lean_ctor_get(x_101, 0); +lean_inc(x_102); +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_87 = x_103; +goto block_100; +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_82, 0); +lean_inc(x_104); +lean_dec(x_82); +x_105 = lean_unsigned_to_nat(1u); +x_106 = lean_nat_sub(x_104, x_105); +x_107 = l_Array_get_x3f___rarg(x_76, x_106); +lean_dec(x_106); +if (lean_obj_tag(x_107) == 0) +{ +lean_dec(x_104); +lean_dec(x_79); +x_87 = x_75; +goto block_100; +} +else +{ +lean_object* x_108; lean_object* x_109; +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l_Array_get_x3f___rarg(x_76, x_104); +lean_dec(x_104); +if (lean_obj_tag(x_109) == 0) +{ +lean_object* x_110; +x_110 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2; +if (lean_obj_tag(x_110) == 0) +{ +lean_dec(x_108); +lean_dec(x_79); +x_87 = x_75; +goto block_100; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint8_t x_118; +lean_dec(x_75); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_108, 1); +lean_inc(x_113); +lean_dec(x_108); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_nat_abs(x_114); +lean_dec(x_114); +x_116 = lean_nat_to_int(x_115); +x_117 = lean_int_sub(x_112, x_116); +lean_dec(x_116); +lean_dec(x_112); +x_118 = lean_int_dec_lt(x_79, x_117); +lean_dec(x_117); +lean_dec(x_79); +if (x_118 == 0) +{ +lean_object* x_119; +lean_dec(x_113); +x_119 = lean_ctor_get(x_111, 1); +lean_inc(x_119); +lean_dec(x_111); +x_87 = x_119; +goto block_100; +} +else +{ +lean_dec(x_111); +x_87 = x_113; +goto block_100; +} +} +} +else +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; +lean_dec(x_75); +x_120 = lean_ctor_get(x_109, 0); +lean_inc(x_120); +lean_dec(x_109); +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_108, 1); +lean_inc(x_122); +lean_dec(x_108); +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_nat_abs(x_123); +lean_dec(x_123); +x_125 = lean_nat_to_int(x_124); +x_126 = lean_int_sub(x_121, x_125); +lean_dec(x_125); +lean_dec(x_121); +x_127 = lean_int_dec_lt(x_79, x_126); +lean_dec(x_126); +lean_dec(x_79); +if (x_127 == 0) +{ +lean_object* x_128; +lean_dec(x_122); +x_128 = lean_ctor_get(x_120, 1); +lean_inc(x_128); +lean_dec(x_120); +x_87 = x_128; +goto block_100; +} +else +{ +lean_dec(x_120); +x_87 = x_122; +goto block_100; +} +} +} +} +block_100: +{ +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_88 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_87); +lean_dec(x_87); +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_int_neg(x_89); +x_91 = lean_int_mul(x_90, x_83); +lean_dec(x_90); +x_92 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1; +x_93 = lean_int_add(x_91, x_92); +lean_dec(x_91); +x_94 = lean_int_add(x_86, x_93); +lean_dec(x_93); +lean_dec(x_86); +x_95 = l_Std_Time_Duration_ofNanoseconds(x_94); +lean_dec(x_94); +lean_inc(x_95); +x_96 = lean_alloc_closure((void*)(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___boxed), 4, 3); +lean_closure_set(x_96, 0, x_95); +lean_closure_set(x_96, 1, x_89); +lean_closure_set(x_96, 2, x_83); +x_97 = lean_mk_thunk(x_96); +x_98 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_95); +lean_ctor_set(x_98, 2, x_77); +lean_ctor_set(x_98, 3, x_88); +if (lean_is_scalar(x_71)) { + x_99 = lean_alloc_ctor(1, 1, 0); +} else { + x_99 = x_71; +} +lean_ctor_set(x_99, 0, x_98); +return x_99; +} +} +} +block_233: +{ +lean_object* x_131; +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_231; +x_231 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16; +x_131 = x_231; +goto block_230; +} +else +{ +lean_object* x_232; +x_232 = lean_ctor_get(x_57, 0); +lean_inc(x_232); +lean_dec(x_57); +x_131 = x_232; +goto block_230; +} +block_230: +{ +lean_object* x_132; +if (lean_obj_tag(x_58) == 0) +{ +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_217; +x_217 = lean_ctor_get(x_1, 1); +lean_inc(x_217); +if (lean_obj_tag(x_217) == 0) +{ +lean_object* x_218; +x_218 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_132 = x_218; +goto block_216; +} +else +{ +lean_object* x_219; uint8_t x_220; lean_object* x_221; +x_219 = lean_ctor_get(x_217, 0); +lean_inc(x_219); +lean_dec(x_217); +x_220 = 1; +x_221 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra(x_219, x_220); +lean_dec(x_219); +x_132 = x_221; +goto block_216; +} +} +else +{ +lean_object* x_222; +x_222 = lean_ctor_get(x_59, 0); +lean_inc(x_222); +lean_dec(x_59); +x_132 = x_222; +goto block_216; +} +} +else +{ +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_223; +x_223 = lean_ctor_get(x_1, 1); +lean_inc(x_223); +if (lean_obj_tag(x_223) == 0) +{ +lean_object* x_224; +lean_dec(x_58); +x_224 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_132 = x_224; +goto block_216; +} +else +{ +lean_object* x_225; lean_object* x_226; uint8_t x_227; lean_object* x_228; +x_225 = lean_ctor_get(x_58, 0); +lean_inc(x_225); +lean_dec(x_58); +x_226 = lean_ctor_get(x_223, 0); +lean_inc(x_226); +lean_dec(x_223); +x_227 = lean_unbox(x_225); +lean_dec(x_225); +x_228 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_convertYearAndEra(x_226, x_227); +lean_dec(x_226); +x_132 = x_228; +goto block_216; +} +} +else +{ +lean_object* x_229; +lean_dec(x_58); +x_229 = lean_ctor_get(x_59, 0); +lean_inc(x_229); +lean_dec(x_59); +x_132 = x_229; +goto block_216; +} +} +block_216: +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_160; lean_object* x_192; +x_133 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1; +x_134 = lean_int_mod(x_132, x_133); +x_135 = l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1; +x_136 = lean_int_dec_eq(x_134, x_135); +lean_dec(x_134); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_202; +x_202 = lean_box(0); +x_192 = x_202; +goto block_201; +} +else +{ +lean_object* x_203; +x_203 = lean_ctor_get(x_1, 13); +lean_inc(x_203); +if (lean_obj_tag(x_203) == 0) +{ +lean_object* x_204; +x_204 = lean_ctor_get(x_1, 14); +lean_inc(x_204); +if (lean_obj_tag(x_204) == 0) +{ +lean_object* x_205; +lean_dec(x_60); +x_205 = lean_box(0); +x_192 = x_205; +goto block_201; +} +else +{ +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; uint8_t x_210; lean_object* x_211; +x_206 = lean_ctor_get(x_60, 0); +lean_inc(x_206); +lean_dec(x_60); +x_207 = lean_ctor_get(x_204, 0); +lean_inc(x_207); +lean_dec(x_204); +x_208 = l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6; +x_209 = lean_int_add(x_207, x_208); +lean_dec(x_207); +x_210 = lean_unbox(x_206); +lean_dec(x_206); +x_211 = l_Std_Time_HourMarker_toAbsolute(x_210, x_209); +lean_dec(x_209); +x_160 = x_211; +goto block_191; +} +} +else +{ +lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; +x_212 = lean_ctor_get(x_60, 0); +lean_inc(x_212); +lean_dec(x_60); +x_213 = lean_ctor_get(x_203, 0); +lean_inc(x_213); +lean_dec(x_203); +x_214 = lean_unbox(x_212); +lean_dec(x_212); +x_215 = l_Std_Time_HourMarker_toAbsolute(x_214, x_213); +lean_dec(x_213); +x_160 = x_215; +goto block_191; +} +} +block_159: +{ +uint8_t x_138; +if (x_136 == 0) +{ +uint8_t x_148; +x_148 = 0; +x_138 = x_148; +goto block_147; +} +else +{ +lean_object* x_149; lean_object* x_150; uint8_t x_151; uint8_t x_152; +x_149 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1; +x_150 = lean_int_mod(x_132, x_149); +x_151 = lean_int_dec_eq(x_150, x_135); +lean_dec(x_150); +x_152 = l_instDecidableNot___rarg(x_151); +if (x_152 == 0) +{ +lean_object* x_153; lean_object* x_154; uint8_t x_155; +x_153 = l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2; +x_154 = lean_int_mod(x_132, x_153); +x_155 = lean_int_dec_eq(x_154, x_135); +lean_dec(x_154); +if (x_155 == 0) +{ +uint8_t x_156; +x_156 = 0; +x_138 = x_156; +goto block_147; +} +else +{ +uint8_t x_157; +x_157 = 1; +x_138 = x_157; +goto block_147; +} +} +else +{ +uint8_t x_158; +x_158 = 1; +x_138 = x_158; +goto block_147; +} +} +block_147: +{ +lean_object* x_139; uint8_t x_140; +x_139 = l_Std_Time_Month_Ordinal_days(x_138, x_130); +x_140 = lean_int_dec_le(x_131, x_139); +lean_dec(x_139); +if (x_140 == 0) +{ +lean_object* x_141; +lean_dec(x_137); +lean_dec(x_132); +lean_dec(x_131); +lean_dec(x_130); +x_141 = lean_box(0); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_142; +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +x_142 = lean_ctor_get(x_2, 0); +x_3 = x_141; +x_4 = x_142; +goto block_52; +} +else +{ +x_68 = x_141; +goto block_129; +} +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_143, 0, x_132); +lean_ctor_set(x_143, 1, x_130); +lean_ctor_set(x_143, 2, x_131); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_137); +x_145 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_145, 0, x_144); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_146; +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_65); +x_146 = lean_ctor_get(x_2, 0); +x_3 = x_145; +x_4 = x_146; +goto block_52; +} +else +{ +x_68 = x_145; +goto block_129; +} +} +} +} +block_191: +{ +lean_object* x_161; +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_189; +x_189 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10; +x_161 = x_189; +goto block_188; +} +else +{ +lean_object* x_190; +x_190 = lean_ctor_get(x_61, 0); +lean_inc(x_190); +lean_dec(x_61); +x_161 = x_190; +goto block_188; +} +block_188: +{ +lean_object* x_162; +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_186; +x_186 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11; +x_162 = x_186; +goto block_185; +} +else +{ +lean_object* x_187; +x_187 = lean_ctor_get(x_62, 0); +lean_inc(x_187); +lean_dec(x_62); +x_162 = x_187; +goto block_185; +} +block_185: +{ +lean_object* x_163; +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_173; +x_173 = lean_ctor_get(x_1, 19); +lean_inc(x_173); +if (lean_obj_tag(x_173) == 0) +{ +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_174; +x_174 = lean_ctor_get(x_1, 20); +lean_inc(x_174); +lean_dec(x_1); +if (lean_obj_tag(x_174) == 0) +{ +lean_object* x_175; lean_object* x_176; +x_175 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5; +x_176 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_176, 0, x_160); +lean_ctor_set(x_176, 1, x_161); +lean_ctor_set(x_176, 2, x_162); +lean_ctor_set(x_176, 3, x_175); +x_137 = x_176; +goto block_159; +} +else +{ +lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; +lean_dec(x_162); +lean_dec(x_161); +lean_dec(x_160); +x_177 = lean_ctor_get(x_174, 0); +lean_inc(x_177); +lean_dec(x_174); +x_178 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3; +x_179 = lean_int_mul(x_177, x_178); +lean_dec(x_177); +x_180 = l_Std_Time_PlainTime_ofNanoseconds(x_179); +lean_dec(x_179); +x_137 = x_180; +goto block_159; +} +} +else +{ +lean_object* x_181; lean_object* x_182; +lean_dec(x_162); +lean_dec(x_161); +lean_dec(x_160); +lean_dec(x_1); +x_181 = lean_ctor_get(x_64, 0); +lean_inc(x_181); +lean_dec(x_64); +x_182 = l_Std_Time_PlainTime_ofNanoseconds(x_181); +lean_dec(x_181); +x_137 = x_182; +goto block_159; +} +} +else +{ +lean_object* x_183; +x_183 = lean_ctor_get(x_173, 0); +lean_inc(x_183); +lean_dec(x_173); +x_163 = x_183; +goto block_172; +} +} +else +{ +lean_object* x_184; +x_184 = lean_ctor_get(x_63, 0); +lean_inc(x_184); +lean_dec(x_63); +x_163 = x_184; +goto block_172; +} +block_172: +{ +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_164; +x_164 = lean_ctor_get(x_1, 20); +lean_inc(x_164); +lean_dec(x_1); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_165, 0, x_160); +lean_ctor_set(x_165, 1, x_161); +lean_ctor_set(x_165, 2, x_162); +lean_ctor_set(x_165, 3, x_163); +x_137 = x_165; +goto block_159; +} +else +{ +lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +lean_dec(x_163); +lean_dec(x_162); +lean_dec(x_161); +lean_dec(x_160); +x_166 = lean_ctor_get(x_164, 0); +lean_inc(x_166); +lean_dec(x_164); +x_167 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3; +x_168 = lean_int_mul(x_166, x_167); +lean_dec(x_166); +x_169 = l_Std_Time_PlainTime_ofNanoseconds(x_168); +lean_dec(x_168); +x_137 = x_169; +goto block_159; +} +} +else +{ +lean_object* x_170; lean_object* x_171; +lean_dec(x_163); +lean_dec(x_162); +lean_dec(x_161); +lean_dec(x_160); +lean_dec(x_1); +x_170 = lean_ctor_get(x_64, 0); +lean_inc(x_170); +lean_dec(x_64); +x_171 = l_Std_Time_PlainTime_ofNanoseconds(x_170); +lean_dec(x_170); +x_137 = x_171; +goto block_159; +} +} +} +} +} +block_201: +{ +lean_object* x_193; +lean_dec(x_192); +x_193 = lean_ctor_get(x_1, 15); +lean_inc(x_193); +if (lean_obj_tag(x_193) == 0) +{ +lean_object* x_194; +x_194 = lean_ctor_get(x_1, 16); +lean_inc(x_194); +if (lean_obj_tag(x_194) == 0) +{ +x_160 = x_135; +goto block_191; +} +else +{ +lean_object* x_195; +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +lean_dec(x_194); +x_160 = x_195; +goto block_191; +} +} +else +{ +lean_object* x_196; +x_196 = lean_ctor_get(x_1, 16); +lean_inc(x_196); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_197; lean_object* x_198; lean_object* x_199; +x_197 = lean_ctor_get(x_193, 0); +lean_inc(x_197); +lean_dec(x_193); +x_198 = l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8; +x_199 = lean_int_add(x_197, x_198); +lean_dec(x_197); +x_160 = x_199; +goto block_191; +} +else +{ +lean_object* x_200; +lean_dec(x_193); +x_200 = lean_ctor_get(x_196, 0); +lean_inc(x_200); +lean_dec(x_196); +x_160 = x_200; +goto block_191; +} +} +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parseWithDate(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Std_Internal_Parsec_String_pstring(x_4, x_3); +lean_dec(x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_5, 1); +lean_dec(x_7); +lean_ctor_set(x_5, 1, x_1); +return x_5; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_1); +return x_9; +} +} +else +{ +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_5); +if (x_10 == 0) +{ +return x_5; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_5, 0); +x_12 = lean_ctor_get(x_5, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_5); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +lean_dec(x_2); +lean_inc(x_14); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith(x_14, x_3); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 1); +x_18 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_insert(x_1, x_14, x_17); +lean_ctor_set(x_15, 1, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_insert(x_1, x_14, x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +else +{ +uint8_t x_23; +lean_dec(x_14); +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_15); +if (x_23 == 0) +{ +return x_15; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_15, 0); +x_25 = lean_ctor_get(x_15, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_15); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1; +x_3 = l_Std_Internal_Parsec_String_Parser_run___rarg(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +return x_3; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_3); +if (x_7 == 0) +{ +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_8); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_spec___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_GenericFormat_spec(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___rarg), 1, 0); +return x_2; +} +} +static lean_object* _init_l_Std_Time_GenericFormat_spec_x21___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Format.Basic", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Std_Time_GenericFormat_spec_x21___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.GenericFormat.spec!", 28, 28); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec_x21(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1; +x_4 = l_Std_Internal_Parsec_String_Parser_run___rarg(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_GenericFormat_spec_x21___closed__1; +x_7 = l_Std_Time_GenericFormat_spec_x21___closed__2; +x_8 = lean_unsigned_to_nat(1380u); +x_9 = lean_unsigned_to_nat(18u); +x_10 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_6, x_7, x_8, x_9, x_5); +lean_dec(x_5); +x_11 = l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___rarg(x_10); +return x_11; +} +else +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_dec(x_4); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_panic___at_Std_Time_GenericFormat_spec_x21___spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_spec_x21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_spec_x21(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_5); +lean_dec(x_7); +x_9 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1; +x_13 = lean_int_mul(x_5, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_7; +lean_dec(x_4); +lean_dec(x_1); +x_7 = l_List_reverse___rarg(x_6); +return x_7; +} +else +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1___boxed), 3, 2); +lean_closure_set(x_13, 0, x_12); +lean_closure_set(x_13, 1, x_11); +x_14 = lean_mk_thunk(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_12); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(x_11, x_15, x_9); +lean_dec(x_15); +lean_dec(x_11); +lean_ctor_set(x_5, 1, x_6); +lean_ctor_set(x_5, 0, x_16); +{ +lean_object* _tmp_4 = x_10; +lean_object* _tmp_5 = x_5; +x_5 = _tmp_4; +x_6 = _tmp_5; +} +goto _start; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_18 = lean_ctor_get(x_5, 0); +x_19 = lean_ctor_get(x_5, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_5); +x_20 = lean_ctor_get(x_1, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_21); +x_22 = lean_alloc_closure((void*)(l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1___boxed), 3, 2); +lean_closure_set(x_22, 0, x_21); +lean_closure_set(x_22, 1, x_20); +x_23 = lean_mk_thunk(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_21); +lean_ctor_set(x_24, 1, x_23); +x_25 = l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(x_20, x_24, x_18); +lean_dec(x_24); +lean_dec(x_20); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_6); +x_5 = x_19; +x_6 = x_26; +goto _start; +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_5); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_5, 0); +x_30 = lean_ctor_get(x_5, 1); +x_31 = l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(x_2, x_4, x_29); +lean_ctor_set(x_5, 1, x_6); +lean_ctor_set(x_5, 0, x_31); +{ +lean_object* _tmp_4 = x_30; +lean_object* _tmp_5 = x_5; +x_5 = _tmp_4; +x_6 = _tmp_5; +} +goto _start; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_5, 0); +x_34 = lean_ctor_get(x_5, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_5); +x_35 = l___private_Std_Time_Format_Basic_0__Std_Time_formatPartWithDate(x_2, x_4, x_33); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_6); +x_5 = x_34; +x_6 = x_36; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_format(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_box(0); +lean_inc(x_3); +x_6 = l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1(x_1, x_2, x_3, x_4, x_3, x_5); +lean_dec(x_3); +x_7 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_8 = l_List_foldl___at_String_join___spec__1(x_7, x_6); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_List_mapTR_loop___at_Std_Time_GenericFormat_format___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_format___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_GenericFormat_format(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("could not parse the date", 24, 24); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build(x_2, x_1); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1; +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_4); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_5, 0); +lean_inc(x_8); +lean_dec(x_5); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_4); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_3, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_dec(x_3); +x_12 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parseWithDate(x_2, x_10, x_4); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_2 = x_14; +x_3 = x_11; +x_4 = x_13; +goto _start; +} +else +{ +uint8_t x_16; +lean_dec(x_11); +x_16 = !lean_is_exclusive(x_12); +if (x_16 == 0) +{ +return x_12; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_12, 0); +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_12); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(0, 30, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +lean_ctor_set(x_2, 4, x_1); +lean_ctor_set(x_2, 5, x_1); +lean_ctor_set(x_2, 6, x_1); +lean_ctor_set(x_2, 7, x_1); +lean_ctor_set(x_2, 8, x_1); +lean_ctor_set(x_2, 9, x_1); +lean_ctor_set(x_2, 10, x_1); +lean_ctor_set(x_2, 11, x_1); +lean_ctor_set(x_2, 12, x_1); +lean_ctor_set(x_2, 13, x_1); +lean_ctor_set(x_2, 14, x_1); +lean_ctor_set(x_2, 15, x_1); +lean_ctor_set(x_2, 16, x_1); +lean_ctor_set(x_2, 17, x_1); +lean_ctor_set(x_2, 18, x_1); +lean_ctor_set(x_2, 19, x_1); +lean_ctor_set(x_2, 20, x_1); +lean_ctor_set(x_2, 21, x_1); +lean_ctor_set(x_2, 22, x_1); +lean_ctor_set(x_2, 23, x_1); +lean_ctor_set(x_2, 24, x_1); +lean_ctor_set(x_2, 25, x_1); +lean_ctor_set(x_2, 26, x_1); +lean_ctor_set(x_2, 27, x_1); +lean_ctor_set(x_2, 28, x_1); +lean_ctor_set(x_2, 29, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go(x_2, x_4, x_1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("invalid date.", 13, 13); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser_go___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1; +x_5 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +lean_dec(x_2); +x_7 = lean_ctor_get(x_3, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +x_9 = lean_string_utf8_byte_size(x_7); +lean_dec(x_7); +x_10 = lean_nat_dec_lt(x_8, x_9); +lean_dec(x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_3); +lean_ctor_set(x_11, 1, x_6); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_6); +x_12 = l_Std_Internal_Parsec_expectedEndOfInput; +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_3); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +} +else +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_1, 0); +lean_inc(x_14); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Std_Internal_Parsec_String_pstring(x_16, x_3); +lean_dec(x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_1 = x_15; +x_3 = x_18; +goto _start; +} +else +{ +uint8_t x_20; +lean_dec(x_15); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) +{ +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_1, 1); +lean_inc(x_24); +lean_dec(x_1); +x_25 = lean_ctor_get(x_14, 0); +lean_inc(x_25); +lean_dec(x_14); +x_26 = l___private_Std_Time_Format_Basic_0__Std_Time_parseWith(x_25, x_3); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_apply_1(x_2, x_28); +x_1 = x_24; +x_2 = x_29; +x_3 = x_27; +goto _start; +} +else +{ +uint8_t x_31; +lean_dec(x_24); +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) +{ +return x_26; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_26); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser_go(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_builderParser_go___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_GenericFormat_builderParser_go___rarg(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_builderParser(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_builderParser___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser(x_1, x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_ctor_get(x_6, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_6, 1); +lean_inc(x_9); +x_10 = lean_string_utf8_byte_size(x_8); +lean_dec(x_8); +x_11 = lean_nat_dec_lt(x_9, x_10); +lean_dec(x_10); +lean_dec(x_9); +if (x_11 == 0) +{ +return x_4; +} +else +{ +lean_object* x_12; +lean_dec(x_7); +x_12 = l_Std_Internal_Parsec_expectedEndOfInput; +lean_ctor_set_tag(x_4, 1); +lean_ctor_set(x_4, 1, x_12); +return x_4; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +x_17 = lean_string_utf8_byte_size(x_15); +lean_dec(x_15); +x_18 = lean_nat_dec_lt(x_16, x_17); +lean_dec(x_17); +lean_dec(x_16); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_14); +x_20 = l_Std_Internal_Parsec_expectedEndOfInput; +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_4); +if (x_22 == 0) +{ +return x_4; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_4, 0); +x_24 = lean_ctor_get(x_4, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_4); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_parse___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_1); +x_5 = l_Std_Internal_Parsec_String_Parser_run___rarg(x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_GenericFormat_parse___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_Awareness_instInhabitedType(x_1); +x_4 = lean_panic_fn(x_3, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_GenericFormat_parse_x21___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.GenericFormat.parse!", 29, 29); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parse_x21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +lean_inc(x_1); +x_4 = l_Std_Time_GenericFormat_parse(x_1, x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_GenericFormat_spec_x21___closed__1; +x_7 = l_Std_Time_GenericFormat_parse_x21___closed__1; +x_8 = lean_unsigned_to_nat(1432u); +x_9 = lean_unsigned_to_nat(18u); +x_10 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_6, x_7, x_8, x_9, x_5); +lean_dec(x_5); +x_11 = l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1(x_1, x_10); +lean_dec(x_1); +return x_11; +} +else +{ +lean_object* x_12; +lean_dec(x_1); +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +lean_dec(x_4); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_panic___at_Std_Time_GenericFormat_parse_x21___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_builderParser___rarg), 3, 2); +lean_closure_set(x_4, 0, x_1); +lean_closure_set(x_4, 1, x_2); +x_5 = l_Std_Internal_Parsec_String_Parser_run___rarg(x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_parseBuilder___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_parseBuilder(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.GenericFormat.parseBuilder!", 36, 36); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Std_Time_GenericFormat_spec_x21___closed__1; +x_8 = l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1; +x_9 = lean_unsigned_to_nat(1446u); +x_10 = lean_unsigned_to_nat(18u); +x_11 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_7, x_8, x_9, x_10, x_6); +lean_dec(x_6); +x_12 = l_panic___rarg(x_1, x_11); +return x_12; +} +else +{ +lean_object* x_13; +lean_dec(x_1); +x_13 = lean_ctor_get(x_5, 0); +lean_inc(x_13); +lean_dec(x_5); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_parseBuilder_x21___rarg), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_parseBuilder_x21___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_GenericFormat_parseBuilder_x21(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric_go(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_1); +x_4 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_4, 0, x_2); +return x_4; +} +else +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_6); +lean_dec(x_3); +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_string_append(x_2, x_7); +lean_dec(x_7); +x_2 = x_8; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_dec(x_3); +x_11 = lean_ctor_get(x_5, 0); +lean_inc(x_11); +lean_dec(x_5); +lean_inc(x_1); +lean_inc(x_11); +x_12 = lean_apply_1(x_1, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +lean_dec(x_1); +x_13 = lean_box(0); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith(x_11, x_14); +x_16 = lean_string_append(x_2, x_15); +lean_dec(x_15); +x_2 = x_16; +x_3 = x_10; +goto _start; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_4 = l_Std_Time_GenericFormat_formatGeneric_go(x_2, x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_formatGeneric___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatGeneric___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_GenericFormat_formatGeneric(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l___private_Std_Time_Format_Basic_0__Std_Time_formatWith(x_1, x_4); +x_6 = lean_string_append(x_2, x_5); +lean_dec(x_5); +x_7 = l_Std_Time_GenericFormat_formatBuilder_go(x_6, x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder_go(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_string_append(x_1, x_5); +lean_dec(x_5); +x_1 = x_6; +x_2 = x_4; +goto _start; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +lean_dec(x_3); +x_10 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_formatBuilder_go___lambda__1), 4, 3); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_1); +lean_closure_set(x_10, 2, x_8); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1; +x_3 = l_Std_Time_GenericFormat_formatBuilder_go(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_formatBuilder___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_GenericFormat_formatBuilder___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_GenericFormat_formatBuilder(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_formatBuilder___rarg), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instFormatGenericFormatFormatTypeString(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_alloc_closure((void*)(l_Std_Time_GenericFormat_parseBuilder___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +x_3 = l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +lean_object* initialize_Std_Internal_Parsec(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Format_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Parsec(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Text_noConfusion___rarg___closed__1 = _init_l_Std_Time_Text_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Text_noConfusion___rarg___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__18); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__19); +l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprText____x40_Std_Time_Format_Basic___hyg_14____closed__20); +l_Std_Time_instReprText___closed__1 = _init_l_Std_Time_instReprText___closed__1(); +lean_mark_persistent(l_Std_Time_instReprText___closed__1); +l_Std_Time_instReprText = _init_l_Std_Time_instReprText(); +lean_mark_persistent(l_Std_Time_instReprText); +l_Std_Time_instInhabitedText = _init_l_Std_Time_instInhabitedText(); +l_Std_Time_Text_classify___closed__1 = _init_l_Std_Time_Text_classify___closed__1(); +lean_mark_persistent(l_Std_Time_Text_classify___closed__1); +l_Std_Time_Text_classify___closed__2 = _init_l_Std_Time_Text_classify___closed__2(); +lean_mark_persistent(l_Std_Time_Text_classify___closed__2); +l_Std_Time_Text_classify___closed__3 = _init_l_Std_Time_Text_classify___closed__3(); +lean_mark_persistent(l_Std_Time_Text_classify___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprNumber____x40_Std_Time_Format_Basic___hyg_204____closed__13); +l_Std_Time_instReprNumber___closed__1 = _init_l_Std_Time_instReprNumber___closed__1(); +lean_mark_persistent(l_Std_Time_instReprNumber___closed__1); +l_Std_Time_instReprNumber = _init_l_Std_Time_instReprNumber(); +lean_mark_persistent(l_Std_Time_instReprNumber); +l_Std_Time_instInhabitedNumber = _init_l_Std_Time_instInhabitedNumber(); +lean_mark_persistent(l_Std_Time_instInhabitedNumber); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFraction____x40_Std_Time_Format_Basic___hyg_313____closed__9); +l_Std_Time_instReprFraction___closed__1 = _init_l_Std_Time_instReprFraction___closed__1(); +lean_mark_persistent(l_Std_Time_instReprFraction___closed__1); +l_Std_Time_instReprFraction = _init_l_Std_Time_instReprFraction(); +lean_mark_persistent(l_Std_Time_instReprFraction); +l_Std_Time_instInhabitedFraction = _init_l_Std_Time_instInhabitedFraction(); +lean_mark_persistent(l_Std_Time_instInhabitedFraction); +l_Std_Time_Fraction_classify___closed__1 = _init_l_Std_Time_Fraction_classify___closed__1(); +lean_mark_persistent(l_Std_Time_Fraction_classify___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprYear____x40_Std_Time_Format_Basic___hyg_466____closed__15); +l_Std_Time_instReprYear___closed__1 = _init_l_Std_Time_instReprYear___closed__1(); +lean_mark_persistent(l_Std_Time_instReprYear___closed__1); +l_Std_Time_instReprYear = _init_l_Std_Time_instReprYear(); +lean_mark_persistent(l_Std_Time_instReprYear); +l_Std_Time_instInhabitedYear = _init_l_Std_Time_instInhabitedYear(); +lean_mark_persistent(l_Std_Time_instInhabitedYear); +l_Std_Time_Year_classify___closed__1 = _init_l_Std_Time_Year_classify___closed__1(); +lean_mark_persistent(l_Std_Time_Year_classify___closed__1); +l_Std_Time_Year_classify___closed__2 = _init_l_Std_Time_Year_classify___closed__2(); +lean_mark_persistent(l_Std_Time_Year_classify___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneId____x40_Std_Time_Format_Basic___hyg_668____closed__12); +l_Std_Time_instReprZoneId___closed__1 = _init_l_Std_Time_instReprZoneId___closed__1(); +lean_mark_persistent(l_Std_Time_instReprZoneId___closed__1); +l_Std_Time_instReprZoneId = _init_l_Std_Time_instReprZoneId(); +lean_mark_persistent(l_Std_Time_instReprZoneId); +l_Std_Time_instInhabitedZoneId = _init_l_Std_Time_instInhabitedZoneId(); +l_Std_Time_ZoneId_classify___closed__1 = _init_l_Std_Time_ZoneId_classify___closed__1(); +lean_mark_persistent(l_Std_Time_ZoneId_classify___closed__1); +l_Std_Time_ZoneId_classify___closed__2 = _init_l_Std_Time_ZoneId_classify___closed__2(); +lean_mark_persistent(l_Std_Time_ZoneId_classify___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprZoneName____x40_Std_Time_Format_Basic___hyg_804____closed__12); +l_Std_Time_instReprZoneName___closed__1 = _init_l_Std_Time_instReprZoneName___closed__1(); +lean_mark_persistent(l_Std_Time_instReprZoneName___closed__1); +l_Std_Time_instReprZoneName = _init_l_Std_Time_instReprZoneName(); +lean_mark_persistent(l_Std_Time_instReprZoneName); +l_Std_Time_instInhabitedZoneName = _init_l_Std_Time_instInhabitedZoneName(); +l_Std_Time_ZoneName_classify___closed__1 = _init_l_Std_Time_ZoneName_classify___closed__1(); +lean_mark_persistent(l_Std_Time_ZoneName_classify___closed__1); +l_Std_Time_ZoneName_classify___closed__2 = _init_l_Std_Time_ZoneName_classify___closed__2(); +lean_mark_persistent(l_Std_Time_ZoneName_classify___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__18); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__19); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__20); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__21); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__22); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__23); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__24); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__25); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__26); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__27); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__28); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__29); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetX____x40_Std_Time_Format_Basic___hyg_995____closed__30); +l_Std_Time_instReprOffsetX___closed__1 = _init_l_Std_Time_instReprOffsetX___closed__1(); +lean_mark_persistent(l_Std_Time_instReprOffsetX___closed__1); +l_Std_Time_instReprOffsetX = _init_l_Std_Time_instReprOffsetX(); +lean_mark_persistent(l_Std_Time_instReprOffsetX); +l_Std_Time_instInhabitedOffsetX = _init_l_Std_Time_instInhabitedOffsetX(); +l_Std_Time_OffsetX_classify___closed__1 = _init_l_Std_Time_OffsetX_classify___closed__1(); +lean_mark_persistent(l_Std_Time_OffsetX_classify___closed__1); +l_Std_Time_OffsetX_classify___closed__2 = _init_l_Std_Time_OffsetX_classify___closed__2(); +lean_mark_persistent(l_Std_Time_OffsetX_classify___closed__2); +l_Std_Time_OffsetX_classify___closed__3 = _init_l_Std_Time_OffsetX_classify___closed__3(); +lean_mark_persistent(l_Std_Time_OffsetX_classify___closed__3); +l_Std_Time_OffsetX_classify___closed__4 = _init_l_Std_Time_OffsetX_classify___closed__4(); +lean_mark_persistent(l_Std_Time_OffsetX_classify___closed__4); +l_Std_Time_OffsetX_classify___closed__5 = _init_l_Std_Time_OffsetX_classify___closed__5(); +lean_mark_persistent(l_Std_Time_OffsetX_classify___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetO____x40_Std_Time_Format_Basic___hyg_1272____closed__12); +l_Std_Time_instReprOffsetO___closed__1 = _init_l_Std_Time_instReprOffsetO___closed__1(); +lean_mark_persistent(l_Std_Time_instReprOffsetO___closed__1); +l_Std_Time_instReprOffsetO = _init_l_Std_Time_instReprOffsetO(); +lean_mark_persistent(l_Std_Time_instReprOffsetO); +l_Std_Time_instInhabitedOffsetO = _init_l_Std_Time_instInhabitedOffsetO(); +l_Std_Time_OffsetO_classify___closed__1 = _init_l_Std_Time_OffsetO_classify___closed__1(); +lean_mark_persistent(l_Std_Time_OffsetO_classify___closed__1); +l_Std_Time_OffsetO_classify___closed__2 = _init_l_Std_Time_OffsetO_classify___closed__2(); +lean_mark_persistent(l_Std_Time_OffsetO_classify___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprOffsetZ____x40_Std_Time_Format_Basic___hyg_1398____closed__18); +l_Std_Time_instReprOffsetZ___closed__1 = _init_l_Std_Time_instReprOffsetZ___closed__1(); +lean_mark_persistent(l_Std_Time_instReprOffsetZ___closed__1); +l_Std_Time_instReprOffsetZ = _init_l_Std_Time_instReprOffsetZ(); +lean_mark_persistent(l_Std_Time_instReprOffsetZ); +l_Std_Time_instInhabitedOffsetZ = _init_l_Std_Time_instInhabitedOffsetZ(); +l_Std_Time_OffsetZ_classify___closed__1 = _init_l_Std_Time_OffsetZ_classify___closed__1(); +lean_mark_persistent(l_Std_Time_OffsetZ_classify___closed__1); +l_Std_Time_OffsetZ_classify___closed__2 = _init_l_Std_Time_OffsetZ_classify___closed__2(); +lean_mark_persistent(l_Std_Time_OffsetZ_classify___closed__2); +l_Std_Time_OffsetZ_classify___closed__3 = _init_l_Std_Time_OffsetZ_classify___closed__3(); +lean_mark_persistent(l_Std_Time_OffsetZ_classify___closed__3); +l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1 = _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1(); +lean_mark_persistent(l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__1); +l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2 = _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2(); +lean_mark_persistent(l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__2); +l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3 = _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3(); +lean_mark_persistent(l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__3); +l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4 = _init_l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4(); +lean_mark_persistent(l_Sum_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____spec__1___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__18); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__19); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__20); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__21); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__22); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__23); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__24); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__25); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__26); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__27); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__28); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__29); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__30); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__31); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__32); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__33); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__34); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__35); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__36); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__37); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__38); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__39); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__40); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__41); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__42); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__43); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__44); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__45); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__46); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__47); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__48); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__49); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__50); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__51); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__52); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__53); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__54); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__55); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__56); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__57); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__58); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__59); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__60); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__61); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__62); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__63); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__64); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__65); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__66); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__67); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__68); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__69); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__70); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__71); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__72); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__73); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__74); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__75); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__76); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__77); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__78); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__79); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__80); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__81); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__82); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__83); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__84); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__85); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__86); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__87); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__88); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__89); +l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprModifier____x40_Std_Time_Format_Basic___hyg_1836____closed__90); +l_Std_Time_instReprModifier___closed__1 = _init_l_Std_Time_instReprModifier___closed__1(); +lean_mark_persistent(l_Std_Time_instReprModifier___closed__1); +l_Std_Time_instReprModifier = _init_l_Std_Time_instReprModifier(); +lean_mark_persistent(l_Std_Time_instReprModifier); +l_Std_Time_instInhabitedModifier___closed__1 = _init_l_Std_Time_instInhabitedModifier___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedModifier___closed__1); +l_Std_Time_instInhabitedModifier = _init_l_Std_Time_instInhabitedModifier(); +lean_mark_persistent(l_Std_Time_instInhabitedModifier); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMod___rarg___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseText___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFraction___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseYear___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetX___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetZ___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffsetO___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseNumberText___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__18); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__19); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__20); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__21); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__22); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__23); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__24); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__25); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__26); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__27); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__28); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__29); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__30); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__31); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__32); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__33); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__34); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__35); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__36); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__37); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__38); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__39); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__40); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__41); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__42); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__43); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__44); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__45); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__46); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__47); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__48); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__49); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__50); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__51); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__52); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__53); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__54); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__55); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__56); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__57); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__58); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__59); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__60); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__61); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__62); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__63); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__64); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__65); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__66); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__67); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__68); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__69); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__70); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__71); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__72); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__73); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__74); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__75); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__76); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__77); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__78); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__79); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__80); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__81); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__82); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__83); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__84); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__85); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__86); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__87); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__88); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseModifier___closed__89); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprFormatPart____x40_Std_Time_Format_Basic___hyg_3701____closed__6); +l_Std_Time_instReprFormatPart___closed__1 = _init_l_Std_Time_instReprFormatPart___closed__1(); +lean_mark_persistent(l_Std_Time_instReprFormatPart___closed__1); +l_Std_Time_instReprFormatPart = _init_l_Std_Time_instReprFormatPart(); +lean_mark_persistent(l_Std_Time_instReprFormatPart); +l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__1); +l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__2); +l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__1___closed__3); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__1); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__2); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__3); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__4); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__5); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__6); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__7); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__8); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__9); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__10); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__11); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__12); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__13); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__14); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__15); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__16); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__17); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__18); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__19); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__20); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__21); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__22); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__23); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__24); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__25); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__26); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__27); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__28); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__29); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__30); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__31); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__32); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__33); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__34); +l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35 = _init_l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___lambda__2___closed__35); +l_Std_Time_Awareness_instInhabitedType___closed__1 = _init_l_Std_Time_Awareness_instInhabitedType___closed__1(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__1); +l_Std_Time_Awareness_instInhabitedType___closed__2 = _init_l_Std_Time_Awareness_instInhabitedType___closed__2(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__2); +l_Std_Time_Awareness_instInhabitedType___closed__3 = _init_l_Std_Time_Awareness_instInhabitedType___closed__3(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__3); +l_Std_Time_Awareness_instInhabitedType___closed__4 = _init_l_Std_Time_Awareness_instInhabitedType___closed__4(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__4); +l_Std_Time_Awareness_instInhabitedType___closed__5 = _init_l_Std_Time_Awareness_instInhabitedType___closed__5(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__5); +l_Std_Time_Awareness_instInhabitedType___closed__6 = _init_l_Std_Time_Awareness_instInhabitedType___closed__6(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__6); +l_Std_Time_Awareness_instInhabitedType___closed__7 = _init_l_Std_Time_Awareness_instInhabitedType___closed__7(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__7); +l_Std_Time_Awareness_instInhabitedType___closed__8 = _init_l_Std_Time_Awareness_instInhabitedType___closed__8(); +lean_mark_persistent(l_Std_Time_Awareness_instInhabitedType___closed__8); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__1); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__2); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__3); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__4); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__5); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__6); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__7); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__8); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__9); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__10); +l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11 = _init_l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11(); +lean_mark_persistent(l_List_repr___at___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____spec__1___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_reprGenericFormat____x40_Std_Time_Format_Basic___hyg_3956____rarg___closed__5); +l_Std_Time_instReprGenericFormat___closed__1 = _init_l_Std_Time_instReprGenericFormat___closed__1(); +lean_mark_persistent(l_Std_Time_instReprGenericFormat___closed__1); +l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1 = _init_l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1(); +lean_mark_persistent(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___spec__1___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__1___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__1(); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___lambda__4___closed__2(); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseFormatPart___boxed__const__2); +l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_specParser___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_specParse___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_pad___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthLong___closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthShort___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMonthNarrow___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayLong___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayShort___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWeekdayNarrow___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraShort___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatEraNarrow___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterNumber___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterShort___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatQuarterLong___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerShort___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatMarkerNarrow___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_toSigned___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_toIsoString___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_formatWith___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_dateFromModifier___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseMonthLong___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseNum___closed__2); +l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1 = _init_l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1(); +lean_mark_persistent(l_Std_Internal_Parsec_manyCharsCore___at___private_Std_Time_Format_Basic_0__Std_Time_parseIdentifier___spec__1___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseNatToBounded___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseOffset___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_parseWith___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___lambda__3___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__3); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__4); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__5); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__6); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__7); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__8); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__9); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__10); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__11); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__12); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__13); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__14); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__15); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__16); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__17); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__18); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__19); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_DateBuilder_build___closed__20); +l_Std_Time_GenericFormat_spec_x21___closed__1 = _init_l_Std_Time_GenericFormat_spec_x21___closed__1(); +lean_mark_persistent(l_Std_Time_GenericFormat_spec_x21___closed__1); +l_Std_Time_GenericFormat_spec_x21___closed__2 = _init_l_Std_Time_GenericFormat_spec_x21___closed__2(); +lean_mark_persistent(l_Std_Time_GenericFormat_spec_x21___closed__2); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser_go___closed__1); +l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1 = _init_l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1(); +lean_mark_persistent(l___private_Std_Time_Format_Basic_0__Std_Time_GenericFormat_parser___closed__1); +l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1 = _init_l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_GenericFormat_builderParser_go___rarg___closed__1); +l_Std_Time_GenericFormat_parse_x21___closed__1 = _init_l_Std_Time_GenericFormat_parse_x21___closed__1(); +lean_mark_persistent(l_Std_Time_GenericFormat_parse_x21___closed__1); +l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1 = _init_l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_GenericFormat_parseBuilder_x21___rarg___closed__1); +l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1 = _init_l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1(); +lean_mark_persistent(l_Std_Time_instFormatGenericFormatFormatTypeString___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Internal.c b/stage0/stdlib/Std/Time/Internal.c new file mode 100644 index 000000000000..740491c2080e --- /dev/null +++ b/stage0/stdlib/Std/Time/Internal.c @@ -0,0 +1,33 @@ +// Lean compiler output +// Module: Std.Time.Internal +// Imports: Std.Time.Internal.Bounded Std.Time.Internal.UnitVal +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Time_Internal_Bounded(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal_UnitVal(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Internal_Bounded(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal_UnitVal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Internal/Bounded.c b/stage0/stdlib/Std/Time/Internal/Bounded.c new file mode 100644 index 000000000000..0f71ba77fdca --- /dev/null +++ b/stage0/stdlib/Std/Time/Internal/Bounded.c @@ -0,0 +1,1958 @@ +// Lean compiler output +// Module: Std.Time.Internal.Bounded +// Imports: Init.Data.Int +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_int_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instInhabitedHAddIntCast___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x3f(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byMod___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instInhabitedHAddIntCast(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___rarg___boxed(lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofInt(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLT___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___rarg___boxed(lean_object*, lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncate___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Internal_Bounded_instDecidableLe___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___rarg(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_clip___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___rarg___boxed(lean_object*, lean_object*); +lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Internal_Bounded_instBEq___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x3f___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_ofInt_x3f(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byEmod___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLT(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncate(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___boxed(lean_object*, lean_object*); +lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_clip(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_eq(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNatWrapping___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___rarg___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_eq___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLE___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byMod(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLE(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___rarg(lean_object*, lean_object*); +lean_object* l_Int_toNat(lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe___boxed(lean_object*, lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_ofInt_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_exact(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod(lean_object*, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___rarg___boxed(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofInt___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq(lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNatWrapping(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byEmod(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLE(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLE___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_instLE(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLT(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instLT___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_instLT(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1; +x_4 = lean_int_dec_lt(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Int_repr(x_1); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Int_repr(x_1); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_instRepr___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instRepr___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_instRepr(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Internal_Bounded_instBEq___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Internal_Bounded_instBEq___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instBEq___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_instBEq(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Internal_Bounded_instDecidableLe___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instDecidableLe___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Internal_Bounded_instDecidableLe___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_instDecidableLe___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_instDecidableLe(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_cast___rarg___boxed), 1, 0); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_cast___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_cast___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Std_Time_Internal_Bounded_cast(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_mk___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_mk___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_mk___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_mk(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_ofInt_x3f___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +lean_inc(x_3); +lean_inc(x_4); +x_5 = lean_apply_2(x_3, x_1, x_4); +x_6 = lean_unbox(x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; +lean_inc(x_4); +x_8 = lean_apply_2(x_3, x_4, x_2); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_box(0); +return x_10; +} +else +{ +lean_object* x_11; +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_ofInt_x3f(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_ofInt_x3f___rarg), 4, 0); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNatWrapping(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_int_sub(x_2, x_1); +x_6 = l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1; +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_sub(x_3, x_1); +x_9 = lean_int_emod(x_8, x_7); +lean_dec(x_8); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = lean_int_emod(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = lean_int_add(x_11, x_1); +lean_dec(x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNatWrapping___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_ofNatWrapping(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_nat_to_int(x_3); +x_5 = lean_int_add(x_1, x_4); +lean_dec(x_4); +x_6 = lean_nat_to_int(x_2); +x_7 = lean_int_sub(x_5, x_1); +lean_dec(x_5); +x_8 = l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1; +x_9 = lean_int_add(x_7, x_8); +lean_dec(x_7); +x_10 = lean_int_sub(x_6, x_1); +lean_dec(x_6); +x_11 = lean_int_emod(x_10, x_9); +lean_dec(x_10); +x_12 = lean_int_add(x_11, x_9); +lean_dec(x_11); +x_13 = lean_int_emod(x_12, x_9); +lean_dec(x_9); +lean_dec(x_12); +x_14 = lean_int_add(x_13, x_1); +lean_dec(x_13); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instInhabitedHAddIntCast(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_3 = lean_nat_to_int(x_2); +x_4 = lean_int_add(x_1, x_3); +lean_dec(x_3); +x_5 = lean_int_sub(x_4, x_1); +lean_dec(x_4); +x_6 = l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1; +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_int_sub(x_1, x_1); +x_9 = lean_int_emod(x_8, x_7); +lean_dec(x_8); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = lean_int_emod(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = lean_int_add(x_11, x_1); +lean_dec(x_11); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_instInhabitedHAddIntCast___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_instInhabitedHAddIntCast(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_mk___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_mk___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mk___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_mk(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_exact(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofInt(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_int_dec_le(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +x_5 = lean_box(0); +return x_5; +} +else +{ +uint8_t x_6; +x_6 = lean_int_dec_le(x_3, x_2); +if (x_6 == 0) +{ +lean_object* x_7; +lean_dec(x_3); +x_7 = lean_box(0); +return x_7; +} +else +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_3); +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofInt___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_ofInt(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_ofNat___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_ofNat(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x3f(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_nat_dec_le(x_2, x_1); +if (x_3 == 0) +{ +lean_object* x_4; +lean_dec(x_2); +x_4 = lean_box(0); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_nat_to_int(x_2); +x_6 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x3f___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_ofNat_x3f(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_ofNat_x27___rarg), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofNat_x27___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_ofNat_x27(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_clip(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_int_dec_le(x_1, x_3); +if (x_5 == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +uint8_t x_6; +x_6 = lean_int_dec_le(x_3, x_2); +if (x_6 == 0) +{ +lean_inc(x_2); +return x_2; +} +else +{ +lean_inc(x_3); +return x_3; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_clip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_clip(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Int_toNat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_toNat___rarg___boxed), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_toNat___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toNat(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_abs(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toNat_x27___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toNat_x27___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toNat_x27(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_toInt___rarg___boxed), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_toInt___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toInt(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Int_toNat(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_toFin___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toFin___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_toFin___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_toFin(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_ofFin___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_ofFin(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_nat_dec_le(x_1, x_2); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_2); +x_5 = lean_nat_to_int(x_1); +return x_5; +} +else +{ +lean_object* x_6; +lean_dec(x_1); +x_6 = lean_nat_to_int(x_2); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_ofFin_x27___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ofFin_x27___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_ofFin_x27(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byEmod(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_emod(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byEmod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_byEmod(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1; +x_8 = lean_int_dec_lt(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +lean_dec(x_6); +lean_dec(x_5); +x_9 = lean_nat_abs(x_1); +x_10 = lean_int_dec_lt(x_2, x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_4); +x_11 = lean_nat_abs(x_2); +x_12 = lean_apply_2(x_3, x_9, x_11); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_3); +x_13 = lean_nat_abs(x_2); +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_sub(x_13, x_14); +lean_dec(x_13); +x_16 = lean_apply_2(x_4, x_9, x_15); +return x_16; +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +lean_dec(x_4); +lean_dec(x_3); +x_17 = lean_nat_abs(x_1); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_17, x_18); +lean_dec(x_17); +x_20 = lean_int_dec_lt(x_2, x_7); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_6); +x_21 = lean_nat_abs(x_2); +x_22 = lean_apply_2(x_5, x_19, x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_5); +x_23 = lean_nat_abs(x_2); +x_24 = lean_nat_sub(x_23, x_18); +lean_dec(x_23); +x_25 = lean_apply_2(x_6, x_19, x_24); +return x_25; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg___boxed), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l___private_Std_Time_Internal_Bounded_0__Int_tdiv_match__1_splitter___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byMod(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_mod(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_byMod___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_byMod(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncate(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_sub(x_3, x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_truncate(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_truncateTop___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_truncateTop___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateTop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_truncateTop(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_truncateBottom___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_truncateBottom___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_truncateBottom(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_neg___rarg___boxed), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_neg___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_neg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_neg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_add___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_add___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_add___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_add(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_int_add(x_2, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_addProven___rarg___boxed), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_addProven___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addProven___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_addProven(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_add(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_addTop___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_addTop___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addTop___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_addTop(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_sub(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_subBottom___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_subBottom___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBottom___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_subBottom(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_addBounds___rarg___boxed), 2, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_addBounds___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_addBounds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_addBounds(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = lean_int_add(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_sub___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_sub___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_sub___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_sub(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = lean_int_add(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_subBounds___rarg___boxed), 2, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_subBounds___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_subBounds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_subBounds(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_emod(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_emod___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_emod___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_emod___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_emod(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_mod(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_mod___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_mod___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mod___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_mod(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_mul(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_mul__pos___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_mul__pos___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__pos___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_mul__pos(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_mul(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_mul__neg___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_mul__neg___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_mul__neg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_mul__neg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_int_ediv(x_1, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_ediv___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_ediv___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_ediv___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_ediv(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_eq(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_eq___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_eq(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_expand___rarg___boxed), 3, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_expand___rarg(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expand___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Internal_Bounded_LE_expand(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_expandTop___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_expandTop___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandTop___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_expandTop(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_expandBottom___rarg___boxed), 2, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_expandBottom___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_expandBottom___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_Bounded_LE_expandBottom(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1; +x_4 = lean_int_add(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_succ___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_succ___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_succ___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_succ(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1; +x_3 = lean_int_dec_le(x_2, x_1); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_int_neg(x_1); +return x_4; +} +else +{ +lean_inc(x_1); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_abs___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_abs___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_abs___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_Bounded_LE_abs(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +if (x_3 == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_LE_max___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_max___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_Bounded_LE_max___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_Bounded_LE_max(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Init_Data_Int(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Internal_Bounded(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Data_Int(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1 = _init_l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_Bounded_instRepr___rarg___closed__1); +l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1 = _init_l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_Bounded_LE_ofNatWrapping___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Internal/UnitVal.c b/stage0/stdlib/Std/Time/Internal/UnitVal.c new file mode 100644 index 000000000000..a9f3b2207791 --- /dev/null +++ b/stage0/stdlib/Std/Time/Internal/UnitVal.c @@ -0,0 +1,832 @@ +// Lean compiler output +// Module: Std.Time.Internal.UnitVal +// Imports: Init.Data Std.Internal.Rat +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLT(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLE(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_Internal_UnitVal_instSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___rarg___boxed(lean_object*); +LEAN_EXPORT uint8_t l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___boxed(lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___rarg(lean_object*, lean_object*); +static lean_object* l_Std_Time_Internal_instInhabitedUnitVal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_convert(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_convert___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instSub___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instInhabitedUnitVal(lean_object*); +lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___boxed(lean_object*); +lean_object* l_Std_Internal_Rat_div(lean_object*, lean_object*); +lean_object* lean_int_div(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Internal_instDecidableLeUnitVal___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLE___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat___boxed(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv(lean_object*); +static lean_object* l_Std_Time_Internal_UnitVal_instAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_instLEUnitVal(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35_(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instLEUnitVal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +lean_object* lean_nat_abs(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instSub(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal___rarg___boxed(lean_object*, lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instBEqUnitVal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instAdd___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instInhabitedUnitVal___boxed(lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString(lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instAdd(lean_object*); +static lean_object* l_Std_Time_Internal_instBEqUnitVal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLT___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___boxed(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___rarg(lean_object*, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___rarg(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___rarg___boxed(lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_instBEqUnitVal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Internal_instInhabitedUnitVal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instInhabitedUnitVal(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instInhabitedUnitVal___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instInhabitedUnitVal___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instInhabitedUnitVal(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35_(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35_(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Internal_instBEqUnitVal___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instBEqUnitVal(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instBEqUnitVal___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instBEqUnitVal___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instBEqUnitVal(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instLEUnitVal(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instLEUnitVal___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instLEUnitVal(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Internal_instDecidableLeUnitVal___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_instDecidableLeUnitVal___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Internal_instDecidableLeUnitVal___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_instDecidableLeUnitVal___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_instDecidableLeUnitVal(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_ofNat___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ofNat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_ofNat(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_toInt___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_toInt___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_toInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_toInt(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_mul___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_mul___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_mul___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_mul(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_ediv(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_ediv___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_ediv___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_ediv___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_ediv(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_div___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_div___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_div___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_div(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_add___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_add___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_add(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_sub___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_sub___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_sub(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_nat_abs(x_1); +x_3 = lean_nat_to_int(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_abs___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_abs___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_abs___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_abs(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_convert(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = l_Std_Internal_Rat_div(x_1, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_5); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_nat_to_int(x_7); +x_9 = lean_int_ediv(x_6, x_8); +lean_dec(x_8); +lean_dec(x_6); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_convert___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Internal_UnitVal_convert(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instOfNat___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instOfNat___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instOfNat(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_Internal_instInhabitedUnitVal___closed__1; +x_4 = lean_int_dec_lt(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Int_repr(x_1); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Int_repr(x_1); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Internal_UnitVal_instRepr___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instRepr___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instRepr(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLE(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLE___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instLE(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLT(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instLT___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instLT(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Internal_UnitVal_instAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instAdd(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instAdd___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instAdd___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instAdd(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Internal_UnitVal_instSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instSub(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instSub___closed__1; +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instSub___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instSub(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instNeg___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instNeg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instNeg(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_Internal_instInhabitedUnitVal___closed__1; +x_3 = lean_int_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_nat_abs(x_1); +x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_nat_abs(x_1); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_6, x_7); +lean_dec(x_6); +x_9 = lean_nat_add(x_8, x_7); +lean_dec(x_8); +x_10 = l___private_Init_Data_Repr_0__Nat_reprFast(x_9); +x_11 = l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1; +x_12 = lean_string_append(x_11, x_10); +lean_dec(x_10); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instToString___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Internal_UnitVal_instToString___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Internal_UnitVal_instToString(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Init_Data(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Internal_UnitVal(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Data(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Internal_instInhabitedUnitVal___closed__1 = _init_l_Std_Time_Internal_instInhabitedUnitVal___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_instInhabitedUnitVal___closed__1); +l_Std_Time_Internal_instBEqUnitVal___closed__1 = _init_l_Std_Time_Internal_instBEqUnitVal___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_instBEqUnitVal___closed__1); +l_Std_Time_Internal_UnitVal_instAdd___closed__1 = _init_l_Std_Time_Internal_UnitVal_instAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_UnitVal_instAdd___closed__1); +l_Std_Time_Internal_UnitVal_instSub___closed__1 = _init_l_Std_Time_Internal_UnitVal_instSub___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_UnitVal_instSub___closed__1); +l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1 = _init_l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Internal_UnitVal_instToString___rarg___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Notation.c b/stage0/stdlib/Std/Time/Notation.c new file mode 100644 index 000000000000..fbf7c84b7c88 --- /dev/null +++ b/stage0/stdlib/Std/Time/Notation.c @@ -0,0 +1,14749 @@ +// Lean compiler output +// Module: Std.Time.Notation +// Imports: Std.Time.Date Std.Time.Time Std.Time.Zoned Std.Time.DateTime Std.Time.Format +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termOffset_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Std_Time_termZoned_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58; +lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_termZoned_x28___x29___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termTimezone_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +static lean_object* l_Std_Time_termTime_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23; +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__3; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159; +static lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6; +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__6; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21; +lean_object* l_Std_Time_PlainDateTime_fromLeanDateTimeString(lean_object*); +lean_object* l_Lean_TSyntax_getString(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82; +lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithIdentifierString(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17; +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__1; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2; +lean_object* l_Std_Time_ZonedDateTime_fromLeanDateTimeWithZoneString(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX(uint8_t, lean_object*, lean_object*); +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16; +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100; +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x2c___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25; +static lean_object* l_Std_Time_termTime_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7; +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__1; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__5; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127; +LEAN_EXPORT lean_object* l_Std_Time_termDate_x28___x29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termTime_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2; +extern lean_object* l_Std_Time_Formats_sqlDate; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166; +static lean_object* l_Std_Time_termDate_x28___x29___closed__2; +static lean_object* l_Std_Time_termTime_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124; +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16; +static lean_object* l_Std_Time_termDate_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2; +lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10; +lean_object* l_Std_Time_PlainTime_fromLeanTime24Hour(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22; +static lean_object* l_Std_Time_termDate_x28___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13; +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221; +static lean_object* l_Std_Time_termDate_x28___x29___closed__1; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12; +static lean_object* l_Std_Time_termDate_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108; +static lean_object* l_Std_Time_termTime_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3; +lean_object* l_Std_Time_TimeZone_Offset_fromOffset(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1; +lean_object* lean_thunk_get_own(lean_object*); +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__9; +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +static lean_object* l_Std_Time_termTime_x28___x29___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156; +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142; +LEAN_EXPORT lean_object* l_Std_Time_termZoned_x28___x29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13; +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDatetime_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210; +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189; +LEAN_EXPORT lean_object* l_Std_Time_termDatetime_x28___x29; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7; +static lean_object* l_Std_Time_termTime_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232; +lean_object* l_Std_Time_PlainDate_ofYearMonthDay_x3f(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240; +lean_object* l_Lean_Name_mkStr6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9; +lean_object* l_Std_Time_GenericFormat_parseBuilder___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17; +lean_object* lean_nat_abs(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7; +lean_object* l_Std_Time_TimeZone_fromTimeZone(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +static lean_object* l_Std_Time_termOffset_x28___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_termTimezone_x28___x29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228; +lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_termOffset_x28___x29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +static lean_object* l_Std_Time_termTime_x28___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16; +static lean_object* l_Std_Time_termDate_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136; +static lean_object* l_Std_Time_termTimezone_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_termZoned_x28___x29___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9; +static lean_object* l_Std_Time_termZoned_x28___x2c___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71; +static lean_object* l_Std_Time_termDate_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20; +LEAN_EXPORT lean_object* l_Std_Time_termTime_x28___x29; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123; +static lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119; +static lean_object* l_Std_Time_termDatetime_x28___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206; +lean_object* l_String_toSubstring_x27(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_termZoned_x28___x2c___x29; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__1; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14; +static lean_object* l_Std_Time_termZoned_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6; +static lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9; +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.short", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Time", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Text", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("short", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.full", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("full", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.narrow", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("narrow", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21; +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertText___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Number.mk", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Number", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("mk", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertNumber(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14; +lean_inc(x_6); +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_15 = lean_box(2); +x_16 = l_Lean_Syntax_mkNumLit(x_14, x_15); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_6); +x_18 = l_Lean_Syntax_node1(x_6, x_17, x_16); +x_19 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_20 = l_Lean_Syntax_node2(x_6, x_19, x_13, x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_3); +return x_21; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.nano", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Fraction", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nano", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.truncated", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("truncated", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFraction(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +lean_dec(x_2); +x_21 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13; +x_22 = l_Lean_addMacroScope(x_20, x_21, x_19); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11; +x_24 = l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17; +lean_inc(x_18); +x_25 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_24); +x_26 = l___private_Init_Data_Repr_0__Nat_reprFast(x_15); +x_27 = lean_box(2); +x_28 = l_Lean_Syntax_mkNumLit(x_26, x_27); +x_29 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_18); +x_30 = l_Lean_Syntax_node1(x_18, x_29, x_28); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_32 = l_Lean_Syntax_node2(x_18, x_31, x_25, x_30); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.twoDigit", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Year", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("twoDigit", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.fourDigit", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("fourDigit", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.extended", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("extended", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertYear(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = lean_ctor_get(x_2, 5); +lean_inc(x_27); +x_28 = 0; +x_29 = l_Lean_SourceInfo_fromRef(x_27, x_28); +lean_dec(x_27); +x_30 = lean_ctor_get(x_2, 2); +lean_inc(x_30); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_32 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21; +x_33 = l_Lean_addMacroScope(x_31, x_32, x_30); +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19; +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25; +lean_inc(x_29); +x_36 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_36, 0, x_29); +lean_ctor_set(x_36, 1, x_34); +lean_ctor_set(x_36, 2, x_33); +lean_ctor_set(x_36, 3, x_35); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_38 = lean_box(2); +x_39 = l_Lean_Syntax_mkNumLit(x_37, x_38); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_29); +x_41 = l_Lean_Syntax_node1(x_29, x_40, x_39); +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_43 = l_Lean_Syntax_node2(x_29, x_42, x_36, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_3); +return x_44; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.short", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ZoneName", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.full", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZoneName___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hour", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetX", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hour", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinute", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteColon", 32, 32); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteColon", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecond", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteSecond", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteSecondColon", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +case 2: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19; +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +case 3: +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_37 = lean_ctor_get(x_2, 5); +lean_inc(x_37); +x_38 = 0; +x_39 = l_Lean_SourceInfo_fromRef(x_37, x_38); +lean_dec(x_37); +x_40 = lean_ctor_get(x_2, 2); +lean_inc(x_40); +x_41 = lean_ctor_get(x_2, 1); +lean_inc(x_41); +lean_dec(x_2); +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_40); +x_44 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27; +x_45 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33; +x_46 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_46, 0, x_39); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set(x_46, 2, x_43); +lean_ctor_set(x_46, 3, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +return x_47; +} +default: +{ +lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_48 = lean_ctor_get(x_2, 5); +lean_inc(x_48); +x_49 = 0; +x_50 = l_Lean_SourceInfo_fromRef(x_48, x_49); +lean_dec(x_48); +x_51 = lean_ctor_get(x_2, 2); +lean_inc(x_51); +x_52 = lean_ctor_get(x_2, 1); +lean_inc(x_52); +lean_dec(x_2); +x_53 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37; +x_54 = l_Lean_addMacroScope(x_52, x_53, x_51); +x_55 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35; +x_56 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41; +x_57 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_57, 0, x_50); +lean_ctor_set(x_57, 1, x_55); +lean_ctor_set(x_57, 2, x_54); +lean_ctor_set(x_57, 3, x_56); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_3); +return x_58; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.short", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetO", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetZ", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2; +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10; +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17; +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.G", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Modifier", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("G", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.y", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.u", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("u", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.D", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("D", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.MorL", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("MorL", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("paren", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dotIdent", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("inl", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("inr", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.d", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("d", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Qorq", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Qorq", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.w", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("w", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.W", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("W", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.E", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("E", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.eorc", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eorc", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.F", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("F", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.a", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("a", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.h", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("h", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.K", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("K", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.k", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("k", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.H", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("H", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.m", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("m", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.s", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("s", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.S", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("S", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.A", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("A", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.n", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("n", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.N", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("N", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.V", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("V", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("z", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.O", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("O", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.X", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("X", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.x", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Z", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertModifier(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_4, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5; +x_14 = l_Lean_addMacroScope(x_12, x_13, x_11); +x_15 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2; +x_16 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9; +lean_inc(x_10); +x_17 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_15); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_16); +x_18 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_10); +x_19 = l_Lean_Syntax_node1(x_10, x_18, x_7); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_21 = l_Lean_Syntax_node2(x_10, x_20, x_17, x_19); +lean_ctor_set(x_5, 0, x_21); +return x_5; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_5, 0); +x_23 = lean_ctor_get(x_5, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_5); +x_24 = lean_ctor_get(x_2, 5); +lean_inc(x_24); +x_25 = 0; +x_26 = l_Lean_SourceInfo_fromRef(x_24, x_25); +lean_dec(x_24); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 1); +lean_inc(x_28); +lean_dec(x_2); +x_29 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5; +x_30 = l_Lean_addMacroScope(x_28, x_29, x_27); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2; +x_32 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9; +lean_inc(x_26); +x_33 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_31); +lean_ctor_set(x_33, 2, x_30); +lean_ctor_set(x_33, 3, x_32); +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_26); +x_35 = l_Lean_Syntax_node1(x_26, x_34, x_22); +x_36 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_37 = l_Lean_Syntax_node2(x_26, x_36, x_33, x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_23); +return x_38; +} +} +case 1: +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +lean_dec(x_1); +lean_inc(x_2); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertYear(x_39, x_2, x_3); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_2, 5); +lean_inc(x_43); +x_44 = 0; +x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); +lean_dec(x_43); +x_46 = lean_ctor_get(x_2, 2); +lean_inc(x_46); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); +lean_dec(x_2); +x_48 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13; +x_49 = l_Lean_addMacroScope(x_47, x_48, x_46); +x_50 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11; +x_51 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17; +lean_inc(x_45); +x_52 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_52, 0, x_45); +lean_ctor_set(x_52, 1, x_50); +lean_ctor_set(x_52, 2, x_49); +lean_ctor_set(x_52, 3, x_51); +x_53 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_45); +x_54 = l_Lean_Syntax_node1(x_45, x_53, x_42); +x_55 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_56 = l_Lean_Syntax_node2(x_45, x_55, x_52, x_54); +lean_ctor_set(x_40, 0, x_56); +return x_40; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_57 = lean_ctor_get(x_40, 0); +x_58 = lean_ctor_get(x_40, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_40); +x_59 = lean_ctor_get(x_2, 5); +lean_inc(x_59); +x_60 = 0; +x_61 = l_Lean_SourceInfo_fromRef(x_59, x_60); +lean_dec(x_59); +x_62 = lean_ctor_get(x_2, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_2, 1); +lean_inc(x_63); +lean_dec(x_2); +x_64 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13; +x_65 = l_Lean_addMacroScope(x_63, x_64, x_62); +x_66 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11; +x_67 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17; +lean_inc(x_61); +x_68 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_68, 0, x_61); +lean_ctor_set(x_68, 1, x_66); +lean_ctor_set(x_68, 2, x_65); +lean_ctor_set(x_68, 3, x_67); +x_69 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_61); +x_70 = l_Lean_Syntax_node1(x_61, x_69, x_57); +x_71 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_72 = l_Lean_Syntax_node2(x_61, x_71, x_68, x_70); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_58); +return x_73; +} +} +case 2: +{ +lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_74 = lean_ctor_get(x_1, 0); +lean_inc(x_74); +lean_dec(x_1); +lean_inc(x_2); +x_75 = l___private_Std_Time_Notation_0__Std_Time_convertYear(x_74, x_2, x_3); +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_77 = lean_ctor_get(x_75, 0); +x_78 = lean_ctor_get(x_2, 5); +lean_inc(x_78); +x_79 = 0; +x_80 = l_Lean_SourceInfo_fromRef(x_78, x_79); +lean_dec(x_78); +x_81 = lean_ctor_get(x_2, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_2, 1); +lean_inc(x_82); +lean_dec(x_2); +x_83 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21; +x_84 = l_Lean_addMacroScope(x_82, x_83, x_81); +x_85 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19; +x_86 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25; +lean_inc(x_80); +x_87 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_87, 0, x_80); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set(x_87, 3, x_86); +x_88 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_80); +x_89 = l_Lean_Syntax_node1(x_80, x_88, x_77); +x_90 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_91 = l_Lean_Syntax_node2(x_80, x_90, x_87, x_89); +lean_ctor_set(x_75, 0, x_91); +return x_75; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_92 = lean_ctor_get(x_75, 0); +x_93 = lean_ctor_get(x_75, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_75); +x_94 = lean_ctor_get(x_2, 5); +lean_inc(x_94); +x_95 = 0; +x_96 = l_Lean_SourceInfo_fromRef(x_94, x_95); +lean_dec(x_94); +x_97 = lean_ctor_get(x_2, 2); +lean_inc(x_97); +x_98 = lean_ctor_get(x_2, 1); +lean_inc(x_98); +lean_dec(x_2); +x_99 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21; +x_100 = l_Lean_addMacroScope(x_98, x_99, x_97); +x_101 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19; +x_102 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25; +lean_inc(x_96); +x_103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_103, 0, x_96); +lean_ctor_set(x_103, 1, x_101); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_102); +x_104 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_96); +x_105 = l_Lean_Syntax_node1(x_96, x_104, x_92); +x_106 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_107 = l_Lean_Syntax_node2(x_96, x_106, x_103, x_105); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_93); +return x_108; +} +} +case 3: +{ +lean_object* x_109; lean_object* x_110; uint8_t x_111; +x_109 = lean_ctor_get(x_1, 0); +lean_inc(x_109); +lean_dec(x_1); +lean_inc(x_2); +x_110 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_109, x_2, x_3); +x_111 = !lean_is_exclusive(x_110); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_112 = lean_ctor_get(x_110, 0); +x_113 = lean_ctor_get(x_2, 5); +lean_inc(x_113); +x_114 = 0; +x_115 = l_Lean_SourceInfo_fromRef(x_113, x_114); +lean_dec(x_113); +x_116 = lean_ctor_get(x_2, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_2, 1); +lean_inc(x_117); +lean_dec(x_2); +x_118 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29; +x_119 = l_Lean_addMacroScope(x_117, x_118, x_116); +x_120 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27; +x_121 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33; +lean_inc(x_115); +x_122 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_122, 0, x_115); +lean_ctor_set(x_122, 1, x_120); +lean_ctor_set(x_122, 2, x_119); +lean_ctor_set(x_122, 3, x_121); +x_123 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_115); +x_124 = l_Lean_Syntax_node1(x_115, x_123, x_112); +x_125 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_126 = l_Lean_Syntax_node2(x_115, x_125, x_122, x_124); +lean_ctor_set(x_110, 0, x_126); +return x_110; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_127 = lean_ctor_get(x_110, 0); +x_128 = lean_ctor_get(x_110, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_110); +x_129 = lean_ctor_get(x_2, 5); +lean_inc(x_129); +x_130 = 0; +x_131 = l_Lean_SourceInfo_fromRef(x_129, x_130); +lean_dec(x_129); +x_132 = lean_ctor_get(x_2, 2); +lean_inc(x_132); +x_133 = lean_ctor_get(x_2, 1); +lean_inc(x_133); +lean_dec(x_2); +x_134 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29; +x_135 = l_Lean_addMacroScope(x_133, x_134, x_132); +x_136 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27; +x_137 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33; +lean_inc(x_131); +x_138 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_138, 0, x_131); +lean_ctor_set(x_138, 1, x_136); +lean_ctor_set(x_138, 2, x_135); +lean_ctor_set(x_138, 3, x_137); +x_139 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_131); +x_140 = l_Lean_Syntax_node1(x_131, x_139, x_127); +x_141 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_142 = l_Lean_Syntax_node2(x_131, x_141, x_138, x_140); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_128); +return x_143; +} +} +case 4: +{ +lean_object* x_144; +x_144 = lean_ctor_get(x_1, 0); +lean_inc(x_144); +lean_dec(x_1); +if (lean_obj_tag(x_144) == 0) +{ +lean_object* x_145; lean_object* x_146; uint8_t x_147; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +lean_dec(x_144); +lean_inc(x_2); +x_146 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_145, x_2, x_3); +x_147 = !lean_is_exclusive(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_148 = lean_ctor_get(x_146, 0); +x_149 = lean_ctor_get(x_2, 5); +lean_inc(x_149); +x_150 = 0; +x_151 = l_Lean_SourceInfo_fromRef(x_149, x_150); +lean_dec(x_149); +x_152 = lean_ctor_get(x_2, 2); +lean_inc(x_152); +x_153 = lean_ctor_get(x_2, 1); +lean_inc(x_153); +lean_dec(x_2); +x_154 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +lean_inc(x_152); +lean_inc(x_153); +x_155 = l_Lean_addMacroScope(x_153, x_154, x_152); +x_156 = lean_box(0); +x_157 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35; +x_158 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41; +lean_inc(x_151); +x_159 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_159, 0, x_151); +lean_ctor_set(x_159, 1, x_157); +lean_ctor_set(x_159, 2, x_155); +lean_ctor_set(x_159, 3, x_158); +x_160 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_151); +x_161 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_161, 0, x_151); +lean_ctor_set(x_161, 1, x_160); +x_162 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_151); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_151); +lean_ctor_set(x_163, 1, x_162); +x_164 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_165 = l_Lean_addMacroScope(x_153, x_164, x_152); +x_166 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_151); +x_167 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_167, 0, x_151); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +lean_ctor_set(x_167, 3, x_156); +x_168 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_151); +x_169 = l_Lean_Syntax_node2(x_151, x_168, x_163, x_167); +x_170 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_151); +x_171 = l_Lean_Syntax_node1(x_151, x_170, x_148); +x_172 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_151); +x_173 = l_Lean_Syntax_node2(x_151, x_172, x_169, x_171); +x_174 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_151); +x_175 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_175, 0, x_151); +lean_ctor_set(x_175, 1, x_174); +x_176 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_151); +x_177 = l_Lean_Syntax_node3(x_151, x_176, x_161, x_173, x_175); +lean_inc(x_151); +x_178 = l_Lean_Syntax_node1(x_151, x_170, x_177); +x_179 = l_Lean_Syntax_node2(x_151, x_172, x_159, x_178); +lean_ctor_set(x_146, 0, x_179); +return x_146; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_180 = lean_ctor_get(x_146, 0); +x_181 = lean_ctor_get(x_146, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_146); +x_182 = lean_ctor_get(x_2, 5); +lean_inc(x_182); +x_183 = 0; +x_184 = l_Lean_SourceInfo_fromRef(x_182, x_183); +lean_dec(x_182); +x_185 = lean_ctor_get(x_2, 2); +lean_inc(x_185); +x_186 = lean_ctor_get(x_2, 1); +lean_inc(x_186); +lean_dec(x_2); +x_187 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +lean_inc(x_185); +lean_inc(x_186); +x_188 = l_Lean_addMacroScope(x_186, x_187, x_185); +x_189 = lean_box(0); +x_190 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35; +x_191 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41; +lean_inc(x_184); +x_192 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_192, 0, x_184); +lean_ctor_set(x_192, 1, x_190); +lean_ctor_set(x_192, 2, x_188); +lean_ctor_set(x_192, 3, x_191); +x_193 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_184); +x_194 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_194, 0, x_184); +lean_ctor_set(x_194, 1, x_193); +x_195 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_184); +x_196 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_196, 0, x_184); +lean_ctor_set(x_196, 1, x_195); +x_197 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_198 = l_Lean_addMacroScope(x_186, x_197, x_185); +x_199 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_184); +x_200 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_200, 0, x_184); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +lean_ctor_set(x_200, 3, x_189); +x_201 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_184); +x_202 = l_Lean_Syntax_node2(x_184, x_201, x_196, x_200); +x_203 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_184); +x_204 = l_Lean_Syntax_node1(x_184, x_203, x_180); +x_205 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_184); +x_206 = l_Lean_Syntax_node2(x_184, x_205, x_202, x_204); +x_207 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_184); +x_208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_208, 0, x_184); +lean_ctor_set(x_208, 1, x_207); +x_209 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_184); +x_210 = l_Lean_Syntax_node3(x_184, x_209, x_194, x_206, x_208); +lean_inc(x_184); +x_211 = l_Lean_Syntax_node1(x_184, x_203, x_210); +x_212 = l_Lean_Syntax_node2(x_184, x_205, x_192, x_211); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_181); +return x_213; +} +} +else +{ +lean_object* x_214; uint8_t x_215; lean_object* x_216; uint8_t x_217; +x_214 = lean_ctor_get(x_144, 0); +lean_inc(x_214); +lean_dec(x_144); +x_215 = lean_unbox(x_214); +lean_dec(x_214); +lean_inc(x_2); +x_216 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_215, x_2, x_3); +x_217 = !lean_is_exclusive(x_216); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_218 = lean_ctor_get(x_216, 0); +x_219 = lean_ctor_get(x_2, 5); +lean_inc(x_219); +x_220 = 0; +x_221 = l_Lean_SourceInfo_fromRef(x_219, x_220); +lean_dec(x_219); +x_222 = lean_ctor_get(x_2, 2); +lean_inc(x_222); +x_223 = lean_ctor_get(x_2, 1); +lean_inc(x_223); +lean_dec(x_2); +x_224 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +lean_inc(x_222); +lean_inc(x_223); +x_225 = l_Lean_addMacroScope(x_223, x_224, x_222); +x_226 = lean_box(0); +x_227 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35; +x_228 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41; +lean_inc(x_221); +x_229 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_229, 0, x_221); +lean_ctor_set(x_229, 1, x_227); +lean_ctor_set(x_229, 2, x_225); +lean_ctor_set(x_229, 3, x_228); +x_230 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_221); +x_231 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_231, 0, x_221); +lean_ctor_set(x_231, 1, x_230); +x_232 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_221); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_221); +lean_ctor_set(x_233, 1, x_232); +x_234 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_235 = l_Lean_addMacroScope(x_223, x_234, x_222); +x_236 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_221); +x_237 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_237, 0, x_221); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set(x_237, 2, x_235); +lean_ctor_set(x_237, 3, x_226); +x_238 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_221); +x_239 = l_Lean_Syntax_node2(x_221, x_238, x_233, x_237); +x_240 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_221); +x_241 = l_Lean_Syntax_node1(x_221, x_240, x_218); +x_242 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_221); +x_243 = l_Lean_Syntax_node2(x_221, x_242, x_239, x_241); +x_244 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_221); +x_245 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_245, 0, x_221); +lean_ctor_set(x_245, 1, x_244); +x_246 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_221); +x_247 = l_Lean_Syntax_node3(x_221, x_246, x_231, x_243, x_245); +lean_inc(x_221); +x_248 = l_Lean_Syntax_node1(x_221, x_240, x_247); +x_249 = l_Lean_Syntax_node2(x_221, x_242, x_229, x_248); +lean_ctor_set(x_216, 0, x_249); +return x_216; +} +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_250 = lean_ctor_get(x_216, 0); +x_251 = lean_ctor_get(x_216, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_216); +x_252 = lean_ctor_get(x_2, 5); +lean_inc(x_252); +x_253 = 0; +x_254 = l_Lean_SourceInfo_fromRef(x_252, x_253); +lean_dec(x_252); +x_255 = lean_ctor_get(x_2, 2); +lean_inc(x_255); +x_256 = lean_ctor_get(x_2, 1); +lean_inc(x_256); +lean_dec(x_2); +x_257 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37; +lean_inc(x_255); +lean_inc(x_256); +x_258 = l_Lean_addMacroScope(x_256, x_257, x_255); +x_259 = lean_box(0); +x_260 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35; +x_261 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41; +lean_inc(x_254); +x_262 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_262, 0, x_254); +lean_ctor_set(x_262, 1, x_260); +lean_ctor_set(x_262, 2, x_258); +lean_ctor_set(x_262, 3, x_261); +x_263 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_254); +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_254); +lean_ctor_set(x_264, 1, x_263); +x_265 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_254); +x_266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_266, 0, x_254); +lean_ctor_set(x_266, 1, x_265); +x_267 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_268 = l_Lean_addMacroScope(x_256, x_267, x_255); +x_269 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_254); +x_270 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_270, 0, x_254); +lean_ctor_set(x_270, 1, x_269); +lean_ctor_set(x_270, 2, x_268); +lean_ctor_set(x_270, 3, x_259); +x_271 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_254); +x_272 = l_Lean_Syntax_node2(x_254, x_271, x_266, x_270); +x_273 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_254); +x_274 = l_Lean_Syntax_node1(x_254, x_273, x_250); +x_275 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_254); +x_276 = l_Lean_Syntax_node2(x_254, x_275, x_272, x_274); +x_277 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_254); +x_278 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_278, 0, x_254); +lean_ctor_set(x_278, 1, x_277); +x_279 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_254); +x_280 = l_Lean_Syntax_node3(x_254, x_279, x_264, x_276, x_278); +lean_inc(x_254); +x_281 = l_Lean_Syntax_node1(x_254, x_273, x_280); +x_282 = l_Lean_Syntax_node2(x_254, x_275, x_262, x_281); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_251); +return x_283; +} +} +} +case 5: +{ +lean_object* x_284; lean_object* x_285; uint8_t x_286; +x_284 = lean_ctor_get(x_1, 0); +lean_inc(x_284); +lean_dec(x_1); +lean_inc(x_2); +x_285 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_284, x_2, x_3); +x_286 = !lean_is_exclusive(x_285); +if (x_286 == 0) +{ +lean_object* x_287; lean_object* x_288; uint8_t x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_287 = lean_ctor_get(x_285, 0); +x_288 = lean_ctor_get(x_2, 5); +lean_inc(x_288); +x_289 = 0; +x_290 = l_Lean_SourceInfo_fromRef(x_288, x_289); +lean_dec(x_288); +x_291 = lean_ctor_get(x_2, 2); +lean_inc(x_291); +x_292 = lean_ctor_get(x_2, 1); +lean_inc(x_292); +lean_dec(x_2); +x_293 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58; +x_294 = l_Lean_addMacroScope(x_292, x_293, x_291); +x_295 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56; +x_296 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62; +lean_inc(x_290); +x_297 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_297, 0, x_290); +lean_ctor_set(x_297, 1, x_295); +lean_ctor_set(x_297, 2, x_294); +lean_ctor_set(x_297, 3, x_296); +x_298 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_290); +x_299 = l_Lean_Syntax_node1(x_290, x_298, x_287); +x_300 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_301 = l_Lean_Syntax_node2(x_290, x_300, x_297, x_299); +lean_ctor_set(x_285, 0, x_301); +return x_285; +} +else +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_302 = lean_ctor_get(x_285, 0); +x_303 = lean_ctor_get(x_285, 1); +lean_inc(x_303); +lean_inc(x_302); +lean_dec(x_285); +x_304 = lean_ctor_get(x_2, 5); +lean_inc(x_304); +x_305 = 0; +x_306 = l_Lean_SourceInfo_fromRef(x_304, x_305); +lean_dec(x_304); +x_307 = lean_ctor_get(x_2, 2); +lean_inc(x_307); +x_308 = lean_ctor_get(x_2, 1); +lean_inc(x_308); +lean_dec(x_2); +x_309 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58; +x_310 = l_Lean_addMacroScope(x_308, x_309, x_307); +x_311 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56; +x_312 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62; +lean_inc(x_306); +x_313 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_313, 0, x_306); +lean_ctor_set(x_313, 1, x_311); +lean_ctor_set(x_313, 2, x_310); +lean_ctor_set(x_313, 3, x_312); +x_314 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_306); +x_315 = l_Lean_Syntax_node1(x_306, x_314, x_302); +x_316 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_317 = l_Lean_Syntax_node2(x_306, x_316, x_313, x_315); +x_318 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_318, 1, x_303); +return x_318; +} +} +case 6: +{ +lean_object* x_319; +x_319 = lean_ctor_get(x_1, 0); +lean_inc(x_319); +lean_dec(x_1); +if (lean_obj_tag(x_319) == 0) +{ +lean_object* x_320; lean_object* x_321; uint8_t x_322; +x_320 = lean_ctor_get(x_319, 0); +lean_inc(x_320); +lean_dec(x_319); +lean_inc(x_2); +x_321 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_320, x_2, x_3); +x_322 = !lean_is_exclusive(x_321); +if (x_322 == 0) +{ +lean_object* x_323; lean_object* x_324; uint8_t x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_323 = lean_ctor_get(x_321, 0); +x_324 = lean_ctor_get(x_2, 5); +lean_inc(x_324); +x_325 = 0; +x_326 = l_Lean_SourceInfo_fromRef(x_324, x_325); +lean_dec(x_324); +x_327 = lean_ctor_get(x_2, 2); +lean_inc(x_327); +x_328 = lean_ctor_get(x_2, 1); +lean_inc(x_328); +lean_dec(x_2); +x_329 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +lean_inc(x_327); +lean_inc(x_328); +x_330 = l_Lean_addMacroScope(x_328, x_329, x_327); +x_331 = lean_box(0); +x_332 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64; +x_333 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70; +lean_inc(x_326); +x_334 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_334, 0, x_326); +lean_ctor_set(x_334, 1, x_332); +lean_ctor_set(x_334, 2, x_330); +lean_ctor_set(x_334, 3, x_333); +x_335 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_326); +x_336 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_336, 0, x_326); +lean_ctor_set(x_336, 1, x_335); +x_337 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_326); +x_338 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_338, 0, x_326); +lean_ctor_set(x_338, 1, x_337); +x_339 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_340 = l_Lean_addMacroScope(x_328, x_339, x_327); +x_341 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_326); +x_342 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_342, 0, x_326); +lean_ctor_set(x_342, 1, x_341); +lean_ctor_set(x_342, 2, x_340); +lean_ctor_set(x_342, 3, x_331); +x_343 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_326); +x_344 = l_Lean_Syntax_node2(x_326, x_343, x_338, x_342); +x_345 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_326); +x_346 = l_Lean_Syntax_node1(x_326, x_345, x_323); +x_347 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_326); +x_348 = l_Lean_Syntax_node2(x_326, x_347, x_344, x_346); +x_349 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_326); +x_350 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_350, 0, x_326); +lean_ctor_set(x_350, 1, x_349); +x_351 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_326); +x_352 = l_Lean_Syntax_node3(x_326, x_351, x_336, x_348, x_350); +lean_inc(x_326); +x_353 = l_Lean_Syntax_node1(x_326, x_345, x_352); +x_354 = l_Lean_Syntax_node2(x_326, x_347, x_334, x_353); +lean_ctor_set(x_321, 0, x_354); +return x_321; +} +else +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; uint8_t x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; +x_355 = lean_ctor_get(x_321, 0); +x_356 = lean_ctor_get(x_321, 1); +lean_inc(x_356); +lean_inc(x_355); +lean_dec(x_321); +x_357 = lean_ctor_get(x_2, 5); +lean_inc(x_357); +x_358 = 0; +x_359 = l_Lean_SourceInfo_fromRef(x_357, x_358); +lean_dec(x_357); +x_360 = lean_ctor_get(x_2, 2); +lean_inc(x_360); +x_361 = lean_ctor_get(x_2, 1); +lean_inc(x_361); +lean_dec(x_2); +x_362 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +lean_inc(x_360); +lean_inc(x_361); +x_363 = l_Lean_addMacroScope(x_361, x_362, x_360); +x_364 = lean_box(0); +x_365 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64; +x_366 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70; +lean_inc(x_359); +x_367 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_367, 0, x_359); +lean_ctor_set(x_367, 1, x_365); +lean_ctor_set(x_367, 2, x_363); +lean_ctor_set(x_367, 3, x_366); +x_368 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_359); +x_369 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_369, 0, x_359); +lean_ctor_set(x_369, 1, x_368); +x_370 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_359); +x_371 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_371, 0, x_359); +lean_ctor_set(x_371, 1, x_370); +x_372 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_373 = l_Lean_addMacroScope(x_361, x_372, x_360); +x_374 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_359); +x_375 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_375, 0, x_359); +lean_ctor_set(x_375, 1, x_374); +lean_ctor_set(x_375, 2, x_373); +lean_ctor_set(x_375, 3, x_364); +x_376 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_359); +x_377 = l_Lean_Syntax_node2(x_359, x_376, x_371, x_375); +x_378 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_359); +x_379 = l_Lean_Syntax_node1(x_359, x_378, x_355); +x_380 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_359); +x_381 = l_Lean_Syntax_node2(x_359, x_380, x_377, x_379); +x_382 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_359); +x_383 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_383, 0, x_359); +lean_ctor_set(x_383, 1, x_382); +x_384 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_359); +x_385 = l_Lean_Syntax_node3(x_359, x_384, x_369, x_381, x_383); +lean_inc(x_359); +x_386 = l_Lean_Syntax_node1(x_359, x_378, x_385); +x_387 = l_Lean_Syntax_node2(x_359, x_380, x_367, x_386); +x_388 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_388, 0, x_387); +lean_ctor_set(x_388, 1, x_356); +return x_388; +} +} +else +{ +lean_object* x_389; uint8_t x_390; lean_object* x_391; uint8_t x_392; +x_389 = lean_ctor_get(x_319, 0); +lean_inc(x_389); +lean_dec(x_319); +x_390 = lean_unbox(x_389); +lean_dec(x_389); +lean_inc(x_2); +x_391 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_390, x_2, x_3); +x_392 = !lean_is_exclusive(x_391); +if (x_392 == 0) +{ +lean_object* x_393; lean_object* x_394; uint8_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_393 = lean_ctor_get(x_391, 0); +x_394 = lean_ctor_get(x_2, 5); +lean_inc(x_394); +x_395 = 0; +x_396 = l_Lean_SourceInfo_fromRef(x_394, x_395); +lean_dec(x_394); +x_397 = lean_ctor_get(x_2, 2); +lean_inc(x_397); +x_398 = lean_ctor_get(x_2, 1); +lean_inc(x_398); +lean_dec(x_2); +x_399 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +lean_inc(x_397); +lean_inc(x_398); +x_400 = l_Lean_addMacroScope(x_398, x_399, x_397); +x_401 = lean_box(0); +x_402 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64; +x_403 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70; +lean_inc(x_396); +x_404 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_404, 0, x_396); +lean_ctor_set(x_404, 1, x_402); +lean_ctor_set(x_404, 2, x_400); +lean_ctor_set(x_404, 3, x_403); +x_405 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_396); +x_406 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_406, 0, x_396); +lean_ctor_set(x_406, 1, x_405); +x_407 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_396); +x_408 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_408, 0, x_396); +lean_ctor_set(x_408, 1, x_407); +x_409 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_410 = l_Lean_addMacroScope(x_398, x_409, x_397); +x_411 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_396); +x_412 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_412, 0, x_396); +lean_ctor_set(x_412, 1, x_411); +lean_ctor_set(x_412, 2, x_410); +lean_ctor_set(x_412, 3, x_401); +x_413 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_396); +x_414 = l_Lean_Syntax_node2(x_396, x_413, x_408, x_412); +x_415 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_396); +x_416 = l_Lean_Syntax_node1(x_396, x_415, x_393); +x_417 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_396); +x_418 = l_Lean_Syntax_node2(x_396, x_417, x_414, x_416); +x_419 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_396); +x_420 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_420, 0, x_396); +lean_ctor_set(x_420, 1, x_419); +x_421 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_396); +x_422 = l_Lean_Syntax_node3(x_396, x_421, x_406, x_418, x_420); +lean_inc(x_396); +x_423 = l_Lean_Syntax_node1(x_396, x_415, x_422); +x_424 = l_Lean_Syntax_node2(x_396, x_417, x_404, x_423); +lean_ctor_set(x_391, 0, x_424); +return x_391; +} +else +{ +lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_425 = lean_ctor_get(x_391, 0); +x_426 = lean_ctor_get(x_391, 1); +lean_inc(x_426); +lean_inc(x_425); +lean_dec(x_391); +x_427 = lean_ctor_get(x_2, 5); +lean_inc(x_427); +x_428 = 0; +x_429 = l_Lean_SourceInfo_fromRef(x_427, x_428); +lean_dec(x_427); +x_430 = lean_ctor_get(x_2, 2); +lean_inc(x_430); +x_431 = lean_ctor_get(x_2, 1); +lean_inc(x_431); +lean_dec(x_2); +x_432 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66; +lean_inc(x_430); +lean_inc(x_431); +x_433 = l_Lean_addMacroScope(x_431, x_432, x_430); +x_434 = lean_box(0); +x_435 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64; +x_436 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70; +lean_inc(x_429); +x_437 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_437, 0, x_429); +lean_ctor_set(x_437, 1, x_435); +lean_ctor_set(x_437, 2, x_433); +lean_ctor_set(x_437, 3, x_436); +x_438 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_429); +x_439 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_439, 0, x_429); +lean_ctor_set(x_439, 1, x_438); +x_440 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_429); +x_441 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_441, 0, x_429); +lean_ctor_set(x_441, 1, x_440); +x_442 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_443 = l_Lean_addMacroScope(x_431, x_442, x_430); +x_444 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_429); +x_445 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_445, 0, x_429); +lean_ctor_set(x_445, 1, x_444); +lean_ctor_set(x_445, 2, x_443); +lean_ctor_set(x_445, 3, x_434); +x_446 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_429); +x_447 = l_Lean_Syntax_node2(x_429, x_446, x_441, x_445); +x_448 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_429); +x_449 = l_Lean_Syntax_node1(x_429, x_448, x_425); +x_450 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_429); +x_451 = l_Lean_Syntax_node2(x_429, x_450, x_447, x_449); +x_452 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_429); +x_453 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_453, 0, x_429); +lean_ctor_set(x_453, 1, x_452); +x_454 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_429); +x_455 = l_Lean_Syntax_node3(x_429, x_454, x_439, x_451, x_453); +lean_inc(x_429); +x_456 = l_Lean_Syntax_node1(x_429, x_448, x_455); +x_457 = l_Lean_Syntax_node2(x_429, x_450, x_437, x_456); +x_458 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_458, 0, x_457); +lean_ctor_set(x_458, 1, x_426); +return x_458; +} +} +} +case 7: +{ +lean_object* x_459; lean_object* x_460; uint8_t x_461; +x_459 = lean_ctor_get(x_1, 0); +lean_inc(x_459); +lean_dec(x_1); +lean_inc(x_2); +x_460 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_459, x_2, x_3); +x_461 = !lean_is_exclusive(x_460); +if (x_461 == 0) +{ +lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +x_462 = lean_ctor_get(x_460, 0); +x_463 = lean_ctor_get(x_2, 5); +lean_inc(x_463); +x_464 = 0; +x_465 = l_Lean_SourceInfo_fromRef(x_463, x_464); +lean_dec(x_463); +x_466 = lean_ctor_get(x_2, 2); +lean_inc(x_466); +x_467 = lean_ctor_get(x_2, 1); +lean_inc(x_467); +lean_dec(x_2); +x_468 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74; +x_469 = l_Lean_addMacroScope(x_467, x_468, x_466); +x_470 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72; +x_471 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78; +lean_inc(x_465); +x_472 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_472, 0, x_465); +lean_ctor_set(x_472, 1, x_470); +lean_ctor_set(x_472, 2, x_469); +lean_ctor_set(x_472, 3, x_471); +x_473 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_465); +x_474 = l_Lean_Syntax_node1(x_465, x_473, x_462); +x_475 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_476 = l_Lean_Syntax_node2(x_465, x_475, x_472, x_474); +lean_ctor_set(x_460, 0, x_476); +return x_460; +} +else +{ +lean_object* x_477; lean_object* x_478; lean_object* x_479; uint8_t x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +x_477 = lean_ctor_get(x_460, 0); +x_478 = lean_ctor_get(x_460, 1); +lean_inc(x_478); +lean_inc(x_477); +lean_dec(x_460); +x_479 = lean_ctor_get(x_2, 5); +lean_inc(x_479); +x_480 = 0; +x_481 = l_Lean_SourceInfo_fromRef(x_479, x_480); +lean_dec(x_479); +x_482 = lean_ctor_get(x_2, 2); +lean_inc(x_482); +x_483 = lean_ctor_get(x_2, 1); +lean_inc(x_483); +lean_dec(x_2); +x_484 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74; +x_485 = l_Lean_addMacroScope(x_483, x_484, x_482); +x_486 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72; +x_487 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78; +lean_inc(x_481); +x_488 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_488, 0, x_481); +lean_ctor_set(x_488, 1, x_486); +lean_ctor_set(x_488, 2, x_485); +lean_ctor_set(x_488, 3, x_487); +x_489 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_481); +x_490 = l_Lean_Syntax_node1(x_481, x_489, x_477); +x_491 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_492 = l_Lean_Syntax_node2(x_481, x_491, x_488, x_490); +x_493 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_493, 0, x_492); +lean_ctor_set(x_493, 1, x_478); +return x_493; +} +} +case 8: +{ +lean_object* x_494; lean_object* x_495; uint8_t x_496; +x_494 = lean_ctor_get(x_1, 0); +lean_inc(x_494); +lean_dec(x_1); +lean_inc(x_2); +x_495 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_494, x_2, x_3); +x_496 = !lean_is_exclusive(x_495); +if (x_496 == 0) +{ +lean_object* x_497; lean_object* x_498; uint8_t x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; +x_497 = lean_ctor_get(x_495, 0); +x_498 = lean_ctor_get(x_2, 5); +lean_inc(x_498); +x_499 = 0; +x_500 = l_Lean_SourceInfo_fromRef(x_498, x_499); +lean_dec(x_498); +x_501 = lean_ctor_get(x_2, 2); +lean_inc(x_501); +x_502 = lean_ctor_get(x_2, 1); +lean_inc(x_502); +lean_dec(x_2); +x_503 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82; +x_504 = l_Lean_addMacroScope(x_502, x_503, x_501); +x_505 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80; +x_506 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86; +lean_inc(x_500); +x_507 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_507, 0, x_500); +lean_ctor_set(x_507, 1, x_505); +lean_ctor_set(x_507, 2, x_504); +lean_ctor_set(x_507, 3, x_506); +x_508 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_500); +x_509 = l_Lean_Syntax_node1(x_500, x_508, x_497); +x_510 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_511 = l_Lean_Syntax_node2(x_500, x_510, x_507, x_509); +lean_ctor_set(x_495, 0, x_511); +return x_495; +} +else +{ +lean_object* x_512; lean_object* x_513; lean_object* x_514; uint8_t x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; +x_512 = lean_ctor_get(x_495, 0); +x_513 = lean_ctor_get(x_495, 1); +lean_inc(x_513); +lean_inc(x_512); +lean_dec(x_495); +x_514 = lean_ctor_get(x_2, 5); +lean_inc(x_514); +x_515 = 0; +x_516 = l_Lean_SourceInfo_fromRef(x_514, x_515); +lean_dec(x_514); +x_517 = lean_ctor_get(x_2, 2); +lean_inc(x_517); +x_518 = lean_ctor_get(x_2, 1); +lean_inc(x_518); +lean_dec(x_2); +x_519 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82; +x_520 = l_Lean_addMacroScope(x_518, x_519, x_517); +x_521 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80; +x_522 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86; +lean_inc(x_516); +x_523 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_523, 0, x_516); +lean_ctor_set(x_523, 1, x_521); +lean_ctor_set(x_523, 2, x_520); +lean_ctor_set(x_523, 3, x_522); +x_524 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_516); +x_525 = l_Lean_Syntax_node1(x_516, x_524, x_512); +x_526 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_527 = l_Lean_Syntax_node2(x_516, x_526, x_523, x_525); +x_528 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_528, 0, x_527); +lean_ctor_set(x_528, 1, x_513); +return x_528; +} +} +case 9: +{ +uint8_t x_529; lean_object* x_530; uint8_t x_531; +x_529 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_530 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_529, x_2, x_3); +x_531 = !lean_is_exclusive(x_530); +if (x_531 == 0) +{ +lean_object* x_532; lean_object* x_533; uint8_t x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_532 = lean_ctor_get(x_530, 0); +x_533 = lean_ctor_get(x_2, 5); +lean_inc(x_533); +x_534 = 0; +x_535 = l_Lean_SourceInfo_fromRef(x_533, x_534); +lean_dec(x_533); +x_536 = lean_ctor_get(x_2, 2); +lean_inc(x_536); +x_537 = lean_ctor_get(x_2, 1); +lean_inc(x_537); +lean_dec(x_2); +x_538 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90; +x_539 = l_Lean_addMacroScope(x_537, x_538, x_536); +x_540 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88; +x_541 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94; +lean_inc(x_535); +x_542 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_542, 0, x_535); +lean_ctor_set(x_542, 1, x_540); +lean_ctor_set(x_542, 2, x_539); +lean_ctor_set(x_542, 3, x_541); +x_543 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_535); +x_544 = l_Lean_Syntax_node1(x_535, x_543, x_532); +x_545 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_546 = l_Lean_Syntax_node2(x_535, x_545, x_542, x_544); +lean_ctor_set(x_530, 0, x_546); +return x_530; +} +else +{ +lean_object* x_547; lean_object* x_548; lean_object* x_549; uint8_t x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; +x_547 = lean_ctor_get(x_530, 0); +x_548 = lean_ctor_get(x_530, 1); +lean_inc(x_548); +lean_inc(x_547); +lean_dec(x_530); +x_549 = lean_ctor_get(x_2, 5); +lean_inc(x_549); +x_550 = 0; +x_551 = l_Lean_SourceInfo_fromRef(x_549, x_550); +lean_dec(x_549); +x_552 = lean_ctor_get(x_2, 2); +lean_inc(x_552); +x_553 = lean_ctor_get(x_2, 1); +lean_inc(x_553); +lean_dec(x_2); +x_554 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90; +x_555 = l_Lean_addMacroScope(x_553, x_554, x_552); +x_556 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88; +x_557 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94; +lean_inc(x_551); +x_558 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_558, 0, x_551); +lean_ctor_set(x_558, 1, x_556); +lean_ctor_set(x_558, 2, x_555); +lean_ctor_set(x_558, 3, x_557); +x_559 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_551); +x_560 = l_Lean_Syntax_node1(x_551, x_559, x_547); +x_561 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_562 = l_Lean_Syntax_node2(x_551, x_561, x_558, x_560); +x_563 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_563, 0, x_562); +lean_ctor_set(x_563, 1, x_548); +return x_563; +} +} +case 10: +{ +lean_object* x_564; +x_564 = lean_ctor_get(x_1, 0); +lean_inc(x_564); +lean_dec(x_1); +if (lean_obj_tag(x_564) == 0) +{ +lean_object* x_565; lean_object* x_566; uint8_t x_567; +x_565 = lean_ctor_get(x_564, 0); +lean_inc(x_565); +lean_dec(x_564); +lean_inc(x_2); +x_566 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_565, x_2, x_3); +x_567 = !lean_is_exclusive(x_566); +if (x_567 == 0) +{ +lean_object* x_568; lean_object* x_569; uint8_t x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; +x_568 = lean_ctor_get(x_566, 0); +x_569 = lean_ctor_get(x_2, 5); +lean_inc(x_569); +x_570 = 0; +x_571 = l_Lean_SourceInfo_fromRef(x_569, x_570); +lean_dec(x_569); +x_572 = lean_ctor_get(x_2, 2); +lean_inc(x_572); +x_573 = lean_ctor_get(x_2, 1); +lean_inc(x_573); +lean_dec(x_2); +x_574 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +lean_inc(x_572); +lean_inc(x_573); +x_575 = l_Lean_addMacroScope(x_573, x_574, x_572); +x_576 = lean_box(0); +x_577 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96; +x_578 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102; +lean_inc(x_571); +x_579 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_579, 0, x_571); +lean_ctor_set(x_579, 1, x_577); +lean_ctor_set(x_579, 2, x_575); +lean_ctor_set(x_579, 3, x_578); +x_580 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_571); +x_581 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_581, 0, x_571); +lean_ctor_set(x_581, 1, x_580); +x_582 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_571); +x_583 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_583, 0, x_571); +lean_ctor_set(x_583, 1, x_582); +x_584 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_585 = l_Lean_addMacroScope(x_573, x_584, x_572); +x_586 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_571); +x_587 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_587, 0, x_571); +lean_ctor_set(x_587, 1, x_586); +lean_ctor_set(x_587, 2, x_585); +lean_ctor_set(x_587, 3, x_576); +x_588 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_571); +x_589 = l_Lean_Syntax_node2(x_571, x_588, x_583, x_587); +x_590 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_571); +x_591 = l_Lean_Syntax_node1(x_571, x_590, x_568); +x_592 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_571); +x_593 = l_Lean_Syntax_node2(x_571, x_592, x_589, x_591); +x_594 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_571); +x_595 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_595, 0, x_571); +lean_ctor_set(x_595, 1, x_594); +x_596 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_571); +x_597 = l_Lean_Syntax_node3(x_571, x_596, x_581, x_593, x_595); +lean_inc(x_571); +x_598 = l_Lean_Syntax_node1(x_571, x_590, x_597); +x_599 = l_Lean_Syntax_node2(x_571, x_592, x_579, x_598); +lean_ctor_set(x_566, 0, x_599); +return x_566; +} +else +{ +lean_object* x_600; lean_object* x_601; lean_object* x_602; uint8_t x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; +x_600 = lean_ctor_get(x_566, 0); +x_601 = lean_ctor_get(x_566, 1); +lean_inc(x_601); +lean_inc(x_600); +lean_dec(x_566); +x_602 = lean_ctor_get(x_2, 5); +lean_inc(x_602); +x_603 = 0; +x_604 = l_Lean_SourceInfo_fromRef(x_602, x_603); +lean_dec(x_602); +x_605 = lean_ctor_get(x_2, 2); +lean_inc(x_605); +x_606 = lean_ctor_get(x_2, 1); +lean_inc(x_606); +lean_dec(x_2); +x_607 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +lean_inc(x_605); +lean_inc(x_606); +x_608 = l_Lean_addMacroScope(x_606, x_607, x_605); +x_609 = lean_box(0); +x_610 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96; +x_611 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102; +lean_inc(x_604); +x_612 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_612, 0, x_604); +lean_ctor_set(x_612, 1, x_610); +lean_ctor_set(x_612, 2, x_608); +lean_ctor_set(x_612, 3, x_611); +x_613 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_604); +x_614 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_614, 0, x_604); +lean_ctor_set(x_614, 1, x_613); +x_615 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_604); +x_616 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_616, 0, x_604); +lean_ctor_set(x_616, 1, x_615); +x_617 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50; +x_618 = l_Lean_addMacroScope(x_606, x_617, x_605); +x_619 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49; +lean_inc(x_604); +x_620 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_620, 0, x_604); +lean_ctor_set(x_620, 1, x_619); +lean_ctor_set(x_620, 2, x_618); +lean_ctor_set(x_620, 3, x_609); +x_621 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_604); +x_622 = l_Lean_Syntax_node2(x_604, x_621, x_616, x_620); +x_623 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_604); +x_624 = l_Lean_Syntax_node1(x_604, x_623, x_600); +x_625 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_604); +x_626 = l_Lean_Syntax_node2(x_604, x_625, x_622, x_624); +x_627 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_604); +x_628 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_628, 0, x_604); +lean_ctor_set(x_628, 1, x_627); +x_629 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_604); +x_630 = l_Lean_Syntax_node3(x_604, x_629, x_614, x_626, x_628); +lean_inc(x_604); +x_631 = l_Lean_Syntax_node1(x_604, x_623, x_630); +x_632 = l_Lean_Syntax_node2(x_604, x_625, x_612, x_631); +x_633 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_633, 0, x_632); +lean_ctor_set(x_633, 1, x_601); +return x_633; +} +} +else +{ +lean_object* x_634; uint8_t x_635; lean_object* x_636; uint8_t x_637; +x_634 = lean_ctor_get(x_564, 0); +lean_inc(x_634); +lean_dec(x_564); +x_635 = lean_unbox(x_634); +lean_dec(x_634); +lean_inc(x_2); +x_636 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_635, x_2, x_3); +x_637 = !lean_is_exclusive(x_636); +if (x_637 == 0) +{ +lean_object* x_638; lean_object* x_639; uint8_t x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; +x_638 = lean_ctor_get(x_636, 0); +x_639 = lean_ctor_get(x_2, 5); +lean_inc(x_639); +x_640 = 0; +x_641 = l_Lean_SourceInfo_fromRef(x_639, x_640); +lean_dec(x_639); +x_642 = lean_ctor_get(x_2, 2); +lean_inc(x_642); +x_643 = lean_ctor_get(x_2, 1); +lean_inc(x_643); +lean_dec(x_2); +x_644 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +lean_inc(x_642); +lean_inc(x_643); +x_645 = l_Lean_addMacroScope(x_643, x_644, x_642); +x_646 = lean_box(0); +x_647 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96; +x_648 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102; +lean_inc(x_641); +x_649 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_649, 0, x_641); +lean_ctor_set(x_649, 1, x_647); +lean_ctor_set(x_649, 2, x_645); +lean_ctor_set(x_649, 3, x_648); +x_650 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_641); +x_651 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_651, 0, x_641); +lean_ctor_set(x_651, 1, x_650); +x_652 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_641); +x_653 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_653, 0, x_641); +lean_ctor_set(x_653, 1, x_652); +x_654 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_655 = l_Lean_addMacroScope(x_643, x_654, x_642); +x_656 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_641); +x_657 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_657, 0, x_641); +lean_ctor_set(x_657, 1, x_656); +lean_ctor_set(x_657, 2, x_655); +lean_ctor_set(x_657, 3, x_646); +x_658 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_641); +x_659 = l_Lean_Syntax_node2(x_641, x_658, x_653, x_657); +x_660 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_641); +x_661 = l_Lean_Syntax_node1(x_641, x_660, x_638); +x_662 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_641); +x_663 = l_Lean_Syntax_node2(x_641, x_662, x_659, x_661); +x_664 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_641); +x_665 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_665, 0, x_641); +lean_ctor_set(x_665, 1, x_664); +x_666 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_641); +x_667 = l_Lean_Syntax_node3(x_641, x_666, x_651, x_663, x_665); +lean_inc(x_641); +x_668 = l_Lean_Syntax_node1(x_641, x_660, x_667); +x_669 = l_Lean_Syntax_node2(x_641, x_662, x_649, x_668); +lean_ctor_set(x_636, 0, x_669); +return x_636; +} +else +{ +lean_object* x_670; lean_object* x_671; lean_object* x_672; uint8_t x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; +x_670 = lean_ctor_get(x_636, 0); +x_671 = lean_ctor_get(x_636, 1); +lean_inc(x_671); +lean_inc(x_670); +lean_dec(x_636); +x_672 = lean_ctor_get(x_2, 5); +lean_inc(x_672); +x_673 = 0; +x_674 = l_Lean_SourceInfo_fromRef(x_672, x_673); +lean_dec(x_672); +x_675 = lean_ctor_get(x_2, 2); +lean_inc(x_675); +x_676 = lean_ctor_get(x_2, 1); +lean_inc(x_676); +lean_dec(x_2); +x_677 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98; +lean_inc(x_675); +lean_inc(x_676); +x_678 = l_Lean_addMacroScope(x_676, x_677, x_675); +x_679 = lean_box(0); +x_680 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96; +x_681 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102; +lean_inc(x_674); +x_682 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_682, 0, x_674); +lean_ctor_set(x_682, 1, x_680); +lean_ctor_set(x_682, 2, x_678); +lean_ctor_set(x_682, 3, x_681); +x_683 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_674); +x_684 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_684, 0, x_674); +lean_ctor_set(x_684, 1, x_683); +x_685 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_674); +x_686 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_686, 0, x_674); +lean_ctor_set(x_686, 1, x_685); +x_687 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54; +x_688 = l_Lean_addMacroScope(x_676, x_687, x_675); +x_689 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53; +lean_inc(x_674); +x_690 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_690, 0, x_674); +lean_ctor_set(x_690, 1, x_689); +lean_ctor_set(x_690, 2, x_688); +lean_ctor_set(x_690, 3, x_679); +x_691 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_674); +x_692 = l_Lean_Syntax_node2(x_674, x_691, x_686, x_690); +x_693 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_674); +x_694 = l_Lean_Syntax_node1(x_674, x_693, x_670); +x_695 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_674); +x_696 = l_Lean_Syntax_node2(x_674, x_695, x_692, x_694); +x_697 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_674); +x_698 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_698, 0, x_674); +lean_ctor_set(x_698, 1, x_697); +x_699 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_674); +x_700 = l_Lean_Syntax_node3(x_674, x_699, x_684, x_696, x_698); +lean_inc(x_674); +x_701 = l_Lean_Syntax_node1(x_674, x_693, x_700); +x_702 = l_Lean_Syntax_node2(x_674, x_695, x_682, x_701); +x_703 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_703, 0, x_702); +lean_ctor_set(x_703, 1, x_671); +return x_703; +} +} +} +case 11: +{ +lean_object* x_704; lean_object* x_705; uint8_t x_706; +x_704 = lean_ctor_get(x_1, 0); +lean_inc(x_704); +lean_dec(x_1); +lean_inc(x_2); +x_705 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_704, x_2, x_3); +x_706 = !lean_is_exclusive(x_705); +if (x_706 == 0) +{ +lean_object* x_707; lean_object* x_708; uint8_t x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; +x_707 = lean_ctor_get(x_705, 0); +x_708 = lean_ctor_get(x_2, 5); +lean_inc(x_708); +x_709 = 0; +x_710 = l_Lean_SourceInfo_fromRef(x_708, x_709); +lean_dec(x_708); +x_711 = lean_ctor_get(x_2, 2); +lean_inc(x_711); +x_712 = lean_ctor_get(x_2, 1); +lean_inc(x_712); +lean_dec(x_2); +x_713 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106; +x_714 = l_Lean_addMacroScope(x_712, x_713, x_711); +x_715 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104; +x_716 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110; +lean_inc(x_710); +x_717 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_717, 0, x_710); +lean_ctor_set(x_717, 1, x_715); +lean_ctor_set(x_717, 2, x_714); +lean_ctor_set(x_717, 3, x_716); +x_718 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_710); +x_719 = l_Lean_Syntax_node1(x_710, x_718, x_707); +x_720 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_721 = l_Lean_Syntax_node2(x_710, x_720, x_717, x_719); +lean_ctor_set(x_705, 0, x_721); +return x_705; +} +else +{ +lean_object* x_722; lean_object* x_723; lean_object* x_724; uint8_t x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; +x_722 = lean_ctor_get(x_705, 0); +x_723 = lean_ctor_get(x_705, 1); +lean_inc(x_723); +lean_inc(x_722); +lean_dec(x_705); +x_724 = lean_ctor_get(x_2, 5); +lean_inc(x_724); +x_725 = 0; +x_726 = l_Lean_SourceInfo_fromRef(x_724, x_725); +lean_dec(x_724); +x_727 = lean_ctor_get(x_2, 2); +lean_inc(x_727); +x_728 = lean_ctor_get(x_2, 1); +lean_inc(x_728); +lean_dec(x_2); +x_729 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106; +x_730 = l_Lean_addMacroScope(x_728, x_729, x_727); +x_731 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104; +x_732 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110; +lean_inc(x_726); +x_733 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_733, 0, x_726); +lean_ctor_set(x_733, 1, x_731); +lean_ctor_set(x_733, 2, x_730); +lean_ctor_set(x_733, 3, x_732); +x_734 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_726); +x_735 = l_Lean_Syntax_node1(x_726, x_734, x_722); +x_736 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_737 = l_Lean_Syntax_node2(x_726, x_736, x_733, x_735); +x_738 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_738, 0, x_737); +lean_ctor_set(x_738, 1, x_723); +return x_738; +} +} +case 12: +{ +uint8_t x_739; lean_object* x_740; uint8_t x_741; +x_739 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_740 = l___private_Std_Time_Notation_0__Std_Time_convertText(x_739, x_2, x_3); +x_741 = !lean_is_exclusive(x_740); +if (x_741 == 0) +{ +lean_object* x_742; lean_object* x_743; uint8_t x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; +x_742 = lean_ctor_get(x_740, 0); +x_743 = lean_ctor_get(x_2, 5); +lean_inc(x_743); +x_744 = 0; +x_745 = l_Lean_SourceInfo_fromRef(x_743, x_744); +lean_dec(x_743); +x_746 = lean_ctor_get(x_2, 2); +lean_inc(x_746); +x_747 = lean_ctor_get(x_2, 1); +lean_inc(x_747); +lean_dec(x_2); +x_748 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114; +x_749 = l_Lean_addMacroScope(x_747, x_748, x_746); +x_750 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112; +x_751 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118; +lean_inc(x_745); +x_752 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_752, 0, x_745); +lean_ctor_set(x_752, 1, x_750); +lean_ctor_set(x_752, 2, x_749); +lean_ctor_set(x_752, 3, x_751); +x_753 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_745); +x_754 = l_Lean_Syntax_node1(x_745, x_753, x_742); +x_755 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_756 = l_Lean_Syntax_node2(x_745, x_755, x_752, x_754); +lean_ctor_set(x_740, 0, x_756); +return x_740; +} +else +{ +lean_object* x_757; lean_object* x_758; lean_object* x_759; uint8_t x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; +x_757 = lean_ctor_get(x_740, 0); +x_758 = lean_ctor_get(x_740, 1); +lean_inc(x_758); +lean_inc(x_757); +lean_dec(x_740); +x_759 = lean_ctor_get(x_2, 5); +lean_inc(x_759); +x_760 = 0; +x_761 = l_Lean_SourceInfo_fromRef(x_759, x_760); +lean_dec(x_759); +x_762 = lean_ctor_get(x_2, 2); +lean_inc(x_762); +x_763 = lean_ctor_get(x_2, 1); +lean_inc(x_763); +lean_dec(x_2); +x_764 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114; +x_765 = l_Lean_addMacroScope(x_763, x_764, x_762); +x_766 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112; +x_767 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118; +lean_inc(x_761); +x_768 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_768, 0, x_761); +lean_ctor_set(x_768, 1, x_766); +lean_ctor_set(x_768, 2, x_765); +lean_ctor_set(x_768, 3, x_767); +x_769 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_761); +x_770 = l_Lean_Syntax_node1(x_761, x_769, x_757); +x_771 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_772 = l_Lean_Syntax_node2(x_761, x_771, x_768, x_770); +x_773 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_773, 0, x_772); +lean_ctor_set(x_773, 1, x_758); +return x_773; +} +} +case 13: +{ +lean_object* x_774; lean_object* x_775; uint8_t x_776; +x_774 = lean_ctor_get(x_1, 0); +lean_inc(x_774); +lean_dec(x_1); +lean_inc(x_2); +x_775 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_774, x_2, x_3); +x_776 = !lean_is_exclusive(x_775); +if (x_776 == 0) +{ +lean_object* x_777; lean_object* x_778; uint8_t x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; +x_777 = lean_ctor_get(x_775, 0); +x_778 = lean_ctor_get(x_2, 5); +lean_inc(x_778); +x_779 = 0; +x_780 = l_Lean_SourceInfo_fromRef(x_778, x_779); +lean_dec(x_778); +x_781 = lean_ctor_get(x_2, 2); +lean_inc(x_781); +x_782 = lean_ctor_get(x_2, 1); +lean_inc(x_782); +lean_dec(x_2); +x_783 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122; +x_784 = l_Lean_addMacroScope(x_782, x_783, x_781); +x_785 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120; +x_786 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126; +lean_inc(x_780); +x_787 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_787, 0, x_780); +lean_ctor_set(x_787, 1, x_785); +lean_ctor_set(x_787, 2, x_784); +lean_ctor_set(x_787, 3, x_786); +x_788 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_780); +x_789 = l_Lean_Syntax_node1(x_780, x_788, x_777); +x_790 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_791 = l_Lean_Syntax_node2(x_780, x_790, x_787, x_789); +lean_ctor_set(x_775, 0, x_791); +return x_775; +} +else +{ +lean_object* x_792; lean_object* x_793; lean_object* x_794; uint8_t x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; +x_792 = lean_ctor_get(x_775, 0); +x_793 = lean_ctor_get(x_775, 1); +lean_inc(x_793); +lean_inc(x_792); +lean_dec(x_775); +x_794 = lean_ctor_get(x_2, 5); +lean_inc(x_794); +x_795 = 0; +x_796 = l_Lean_SourceInfo_fromRef(x_794, x_795); +lean_dec(x_794); +x_797 = lean_ctor_get(x_2, 2); +lean_inc(x_797); +x_798 = lean_ctor_get(x_2, 1); +lean_inc(x_798); +lean_dec(x_2); +x_799 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122; +x_800 = l_Lean_addMacroScope(x_798, x_799, x_797); +x_801 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120; +x_802 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126; +lean_inc(x_796); +x_803 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_803, 0, x_796); +lean_ctor_set(x_803, 1, x_801); +lean_ctor_set(x_803, 2, x_800); +lean_ctor_set(x_803, 3, x_802); +x_804 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_796); +x_805 = l_Lean_Syntax_node1(x_796, x_804, x_792); +x_806 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_807 = l_Lean_Syntax_node2(x_796, x_806, x_803, x_805); +x_808 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_808, 0, x_807); +lean_ctor_set(x_808, 1, x_793); +return x_808; +} +} +case 14: +{ +lean_object* x_809; lean_object* x_810; uint8_t x_811; +x_809 = lean_ctor_get(x_1, 0); +lean_inc(x_809); +lean_dec(x_1); +lean_inc(x_2); +x_810 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_809, x_2, x_3); +x_811 = !lean_is_exclusive(x_810); +if (x_811 == 0) +{ +lean_object* x_812; lean_object* x_813; uint8_t x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; +x_812 = lean_ctor_get(x_810, 0); +x_813 = lean_ctor_get(x_2, 5); +lean_inc(x_813); +x_814 = 0; +x_815 = l_Lean_SourceInfo_fromRef(x_813, x_814); +lean_dec(x_813); +x_816 = lean_ctor_get(x_2, 2); +lean_inc(x_816); +x_817 = lean_ctor_get(x_2, 1); +lean_inc(x_817); +lean_dec(x_2); +x_818 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130; +x_819 = l_Lean_addMacroScope(x_817, x_818, x_816); +x_820 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128; +x_821 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134; +lean_inc(x_815); +x_822 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_822, 0, x_815); +lean_ctor_set(x_822, 1, x_820); +lean_ctor_set(x_822, 2, x_819); +lean_ctor_set(x_822, 3, x_821); +x_823 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_815); +x_824 = l_Lean_Syntax_node1(x_815, x_823, x_812); +x_825 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_826 = l_Lean_Syntax_node2(x_815, x_825, x_822, x_824); +lean_ctor_set(x_810, 0, x_826); +return x_810; +} +else +{ +lean_object* x_827; lean_object* x_828; lean_object* x_829; uint8_t x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; +x_827 = lean_ctor_get(x_810, 0); +x_828 = lean_ctor_get(x_810, 1); +lean_inc(x_828); +lean_inc(x_827); +lean_dec(x_810); +x_829 = lean_ctor_get(x_2, 5); +lean_inc(x_829); +x_830 = 0; +x_831 = l_Lean_SourceInfo_fromRef(x_829, x_830); +lean_dec(x_829); +x_832 = lean_ctor_get(x_2, 2); +lean_inc(x_832); +x_833 = lean_ctor_get(x_2, 1); +lean_inc(x_833); +lean_dec(x_2); +x_834 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130; +x_835 = l_Lean_addMacroScope(x_833, x_834, x_832); +x_836 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128; +x_837 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134; +lean_inc(x_831); +x_838 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_838, 0, x_831); +lean_ctor_set(x_838, 1, x_836); +lean_ctor_set(x_838, 2, x_835); +lean_ctor_set(x_838, 3, x_837); +x_839 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_831); +x_840 = l_Lean_Syntax_node1(x_831, x_839, x_827); +x_841 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_842 = l_Lean_Syntax_node2(x_831, x_841, x_838, x_840); +x_843 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_843, 0, x_842); +lean_ctor_set(x_843, 1, x_828); +return x_843; +} +} +case 15: +{ +lean_object* x_844; lean_object* x_845; uint8_t x_846; +x_844 = lean_ctor_get(x_1, 0); +lean_inc(x_844); +lean_dec(x_1); +lean_inc(x_2); +x_845 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_844, x_2, x_3); +x_846 = !lean_is_exclusive(x_845); +if (x_846 == 0) +{ +lean_object* x_847; lean_object* x_848; uint8_t x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; +x_847 = lean_ctor_get(x_845, 0); +x_848 = lean_ctor_get(x_2, 5); +lean_inc(x_848); +x_849 = 0; +x_850 = l_Lean_SourceInfo_fromRef(x_848, x_849); +lean_dec(x_848); +x_851 = lean_ctor_get(x_2, 2); +lean_inc(x_851); +x_852 = lean_ctor_get(x_2, 1); +lean_inc(x_852); +lean_dec(x_2); +x_853 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138; +x_854 = l_Lean_addMacroScope(x_852, x_853, x_851); +x_855 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136; +x_856 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142; +lean_inc(x_850); +x_857 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_857, 0, x_850); +lean_ctor_set(x_857, 1, x_855); +lean_ctor_set(x_857, 2, x_854); +lean_ctor_set(x_857, 3, x_856); +x_858 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_850); +x_859 = l_Lean_Syntax_node1(x_850, x_858, x_847); +x_860 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_861 = l_Lean_Syntax_node2(x_850, x_860, x_857, x_859); +lean_ctor_set(x_845, 0, x_861); +return x_845; +} +else +{ +lean_object* x_862; lean_object* x_863; lean_object* x_864; uint8_t x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; +x_862 = lean_ctor_get(x_845, 0); +x_863 = lean_ctor_get(x_845, 1); +lean_inc(x_863); +lean_inc(x_862); +lean_dec(x_845); +x_864 = lean_ctor_get(x_2, 5); +lean_inc(x_864); +x_865 = 0; +x_866 = l_Lean_SourceInfo_fromRef(x_864, x_865); +lean_dec(x_864); +x_867 = lean_ctor_get(x_2, 2); +lean_inc(x_867); +x_868 = lean_ctor_get(x_2, 1); +lean_inc(x_868); +lean_dec(x_2); +x_869 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138; +x_870 = l_Lean_addMacroScope(x_868, x_869, x_867); +x_871 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136; +x_872 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142; +lean_inc(x_866); +x_873 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_873, 0, x_866); +lean_ctor_set(x_873, 1, x_871); +lean_ctor_set(x_873, 2, x_870); +lean_ctor_set(x_873, 3, x_872); +x_874 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_866); +x_875 = l_Lean_Syntax_node1(x_866, x_874, x_862); +x_876 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_877 = l_Lean_Syntax_node2(x_866, x_876, x_873, x_875); +x_878 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_878, 0, x_877); +lean_ctor_set(x_878, 1, x_863); +return x_878; +} +} +case 16: +{ +lean_object* x_879; lean_object* x_880; uint8_t x_881; +x_879 = lean_ctor_get(x_1, 0); +lean_inc(x_879); +lean_dec(x_1); +lean_inc(x_2); +x_880 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_879, x_2, x_3); +x_881 = !lean_is_exclusive(x_880); +if (x_881 == 0) +{ +lean_object* x_882; lean_object* x_883; uint8_t x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; +x_882 = lean_ctor_get(x_880, 0); +x_883 = lean_ctor_get(x_2, 5); +lean_inc(x_883); +x_884 = 0; +x_885 = l_Lean_SourceInfo_fromRef(x_883, x_884); +lean_dec(x_883); +x_886 = lean_ctor_get(x_2, 2); +lean_inc(x_886); +x_887 = lean_ctor_get(x_2, 1); +lean_inc(x_887); +lean_dec(x_2); +x_888 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146; +x_889 = l_Lean_addMacroScope(x_887, x_888, x_886); +x_890 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144; +x_891 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150; +lean_inc(x_885); +x_892 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_892, 0, x_885); +lean_ctor_set(x_892, 1, x_890); +lean_ctor_set(x_892, 2, x_889); +lean_ctor_set(x_892, 3, x_891); +x_893 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_885); +x_894 = l_Lean_Syntax_node1(x_885, x_893, x_882); +x_895 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_896 = l_Lean_Syntax_node2(x_885, x_895, x_892, x_894); +lean_ctor_set(x_880, 0, x_896); +return x_880; +} +else +{ +lean_object* x_897; lean_object* x_898; lean_object* x_899; uint8_t x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; +x_897 = lean_ctor_get(x_880, 0); +x_898 = lean_ctor_get(x_880, 1); +lean_inc(x_898); +lean_inc(x_897); +lean_dec(x_880); +x_899 = lean_ctor_get(x_2, 5); +lean_inc(x_899); +x_900 = 0; +x_901 = l_Lean_SourceInfo_fromRef(x_899, x_900); +lean_dec(x_899); +x_902 = lean_ctor_get(x_2, 2); +lean_inc(x_902); +x_903 = lean_ctor_get(x_2, 1); +lean_inc(x_903); +lean_dec(x_2); +x_904 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146; +x_905 = l_Lean_addMacroScope(x_903, x_904, x_902); +x_906 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144; +x_907 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150; +lean_inc(x_901); +x_908 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_908, 0, x_901); +lean_ctor_set(x_908, 1, x_906); +lean_ctor_set(x_908, 2, x_905); +lean_ctor_set(x_908, 3, x_907); +x_909 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_901); +x_910 = l_Lean_Syntax_node1(x_901, x_909, x_897); +x_911 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_912 = l_Lean_Syntax_node2(x_901, x_911, x_908, x_910); +x_913 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_913, 0, x_912); +lean_ctor_set(x_913, 1, x_898); +return x_913; +} +} +case 17: +{ +lean_object* x_914; lean_object* x_915; uint8_t x_916; +x_914 = lean_ctor_get(x_1, 0); +lean_inc(x_914); +lean_dec(x_1); +lean_inc(x_2); +x_915 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_914, x_2, x_3); +x_916 = !lean_is_exclusive(x_915); +if (x_916 == 0) +{ +lean_object* x_917; lean_object* x_918; uint8_t x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +x_917 = lean_ctor_get(x_915, 0); +x_918 = lean_ctor_get(x_2, 5); +lean_inc(x_918); +x_919 = 0; +x_920 = l_Lean_SourceInfo_fromRef(x_918, x_919); +lean_dec(x_918); +x_921 = lean_ctor_get(x_2, 2); +lean_inc(x_921); +x_922 = lean_ctor_get(x_2, 1); +lean_inc(x_922); +lean_dec(x_2); +x_923 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154; +x_924 = l_Lean_addMacroScope(x_922, x_923, x_921); +x_925 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152; +x_926 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158; +lean_inc(x_920); +x_927 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_927, 0, x_920); +lean_ctor_set(x_927, 1, x_925); +lean_ctor_set(x_927, 2, x_924); +lean_ctor_set(x_927, 3, x_926); +x_928 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_920); +x_929 = l_Lean_Syntax_node1(x_920, x_928, x_917); +x_930 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_931 = l_Lean_Syntax_node2(x_920, x_930, x_927, x_929); +lean_ctor_set(x_915, 0, x_931); +return x_915; +} +else +{ +lean_object* x_932; lean_object* x_933; lean_object* x_934; uint8_t x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; +x_932 = lean_ctor_get(x_915, 0); +x_933 = lean_ctor_get(x_915, 1); +lean_inc(x_933); +lean_inc(x_932); +lean_dec(x_915); +x_934 = lean_ctor_get(x_2, 5); +lean_inc(x_934); +x_935 = 0; +x_936 = l_Lean_SourceInfo_fromRef(x_934, x_935); +lean_dec(x_934); +x_937 = lean_ctor_get(x_2, 2); +lean_inc(x_937); +x_938 = lean_ctor_get(x_2, 1); +lean_inc(x_938); +lean_dec(x_2); +x_939 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154; +x_940 = l_Lean_addMacroScope(x_938, x_939, x_937); +x_941 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152; +x_942 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158; +lean_inc(x_936); +x_943 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_943, 0, x_936); +lean_ctor_set(x_943, 1, x_941); +lean_ctor_set(x_943, 2, x_940); +lean_ctor_set(x_943, 3, x_942); +x_944 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_936); +x_945 = l_Lean_Syntax_node1(x_936, x_944, x_932); +x_946 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_947 = l_Lean_Syntax_node2(x_936, x_946, x_943, x_945); +x_948 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_948, 0, x_947); +lean_ctor_set(x_948, 1, x_933); +return x_948; +} +} +case 18: +{ +lean_object* x_949; lean_object* x_950; uint8_t x_951; +x_949 = lean_ctor_get(x_1, 0); +lean_inc(x_949); +lean_dec(x_1); +lean_inc(x_2); +x_950 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_949, x_2, x_3); +x_951 = !lean_is_exclusive(x_950); +if (x_951 == 0) +{ +lean_object* x_952; lean_object* x_953; uint8_t x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; +x_952 = lean_ctor_get(x_950, 0); +x_953 = lean_ctor_get(x_2, 5); +lean_inc(x_953); +x_954 = 0; +x_955 = l_Lean_SourceInfo_fromRef(x_953, x_954); +lean_dec(x_953); +x_956 = lean_ctor_get(x_2, 2); +lean_inc(x_956); +x_957 = lean_ctor_get(x_2, 1); +lean_inc(x_957); +lean_dec(x_2); +x_958 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162; +x_959 = l_Lean_addMacroScope(x_957, x_958, x_956); +x_960 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160; +x_961 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166; +lean_inc(x_955); +x_962 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_962, 0, x_955); +lean_ctor_set(x_962, 1, x_960); +lean_ctor_set(x_962, 2, x_959); +lean_ctor_set(x_962, 3, x_961); +x_963 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_955); +x_964 = l_Lean_Syntax_node1(x_955, x_963, x_952); +x_965 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_966 = l_Lean_Syntax_node2(x_955, x_965, x_962, x_964); +lean_ctor_set(x_950, 0, x_966); +return x_950; +} +else +{ +lean_object* x_967; lean_object* x_968; lean_object* x_969; uint8_t x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +x_967 = lean_ctor_get(x_950, 0); +x_968 = lean_ctor_get(x_950, 1); +lean_inc(x_968); +lean_inc(x_967); +lean_dec(x_950); +x_969 = lean_ctor_get(x_2, 5); +lean_inc(x_969); +x_970 = 0; +x_971 = l_Lean_SourceInfo_fromRef(x_969, x_970); +lean_dec(x_969); +x_972 = lean_ctor_get(x_2, 2); +lean_inc(x_972); +x_973 = lean_ctor_get(x_2, 1); +lean_inc(x_973); +lean_dec(x_2); +x_974 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162; +x_975 = l_Lean_addMacroScope(x_973, x_974, x_972); +x_976 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160; +x_977 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166; +lean_inc(x_971); +x_978 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_978, 0, x_971); +lean_ctor_set(x_978, 1, x_976); +lean_ctor_set(x_978, 2, x_975); +lean_ctor_set(x_978, 3, x_977); +x_979 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_971); +x_980 = l_Lean_Syntax_node1(x_971, x_979, x_967); +x_981 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_982 = l_Lean_Syntax_node2(x_971, x_981, x_978, x_980); +x_983 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_983, 0, x_982); +lean_ctor_set(x_983, 1, x_968); +return x_983; +} +} +case 19: +{ +lean_object* x_984; lean_object* x_985; uint8_t x_986; +x_984 = lean_ctor_get(x_1, 0); +lean_inc(x_984); +lean_dec(x_1); +lean_inc(x_2); +x_985 = l___private_Std_Time_Notation_0__Std_Time_convertFraction(x_984, x_2, x_3); +x_986 = !lean_is_exclusive(x_985); +if (x_986 == 0) +{ +lean_object* x_987; lean_object* x_988; uint8_t x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; +x_987 = lean_ctor_get(x_985, 0); +x_988 = lean_ctor_get(x_2, 5); +lean_inc(x_988); +x_989 = 0; +x_990 = l_Lean_SourceInfo_fromRef(x_988, x_989); +lean_dec(x_988); +x_991 = lean_ctor_get(x_2, 2); +lean_inc(x_991); +x_992 = lean_ctor_get(x_2, 1); +lean_inc(x_992); +lean_dec(x_2); +x_993 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170; +x_994 = l_Lean_addMacroScope(x_992, x_993, x_991); +x_995 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168; +x_996 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174; +lean_inc(x_990); +x_997 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_997, 0, x_990); +lean_ctor_set(x_997, 1, x_995); +lean_ctor_set(x_997, 2, x_994); +lean_ctor_set(x_997, 3, x_996); +x_998 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_990); +x_999 = l_Lean_Syntax_node1(x_990, x_998, x_987); +x_1000 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1001 = l_Lean_Syntax_node2(x_990, x_1000, x_997, x_999); +lean_ctor_set(x_985, 0, x_1001); +return x_985; +} +else +{ +lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; uint8_t x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; +x_1002 = lean_ctor_get(x_985, 0); +x_1003 = lean_ctor_get(x_985, 1); +lean_inc(x_1003); +lean_inc(x_1002); +lean_dec(x_985); +x_1004 = lean_ctor_get(x_2, 5); +lean_inc(x_1004); +x_1005 = 0; +x_1006 = l_Lean_SourceInfo_fromRef(x_1004, x_1005); +lean_dec(x_1004); +x_1007 = lean_ctor_get(x_2, 2); +lean_inc(x_1007); +x_1008 = lean_ctor_get(x_2, 1); +lean_inc(x_1008); +lean_dec(x_2); +x_1009 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170; +x_1010 = l_Lean_addMacroScope(x_1008, x_1009, x_1007); +x_1011 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168; +x_1012 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174; +lean_inc(x_1006); +x_1013 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1013, 0, x_1006); +lean_ctor_set(x_1013, 1, x_1011); +lean_ctor_set(x_1013, 2, x_1010); +lean_ctor_set(x_1013, 3, x_1012); +x_1014 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1006); +x_1015 = l_Lean_Syntax_node1(x_1006, x_1014, x_1002); +x_1016 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1017 = l_Lean_Syntax_node2(x_1006, x_1016, x_1013, x_1015); +x_1018 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1018, 0, x_1017); +lean_ctor_set(x_1018, 1, x_1003); +return x_1018; +} +} +case 20: +{ +lean_object* x_1019; lean_object* x_1020; uint8_t x_1021; +x_1019 = lean_ctor_get(x_1, 0); +lean_inc(x_1019); +lean_dec(x_1); +lean_inc(x_2); +x_1020 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_1019, x_2, x_3); +x_1021 = !lean_is_exclusive(x_1020); +if (x_1021 == 0) +{ +lean_object* x_1022; lean_object* x_1023; uint8_t x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; +x_1022 = lean_ctor_get(x_1020, 0); +x_1023 = lean_ctor_get(x_2, 5); +lean_inc(x_1023); +x_1024 = 0; +x_1025 = l_Lean_SourceInfo_fromRef(x_1023, x_1024); +lean_dec(x_1023); +x_1026 = lean_ctor_get(x_2, 2); +lean_inc(x_1026); +x_1027 = lean_ctor_get(x_2, 1); +lean_inc(x_1027); +lean_dec(x_2); +x_1028 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178; +x_1029 = l_Lean_addMacroScope(x_1027, x_1028, x_1026); +x_1030 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176; +x_1031 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182; +lean_inc(x_1025); +x_1032 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1032, 0, x_1025); +lean_ctor_set(x_1032, 1, x_1030); +lean_ctor_set(x_1032, 2, x_1029); +lean_ctor_set(x_1032, 3, x_1031); +x_1033 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1025); +x_1034 = l_Lean_Syntax_node1(x_1025, x_1033, x_1022); +x_1035 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1036 = l_Lean_Syntax_node2(x_1025, x_1035, x_1032, x_1034); +lean_ctor_set(x_1020, 0, x_1036); +return x_1020; +} +else +{ +lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; uint8_t x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; +x_1037 = lean_ctor_get(x_1020, 0); +x_1038 = lean_ctor_get(x_1020, 1); +lean_inc(x_1038); +lean_inc(x_1037); +lean_dec(x_1020); +x_1039 = lean_ctor_get(x_2, 5); +lean_inc(x_1039); +x_1040 = 0; +x_1041 = l_Lean_SourceInfo_fromRef(x_1039, x_1040); +lean_dec(x_1039); +x_1042 = lean_ctor_get(x_2, 2); +lean_inc(x_1042); +x_1043 = lean_ctor_get(x_2, 1); +lean_inc(x_1043); +lean_dec(x_2); +x_1044 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178; +x_1045 = l_Lean_addMacroScope(x_1043, x_1044, x_1042); +x_1046 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176; +x_1047 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182; +lean_inc(x_1041); +x_1048 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1048, 0, x_1041); +lean_ctor_set(x_1048, 1, x_1046); +lean_ctor_set(x_1048, 2, x_1045); +lean_ctor_set(x_1048, 3, x_1047); +x_1049 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1041); +x_1050 = l_Lean_Syntax_node1(x_1041, x_1049, x_1037); +x_1051 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1052 = l_Lean_Syntax_node2(x_1041, x_1051, x_1048, x_1050); +x_1053 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1053, 0, x_1052); +lean_ctor_set(x_1053, 1, x_1038); +return x_1053; +} +} +case 21: +{ +lean_object* x_1054; lean_object* x_1055; uint8_t x_1056; +x_1054 = lean_ctor_get(x_1, 0); +lean_inc(x_1054); +lean_dec(x_1); +lean_inc(x_2); +x_1055 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_1054, x_2, x_3); +x_1056 = !lean_is_exclusive(x_1055); +if (x_1056 == 0) +{ +lean_object* x_1057; lean_object* x_1058; uint8_t x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; +x_1057 = lean_ctor_get(x_1055, 0); +x_1058 = lean_ctor_get(x_2, 5); +lean_inc(x_1058); +x_1059 = 0; +x_1060 = l_Lean_SourceInfo_fromRef(x_1058, x_1059); +lean_dec(x_1058); +x_1061 = lean_ctor_get(x_2, 2); +lean_inc(x_1061); +x_1062 = lean_ctor_get(x_2, 1); +lean_inc(x_1062); +lean_dec(x_2); +x_1063 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186; +x_1064 = l_Lean_addMacroScope(x_1062, x_1063, x_1061); +x_1065 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184; +x_1066 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190; +lean_inc(x_1060); +x_1067 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1067, 0, x_1060); +lean_ctor_set(x_1067, 1, x_1065); +lean_ctor_set(x_1067, 2, x_1064); +lean_ctor_set(x_1067, 3, x_1066); +x_1068 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1060); +x_1069 = l_Lean_Syntax_node1(x_1060, x_1068, x_1057); +x_1070 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1071 = l_Lean_Syntax_node2(x_1060, x_1070, x_1067, x_1069); +lean_ctor_set(x_1055, 0, x_1071); +return x_1055; +} +else +{ +lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; uint8_t x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; +x_1072 = lean_ctor_get(x_1055, 0); +x_1073 = lean_ctor_get(x_1055, 1); +lean_inc(x_1073); +lean_inc(x_1072); +lean_dec(x_1055); +x_1074 = lean_ctor_get(x_2, 5); +lean_inc(x_1074); +x_1075 = 0; +x_1076 = l_Lean_SourceInfo_fromRef(x_1074, x_1075); +lean_dec(x_1074); +x_1077 = lean_ctor_get(x_2, 2); +lean_inc(x_1077); +x_1078 = lean_ctor_get(x_2, 1); +lean_inc(x_1078); +lean_dec(x_2); +x_1079 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186; +x_1080 = l_Lean_addMacroScope(x_1078, x_1079, x_1077); +x_1081 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184; +x_1082 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190; +lean_inc(x_1076); +x_1083 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1083, 0, x_1076); +lean_ctor_set(x_1083, 1, x_1081); +lean_ctor_set(x_1083, 2, x_1080); +lean_ctor_set(x_1083, 3, x_1082); +x_1084 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1076); +x_1085 = l_Lean_Syntax_node1(x_1076, x_1084, x_1072); +x_1086 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1087 = l_Lean_Syntax_node2(x_1076, x_1086, x_1083, x_1085); +x_1088 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1088, 0, x_1087); +lean_ctor_set(x_1088, 1, x_1073); +return x_1088; +} +} +case 22: +{ +lean_object* x_1089; lean_object* x_1090; uint8_t x_1091; +x_1089 = lean_ctor_get(x_1, 0); +lean_inc(x_1089); +lean_dec(x_1); +lean_inc(x_2); +x_1090 = l___private_Std_Time_Notation_0__Std_Time_convertNumber(x_1089, x_2, x_3); +x_1091 = !lean_is_exclusive(x_1090); +if (x_1091 == 0) +{ +lean_object* x_1092; lean_object* x_1093; uint8_t x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; +x_1092 = lean_ctor_get(x_1090, 0); +x_1093 = lean_ctor_get(x_2, 5); +lean_inc(x_1093); +x_1094 = 0; +x_1095 = l_Lean_SourceInfo_fromRef(x_1093, x_1094); +lean_dec(x_1093); +x_1096 = lean_ctor_get(x_2, 2); +lean_inc(x_1096); +x_1097 = lean_ctor_get(x_2, 1); +lean_inc(x_1097); +lean_dec(x_2); +x_1098 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194; +x_1099 = l_Lean_addMacroScope(x_1097, x_1098, x_1096); +x_1100 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192; +x_1101 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198; +lean_inc(x_1095); +x_1102 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1102, 0, x_1095); +lean_ctor_set(x_1102, 1, x_1100); +lean_ctor_set(x_1102, 2, x_1099); +lean_ctor_set(x_1102, 3, x_1101); +x_1103 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1095); +x_1104 = l_Lean_Syntax_node1(x_1095, x_1103, x_1092); +x_1105 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1106 = l_Lean_Syntax_node2(x_1095, x_1105, x_1102, x_1104); +lean_ctor_set(x_1090, 0, x_1106); +return x_1090; +} +else +{ +lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; uint8_t x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; +x_1107 = lean_ctor_get(x_1090, 0); +x_1108 = lean_ctor_get(x_1090, 1); +lean_inc(x_1108); +lean_inc(x_1107); +lean_dec(x_1090); +x_1109 = lean_ctor_get(x_2, 5); +lean_inc(x_1109); +x_1110 = 0; +x_1111 = l_Lean_SourceInfo_fromRef(x_1109, x_1110); +lean_dec(x_1109); +x_1112 = lean_ctor_get(x_2, 2); +lean_inc(x_1112); +x_1113 = lean_ctor_get(x_2, 1); +lean_inc(x_1113); +lean_dec(x_2); +x_1114 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194; +x_1115 = l_Lean_addMacroScope(x_1113, x_1114, x_1112); +x_1116 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192; +x_1117 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198; +lean_inc(x_1111); +x_1118 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1118, 0, x_1111); +lean_ctor_set(x_1118, 1, x_1116); +lean_ctor_set(x_1118, 2, x_1115); +lean_ctor_set(x_1118, 3, x_1117); +x_1119 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1111); +x_1120 = l_Lean_Syntax_node1(x_1111, x_1119, x_1107); +x_1121 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1122 = l_Lean_Syntax_node2(x_1111, x_1121, x_1118, x_1120); +x_1123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1123, 0, x_1122); +lean_ctor_set(x_1123, 1, x_1108); +return x_1123; +} +} +case 23: +{ +lean_object* x_1124; uint8_t x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; +x_1124 = lean_ctor_get(x_2, 5); +lean_inc(x_1124); +x_1125 = 0; +x_1126 = l_Lean_SourceInfo_fromRef(x_1124, x_1125); +lean_dec(x_1124); +x_1127 = lean_ctor_get(x_2, 2); +lean_inc(x_1127); +x_1128 = lean_ctor_get(x_2, 1); +lean_inc(x_1128); +lean_dec(x_2); +x_1129 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202; +x_1130 = l_Lean_addMacroScope(x_1128, x_1129, x_1127); +x_1131 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200; +x_1132 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206; +x_1133 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1133, 0, x_1126); +lean_ctor_set(x_1133, 1, x_1131); +lean_ctor_set(x_1133, 2, x_1130); +lean_ctor_set(x_1133, 3, x_1132); +x_1134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1134, 0, x_1133); +lean_ctor_set(x_1134, 1, x_3); +return x_1134; +} +case 24: +{ +uint8_t x_1135; lean_object* x_1136; uint8_t x_1137; +x_1135 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1136 = l___private_Std_Time_Notation_0__Std_Time_convertZoneName(x_1135, x_2, x_3); +x_1137 = !lean_is_exclusive(x_1136); +if (x_1137 == 0) +{ +lean_object* x_1138; lean_object* x_1139; uint8_t x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; +x_1138 = lean_ctor_get(x_1136, 0); +x_1139 = lean_ctor_get(x_2, 5); +lean_inc(x_1139); +x_1140 = 0; +x_1141 = l_Lean_SourceInfo_fromRef(x_1139, x_1140); +lean_dec(x_1139); +x_1142 = lean_ctor_get(x_2, 2); +lean_inc(x_1142); +x_1143 = lean_ctor_get(x_2, 1); +lean_inc(x_1143); +lean_dec(x_2); +x_1144 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210; +x_1145 = l_Lean_addMacroScope(x_1143, x_1144, x_1142); +x_1146 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208; +x_1147 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214; +lean_inc(x_1141); +x_1148 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1148, 0, x_1141); +lean_ctor_set(x_1148, 1, x_1146); +lean_ctor_set(x_1148, 2, x_1145); +lean_ctor_set(x_1148, 3, x_1147); +x_1149 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1141); +x_1150 = l_Lean_Syntax_node1(x_1141, x_1149, x_1138); +x_1151 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1152 = l_Lean_Syntax_node2(x_1141, x_1151, x_1148, x_1150); +lean_ctor_set(x_1136, 0, x_1152); +return x_1136; +} +else +{ +lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; uint8_t x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; +x_1153 = lean_ctor_get(x_1136, 0); +x_1154 = lean_ctor_get(x_1136, 1); +lean_inc(x_1154); +lean_inc(x_1153); +lean_dec(x_1136); +x_1155 = lean_ctor_get(x_2, 5); +lean_inc(x_1155); +x_1156 = 0; +x_1157 = l_Lean_SourceInfo_fromRef(x_1155, x_1156); +lean_dec(x_1155); +x_1158 = lean_ctor_get(x_2, 2); +lean_inc(x_1158); +x_1159 = lean_ctor_get(x_2, 1); +lean_inc(x_1159); +lean_dec(x_2); +x_1160 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210; +x_1161 = l_Lean_addMacroScope(x_1159, x_1160, x_1158); +x_1162 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208; +x_1163 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214; +lean_inc(x_1157); +x_1164 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1164, 0, x_1157); +lean_ctor_set(x_1164, 1, x_1162); +lean_ctor_set(x_1164, 2, x_1161); +lean_ctor_set(x_1164, 3, x_1163); +x_1165 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1157); +x_1166 = l_Lean_Syntax_node1(x_1157, x_1165, x_1153); +x_1167 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1168 = l_Lean_Syntax_node2(x_1157, x_1167, x_1164, x_1166); +x_1169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1169, 0, x_1168); +lean_ctor_set(x_1169, 1, x_1154); +return x_1169; +} +} +case 25: +{ +uint8_t x_1170; lean_object* x_1171; uint8_t x_1172; +x_1170 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1171 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetO(x_1170, x_2, x_3); +x_1172 = !lean_is_exclusive(x_1171); +if (x_1172 == 0) +{ +lean_object* x_1173; lean_object* x_1174; uint8_t x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; +x_1173 = lean_ctor_get(x_1171, 0); +x_1174 = lean_ctor_get(x_2, 5); +lean_inc(x_1174); +x_1175 = 0; +x_1176 = l_Lean_SourceInfo_fromRef(x_1174, x_1175); +lean_dec(x_1174); +x_1177 = lean_ctor_get(x_2, 2); +lean_inc(x_1177); +x_1178 = lean_ctor_get(x_2, 1); +lean_inc(x_1178); +lean_dec(x_2); +x_1179 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218; +x_1180 = l_Lean_addMacroScope(x_1178, x_1179, x_1177); +x_1181 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216; +x_1182 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222; +lean_inc(x_1176); +x_1183 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1183, 0, x_1176); +lean_ctor_set(x_1183, 1, x_1181); +lean_ctor_set(x_1183, 2, x_1180); +lean_ctor_set(x_1183, 3, x_1182); +x_1184 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1176); +x_1185 = l_Lean_Syntax_node1(x_1176, x_1184, x_1173); +x_1186 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1187 = l_Lean_Syntax_node2(x_1176, x_1186, x_1183, x_1185); +lean_ctor_set(x_1171, 0, x_1187); +return x_1171; +} +else +{ +lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; uint8_t x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; +x_1188 = lean_ctor_get(x_1171, 0); +x_1189 = lean_ctor_get(x_1171, 1); +lean_inc(x_1189); +lean_inc(x_1188); +lean_dec(x_1171); +x_1190 = lean_ctor_get(x_2, 5); +lean_inc(x_1190); +x_1191 = 0; +x_1192 = l_Lean_SourceInfo_fromRef(x_1190, x_1191); +lean_dec(x_1190); +x_1193 = lean_ctor_get(x_2, 2); +lean_inc(x_1193); +x_1194 = lean_ctor_get(x_2, 1); +lean_inc(x_1194); +lean_dec(x_2); +x_1195 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218; +x_1196 = l_Lean_addMacroScope(x_1194, x_1195, x_1193); +x_1197 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216; +x_1198 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222; +lean_inc(x_1192); +x_1199 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1199, 0, x_1192); +lean_ctor_set(x_1199, 1, x_1197); +lean_ctor_set(x_1199, 2, x_1196); +lean_ctor_set(x_1199, 3, x_1198); +x_1200 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1192); +x_1201 = l_Lean_Syntax_node1(x_1192, x_1200, x_1188); +x_1202 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1203 = l_Lean_Syntax_node2(x_1192, x_1202, x_1199, x_1201); +x_1204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1204, 0, x_1203); +lean_ctor_set(x_1204, 1, x_1189); +return x_1204; +} +} +case 26: +{ +uint8_t x_1205; lean_object* x_1206; uint8_t x_1207; +x_1205 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1206 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX(x_1205, x_2, x_3); +x_1207 = !lean_is_exclusive(x_1206); +if (x_1207 == 0) +{ +lean_object* x_1208; lean_object* x_1209; uint8_t x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; +x_1208 = lean_ctor_get(x_1206, 0); +x_1209 = lean_ctor_get(x_2, 5); +lean_inc(x_1209); +x_1210 = 0; +x_1211 = l_Lean_SourceInfo_fromRef(x_1209, x_1210); +lean_dec(x_1209); +x_1212 = lean_ctor_get(x_2, 2); +lean_inc(x_1212); +x_1213 = lean_ctor_get(x_2, 1); +lean_inc(x_1213); +lean_dec(x_2); +x_1214 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226; +x_1215 = l_Lean_addMacroScope(x_1213, x_1214, x_1212); +x_1216 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224; +x_1217 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230; +lean_inc(x_1211); +x_1218 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1218, 0, x_1211); +lean_ctor_set(x_1218, 1, x_1216); +lean_ctor_set(x_1218, 2, x_1215); +lean_ctor_set(x_1218, 3, x_1217); +x_1219 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1211); +x_1220 = l_Lean_Syntax_node1(x_1211, x_1219, x_1208); +x_1221 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1222 = l_Lean_Syntax_node2(x_1211, x_1221, x_1218, x_1220); +lean_ctor_set(x_1206, 0, x_1222); +return x_1206; +} +else +{ +lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; uint8_t x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; +x_1223 = lean_ctor_get(x_1206, 0); +x_1224 = lean_ctor_get(x_1206, 1); +lean_inc(x_1224); +lean_inc(x_1223); +lean_dec(x_1206); +x_1225 = lean_ctor_get(x_2, 5); +lean_inc(x_1225); +x_1226 = 0; +x_1227 = l_Lean_SourceInfo_fromRef(x_1225, x_1226); +lean_dec(x_1225); +x_1228 = lean_ctor_get(x_2, 2); +lean_inc(x_1228); +x_1229 = lean_ctor_get(x_2, 1); +lean_inc(x_1229); +lean_dec(x_2); +x_1230 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226; +x_1231 = l_Lean_addMacroScope(x_1229, x_1230, x_1228); +x_1232 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224; +x_1233 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230; +lean_inc(x_1227); +x_1234 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1234, 0, x_1227); +lean_ctor_set(x_1234, 1, x_1232); +lean_ctor_set(x_1234, 2, x_1231); +lean_ctor_set(x_1234, 3, x_1233); +x_1235 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1227); +x_1236 = l_Lean_Syntax_node1(x_1227, x_1235, x_1223); +x_1237 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1238 = l_Lean_Syntax_node2(x_1227, x_1237, x_1234, x_1236); +x_1239 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1239, 0, x_1238); +lean_ctor_set(x_1239, 1, x_1224); +return x_1239; +} +} +case 27: +{ +uint8_t x_1240; lean_object* x_1241; uint8_t x_1242; +x_1240 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1241 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetX(x_1240, x_2, x_3); +x_1242 = !lean_is_exclusive(x_1241); +if (x_1242 == 0) +{ +lean_object* x_1243; lean_object* x_1244; uint8_t x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; +x_1243 = lean_ctor_get(x_1241, 0); +x_1244 = lean_ctor_get(x_2, 5); +lean_inc(x_1244); +x_1245 = 0; +x_1246 = l_Lean_SourceInfo_fromRef(x_1244, x_1245); +lean_dec(x_1244); +x_1247 = lean_ctor_get(x_2, 2); +lean_inc(x_1247); +x_1248 = lean_ctor_get(x_2, 1); +lean_inc(x_1248); +lean_dec(x_2); +x_1249 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234; +x_1250 = l_Lean_addMacroScope(x_1248, x_1249, x_1247); +x_1251 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232; +x_1252 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238; +lean_inc(x_1246); +x_1253 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1253, 0, x_1246); +lean_ctor_set(x_1253, 1, x_1251); +lean_ctor_set(x_1253, 2, x_1250); +lean_ctor_set(x_1253, 3, x_1252); +x_1254 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1246); +x_1255 = l_Lean_Syntax_node1(x_1246, x_1254, x_1243); +x_1256 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1257 = l_Lean_Syntax_node2(x_1246, x_1256, x_1253, x_1255); +lean_ctor_set(x_1241, 0, x_1257); +return x_1241; +} +else +{ +lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; uint8_t x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; +x_1258 = lean_ctor_get(x_1241, 0); +x_1259 = lean_ctor_get(x_1241, 1); +lean_inc(x_1259); +lean_inc(x_1258); +lean_dec(x_1241); +x_1260 = lean_ctor_get(x_2, 5); +lean_inc(x_1260); +x_1261 = 0; +x_1262 = l_Lean_SourceInfo_fromRef(x_1260, x_1261); +lean_dec(x_1260); +x_1263 = lean_ctor_get(x_2, 2); +lean_inc(x_1263); +x_1264 = lean_ctor_get(x_2, 1); +lean_inc(x_1264); +lean_dec(x_2); +x_1265 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234; +x_1266 = l_Lean_addMacroScope(x_1264, x_1265, x_1263); +x_1267 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232; +x_1268 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238; +lean_inc(x_1262); +x_1269 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1269, 0, x_1262); +lean_ctor_set(x_1269, 1, x_1267); +lean_ctor_set(x_1269, 2, x_1266); +lean_ctor_set(x_1269, 3, x_1268); +x_1270 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1262); +x_1271 = l_Lean_Syntax_node1(x_1262, x_1270, x_1258); +x_1272 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1273 = l_Lean_Syntax_node2(x_1262, x_1272, x_1269, x_1271); +x_1274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1274, 0, x_1273); +lean_ctor_set(x_1274, 1, x_1259); +return x_1274; +} +} +default: +{ +uint8_t x_1275; lean_object* x_1276; uint8_t x_1277; +x_1275 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1276 = l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ(x_1275, x_2, x_3); +x_1277 = !lean_is_exclusive(x_1276); +if (x_1277 == 0) +{ +lean_object* x_1278; lean_object* x_1279; uint8_t x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; +x_1278 = lean_ctor_get(x_1276, 0); +x_1279 = lean_ctor_get(x_2, 5); +lean_inc(x_1279); +x_1280 = 0; +x_1281 = l_Lean_SourceInfo_fromRef(x_1279, x_1280); +lean_dec(x_1279); +x_1282 = lean_ctor_get(x_2, 2); +lean_inc(x_1282); +x_1283 = lean_ctor_get(x_2, 1); +lean_inc(x_1283); +lean_dec(x_2); +x_1284 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242; +x_1285 = l_Lean_addMacroScope(x_1283, x_1284, x_1282); +x_1286 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240; +x_1287 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246; +lean_inc(x_1281); +x_1288 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1288, 0, x_1281); +lean_ctor_set(x_1288, 1, x_1286); +lean_ctor_set(x_1288, 2, x_1285); +lean_ctor_set(x_1288, 3, x_1287); +x_1289 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1281); +x_1290 = l_Lean_Syntax_node1(x_1281, x_1289, x_1278); +x_1291 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1292 = l_Lean_Syntax_node2(x_1281, x_1291, x_1288, x_1290); +lean_ctor_set(x_1276, 0, x_1292); +return x_1276; +} +else +{ +lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; uint8_t x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; +x_1293 = lean_ctor_get(x_1276, 0); +x_1294 = lean_ctor_get(x_1276, 1); +lean_inc(x_1294); +lean_inc(x_1293); +lean_dec(x_1276); +x_1295 = lean_ctor_get(x_2, 5); +lean_inc(x_1295); +x_1296 = 0; +x_1297 = l_Lean_SourceInfo_fromRef(x_1295, x_1296); +lean_dec(x_1295); +x_1298 = lean_ctor_get(x_2, 2); +lean_inc(x_1298); +x_1299 = lean_ctor_get(x_2, 1); +lean_inc(x_1299); +lean_dec(x_2); +x_1300 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242; +x_1301 = l_Lean_addMacroScope(x_1299, x_1300, x_1298); +x_1302 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240; +x_1303 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246; +lean_inc(x_1297); +x_1304 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1304, 0, x_1297); +lean_ctor_set(x_1304, 1, x_1302); +lean_ctor_set(x_1304, 2, x_1301); +lean_ctor_set(x_1304, 3, x_1303); +x_1305 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1297); +x_1306 = l_Lean_Syntax_node1(x_1297, x_1305, x_1293); +x_1307 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_1308 = l_Lean_Syntax_node2(x_1297, x_1307, x_1304, x_1306); +x_1309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1309, 0, x_1308); +lean_ctor_set(x_1309, 1, x_1294); +return x_1309; +} +} +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("string", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("modifier", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertFormatPart(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_ctor_get(x_2, 5); +lean_inc(x_5); +x_6 = 0; +x_7 = l_Lean_SourceInfo_fromRef(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec(x_2); +x_10 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_7); +x_11 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3; +x_13 = l_Lean_addMacroScope(x_9, x_12, x_8); +x_14 = lean_box(0); +x_15 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2; +lean_inc(x_7); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_14); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_7); +x_18 = l_Lean_Syntax_node2(x_7, x_17, x_11, x_16); +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkStrLit(x_4, x_19); +lean_dec(x_4); +x_21 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_7); +x_22 = l_Lean_Syntax_node1(x_7, x_21, x_20); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_24 = l_Lean_Syntax_node2(x_7, x_23, x_18, x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +lean_inc(x_2); +x_27 = l___private_Std_Time_Notation_0__Std_Time_convertModifier(x_26, x_2, x_3); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_2, 5); +lean_inc(x_30); +x_31 = 0; +x_32 = l_Lean_SourceInfo_fromRef(x_30, x_31); +lean_dec(x_30); +x_33 = lean_ctor_get(x_2, 2); +lean_inc(x_33); +x_34 = lean_ctor_get(x_2, 1); +lean_inc(x_34); +lean_dec(x_2); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_32); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6; +x_38 = l_Lean_addMacroScope(x_34, x_37, x_33); +x_39 = lean_box(0); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5; +lean_inc(x_32); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_39); +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_32); +x_43 = l_Lean_Syntax_node2(x_32, x_42, x_36, x_41); +x_44 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_32); +x_45 = l_Lean_Syntax_node1(x_32, x_44, x_29); +x_46 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_47 = l_Lean_Syntax_node2(x_32, x_46, x_43, x_45); +lean_ctor_set(x_27, 0, x_47); +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +x_50 = lean_ctor_get(x_2, 5); +lean_inc(x_50); +x_51 = 0; +x_52 = l_Lean_SourceInfo_fromRef(x_50, x_51); +lean_dec(x_50); +x_53 = lean_ctor_get(x_2, 2); +lean_inc(x_53); +x_54 = lean_ctor_get(x_2, 1); +lean_inc(x_54); +lean_dec(x_2); +x_55 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47; +lean_inc(x_52); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_52); +lean_ctor_set(x_56, 1, x_55); +x_57 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6; +x_58 = l_Lean_addMacroScope(x_54, x_57, x_53); +x_59 = lean_box(0); +x_60 = l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5; +lean_inc(x_52); +x_61 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_61, 0, x_52); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_61, 3, x_59); +x_62 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46; +lean_inc(x_52); +x_63 = l_Lean_Syntax_node2(x_52, x_62, x_56, x_61); +x_64 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_52); +x_65 = l_Lean_Syntax_node1(x_52, x_64, x_48); +x_66 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_67 = l_Lean_Syntax_node2(x_52, x_66, x_63, x_65); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_49); +return x_68; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("num", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_4 = lean_ctor_get(x_2, 5); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +x_7 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +lean_inc(x_6); +x_8 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_7); +x_9 = l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2; +x_10 = l_Lean_Syntax_node1(x_6, x_9, x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxNat___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxNat(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("str", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_2, 5); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_inc(x_6); +x_7 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_1); +x_8 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +x_9 = l_Lean_Syntax_node1(x_6, x_8, x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_3); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxString___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxString(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Int.ofNat", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Int", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofNat", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Int.negSucc", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("negSucc", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1; +x_5 = lean_int_dec_lt(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_6 = lean_nat_abs(x_1); +x_7 = lean_ctor_get(x_2, 5); +lean_inc(x_7); +x_8 = 0; +x_9 = l_Lean_SourceInfo_fromRef(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6; +x_13 = l_Lean_addMacroScope(x_11, x_12, x_10); +x_14 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3; +x_15 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10; +lean_inc(x_9); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_15); +x_17 = l___private_Init_Data_Repr_0__Nat_reprFast(x_6); +x_18 = lean_box(2); +x_19 = l_Lean_Syntax_mkNumLit(x_17, x_18); +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_9); +x_21 = l_Lean_Syntax_node1(x_9, x_20, x_19); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_23 = l_Lean_Syntax_node2(x_9, x_22, x_16, x_21); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_3); +return x_24; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_25 = lean_nat_abs(x_1); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_sub(x_25, x_26); +lean_dec(x_25); +x_28 = lean_ctor_get(x_2, 5); +lean_inc(x_28); +x_29 = 0; +x_30 = l_Lean_SourceInfo_fromRef(x_28, x_29); +lean_dec(x_28); +x_31 = lean_ctor_get(x_2, 2); +lean_inc(x_31); +x_32 = lean_ctor_get(x_2, 1); +lean_inc(x_32); +lean_dec(x_2); +x_33 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14; +x_34 = l_Lean_addMacroScope(x_32, x_33, x_31); +x_35 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12; +x_36 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18; +lean_inc(x_30); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +x_38 = l___private_Init_Data_Repr_0__Nat_reprFast(x_27); +x_39 = lean_box(2); +x_40 = l_Lean_Syntax_mkNumLit(x_38, x_39); +x_41 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_30); +x_42 = l_Lean_Syntax_node1(x_30, x_41, x_40); +x_43 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_44 = l_Lean_Syntax_node2(x_30, x_43, x_37, x_42); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_3); +return x_45; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxInt___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Internal.Bounded.LE.ofNatWrapping", 42, 42); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Internal", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Bounded", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("LE", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofNatWrapping", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4; +x_5 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5; +x_6 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6; +x_7 = l_Lean_Name_mkStr6(x_1, x_2, x_3, x_4, x_5, x_6); +return x_7; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("byTactic", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("by", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Tactic", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("tacticSeq1Indented", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("decide", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("optConfig", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +lean_inc(x_2); +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt(x_1, x_2, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_2, 5); +lean_inc(x_7); +x_8 = 0; +x_9 = l_Lean_SourceInfo_fromRef(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7; +x_13 = l_Lean_addMacroScope(x_11, x_12, x_10); +x_14 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2; +x_15 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11; +lean_inc(x_9); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_15); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_9); +x_18 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_18, 0, x_9); +lean_ctor_set(x_18, 1, x_17); +x_19 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14; +lean_inc(x_9); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_9); +lean_ctor_set(x_20, 1, x_19); +x_21 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20; +lean_inc(x_9); +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_9); +lean_ctor_set(x_22, 1, x_21); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +x_24 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24; +lean_inc(x_9); +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_9); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_24); +x_26 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23; +lean_inc(x_9); +x_27 = l_Lean_Syntax_node1(x_9, x_26, x_25); +x_28 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21; +lean_inc(x_9); +x_29 = l_Lean_Syntax_node2(x_9, x_28, x_22, x_27); +lean_inc(x_9); +x_30 = l_Lean_Syntax_node1(x_9, x_23, x_29); +x_31 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19; +lean_inc(x_9); +x_32 = l_Lean_Syntax_node1(x_9, x_31, x_30); +x_33 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17; +lean_inc(x_9); +x_34 = l_Lean_Syntax_node1(x_9, x_33, x_32); +x_35 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13; +lean_inc(x_9); +x_36 = l_Lean_Syntax_node2(x_9, x_35, x_20, x_34); +x_37 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_9); +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_9); +lean_ctor_set(x_38, 1, x_37); +x_39 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_9); +x_40 = l_Lean_Syntax_node3(x_9, x_39, x_18, x_36, x_38); +lean_inc(x_9); +x_41 = l_Lean_Syntax_node2(x_9, x_23, x_6, x_40); +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_43 = l_Lean_Syntax_node2(x_9, x_42, x_16, x_41); +lean_ctor_set(x_4, 0, x_43); +return x_4; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_44 = lean_ctor_get(x_4, 0); +x_45 = lean_ctor_get(x_4, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_4); +x_46 = lean_ctor_get(x_2, 5); +lean_inc(x_46); +x_47 = 0; +x_48 = l_Lean_SourceInfo_fromRef(x_46, x_47); +lean_dec(x_46); +x_49 = lean_ctor_get(x_2, 2); +lean_inc(x_49); +x_50 = lean_ctor_get(x_2, 1); +lean_inc(x_50); +lean_dec(x_2); +x_51 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7; +x_52 = l_Lean_addMacroScope(x_50, x_51, x_49); +x_53 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2; +x_54 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11; +lean_inc(x_48); +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_48); +lean_ctor_set(x_55, 1, x_53); +lean_ctor_set(x_55, 2, x_52); +lean_ctor_set(x_55, 3, x_54); +x_56 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_48); +x_57 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_57, 0, x_48); +lean_ctor_set(x_57, 1, x_56); +x_58 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14; +lean_inc(x_48); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_48); +lean_ctor_set(x_59, 1, x_58); +x_60 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20; +lean_inc(x_48); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_48); +lean_ctor_set(x_61, 1, x_60); +x_62 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +x_63 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24; +lean_inc(x_48); +x_64 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_64, 0, x_48); +lean_ctor_set(x_64, 1, x_62); +lean_ctor_set(x_64, 2, x_63); +x_65 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23; +lean_inc(x_48); +x_66 = l_Lean_Syntax_node1(x_48, x_65, x_64); +x_67 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21; +lean_inc(x_48); +x_68 = l_Lean_Syntax_node2(x_48, x_67, x_61, x_66); +lean_inc(x_48); +x_69 = l_Lean_Syntax_node1(x_48, x_62, x_68); +x_70 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19; +lean_inc(x_48); +x_71 = l_Lean_Syntax_node1(x_48, x_70, x_69); +x_72 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17; +lean_inc(x_48); +x_73 = l_Lean_Syntax_node1(x_48, x_72, x_71); +x_74 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13; +lean_inc(x_48); +x_75 = l_Lean_Syntax_node2(x_48, x_74, x_59, x_73); +x_76 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_48); +x_77 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_77, 0, x_48); +lean_ctor_set(x_77, 1, x_76); +x_78 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_48); +x_79 = l_Lean_Syntax_node3(x_48, x_78, x_57, x_75, x_77); +lean_inc(x_48); +x_80 = l_Lean_Syntax_node2(x_48, x_62, x_44, x_79); +x_81 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_82 = l_Lean_Syntax_node2(x_48, x_81, x_55, x_80); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_45); +return x_83; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Internal.UnitVal.ofInt", 31, 31); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UnitVal", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofInt", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3; +x_5 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +lean_inc(x_2); +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt(x_1, x_2, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_2, 5); +lean_inc(x_7); +x_8 = 0; +x_9 = l_Lean_SourceInfo_fromRef(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5; +x_13 = l_Lean_addMacroScope(x_11, x_12, x_10); +x_14 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2; +x_15 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9; +lean_inc(x_9); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_15); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_9); +x_18 = l_Lean_Syntax_node1(x_9, x_17, x_6); +x_19 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_20 = l_Lean_Syntax_node2(x_9, x_19, x_16, x_18); +lean_ctor_set(x_4, 0, x_20); +return x_4; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +x_23 = lean_ctor_get(x_2, 5); +lean_inc(x_23); +x_24 = 0; +x_25 = l_Lean_SourceInfo_fromRef(x_23, x_24); +lean_dec(x_23); +x_26 = lean_ctor_get(x_2, 2); +lean_inc(x_26); +x_27 = lean_ctor_get(x_2, 1); +lean_inc(x_27); +lean_dec(x_2); +x_28 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5; +x_29 = l_Lean_addMacroScope(x_27, x_28, x_26); +x_30 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2; +x_31 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9; +lean_inc(x_25); +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_25); +lean_ctor_set(x_32, 1, x_30); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_31); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_25); +x_34 = l_Lean_Syntax_node1(x_25, x_33, x_21); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_36 = l_Lean_Syntax_node2(x_25, x_35, x_32, x_34); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_22); +return x_37; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_syntaxVal___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.Offset.ofSeconds", 34, 34); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("TimeZone", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Offset", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofSeconds", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4; +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +lean_inc(x_2); +x_4 = l___private_Std_Time_Notation_0__Std_Time_syntaxVal(x_1, x_2, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_2, 5); +lean_inc(x_7); +x_8 = 0; +x_9 = l_Lean_SourceInfo_fromRef(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6; +x_13 = l_Lean_addMacroScope(x_11, x_12, x_10); +x_14 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2; +x_15 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10; +lean_inc(x_9); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_9); +lean_ctor_set(x_16, 1, x_14); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_15); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_9); +x_18 = l_Lean_Syntax_node1(x_9, x_17, x_6); +x_19 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_20 = l_Lean_Syntax_node2(x_9, x_19, x_16, x_18); +lean_ctor_set(x_4, 0, x_20); +return x_4; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_21 = lean_ctor_get(x_4, 0); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_4); +x_23 = lean_ctor_get(x_2, 5); +lean_inc(x_23); +x_24 = 0; +x_25 = l_Lean_SourceInfo_fromRef(x_23, x_24); +lean_dec(x_23); +x_26 = lean_ctor_get(x_2, 2); +lean_inc(x_26); +x_27 = lean_ctor_get(x_2, 1); +lean_inc(x_27); +lean_dec(x_2); +x_28 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6; +x_29 = l_Lean_addMacroScope(x_27, x_28, x_26); +x_30 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2; +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10; +lean_inc(x_25); +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_25); +lean_ctor_set(x_32, 1, x_30); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_31); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_25); +x_34 = l_Lean_Syntax_node1(x_25, x_33, x_21); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_36 = l_Lean_Syntax_node2(x_25, x_35, x_32, x_34); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_22); +return x_37; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertOffset___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertOffset(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.mk", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Bool", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertOffset(x_4, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3; +lean_inc(x_11); +lean_inc(x_12); +x_14 = l_Lean_addMacroScope(x_12, x_13, x_11); +x_15 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2; +x_16 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7; +lean_inc(x_10); +x_17 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_15); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_16); +x_18 = lean_ctor_get(x_1, 1); +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkStrLit(x_18, x_19); +x_21 = lean_ctor_get(x_1, 2); +x_22 = l_Lean_Syntax_mkStrLit(x_21, x_19); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10; +x_24 = l_Lean_addMacroScope(x_12, x_23, x_11); +x_25 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9; +x_26 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14; +lean_inc(x_10); +x_27 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_26); +x_28 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_10); +x_29 = l_Lean_Syntax_node4(x_10, x_28, x_7, x_20, x_22, x_27); +x_30 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_31 = l_Lean_Syntax_node2(x_10, x_30, x_17, x_29); +lean_ctor_set(x_5, 0, x_31); +return x_5; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = lean_ctor_get(x_2, 5); +lean_inc(x_34); +x_35 = 0; +x_36 = l_Lean_SourceInfo_fromRef(x_34, x_35); +lean_dec(x_34); +x_37 = lean_ctor_get(x_2, 2); +lean_inc(x_37); +x_38 = lean_ctor_get(x_2, 1); +lean_inc(x_38); +lean_dec(x_2); +x_39 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3; +lean_inc(x_37); +lean_inc(x_38); +x_40 = l_Lean_addMacroScope(x_38, x_39, x_37); +x_41 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2; +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7; +lean_inc(x_36); +x_43 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_43, 0, x_36); +lean_ctor_set(x_43, 1, x_41); +lean_ctor_set(x_43, 2, x_40); +lean_ctor_set(x_43, 3, x_42); +x_44 = lean_ctor_get(x_1, 1); +x_45 = lean_box(2); +x_46 = l_Lean_Syntax_mkStrLit(x_44, x_45); +x_47 = lean_ctor_get(x_1, 2); +x_48 = l_Lean_Syntax_mkStrLit(x_47, x_45); +x_49 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10; +x_50 = l_Lean_addMacroScope(x_38, x_49, x_37); +x_51 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9; +x_52 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14; +lean_inc(x_36); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_36); +lean_ctor_set(x_53, 1, x_51); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_52); +x_54 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_36); +x_55 = l_Lean_Syntax_node4(x_36, x_54, x_32, x_46, x_48, x_53); +x_56 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_57 = l_Lean_Syntax_node2(x_36, x_56, x_43, x_55); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_33); +return x_58; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertTimezone___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.PlainDate.ofYearMonthDayClip", 37, 37); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PlainDate", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofYearMonthDayClip", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_0__Std_Time_syntaxInt(x_4, x_2, x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_8, x_2, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_1, 2); +lean_inc(x_2); +x_13 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_12, x_2, x_11); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +lean_dec(x_2); +x_21 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5; +x_22 = l_Lean_addMacroScope(x_20, x_21, x_19); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2; +x_24 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9; +lean_inc(x_18); +x_25 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_24); +x_26 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_18); +x_27 = l_Lean_Syntax_node3(x_18, x_26, x_6, x_10, x_15); +x_28 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_29 = l_Lean_Syntax_node2(x_18, x_28, x_25, x_27); +lean_ctor_set(x_13, 0, x_29); +return x_13; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_30 = lean_ctor_get(x_13, 0); +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_13); +x_32 = lean_ctor_get(x_2, 5); +lean_inc(x_32); +x_33 = 0; +x_34 = l_Lean_SourceInfo_fromRef(x_32, x_33); +lean_dec(x_32); +x_35 = lean_ctor_get(x_2, 2); +lean_inc(x_35); +x_36 = lean_ctor_get(x_2, 1); +lean_inc(x_36); +lean_dec(x_2); +x_37 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5; +x_38 = l_Lean_addMacroScope(x_36, x_37, x_35); +x_39 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2; +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9; +lean_inc(x_34); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_34); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_40); +x_42 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_34); +x_43 = l_Lean_Syntax_node3(x_34, x_42, x_6, x_10, x_30); +x_44 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_45 = l_Lean_Syntax_node2(x_34, x_44, x_41, x_43); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_31); +return x_46; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.PlainTime.mk", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PlainTime", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("anonymousCtor", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟨", 3, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟩", 3, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_4, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_9, x_2, x_8); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_ctor_get(x_1, 2); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_2); +x_16 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_15, x_2, x_13); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +x_20 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_21 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_20, x_2, x_19); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_23 = lean_ctor_get(x_21, 0); +x_24 = lean_ctor_get(x_2, 5); +lean_inc(x_24); +x_25 = 0; +x_26 = l_Lean_SourceInfo_fromRef(x_24, x_25); +lean_dec(x_24); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 1); +lean_inc(x_28); +lean_dec(x_2); +x_29 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +lean_inc(x_27); +lean_inc(x_28); +x_30 = l_Lean_addMacroScope(x_28, x_29, x_27); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +x_32 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +lean_inc(x_26); +x_33 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_31); +lean_ctor_set(x_33, 2, x_30); +lean_ctor_set(x_33, 3, x_32); +x_34 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +lean_inc(x_26); +lean_ctor_set_tag(x_16, 2); +lean_ctor_set(x_16, 1, x_34); +lean_ctor_set(x_16, 0, x_26); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_27); +x_37 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +x_38 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +lean_inc(x_26); +x_39 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_39, 0, x_26); +lean_ctor_set(x_39, 1, x_37); +lean_ctor_set(x_39, 2, x_36); +lean_ctor_set(x_39, 3, x_38); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +lean_inc(x_26); +lean_ctor_set_tag(x_10, 2); +lean_ctor_set(x_10, 1, x_40); +lean_ctor_set(x_10, 0, x_26); +x_41 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_26); +x_42 = l_Lean_Syntax_node3(x_26, x_41, x_39, x_10, x_18); +x_43 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +lean_inc(x_26); +lean_ctor_set_tag(x_5, 2); +lean_ctor_set(x_5, 1, x_43); +lean_ctor_set(x_5, 0, x_26); +x_44 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +lean_inc(x_26); +x_45 = l_Lean_Syntax_node3(x_26, x_44, x_16, x_42, x_5); +lean_inc(x_26); +x_46 = l_Lean_Syntax_node4(x_26, x_41, x_7, x_12, x_45, x_23); +x_47 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_48 = l_Lean_Syntax_node2(x_26, x_47, x_33, x_46); +lean_ctor_set(x_21, 0, x_48); +return x_21; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_49 = lean_ctor_get(x_21, 0); +x_50 = lean_ctor_get(x_21, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_21); +x_51 = lean_ctor_get(x_2, 5); +lean_inc(x_51); +x_52 = 0; +x_53 = l_Lean_SourceInfo_fromRef(x_51, x_52); +lean_dec(x_51); +x_54 = lean_ctor_get(x_2, 2); +lean_inc(x_54); +x_55 = lean_ctor_get(x_2, 1); +lean_inc(x_55); +lean_dec(x_2); +x_56 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +lean_inc(x_54); +lean_inc(x_55); +x_57 = l_Lean_addMacroScope(x_55, x_56, x_54); +x_58 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +x_59 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +lean_inc(x_53); +x_60 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_60, 0, x_53); +lean_ctor_set(x_60, 1, x_58); +lean_ctor_set(x_60, 2, x_57); +lean_ctor_set(x_60, 3, x_59); +x_61 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +lean_inc(x_53); +lean_ctor_set_tag(x_16, 2); +lean_ctor_set(x_16, 1, x_61); +lean_ctor_set(x_16, 0, x_53); +x_62 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +x_63 = l_Lean_addMacroScope(x_55, x_62, x_54); +x_64 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +x_65 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +lean_inc(x_53); +x_66 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_66, 0, x_53); +lean_ctor_set(x_66, 1, x_64); +lean_ctor_set(x_66, 2, x_63); +lean_ctor_set(x_66, 3, x_65); +x_67 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +lean_inc(x_53); +lean_ctor_set_tag(x_10, 2); +lean_ctor_set(x_10, 1, x_67); +lean_ctor_set(x_10, 0, x_53); +x_68 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_53); +x_69 = l_Lean_Syntax_node3(x_53, x_68, x_66, x_10, x_18); +x_70 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +lean_inc(x_53); +lean_ctor_set_tag(x_5, 2); +lean_ctor_set(x_5, 1, x_70); +lean_ctor_set(x_5, 0, x_53); +x_71 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +lean_inc(x_53); +x_72 = l_Lean_Syntax_node3(x_53, x_71, x_16, x_69, x_5); +lean_inc(x_53); +x_73 = l_Lean_Syntax_node4(x_53, x_68, x_7, x_12, x_72, x_49); +x_74 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_75 = l_Lean_Syntax_node2(x_53, x_74, x_60, x_73); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_50); +return x_76; +} +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_77 = lean_ctor_get(x_16, 0); +x_78 = lean_ctor_get(x_16, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_16); +x_79 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_80 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_79, x_2, x_78); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; +} else { + lean_dec_ref(x_80); + x_83 = lean_box(0); +} +x_84 = lean_ctor_get(x_2, 5); +lean_inc(x_84); +x_85 = 0; +x_86 = l_Lean_SourceInfo_fromRef(x_84, x_85); +lean_dec(x_84); +x_87 = lean_ctor_get(x_2, 2); +lean_inc(x_87); +x_88 = lean_ctor_get(x_2, 1); +lean_inc(x_88); +lean_dec(x_2); +x_89 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +lean_inc(x_87); +lean_inc(x_88); +x_90 = l_Lean_addMacroScope(x_88, x_89, x_87); +x_91 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +x_92 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +lean_inc(x_86); +x_93 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_93, 0, x_86); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_93, 2, x_90); +lean_ctor_set(x_93, 3, x_92); +x_94 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +lean_inc(x_86); +x_95 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_95, 0, x_86); +lean_ctor_set(x_95, 1, x_94); +x_96 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +x_97 = l_Lean_addMacroScope(x_88, x_96, x_87); +x_98 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +x_99 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +lean_inc(x_86); +x_100 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_100, 0, x_86); +lean_ctor_set(x_100, 1, x_98); +lean_ctor_set(x_100, 2, x_97); +lean_ctor_set(x_100, 3, x_99); +x_101 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +lean_inc(x_86); +lean_ctor_set_tag(x_10, 2); +lean_ctor_set(x_10, 1, x_101); +lean_ctor_set(x_10, 0, x_86); +x_102 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_86); +x_103 = l_Lean_Syntax_node3(x_86, x_102, x_100, x_10, x_77); +x_104 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +lean_inc(x_86); +lean_ctor_set_tag(x_5, 2); +lean_ctor_set(x_5, 1, x_104); +lean_ctor_set(x_5, 0, x_86); +x_105 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +lean_inc(x_86); +x_106 = l_Lean_Syntax_node3(x_86, x_105, x_95, x_103, x_5); +lean_inc(x_86); +x_107 = l_Lean_Syntax_node4(x_86, x_102, x_7, x_12, x_106, x_81); +x_108 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_109 = l_Lean_Syntax_node2(x_86, x_108, x_93, x_107); +if (lean_is_scalar(x_83)) { + x_110 = lean_alloc_ctor(0, 2, 0); +} else { + x_110 = x_83; +} +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_82); +return x_110; +} +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_111 = lean_ctor_get(x_10, 0); +x_112 = lean_ctor_get(x_10, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_10); +x_113 = lean_ctor_get(x_1, 2); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_2); +x_115 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_114, x_2, x_112); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + x_118 = x_115; +} else { + lean_dec_ref(x_115); + x_118 = lean_box(0); +} +x_119 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_120 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_119, x_2, x_117); +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_123 = x_120; +} else { + lean_dec_ref(x_120); + x_123 = lean_box(0); +} +x_124 = lean_ctor_get(x_2, 5); +lean_inc(x_124); +x_125 = 0; +x_126 = l_Lean_SourceInfo_fromRef(x_124, x_125); +lean_dec(x_124); +x_127 = lean_ctor_get(x_2, 2); +lean_inc(x_127); +x_128 = lean_ctor_get(x_2, 1); +lean_inc(x_128); +lean_dec(x_2); +x_129 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +lean_inc(x_127); +lean_inc(x_128); +x_130 = l_Lean_addMacroScope(x_128, x_129, x_127); +x_131 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +x_132 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +lean_inc(x_126); +x_133 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_133, 0, x_126); +lean_ctor_set(x_133, 1, x_131); +lean_ctor_set(x_133, 2, x_130); +lean_ctor_set(x_133, 3, x_132); +x_134 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +lean_inc(x_126); +if (lean_is_scalar(x_118)) { + x_135 = lean_alloc_ctor(2, 2, 0); +} else { + x_135 = x_118; + lean_ctor_set_tag(x_135, 2); +} +lean_ctor_set(x_135, 0, x_126); +lean_ctor_set(x_135, 1, x_134); +x_136 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +x_137 = l_Lean_addMacroScope(x_128, x_136, x_127); +x_138 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +x_139 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +lean_inc(x_126); +x_140 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_140, 0, x_126); +lean_ctor_set(x_140, 1, x_138); +lean_ctor_set(x_140, 2, x_137); +lean_ctor_set(x_140, 3, x_139); +x_141 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +lean_inc(x_126); +x_142 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_142, 0, x_126); +lean_ctor_set(x_142, 1, x_141); +x_143 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_126); +x_144 = l_Lean_Syntax_node3(x_126, x_143, x_140, x_142, x_116); +x_145 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +lean_inc(x_126); +lean_ctor_set_tag(x_5, 2); +lean_ctor_set(x_5, 1, x_145); +lean_ctor_set(x_5, 0, x_126); +x_146 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +lean_inc(x_126); +x_147 = l_Lean_Syntax_node3(x_126, x_146, x_135, x_144, x_5); +lean_inc(x_126); +x_148 = l_Lean_Syntax_node4(x_126, x_143, x_7, x_111, x_147, x_121); +x_149 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_150 = l_Lean_Syntax_node2(x_126, x_149, x_133, x_148); +if (lean_is_scalar(x_123)) { + x_151 = lean_alloc_ctor(0, 2, 0); +} else { + x_151 = x_123; +} +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_122); +return x_151; +} +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; uint8_t x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +x_152 = lean_ctor_get(x_5, 0); +x_153 = lean_ctor_get(x_5, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_5); +x_154 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_155 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_154, x_2, x_153); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_158 = x_155; +} else { + lean_dec_ref(x_155); + x_158 = lean_box(0); +} +x_159 = lean_ctor_get(x_1, 2); +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_2); +x_161 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_160, x_2, x_157); +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_161, 1); +lean_inc(x_163); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_164 = x_161; +} else { + lean_dec_ref(x_161); + x_164 = lean_box(0); +} +x_165 = lean_ctor_get(x_1, 3); +lean_inc(x_2); +x_166 = l___private_Std_Time_Notation_0__Std_Time_syntaxBounded(x_165, x_2, x_163); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_169 = x_166; +} else { + lean_dec_ref(x_166); + x_169 = lean_box(0); +} +x_170 = lean_ctor_get(x_2, 5); +lean_inc(x_170); +x_171 = 0; +x_172 = l_Lean_SourceInfo_fromRef(x_170, x_171); +lean_dec(x_170); +x_173 = lean_ctor_get(x_2, 2); +lean_inc(x_173); +x_174 = lean_ctor_get(x_2, 1); +lean_inc(x_174); +lean_dec(x_2); +x_175 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4; +lean_inc(x_173); +lean_inc(x_174); +x_176 = l_Lean_addMacroScope(x_174, x_175, x_173); +x_177 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2; +x_178 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8; +lean_inc(x_172); +x_179 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_179, 0, x_172); +lean_ctor_set(x_179, 1, x_177); +lean_ctor_set(x_179, 2, x_176); +lean_ctor_set(x_179, 3, x_178); +x_180 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11; +lean_inc(x_172); +if (lean_is_scalar(x_164)) { + x_181 = lean_alloc_ctor(2, 2, 0); +} else { + x_181 = x_164; + lean_ctor_set_tag(x_181, 2); +} +lean_ctor_set(x_181, 0, x_172); +lean_ctor_set(x_181, 1, x_180); +x_182 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14; +x_183 = l_Lean_addMacroScope(x_174, x_182, x_173); +x_184 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13; +x_185 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17; +lean_inc(x_172); +x_186 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_186, 0, x_172); +lean_ctor_set(x_186, 1, x_184); +lean_ctor_set(x_186, 2, x_183); +lean_ctor_set(x_186, 3, x_185); +x_187 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +lean_inc(x_172); +if (lean_is_scalar(x_158)) { + x_188 = lean_alloc_ctor(2, 2, 0); +} else { + x_188 = x_158; + lean_ctor_set_tag(x_188, 2); +} +lean_ctor_set(x_188, 0, x_172); +lean_ctor_set(x_188, 1, x_187); +x_189 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_172); +x_190 = l_Lean_Syntax_node3(x_172, x_189, x_186, x_188, x_162); +x_191 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19; +lean_inc(x_172); +x_192 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_192, 0, x_172); +lean_ctor_set(x_192, 1, x_191); +x_193 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10; +lean_inc(x_172); +x_194 = l_Lean_Syntax_node3(x_172, x_193, x_181, x_190, x_192); +lean_inc(x_172); +x_195 = l_Lean_Syntax_node4(x_172, x_189, x_152, x_156, x_194, x_167); +x_196 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_197 = l_Lean_Syntax_node2(x_172, x_196, x_179, x_195); +if (lean_is_scalar(x_169)) { + x_198 = lean_alloc_ctor(0, 2, 0); +} else { + x_198 = x_169; +} +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_168); +return x_198; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.PlainDateTime.mk", 25, 25); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("PlainDateTime", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate(x_4, x_2, x_3); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_9 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime(x_8, x_2, x_7); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_2, 5); +lean_inc(x_12); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +lean_dec(x_12); +x_15 = lean_ctor_get(x_2, 2); +lean_inc(x_15); +x_16 = lean_ctor_get(x_2, 1); +lean_inc(x_16); +lean_dec(x_2); +x_17 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4; +x_18 = l_Lean_addMacroScope(x_16, x_17, x_15); +x_19 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2; +x_20 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8; +lean_inc(x_14); +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_14); +x_23 = l_Lean_Syntax_node2(x_14, x_22, x_6, x_11); +x_24 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_25 = l_Lean_Syntax_node2(x_14, x_24, x_21, x_23); +lean_ctor_set(x_9, 0, x_25); +return x_9; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_26 = lean_ctor_get(x_9, 0); +x_27 = lean_ctor_get(x_9, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_9); +x_28 = lean_ctor_get(x_2, 5); +lean_inc(x_28); +x_29 = 0; +x_30 = l_Lean_SourceInfo_fromRef(x_28, x_29); +lean_dec(x_28); +x_31 = lean_ctor_get(x_2, 2); +lean_inc(x_31); +x_32 = lean_ctor_get(x_2, 1); +lean_inc(x_32); +lean_dec(x_2); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4; +x_34 = l_Lean_addMacroScope(x_32, x_33, x_31); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2; +x_36 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8; +lean_inc(x_30); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +x_38 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_30); +x_39 = l_Lean_Syntax_node2(x_30, x_38, x_6, x_26); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_41 = l_Lean_Syntax_node2(x_30, x_40, x_37, x_39); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_27); +return x_42; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZonedDateTime.ofPlainDateTime", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ZonedDateTime", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofPlainDateTime", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.ZoneRules.ofTimeZone", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ZoneRules", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ofTimeZone", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12; +x_5 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13; +x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term_<$>_", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("<$>", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Database.defaultGetZoneRules", 37, 37); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Database", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("defaultGetZoneRules", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22; +x_4 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_thunk_get_own(x_5); +lean_inc(x_3); +x_7 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(x_6, x_3, x_4); +lean_dec(x_6); +if (x_2 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_1, 3); +lean_inc(x_3); +x_11 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone(x_10, x_3, x_9); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_3, 5); +lean_inc(x_14); +x_15 = 0; +x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_3, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_3, 1); +lean_inc(x_18); +lean_dec(x_3); +x_19 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +lean_inc(x_17); +lean_inc(x_18); +x_20 = l_Lean_addMacroScope(x_18, x_19, x_17); +x_21 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_22 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_16); +x_23 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_23, 0, x_16); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_20); +lean_ctor_set(x_23, 3, x_22); +x_24 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_16); +x_25 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_25, 0, x_16); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14; +x_27 = l_Lean_addMacroScope(x_18, x_26, x_17); +x_28 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11; +x_29 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16; +lean_inc(x_16); +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_16); +lean_ctor_set(x_30, 1, x_28); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_29); +x_31 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_16); +x_32 = l_Lean_Syntax_node1(x_16, x_31, x_13); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_16); +x_34 = l_Lean_Syntax_node2(x_16, x_33, x_30, x_32); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_16); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_16); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_16); +x_38 = l_Lean_Syntax_node3(x_16, x_37, x_25, x_34, x_36); +lean_inc(x_16); +x_39 = l_Lean_Syntax_node2(x_16, x_31, x_8, x_38); +x_40 = l_Lean_Syntax_node2(x_16, x_33, x_23, x_39); +lean_ctor_set(x_11, 0, x_40); +return x_11; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_41 = lean_ctor_get(x_11, 0); +x_42 = lean_ctor_get(x_11, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_11); +x_43 = lean_ctor_get(x_3, 5); +lean_inc(x_43); +x_44 = 0; +x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); +lean_dec(x_43); +x_46 = lean_ctor_get(x_3, 2); +lean_inc(x_46); +x_47 = lean_ctor_get(x_3, 1); +lean_inc(x_47); +lean_dec(x_3); +x_48 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +lean_inc(x_46); +lean_inc(x_47); +x_49 = l_Lean_addMacroScope(x_47, x_48, x_46); +x_50 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_51 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_45); +x_52 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_52, 0, x_45); +lean_ctor_set(x_52, 1, x_50); +lean_ctor_set(x_52, 2, x_49); +lean_ctor_set(x_52, 3, x_51); +x_53 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44; +lean_inc(x_45); +x_54 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_54, 0, x_45); +lean_ctor_set(x_54, 1, x_53); +x_55 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14; +x_56 = l_Lean_addMacroScope(x_47, x_55, x_46); +x_57 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11; +x_58 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16; +lean_inc(x_45); +x_59 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_59, 0, x_45); +lean_ctor_set(x_59, 1, x_57); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_59, 3, x_58); +x_60 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_45); +x_61 = l_Lean_Syntax_node1(x_45, x_60, x_41); +x_62 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_45); +x_63 = l_Lean_Syntax_node2(x_45, x_62, x_59, x_61); +x_64 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +lean_inc(x_45); +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_45); +lean_ctor_set(x_65, 1, x_64); +x_66 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43; +lean_inc(x_45); +x_67 = l_Lean_Syntax_node3(x_45, x_66, x_54, x_63, x_65); +lean_inc(x_45); +x_68 = l_Lean_Syntax_node2(x_45, x_60, x_8, x_67); +x_69 = l_Lean_Syntax_node2(x_45, x_62, x_52, x_68); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_42); +return x_70; +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_7); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_72 = lean_ctor_get(x_7, 0); +x_73 = lean_ctor_get(x_3, 5); +lean_inc(x_73); +x_74 = 0; +x_75 = l_Lean_SourceInfo_fromRef(x_73, x_74); +lean_dec(x_73); +x_76 = lean_ctor_get(x_3, 2); +lean_inc(x_76); +x_77 = lean_ctor_get(x_3, 1); +lean_inc(x_77); +lean_dec(x_3); +x_78 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +lean_inc(x_76); +lean_inc(x_77); +x_79 = l_Lean_addMacroScope(x_77, x_78, x_76); +x_80 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_81 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_75); +x_82 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_82, 0, x_75); +lean_ctor_set(x_82, 1, x_80); +lean_ctor_set(x_82, 2, x_79); +lean_ctor_set(x_82, 3, x_81); +x_83 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_75); +x_84 = l_Lean_Syntax_node1(x_75, x_83, x_72); +x_85 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_75); +x_86 = l_Lean_Syntax_node2(x_75, x_85, x_82, x_84); +x_87 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19; +lean_inc(x_75); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_75); +lean_ctor_set(x_88, 1, x_87); +x_89 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24; +x_90 = l_Lean_addMacroScope(x_77, x_89, x_76); +x_91 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21; +x_92 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26; +lean_inc(x_75); +x_93 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_93, 0, x_75); +lean_ctor_set(x_93, 1, x_91); +lean_ctor_set(x_93, 2, x_90); +lean_ctor_set(x_93, 3, x_92); +x_94 = lean_ctor_get(x_1, 3); +x_95 = lean_ctor_get(x_94, 1); +x_96 = lean_box(2); +x_97 = l_Lean_Syntax_mkStrLit(x_95, x_96); +lean_inc(x_75); +x_98 = l_Lean_Syntax_node1(x_75, x_83, x_97); +lean_inc(x_75); +x_99 = l_Lean_Syntax_node2(x_75, x_85, x_93, x_98); +x_100 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18; +x_101 = l_Lean_Syntax_node3(x_75, x_100, x_86, x_88, x_99); +lean_ctor_set(x_7, 0, x_101); +return x_7; +} +else +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_102 = lean_ctor_get(x_7, 0); +x_103 = lean_ctor_get(x_7, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_7); +x_104 = lean_ctor_get(x_3, 5); +lean_inc(x_104); +x_105 = 0; +x_106 = l_Lean_SourceInfo_fromRef(x_104, x_105); +lean_dec(x_104); +x_107 = lean_ctor_get(x_3, 2); +lean_inc(x_107); +x_108 = lean_ctor_get(x_3, 1); +lean_inc(x_108); +lean_dec(x_3); +x_109 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +lean_inc(x_107); +lean_inc(x_108); +x_110 = l_Lean_addMacroScope(x_108, x_109, x_107); +x_111 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_112 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_106); +x_113 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_113, 0, x_106); +lean_ctor_set(x_113, 1, x_111); +lean_ctor_set(x_113, 2, x_110); +lean_ctor_set(x_113, 3, x_112); +x_114 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_106); +x_115 = l_Lean_Syntax_node1(x_106, x_114, x_102); +x_116 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +lean_inc(x_106); +x_117 = l_Lean_Syntax_node2(x_106, x_116, x_113, x_115); +x_118 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19; +lean_inc(x_106); +x_119 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_119, 0, x_106); +lean_ctor_set(x_119, 1, x_118); +x_120 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24; +x_121 = l_Lean_addMacroScope(x_108, x_120, x_107); +x_122 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21; +x_123 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26; +lean_inc(x_106); +x_124 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_124, 0, x_106); +lean_ctor_set(x_124, 1, x_122); +lean_ctor_set(x_124, 2, x_121); +lean_ctor_set(x_124, 3, x_123); +x_125 = lean_ctor_get(x_1, 3); +x_126 = lean_ctor_get(x_125, 1); +x_127 = lean_box(2); +x_128 = l_Lean_Syntax_mkStrLit(x_126, x_127); +lean_inc(x_106); +x_129 = l_Lean_Syntax_node1(x_106, x_114, x_128); +lean_inc(x_106); +x_130 = l_Lean_Syntax_node2(x_106, x_116, x_124, x_129); +x_131 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18; +x_132 = l_Lean_Syntax_node3(x_106, x_131, x_117, x_119, x_130); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_103); +return x_133; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime(x_1, x_5, x_3, x_4); +lean_dec(x_1); +return x_6; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termZoned(_)", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_termZoned_x28___x29___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("zoned(", 6, 6); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termZoned_x28___x29___closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termZoned_x28___x29___closed__6; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termZoned_x28___x29___closed__8; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termZoned_x28___x29___closed__10; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termZoned_x28___x29___closed__11; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termZoned(_,_)", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termZoned_x28___x2c___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termZoned_x28___x29___closed__8; +x_3 = l_Std_Time_termZoned_x28___x2c___x29___closed__3; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_termZoned_x28___x2c___x29___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_termZoned_x28___x2c___x29___closed__6; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termZoned_x28___x2c___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x2c___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termZoned_x28___x2c___x29___closed__8; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x2c___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termZoned_x28___x2c___x29___closed__9; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termZoned_x28___x2c___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termZoned_x28___x2c___x29___closed__10; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termDatetime(_)", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termDatetime_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("datetime(", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termDatetime_x28___x29___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termDatetime_x28___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termDatetime_x28___x29___closed__5; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termDatetime_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termDatetime_x28___x29___closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatetime_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termDatetime_x28___x29___closed__7; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termDate(_)", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termDate_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("date(", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termDate_x28___x29___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termDate_x28___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termDate_x28___x29___closed__5; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termDate_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termDate_x28___x29___closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDate_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termDate_x28___x29___closed__7; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termTime(_)", 11, 11); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termTime_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("time(", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termTime_x28___x29___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termTime_x28___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termTime_x28___x29___closed__5; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termTime_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termTime_x28___x29___closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTime_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termTime_x28___x29___closed__7; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termOffset(_)", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termOffset_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("offset(", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termOffset_x28___x29___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termOffset_x28___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termOffset_x28___x29___closed__5; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termOffset_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termOffset_x28___x29___closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termOffset_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termOffset_x28___x29___closed__7; +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termTimezone(_)", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termTimezone_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("timezone(", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termTimezone_x28___x29___closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termTimezone_x28___x29___closed__4; +x_3 = l_Std_Time_termZoned_x28___x29___closed__7; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termZoned_x28___x29___closed__4; +x_2 = l_Std_Time_termTimezone_x28___x29___closed__5; +x_3 = l_Std_Time_termZoned_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termTimezone_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termTimezone_x28___x29___closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termTimezone_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termTimezone_x28___x29___closed__7; +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("error: ", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termZoned_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_TSyntax_getString(x_9); +lean_inc(x_14); +x_15 = l_Std_Time_ZonedDateTime_fromLeanDateTimeWithZoneString(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +lean_dec(x_15); +x_16 = l_Std_Time_ZonedDateTime_fromLeanDateTimeWithIdentifierString(x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_21 = lean_string_append(x_19, x_20); +x_22 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_21, x_2, x_3); +lean_dec(x_9); +return x_22; +} +else +{ +lean_object* x_23; uint8_t x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_9); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +x_24 = 1; +x_25 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime(x_23, x_24, x_2, x_3); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +return x_25; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +else +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_14); +lean_dec(x_9); +x_30 = lean_ctor_get(x_15, 0); +lean_inc(x_30); +lean_dec(x_15); +x_31 = 0; +x_32 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime(x_30, x_31, x_2, x_3); +lean_dec(x_30); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +return x_32; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_32, 0); +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_32); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x2c___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termZoned_x28___x2c___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_unsigned_to_nat(3u); +x_15 = l_Lean_Syntax_getArg(x_1, x_14); +lean_dec(x_1); +x_16 = l_Lean_TSyntax_getString(x_9); +x_17 = l_Std_Time_PlainDateTime_fromLeanDateTimeString(x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_15); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_22, x_2, x_3); +lean_dec(x_9); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_9); +x_24 = lean_ctor_get(x_17, 0); +lean_inc(x_24); +lean_dec(x_17); +lean_inc(x_2); +x_25 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(x_24, x_2, x_3); +lean_dec(x_24); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_2, 5); +lean_inc(x_28); +x_29 = 0; +x_30 = l_Lean_SourceInfo_fromRef(x_28, x_29); +lean_dec(x_28); +x_31 = lean_ctor_get(x_2, 2); +lean_inc(x_31); +x_32 = lean_ctor_get(x_2, 1); +lean_inc(x_32); +lean_dec(x_2); +x_33 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +x_34 = l_Lean_addMacroScope(x_32, x_33, x_31); +x_35 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_36 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_30); +x_37 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_35); +lean_ctor_set(x_37, 2, x_34); +lean_ctor_set(x_37, 3, x_36); +x_38 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_30); +x_39 = l_Lean_Syntax_node2(x_30, x_38, x_27, x_15); +x_40 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_41 = l_Lean_Syntax_node2(x_30, x_40, x_37, x_39); +lean_ctor_set(x_25, 0, x_41); +return x_25; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_42 = lean_ctor_get(x_25, 0); +x_43 = lean_ctor_get(x_25, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_25); +x_44 = lean_ctor_get(x_2, 5); +lean_inc(x_44); +x_45 = 0; +x_46 = l_Lean_SourceInfo_fromRef(x_44, x_45); +lean_dec(x_44); +x_47 = lean_ctor_get(x_2, 2); +lean_inc(x_47); +x_48 = lean_ctor_get(x_2, 1); +lean_inc(x_48); +lean_dec(x_2); +x_49 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5; +x_50 = l_Lean_addMacroScope(x_48, x_49, x_47); +x_51 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2; +x_52 = l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9; +lean_inc(x_46); +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_46); +lean_ctor_set(x_53, 1, x_51); +lean_ctor_set(x_53, 2, x_50); +lean_ctor_set(x_53, 3, x_52); +x_54 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16; +lean_inc(x_46); +x_55 = l_Lean_Syntax_node2(x_46, x_54, x_42, x_15); +x_56 = l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5; +x_57 = l_Lean_Syntax_node2(x_46, x_56, x_53, x_55); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_43); +return x_58; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDatetime_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termDatetime_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_TSyntax_getString(x_9); +x_15 = l_Std_Time_PlainDateTime_fromLeanDateTimeString(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_18 = lean_string_append(x_17, x_16); +lean_dec(x_16); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_20, x_2, x_3); +lean_dec(x_9); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_9); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime(x_22, x_2, x_3); +lean_dec(x_22); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +return x_23; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainDate_ofYearMonthDay_x3f), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termDate_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = l_Lean_TSyntax_getString(x_9); +x_15 = l_Std_Time_Formats_sqlDate; +x_16 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1; +x_17 = l_Std_Time_GenericFormat_parseBuilder___rarg(x_15, x_16, x_14); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_22, x_2, x_3); +lean_dec(x_9); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_9); +x_24 = lean_ctor_get(x_17, 0); +lean_inc(x_24); +lean_dec(x_17); +x_25 = l___private_Std_Time_Notation_0__Std_Time_convertPlainDate(x_24, x_2, x_3); +lean_dec(x_24); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +return x_25; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termTime_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termTime_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_TSyntax_getString(x_9); +x_15 = l_Std_Time_PlainTime_fromLeanTime24Hour(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_18 = lean_string_append(x_17, x_16); +lean_dec(x_16); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_20, x_2, x_3); +lean_dec(x_9); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_9); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertPlainTime(x_22, x_2, x_3); +lean_dec(x_22); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +return x_23; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termOffset_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termOffset_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_TSyntax_getString(x_9); +x_15 = l_Std_Time_TimeZone_Offset_fromOffset(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_18 = lean_string_append(x_17, x_16); +lean_dec(x_16); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_20, x_2, x_3); +lean_dec(x_9); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_9); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertOffset(x_22, x_2, x_3); +lean_dec(x_22); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +return x_23; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termTimezone_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termTimezone_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = l_Lean_TSyntax_getString(x_9); +x_15 = l_Std_Time_TimeZone_fromTimeZone(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1; +x_18 = lean_string_append(x_17, x_16); +lean_dec(x_16); +x_19 = l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_20, x_2, x_3); +lean_dec(x_9); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_9); +x_22 = lean_ctor_get(x_15, 0); +lean_inc(x_22); +lean_dec(x_15); +x_23 = l___private_Std_Time_Notation_0__Std_Time_convertTimezone(x_22, x_2, x_3); +lean_dec(x_22); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) +{ +return x_23; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +} +} +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Format(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Notation(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Format(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__23); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__24); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__25); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__26); +l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27 = _init_l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertText___closed__27); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertNumber___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFraction___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__23); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__24); +l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25 = _init_l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertYear___closed__25); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZoneName___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__23); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__24); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__25); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__26); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__27); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__28); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__29); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__30); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__31); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__32); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__33); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__34); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__35); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__36); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__37); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__38); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__39); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__40); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetX___closed__41); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetO___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffsetZ___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__23); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__24); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__25); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__26); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__27); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__28); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__29); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__30); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__31); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__32); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__33); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__34); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__35); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__36); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__37); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__38); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__39); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__40); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__41); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__42); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__43); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__44); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__45); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__46); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__47); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__48); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__49); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__50); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__51); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__52); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__53); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__54); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__55); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__56); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__57); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__58); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__59); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__60); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__61); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__62); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__63); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__64); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__65); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__66); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__67); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__68); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__69); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__70); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__71); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__72); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__73); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__74); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__75); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__76); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__77); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__78); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__79); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__80); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__81); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__82); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__83); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__84); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__85); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__86); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__87); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__88); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__89); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__90); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__91); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__92); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__93); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__94); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__95); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__96); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__97); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__98); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__99); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__100); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__101); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__102); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__103); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__104); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__105); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__106); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__107); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__108); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__109); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__110); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__111); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__112); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__113); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__114); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__115); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__116); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__117); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__118); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__119); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__120); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__121); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__122); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__123); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__124); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__125); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__126); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__127); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__128); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__129); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__130); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__131); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__132); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__133); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__134); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__135); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__136); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__137); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__138); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__139); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__140); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__141); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__142); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__143); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__144); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__145); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__146); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__147); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__148); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__149); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__150); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__151); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__152); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__153); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__154); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__155); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__156); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__157); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__158); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__159); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__160); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__161); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__162); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__163); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__164); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__165); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__166); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__167); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__168); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__169); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__170); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__171); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__172); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__173); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__174); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__175); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__176); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__177); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__178); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__179); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__180); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__181); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__182); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__183); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__184); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__185); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__186); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__187); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__188); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__189); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__190); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__191); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__192); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__193); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__194); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__195); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__196); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__197); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__198); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__199); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__200); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__201); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__202); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__203); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__204); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__205); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__206); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__207); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__208); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__209); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__210); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__211); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__212); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__213); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__214); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__215); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__216); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__217); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__218); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__219); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__220); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__221); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__222); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__223); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__224); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__225); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__226); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__227); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__228); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__229); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__230); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__231); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__232); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__233); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__234); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__235); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__236); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__237); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__238); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__239); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__240); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__241); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__242); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__243); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__244); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__245); +l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246 = _init_l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertModifier___closed__246); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertFormatPart___closed__6); +l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__1); +l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxNat___closed__2); +l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__1); +l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxString___closed__2); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__1); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__2); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__3); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__4); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__5); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__6); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__7); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__8); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__9); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__10); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__11); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__12); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__13); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__14); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__15); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__16); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__17); +l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxInt___closed__18); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__1); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__2); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__3); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__4); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__5); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__6); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__7); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__8); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__9); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__10); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__11); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__12); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__13); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__14); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__15); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__16); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__17); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__18); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__19); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__20); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__21); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__22); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__23); +l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxBounded___closed__24); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__1); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__2); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__3); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__4); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__5); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__6); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__7); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__8); +l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_syntaxVal___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertOffset___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertTimezone___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDate___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainTime___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertPlainDateTime___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__1); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__2); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__3); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__4); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__5); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__6); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__7); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__8); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__9); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__10); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__11); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__12); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__13); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__14); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__15); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__16); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__17); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__18); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__19); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__20); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__21); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__22); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__23); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__24); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__25); +l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26 = _init_l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_0__Std_Time_convertZonedDateTime___closed__26); +l_Std_Time_termZoned_x28___x29___closed__1 = _init_l_Std_Time_termZoned_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__1); +l_Std_Time_termZoned_x28___x29___closed__2 = _init_l_Std_Time_termZoned_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__2); +l_Std_Time_termZoned_x28___x29___closed__3 = _init_l_Std_Time_termZoned_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__3); +l_Std_Time_termZoned_x28___x29___closed__4 = _init_l_Std_Time_termZoned_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__4); +l_Std_Time_termZoned_x28___x29___closed__5 = _init_l_Std_Time_termZoned_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__5); +l_Std_Time_termZoned_x28___x29___closed__6 = _init_l_Std_Time_termZoned_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__6); +l_Std_Time_termZoned_x28___x29___closed__7 = _init_l_Std_Time_termZoned_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__7); +l_Std_Time_termZoned_x28___x29___closed__8 = _init_l_Std_Time_termZoned_x28___x29___closed__8(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__8); +l_Std_Time_termZoned_x28___x29___closed__9 = _init_l_Std_Time_termZoned_x28___x29___closed__9(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__9); +l_Std_Time_termZoned_x28___x29___closed__10 = _init_l_Std_Time_termZoned_x28___x29___closed__10(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__10); +l_Std_Time_termZoned_x28___x29___closed__11 = _init_l_Std_Time_termZoned_x28___x29___closed__11(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29___closed__11); +l_Std_Time_termZoned_x28___x29 = _init_l_Std_Time_termZoned_x28___x29(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x29); +l_Std_Time_termZoned_x28___x2c___x29___closed__1 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__1); +l_Std_Time_termZoned_x28___x2c___x29___closed__2 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__2); +l_Std_Time_termZoned_x28___x2c___x29___closed__3 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__3); +l_Std_Time_termZoned_x28___x2c___x29___closed__4 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__4); +l_Std_Time_termZoned_x28___x2c___x29___closed__5 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__5); +l_Std_Time_termZoned_x28___x2c___x29___closed__6 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__6); +l_Std_Time_termZoned_x28___x2c___x29___closed__7 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__7); +l_Std_Time_termZoned_x28___x2c___x29___closed__8 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__8(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__8); +l_Std_Time_termZoned_x28___x2c___x29___closed__9 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__9(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__9); +l_Std_Time_termZoned_x28___x2c___x29___closed__10 = _init_l_Std_Time_termZoned_x28___x2c___x29___closed__10(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29___closed__10); +l_Std_Time_termZoned_x28___x2c___x29 = _init_l_Std_Time_termZoned_x28___x2c___x29(); +lean_mark_persistent(l_Std_Time_termZoned_x28___x2c___x29); +l_Std_Time_termDatetime_x28___x29___closed__1 = _init_l_Std_Time_termDatetime_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__1); +l_Std_Time_termDatetime_x28___x29___closed__2 = _init_l_Std_Time_termDatetime_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__2); +l_Std_Time_termDatetime_x28___x29___closed__3 = _init_l_Std_Time_termDatetime_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__3); +l_Std_Time_termDatetime_x28___x29___closed__4 = _init_l_Std_Time_termDatetime_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__4); +l_Std_Time_termDatetime_x28___x29___closed__5 = _init_l_Std_Time_termDatetime_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__5); +l_Std_Time_termDatetime_x28___x29___closed__6 = _init_l_Std_Time_termDatetime_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__6); +l_Std_Time_termDatetime_x28___x29___closed__7 = _init_l_Std_Time_termDatetime_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29___closed__7); +l_Std_Time_termDatetime_x28___x29 = _init_l_Std_Time_termDatetime_x28___x29(); +lean_mark_persistent(l_Std_Time_termDatetime_x28___x29); +l_Std_Time_termDate_x28___x29___closed__1 = _init_l_Std_Time_termDate_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__1); +l_Std_Time_termDate_x28___x29___closed__2 = _init_l_Std_Time_termDate_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__2); +l_Std_Time_termDate_x28___x29___closed__3 = _init_l_Std_Time_termDate_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__3); +l_Std_Time_termDate_x28___x29___closed__4 = _init_l_Std_Time_termDate_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__4); +l_Std_Time_termDate_x28___x29___closed__5 = _init_l_Std_Time_termDate_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__5); +l_Std_Time_termDate_x28___x29___closed__6 = _init_l_Std_Time_termDate_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__6); +l_Std_Time_termDate_x28___x29___closed__7 = _init_l_Std_Time_termDate_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29___closed__7); +l_Std_Time_termDate_x28___x29 = _init_l_Std_Time_termDate_x28___x29(); +lean_mark_persistent(l_Std_Time_termDate_x28___x29); +l_Std_Time_termTime_x28___x29___closed__1 = _init_l_Std_Time_termTime_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__1); +l_Std_Time_termTime_x28___x29___closed__2 = _init_l_Std_Time_termTime_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__2); +l_Std_Time_termTime_x28___x29___closed__3 = _init_l_Std_Time_termTime_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__3); +l_Std_Time_termTime_x28___x29___closed__4 = _init_l_Std_Time_termTime_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__4); +l_Std_Time_termTime_x28___x29___closed__5 = _init_l_Std_Time_termTime_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__5); +l_Std_Time_termTime_x28___x29___closed__6 = _init_l_Std_Time_termTime_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__6); +l_Std_Time_termTime_x28___x29___closed__7 = _init_l_Std_Time_termTime_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29___closed__7); +l_Std_Time_termTime_x28___x29 = _init_l_Std_Time_termTime_x28___x29(); +lean_mark_persistent(l_Std_Time_termTime_x28___x29); +l_Std_Time_termOffset_x28___x29___closed__1 = _init_l_Std_Time_termOffset_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__1); +l_Std_Time_termOffset_x28___x29___closed__2 = _init_l_Std_Time_termOffset_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__2); +l_Std_Time_termOffset_x28___x29___closed__3 = _init_l_Std_Time_termOffset_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__3); +l_Std_Time_termOffset_x28___x29___closed__4 = _init_l_Std_Time_termOffset_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__4); +l_Std_Time_termOffset_x28___x29___closed__5 = _init_l_Std_Time_termOffset_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__5); +l_Std_Time_termOffset_x28___x29___closed__6 = _init_l_Std_Time_termOffset_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__6); +l_Std_Time_termOffset_x28___x29___closed__7 = _init_l_Std_Time_termOffset_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29___closed__7); +l_Std_Time_termOffset_x28___x29 = _init_l_Std_Time_termOffset_x28___x29(); +lean_mark_persistent(l_Std_Time_termOffset_x28___x29); +l_Std_Time_termTimezone_x28___x29___closed__1 = _init_l_Std_Time_termTimezone_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__1); +l_Std_Time_termTimezone_x28___x29___closed__2 = _init_l_Std_Time_termTimezone_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__2); +l_Std_Time_termTimezone_x28___x29___closed__3 = _init_l_Std_Time_termTimezone_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__3); +l_Std_Time_termTimezone_x28___x29___closed__4 = _init_l_Std_Time_termTimezone_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__4); +l_Std_Time_termTimezone_x28___x29___closed__5 = _init_l_Std_Time_termTimezone_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__5); +l_Std_Time_termTimezone_x28___x29___closed__6 = _init_l_Std_Time_termTimezone_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__6); +l_Std_Time_termTimezone_x28___x29___closed__7 = _init_l_Std_Time_termTimezone_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29___closed__7); +l_Std_Time_termTimezone_x28___x29 = _init_l_Std_Time_termTimezone_x28___x29(); +lean_mark_persistent(l_Std_Time_termTimezone_x28___x29); +l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1 = _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__1); +l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2 = _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termZoned_x28___x29__1___closed__2); +l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1 = _init_l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation______macroRules__Std__Time__termDate_x28___x29__1___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Notation/Spec.c b/stage0/stdlib/Std/Time/Notation/Spec.c new file mode 100644 index 000000000000..1de84564cf02 --- /dev/null +++ b/stage0/stdlib/Std/Time/Notation/Spec.c @@ -0,0 +1,10271 @@ +// Lean compiler output +// Module: Std.Time.Notation.Spec +// Imports: Std.Time.Date Std.Time.Time Std.Time.Zoned Std.Time.DateTime Std.Time.Format.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1; +static lean_object* l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2; +lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_Syntax_TSepArray_push___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22; +lean_object* l_Lean_TSyntax_getString(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__8; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ(uint8_t, lean_object*, lean_object*); +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144; +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___boxed(lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161; +lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10; +lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173; +LEAN_EXPORT lean_object* l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29; +lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2; +lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15; +lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130; +lean_object* l_Array_append___rarg(lean_object*, lean_object*); +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12; +lean_object* l_Std_Time_GenericFormat_spec___rarg(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222; +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7; +lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171; +LEAN_EXPORT lean_object* l_Std_Time_termDatespec_x28___x29; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42; +lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24; +lean_object* lean_array_mk(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5; +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109; +lean_object* l_String_toSubstring_x27(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200; +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName(uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2; +static lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11; +lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20; +static lean_object* l_Std_Time_termDatespec_x28___x29___closed__7; +static lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16; +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.short", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Time", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Text", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("short", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.full", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("full", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Text.narrow", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("narrow", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21; +x_34 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Lean", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Parser", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Term", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("app", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Number.mk", 18, 18); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Number", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("mk", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("null", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14; +lean_inc(x_6); +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_15 = lean_box(2); +x_16 = l_Lean_Syntax_mkNumLit(x_14, x_15); +x_17 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_6); +x_18 = l_Lean_Syntax_node1(x_6, x_17, x_16); +x_19 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_20 = l_Lean_Syntax_node2(x_6, x_19, x_13, x_18); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_3); +return x_21; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.nano", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Fraction", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nano", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Fraction.truncated", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("truncated", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +lean_dec(x_1); +x_16 = lean_ctor_get(x_2, 5); +lean_inc(x_16); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +lean_dec(x_16); +x_19 = lean_ctor_get(x_2, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 1); +lean_inc(x_20); +lean_dec(x_2); +x_21 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13; +x_22 = l_Lean_addMacroScope(x_20, x_21, x_19); +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11; +x_24 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17; +lean_inc(x_18); +x_25 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_24); +x_26 = l___private_Init_Data_Repr_0__Nat_reprFast(x_15); +x_27 = lean_box(2); +x_28 = l_Lean_Syntax_mkNumLit(x_26, x_27); +x_29 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_18); +x_30 = l_Lean_Syntax_node1(x_18, x_29, x_28); +x_31 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_32 = l_Lean_Syntax_node2(x_18, x_31, x_25, x_30); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_3); +return x_33; +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.twoDigit", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Year", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("twoDigit", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.fourDigit", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("fourDigit", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Year.extended", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("extended", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = lean_ctor_get(x_2, 5); +lean_inc(x_27); +x_28 = 0; +x_29 = l_Lean_SourceInfo_fromRef(x_27, x_28); +lean_dec(x_27); +x_30 = lean_ctor_get(x_2, 2); +lean_inc(x_30); +x_31 = lean_ctor_get(x_2, 1); +lean_inc(x_31); +lean_dec(x_2); +x_32 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21; +x_33 = l_Lean_addMacroScope(x_31, x_32, x_30); +x_34 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19; +x_35 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25; +lean_inc(x_29); +x_36 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_36, 0, x_29); +lean_ctor_set(x_36, 1, x_34); +lean_ctor_set(x_36, 2, x_33); +lean_ctor_set(x_36, 3, x_35); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_38 = lean_box(2); +x_39 = l_Lean_Syntax_mkNumLit(x_37, x_38); +x_40 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_29); +x_41 = l_Lean_Syntax_node1(x_29, x_40, x_39); +x_42 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_43 = l_Lean_Syntax_node2(x_29, x_42, x_36, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_3); +return x_44; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.short", 23, 23); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("ZoneName", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.ZoneName.full", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hour", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetX", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hour", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinute", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteColon", 32, 32); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteColon", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecond", 33, 33); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteSecond", 16, 16); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetX.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hourMinuteSecondColon", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +case 2: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19; +x_34 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +case 3: +{ +lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_37 = lean_ctor_get(x_2, 5); +lean_inc(x_37); +x_38 = 0; +x_39 = l_Lean_SourceInfo_fromRef(x_37, x_38); +lean_dec(x_37); +x_40 = lean_ctor_get(x_2, 2); +lean_inc(x_40); +x_41 = lean_ctor_get(x_2, 1); +lean_inc(x_41); +lean_dec(x_2); +x_42 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29; +x_43 = l_Lean_addMacroScope(x_41, x_42, x_40); +x_44 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27; +x_45 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33; +x_46 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_46, 0, x_39); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set(x_46, 2, x_43); +lean_ctor_set(x_46, 3, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_3); +return x_47; +} +default: +{ +lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_48 = lean_ctor_get(x_2, 5); +lean_inc(x_48); +x_49 = 0; +x_50 = l_Lean_SourceInfo_fromRef(x_48, x_49); +lean_dec(x_48); +x_51 = lean_ctor_get(x_2, 2); +lean_inc(x_51); +x_52 = lean_ctor_get(x_2, 1); +lean_inc(x_52); +lean_dec(x_2); +x_53 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37; +x_54 = l_Lean_addMacroScope(x_52, x_53, x_51); +x_55 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35; +x_56 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41; +x_57 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_57, 0, x_50); +lean_ctor_set(x_57, 1, x_55); +lean_ctor_set(x_57, 2, x_54); +lean_ctor_set(x_57, 3, x_56); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_3); +return x_58; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.short", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetO", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetO.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinute", 27, 27); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("OffsetZ", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.full", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.OffsetZ.hourMinuteSecondColon", 38, 38); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_4 = lean_ctor_get(x_2, 5); +lean_inc(x_4); +x_5 = 0; +x_6 = l_Lean_SourceInfo_fromRef(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4; +x_10 = l_Lean_addMacroScope(x_8, x_9, x_7); +x_11 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2; +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8; +x_13 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_13, 0, x_6); +lean_ctor_set(x_13, 1, x_11); +lean_ctor_set(x_13, 2, x_10); +lean_ctor_set(x_13, 3, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_3); +return x_14; +} +case 1: +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_15 = lean_ctor_get(x_2, 5); +lean_inc(x_15); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 2); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_dec(x_2); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11; +x_21 = l_Lean_addMacroScope(x_19, x_20, x_18); +x_22 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10; +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15; +x_24 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_24, 0, x_17); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_21); +lean_ctor_set(x_24, 3, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +default: +{ +lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_2, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +lean_dec(x_26); +x_29 = lean_ctor_get(x_2, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +x_31 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18; +x_32 = l_Lean_addMacroScope(x_30, x_31, x_29); +x_33 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17; +x_34 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22; +x_35 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_35, 0, x_28); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_32); +lean_ctor_set(x_35, 3, x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_3); +return x_36; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ(x_4, x_2, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.G", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Modifier", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("G", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.y", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("y", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.u", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("u", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.D", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("D", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.MorL", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("MorL", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("paren", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("(", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("dotIdent", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(".", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("inl", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(")", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("inr", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.d", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("d", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Qorq", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Qorq", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.w", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("w", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.W", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("W", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.E", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("E", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.eorc", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("eorc", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.F", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("F", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.a", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("a", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.h", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("h", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.K", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("K", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.k", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("k", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.H", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("H", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.m", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("m", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.s", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("s", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.S", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("S", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.A", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("A", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.n", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("n", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.N", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("N", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.V", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("V", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("z", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.O", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("O", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.X", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("X", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.x", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("x", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Modifier.Z", 19, 19); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Z", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3; +x_4 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +uint8_t x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_5 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_4, x_2, x_3); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_2, 5); +lean_inc(x_8); +x_9 = 0; +x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 2); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5; +x_14 = l_Lean_addMacroScope(x_12, x_13, x_11); +x_15 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2; +x_16 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9; +lean_inc(x_10); +x_17 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_17, 0, x_10); +lean_ctor_set(x_17, 1, x_15); +lean_ctor_set(x_17, 2, x_14); +lean_ctor_set(x_17, 3, x_16); +x_18 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_10); +x_19 = l_Lean_Syntax_node1(x_10, x_18, x_7); +x_20 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_21 = l_Lean_Syntax_node2(x_10, x_20, x_17, x_19); +lean_ctor_set(x_5, 0, x_21); +return x_5; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_5, 0); +x_23 = lean_ctor_get(x_5, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_5); +x_24 = lean_ctor_get(x_2, 5); +lean_inc(x_24); +x_25 = 0; +x_26 = l_Lean_SourceInfo_fromRef(x_24, x_25); +lean_dec(x_24); +x_27 = lean_ctor_get(x_2, 2); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 1); +lean_inc(x_28); +lean_dec(x_2); +x_29 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5; +x_30 = l_Lean_addMacroScope(x_28, x_29, x_27); +x_31 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2; +x_32 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9; +lean_inc(x_26); +x_33 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_31); +lean_ctor_set(x_33, 2, x_30); +lean_ctor_set(x_33, 3, x_32); +x_34 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_26); +x_35 = l_Lean_Syntax_node1(x_26, x_34, x_22); +x_36 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_37 = l_Lean_Syntax_node2(x_26, x_36, x_33, x_35); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_23); +return x_38; +} +} +case 1: +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +lean_dec(x_1); +lean_inc(x_2); +x_40 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear(x_39, x_2, x_3); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_42 = lean_ctor_get(x_40, 0); +x_43 = lean_ctor_get(x_2, 5); +lean_inc(x_43); +x_44 = 0; +x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); +lean_dec(x_43); +x_46 = lean_ctor_get(x_2, 2); +lean_inc(x_46); +x_47 = lean_ctor_get(x_2, 1); +lean_inc(x_47); +lean_dec(x_2); +x_48 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13; +x_49 = l_Lean_addMacroScope(x_47, x_48, x_46); +x_50 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11; +x_51 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17; +lean_inc(x_45); +x_52 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_52, 0, x_45); +lean_ctor_set(x_52, 1, x_50); +lean_ctor_set(x_52, 2, x_49); +lean_ctor_set(x_52, 3, x_51); +x_53 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_45); +x_54 = l_Lean_Syntax_node1(x_45, x_53, x_42); +x_55 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_56 = l_Lean_Syntax_node2(x_45, x_55, x_52, x_54); +lean_ctor_set(x_40, 0, x_56); +return x_40; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_57 = lean_ctor_get(x_40, 0); +x_58 = lean_ctor_get(x_40, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_40); +x_59 = lean_ctor_get(x_2, 5); +lean_inc(x_59); +x_60 = 0; +x_61 = l_Lean_SourceInfo_fromRef(x_59, x_60); +lean_dec(x_59); +x_62 = lean_ctor_get(x_2, 2); +lean_inc(x_62); +x_63 = lean_ctor_get(x_2, 1); +lean_inc(x_63); +lean_dec(x_2); +x_64 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13; +x_65 = l_Lean_addMacroScope(x_63, x_64, x_62); +x_66 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11; +x_67 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17; +lean_inc(x_61); +x_68 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_68, 0, x_61); +lean_ctor_set(x_68, 1, x_66); +lean_ctor_set(x_68, 2, x_65); +lean_ctor_set(x_68, 3, x_67); +x_69 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_61); +x_70 = l_Lean_Syntax_node1(x_61, x_69, x_57); +x_71 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_72 = l_Lean_Syntax_node2(x_61, x_71, x_68, x_70); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_58); +return x_73; +} +} +case 2: +{ +lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_74 = lean_ctor_get(x_1, 0); +lean_inc(x_74); +lean_dec(x_1); +lean_inc(x_2); +x_75 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear(x_74, x_2, x_3); +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_77 = lean_ctor_get(x_75, 0); +x_78 = lean_ctor_get(x_2, 5); +lean_inc(x_78); +x_79 = 0; +x_80 = l_Lean_SourceInfo_fromRef(x_78, x_79); +lean_dec(x_78); +x_81 = lean_ctor_get(x_2, 2); +lean_inc(x_81); +x_82 = lean_ctor_get(x_2, 1); +lean_inc(x_82); +lean_dec(x_2); +x_83 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21; +x_84 = l_Lean_addMacroScope(x_82, x_83, x_81); +x_85 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19; +x_86 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25; +lean_inc(x_80); +x_87 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_87, 0, x_80); +lean_ctor_set(x_87, 1, x_85); +lean_ctor_set(x_87, 2, x_84); +lean_ctor_set(x_87, 3, x_86); +x_88 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_80); +x_89 = l_Lean_Syntax_node1(x_80, x_88, x_77); +x_90 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_91 = l_Lean_Syntax_node2(x_80, x_90, x_87, x_89); +lean_ctor_set(x_75, 0, x_91); +return x_75; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_92 = lean_ctor_get(x_75, 0); +x_93 = lean_ctor_get(x_75, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_75); +x_94 = lean_ctor_get(x_2, 5); +lean_inc(x_94); +x_95 = 0; +x_96 = l_Lean_SourceInfo_fromRef(x_94, x_95); +lean_dec(x_94); +x_97 = lean_ctor_get(x_2, 2); +lean_inc(x_97); +x_98 = lean_ctor_get(x_2, 1); +lean_inc(x_98); +lean_dec(x_2); +x_99 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21; +x_100 = l_Lean_addMacroScope(x_98, x_99, x_97); +x_101 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19; +x_102 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25; +lean_inc(x_96); +x_103 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_103, 0, x_96); +lean_ctor_set(x_103, 1, x_101); +lean_ctor_set(x_103, 2, x_100); +lean_ctor_set(x_103, 3, x_102); +x_104 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_96); +x_105 = l_Lean_Syntax_node1(x_96, x_104, x_92); +x_106 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_107 = l_Lean_Syntax_node2(x_96, x_106, x_103, x_105); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_93); +return x_108; +} +} +case 3: +{ +lean_object* x_109; lean_object* x_110; uint8_t x_111; +x_109 = lean_ctor_get(x_1, 0); +lean_inc(x_109); +lean_dec(x_1); +lean_inc(x_2); +x_110 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_109, x_2, x_3); +x_111 = !lean_is_exclusive(x_110); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_112 = lean_ctor_get(x_110, 0); +x_113 = lean_ctor_get(x_2, 5); +lean_inc(x_113); +x_114 = 0; +x_115 = l_Lean_SourceInfo_fromRef(x_113, x_114); +lean_dec(x_113); +x_116 = lean_ctor_get(x_2, 2); +lean_inc(x_116); +x_117 = lean_ctor_get(x_2, 1); +lean_inc(x_117); +lean_dec(x_2); +x_118 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29; +x_119 = l_Lean_addMacroScope(x_117, x_118, x_116); +x_120 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27; +x_121 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33; +lean_inc(x_115); +x_122 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_122, 0, x_115); +lean_ctor_set(x_122, 1, x_120); +lean_ctor_set(x_122, 2, x_119); +lean_ctor_set(x_122, 3, x_121); +x_123 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_115); +x_124 = l_Lean_Syntax_node1(x_115, x_123, x_112); +x_125 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_126 = l_Lean_Syntax_node2(x_115, x_125, x_122, x_124); +lean_ctor_set(x_110, 0, x_126); +return x_110; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_127 = lean_ctor_get(x_110, 0); +x_128 = lean_ctor_get(x_110, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_110); +x_129 = lean_ctor_get(x_2, 5); +lean_inc(x_129); +x_130 = 0; +x_131 = l_Lean_SourceInfo_fromRef(x_129, x_130); +lean_dec(x_129); +x_132 = lean_ctor_get(x_2, 2); +lean_inc(x_132); +x_133 = lean_ctor_get(x_2, 1); +lean_inc(x_133); +lean_dec(x_2); +x_134 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29; +x_135 = l_Lean_addMacroScope(x_133, x_134, x_132); +x_136 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27; +x_137 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33; +lean_inc(x_131); +x_138 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_138, 0, x_131); +lean_ctor_set(x_138, 1, x_136); +lean_ctor_set(x_138, 2, x_135); +lean_ctor_set(x_138, 3, x_137); +x_139 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_131); +x_140 = l_Lean_Syntax_node1(x_131, x_139, x_127); +x_141 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_142 = l_Lean_Syntax_node2(x_131, x_141, x_138, x_140); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_128); +return x_143; +} +} +case 4: +{ +lean_object* x_144; +x_144 = lean_ctor_get(x_1, 0); +lean_inc(x_144); +lean_dec(x_1); +if (lean_obj_tag(x_144) == 0) +{ +lean_object* x_145; lean_object* x_146; uint8_t x_147; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +lean_dec(x_144); +lean_inc(x_2); +x_146 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_145, x_2, x_3); +x_147 = !lean_is_exclusive(x_146); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; uint8_t x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_148 = lean_ctor_get(x_146, 0); +x_149 = lean_ctor_get(x_2, 5); +lean_inc(x_149); +x_150 = 0; +x_151 = l_Lean_SourceInfo_fromRef(x_149, x_150); +lean_dec(x_149); +x_152 = lean_ctor_get(x_2, 2); +lean_inc(x_152); +x_153 = lean_ctor_get(x_2, 1); +lean_inc(x_153); +lean_dec(x_2); +x_154 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +lean_inc(x_152); +lean_inc(x_153); +x_155 = l_Lean_addMacroScope(x_153, x_154, x_152); +x_156 = lean_box(0); +x_157 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35; +x_158 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41; +lean_inc(x_151); +x_159 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_159, 0, x_151); +lean_ctor_set(x_159, 1, x_157); +lean_ctor_set(x_159, 2, x_155); +lean_ctor_set(x_159, 3, x_158); +x_160 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_151); +x_161 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_161, 0, x_151); +lean_ctor_set(x_161, 1, x_160); +x_162 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_151); +x_163 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_163, 0, x_151); +lean_ctor_set(x_163, 1, x_162); +x_164 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_165 = l_Lean_addMacroScope(x_153, x_164, x_152); +x_166 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_151); +x_167 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_167, 0, x_151); +lean_ctor_set(x_167, 1, x_166); +lean_ctor_set(x_167, 2, x_165); +lean_ctor_set(x_167, 3, x_156); +x_168 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_151); +x_169 = l_Lean_Syntax_node2(x_151, x_168, x_163, x_167); +x_170 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_151); +x_171 = l_Lean_Syntax_node1(x_151, x_170, x_148); +x_172 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_151); +x_173 = l_Lean_Syntax_node2(x_151, x_172, x_169, x_171); +x_174 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_151); +x_175 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_175, 0, x_151); +lean_ctor_set(x_175, 1, x_174); +x_176 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_151); +x_177 = l_Lean_Syntax_node3(x_151, x_176, x_161, x_173, x_175); +lean_inc(x_151); +x_178 = l_Lean_Syntax_node1(x_151, x_170, x_177); +x_179 = l_Lean_Syntax_node2(x_151, x_172, x_159, x_178); +lean_ctor_set(x_146, 0, x_179); +return x_146; +} +else +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_180 = lean_ctor_get(x_146, 0); +x_181 = lean_ctor_get(x_146, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_146); +x_182 = lean_ctor_get(x_2, 5); +lean_inc(x_182); +x_183 = 0; +x_184 = l_Lean_SourceInfo_fromRef(x_182, x_183); +lean_dec(x_182); +x_185 = lean_ctor_get(x_2, 2); +lean_inc(x_185); +x_186 = lean_ctor_get(x_2, 1); +lean_inc(x_186); +lean_dec(x_2); +x_187 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +lean_inc(x_185); +lean_inc(x_186); +x_188 = l_Lean_addMacroScope(x_186, x_187, x_185); +x_189 = lean_box(0); +x_190 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35; +x_191 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41; +lean_inc(x_184); +x_192 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_192, 0, x_184); +lean_ctor_set(x_192, 1, x_190); +lean_ctor_set(x_192, 2, x_188); +lean_ctor_set(x_192, 3, x_191); +x_193 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_184); +x_194 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_194, 0, x_184); +lean_ctor_set(x_194, 1, x_193); +x_195 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_184); +x_196 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_196, 0, x_184); +lean_ctor_set(x_196, 1, x_195); +x_197 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_198 = l_Lean_addMacroScope(x_186, x_197, x_185); +x_199 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_184); +x_200 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_200, 0, x_184); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_198); +lean_ctor_set(x_200, 3, x_189); +x_201 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_184); +x_202 = l_Lean_Syntax_node2(x_184, x_201, x_196, x_200); +x_203 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_184); +x_204 = l_Lean_Syntax_node1(x_184, x_203, x_180); +x_205 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_184); +x_206 = l_Lean_Syntax_node2(x_184, x_205, x_202, x_204); +x_207 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_184); +x_208 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_208, 0, x_184); +lean_ctor_set(x_208, 1, x_207); +x_209 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_184); +x_210 = l_Lean_Syntax_node3(x_184, x_209, x_194, x_206, x_208); +lean_inc(x_184); +x_211 = l_Lean_Syntax_node1(x_184, x_203, x_210); +x_212 = l_Lean_Syntax_node2(x_184, x_205, x_192, x_211); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_181); +return x_213; +} +} +else +{ +lean_object* x_214; uint8_t x_215; lean_object* x_216; uint8_t x_217; +x_214 = lean_ctor_get(x_144, 0); +lean_inc(x_214); +lean_dec(x_144); +x_215 = lean_unbox(x_214); +lean_dec(x_214); +lean_inc(x_2); +x_216 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_215, x_2, x_3); +x_217 = !lean_is_exclusive(x_216); +if (x_217 == 0) +{ +lean_object* x_218; lean_object* x_219; uint8_t x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_218 = lean_ctor_get(x_216, 0); +x_219 = lean_ctor_get(x_2, 5); +lean_inc(x_219); +x_220 = 0; +x_221 = l_Lean_SourceInfo_fromRef(x_219, x_220); +lean_dec(x_219); +x_222 = lean_ctor_get(x_2, 2); +lean_inc(x_222); +x_223 = lean_ctor_get(x_2, 1); +lean_inc(x_223); +lean_dec(x_2); +x_224 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +lean_inc(x_222); +lean_inc(x_223); +x_225 = l_Lean_addMacroScope(x_223, x_224, x_222); +x_226 = lean_box(0); +x_227 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35; +x_228 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41; +lean_inc(x_221); +x_229 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_229, 0, x_221); +lean_ctor_set(x_229, 1, x_227); +lean_ctor_set(x_229, 2, x_225); +lean_ctor_set(x_229, 3, x_228); +x_230 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_221); +x_231 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_231, 0, x_221); +lean_ctor_set(x_231, 1, x_230); +x_232 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_221); +x_233 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_233, 0, x_221); +lean_ctor_set(x_233, 1, x_232); +x_234 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_235 = l_Lean_addMacroScope(x_223, x_234, x_222); +x_236 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_221); +x_237 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_237, 0, x_221); +lean_ctor_set(x_237, 1, x_236); +lean_ctor_set(x_237, 2, x_235); +lean_ctor_set(x_237, 3, x_226); +x_238 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_221); +x_239 = l_Lean_Syntax_node2(x_221, x_238, x_233, x_237); +x_240 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_221); +x_241 = l_Lean_Syntax_node1(x_221, x_240, x_218); +x_242 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_221); +x_243 = l_Lean_Syntax_node2(x_221, x_242, x_239, x_241); +x_244 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_221); +x_245 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_245, 0, x_221); +lean_ctor_set(x_245, 1, x_244); +x_246 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_221); +x_247 = l_Lean_Syntax_node3(x_221, x_246, x_231, x_243, x_245); +lean_inc(x_221); +x_248 = l_Lean_Syntax_node1(x_221, x_240, x_247); +x_249 = l_Lean_Syntax_node2(x_221, x_242, x_229, x_248); +lean_ctor_set(x_216, 0, x_249); +return x_216; +} +else +{ +lean_object* x_250; lean_object* x_251; lean_object* x_252; uint8_t x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; +x_250 = lean_ctor_get(x_216, 0); +x_251 = lean_ctor_get(x_216, 1); +lean_inc(x_251); +lean_inc(x_250); +lean_dec(x_216); +x_252 = lean_ctor_get(x_2, 5); +lean_inc(x_252); +x_253 = 0; +x_254 = l_Lean_SourceInfo_fromRef(x_252, x_253); +lean_dec(x_252); +x_255 = lean_ctor_get(x_2, 2); +lean_inc(x_255); +x_256 = lean_ctor_get(x_2, 1); +lean_inc(x_256); +lean_dec(x_2); +x_257 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37; +lean_inc(x_255); +lean_inc(x_256); +x_258 = l_Lean_addMacroScope(x_256, x_257, x_255); +x_259 = lean_box(0); +x_260 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35; +x_261 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41; +lean_inc(x_254); +x_262 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_262, 0, x_254); +lean_ctor_set(x_262, 1, x_260); +lean_ctor_set(x_262, 2, x_258); +lean_ctor_set(x_262, 3, x_261); +x_263 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_254); +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_254); +lean_ctor_set(x_264, 1, x_263); +x_265 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_254); +x_266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_266, 0, x_254); +lean_ctor_set(x_266, 1, x_265); +x_267 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_268 = l_Lean_addMacroScope(x_256, x_267, x_255); +x_269 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_254); +x_270 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_270, 0, x_254); +lean_ctor_set(x_270, 1, x_269); +lean_ctor_set(x_270, 2, x_268); +lean_ctor_set(x_270, 3, x_259); +x_271 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_254); +x_272 = l_Lean_Syntax_node2(x_254, x_271, x_266, x_270); +x_273 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_254); +x_274 = l_Lean_Syntax_node1(x_254, x_273, x_250); +x_275 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_254); +x_276 = l_Lean_Syntax_node2(x_254, x_275, x_272, x_274); +x_277 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_254); +x_278 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_278, 0, x_254); +lean_ctor_set(x_278, 1, x_277); +x_279 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_254); +x_280 = l_Lean_Syntax_node3(x_254, x_279, x_264, x_276, x_278); +lean_inc(x_254); +x_281 = l_Lean_Syntax_node1(x_254, x_273, x_280); +x_282 = l_Lean_Syntax_node2(x_254, x_275, x_262, x_281); +x_283 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_251); +return x_283; +} +} +} +case 5: +{ +lean_object* x_284; lean_object* x_285; uint8_t x_286; +x_284 = lean_ctor_get(x_1, 0); +lean_inc(x_284); +lean_dec(x_1); +lean_inc(x_2); +x_285 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_284, x_2, x_3); +x_286 = !lean_is_exclusive(x_285); +if (x_286 == 0) +{ +lean_object* x_287; lean_object* x_288; uint8_t x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_287 = lean_ctor_get(x_285, 0); +x_288 = lean_ctor_get(x_2, 5); +lean_inc(x_288); +x_289 = 0; +x_290 = l_Lean_SourceInfo_fromRef(x_288, x_289); +lean_dec(x_288); +x_291 = lean_ctor_get(x_2, 2); +lean_inc(x_291); +x_292 = lean_ctor_get(x_2, 1); +lean_inc(x_292); +lean_dec(x_2); +x_293 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58; +x_294 = l_Lean_addMacroScope(x_292, x_293, x_291); +x_295 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56; +x_296 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62; +lean_inc(x_290); +x_297 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_297, 0, x_290); +lean_ctor_set(x_297, 1, x_295); +lean_ctor_set(x_297, 2, x_294); +lean_ctor_set(x_297, 3, x_296); +x_298 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_290); +x_299 = l_Lean_Syntax_node1(x_290, x_298, x_287); +x_300 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_301 = l_Lean_Syntax_node2(x_290, x_300, x_297, x_299); +lean_ctor_set(x_285, 0, x_301); +return x_285; +} +else +{ +lean_object* x_302; lean_object* x_303; lean_object* x_304; uint8_t x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_302 = lean_ctor_get(x_285, 0); +x_303 = lean_ctor_get(x_285, 1); +lean_inc(x_303); +lean_inc(x_302); +lean_dec(x_285); +x_304 = lean_ctor_get(x_2, 5); +lean_inc(x_304); +x_305 = 0; +x_306 = l_Lean_SourceInfo_fromRef(x_304, x_305); +lean_dec(x_304); +x_307 = lean_ctor_get(x_2, 2); +lean_inc(x_307); +x_308 = lean_ctor_get(x_2, 1); +lean_inc(x_308); +lean_dec(x_2); +x_309 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58; +x_310 = l_Lean_addMacroScope(x_308, x_309, x_307); +x_311 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56; +x_312 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62; +lean_inc(x_306); +x_313 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_313, 0, x_306); +lean_ctor_set(x_313, 1, x_311); +lean_ctor_set(x_313, 2, x_310); +lean_ctor_set(x_313, 3, x_312); +x_314 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_306); +x_315 = l_Lean_Syntax_node1(x_306, x_314, x_302); +x_316 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_317 = l_Lean_Syntax_node2(x_306, x_316, x_313, x_315); +x_318 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_318, 1, x_303); +return x_318; +} +} +case 6: +{ +lean_object* x_319; +x_319 = lean_ctor_get(x_1, 0); +lean_inc(x_319); +lean_dec(x_1); +if (lean_obj_tag(x_319) == 0) +{ +lean_object* x_320; lean_object* x_321; uint8_t x_322; +x_320 = lean_ctor_get(x_319, 0); +lean_inc(x_320); +lean_dec(x_319); +lean_inc(x_2); +x_321 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_320, x_2, x_3); +x_322 = !lean_is_exclusive(x_321); +if (x_322 == 0) +{ +lean_object* x_323; lean_object* x_324; uint8_t x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; +x_323 = lean_ctor_get(x_321, 0); +x_324 = lean_ctor_get(x_2, 5); +lean_inc(x_324); +x_325 = 0; +x_326 = l_Lean_SourceInfo_fromRef(x_324, x_325); +lean_dec(x_324); +x_327 = lean_ctor_get(x_2, 2); +lean_inc(x_327); +x_328 = lean_ctor_get(x_2, 1); +lean_inc(x_328); +lean_dec(x_2); +x_329 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +lean_inc(x_327); +lean_inc(x_328); +x_330 = l_Lean_addMacroScope(x_328, x_329, x_327); +x_331 = lean_box(0); +x_332 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64; +x_333 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70; +lean_inc(x_326); +x_334 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_334, 0, x_326); +lean_ctor_set(x_334, 1, x_332); +lean_ctor_set(x_334, 2, x_330); +lean_ctor_set(x_334, 3, x_333); +x_335 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_326); +x_336 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_336, 0, x_326); +lean_ctor_set(x_336, 1, x_335); +x_337 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_326); +x_338 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_338, 0, x_326); +lean_ctor_set(x_338, 1, x_337); +x_339 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_340 = l_Lean_addMacroScope(x_328, x_339, x_327); +x_341 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_326); +x_342 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_342, 0, x_326); +lean_ctor_set(x_342, 1, x_341); +lean_ctor_set(x_342, 2, x_340); +lean_ctor_set(x_342, 3, x_331); +x_343 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_326); +x_344 = l_Lean_Syntax_node2(x_326, x_343, x_338, x_342); +x_345 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_326); +x_346 = l_Lean_Syntax_node1(x_326, x_345, x_323); +x_347 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_326); +x_348 = l_Lean_Syntax_node2(x_326, x_347, x_344, x_346); +x_349 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_326); +x_350 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_350, 0, x_326); +lean_ctor_set(x_350, 1, x_349); +x_351 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_326); +x_352 = l_Lean_Syntax_node3(x_326, x_351, x_336, x_348, x_350); +lean_inc(x_326); +x_353 = l_Lean_Syntax_node1(x_326, x_345, x_352); +x_354 = l_Lean_Syntax_node2(x_326, x_347, x_334, x_353); +lean_ctor_set(x_321, 0, x_354); +return x_321; +} +else +{ +lean_object* x_355; lean_object* x_356; lean_object* x_357; uint8_t x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; +x_355 = lean_ctor_get(x_321, 0); +x_356 = lean_ctor_get(x_321, 1); +lean_inc(x_356); +lean_inc(x_355); +lean_dec(x_321); +x_357 = lean_ctor_get(x_2, 5); +lean_inc(x_357); +x_358 = 0; +x_359 = l_Lean_SourceInfo_fromRef(x_357, x_358); +lean_dec(x_357); +x_360 = lean_ctor_get(x_2, 2); +lean_inc(x_360); +x_361 = lean_ctor_get(x_2, 1); +lean_inc(x_361); +lean_dec(x_2); +x_362 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +lean_inc(x_360); +lean_inc(x_361); +x_363 = l_Lean_addMacroScope(x_361, x_362, x_360); +x_364 = lean_box(0); +x_365 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64; +x_366 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70; +lean_inc(x_359); +x_367 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_367, 0, x_359); +lean_ctor_set(x_367, 1, x_365); +lean_ctor_set(x_367, 2, x_363); +lean_ctor_set(x_367, 3, x_366); +x_368 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_359); +x_369 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_369, 0, x_359); +lean_ctor_set(x_369, 1, x_368); +x_370 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_359); +x_371 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_371, 0, x_359); +lean_ctor_set(x_371, 1, x_370); +x_372 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_373 = l_Lean_addMacroScope(x_361, x_372, x_360); +x_374 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_359); +x_375 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_375, 0, x_359); +lean_ctor_set(x_375, 1, x_374); +lean_ctor_set(x_375, 2, x_373); +lean_ctor_set(x_375, 3, x_364); +x_376 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_359); +x_377 = l_Lean_Syntax_node2(x_359, x_376, x_371, x_375); +x_378 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_359); +x_379 = l_Lean_Syntax_node1(x_359, x_378, x_355); +x_380 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_359); +x_381 = l_Lean_Syntax_node2(x_359, x_380, x_377, x_379); +x_382 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_359); +x_383 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_383, 0, x_359); +lean_ctor_set(x_383, 1, x_382); +x_384 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_359); +x_385 = l_Lean_Syntax_node3(x_359, x_384, x_369, x_381, x_383); +lean_inc(x_359); +x_386 = l_Lean_Syntax_node1(x_359, x_378, x_385); +x_387 = l_Lean_Syntax_node2(x_359, x_380, x_367, x_386); +x_388 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_388, 0, x_387); +lean_ctor_set(x_388, 1, x_356); +return x_388; +} +} +else +{ +lean_object* x_389; uint8_t x_390; lean_object* x_391; uint8_t x_392; +x_389 = lean_ctor_get(x_319, 0); +lean_inc(x_389); +lean_dec(x_319); +x_390 = lean_unbox(x_389); +lean_dec(x_389); +lean_inc(x_2); +x_391 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_390, x_2, x_3); +x_392 = !lean_is_exclusive(x_391); +if (x_392 == 0) +{ +lean_object* x_393; lean_object* x_394; uint8_t x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +x_393 = lean_ctor_get(x_391, 0); +x_394 = lean_ctor_get(x_2, 5); +lean_inc(x_394); +x_395 = 0; +x_396 = l_Lean_SourceInfo_fromRef(x_394, x_395); +lean_dec(x_394); +x_397 = lean_ctor_get(x_2, 2); +lean_inc(x_397); +x_398 = lean_ctor_get(x_2, 1); +lean_inc(x_398); +lean_dec(x_2); +x_399 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +lean_inc(x_397); +lean_inc(x_398); +x_400 = l_Lean_addMacroScope(x_398, x_399, x_397); +x_401 = lean_box(0); +x_402 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64; +x_403 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70; +lean_inc(x_396); +x_404 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_404, 0, x_396); +lean_ctor_set(x_404, 1, x_402); +lean_ctor_set(x_404, 2, x_400); +lean_ctor_set(x_404, 3, x_403); +x_405 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_396); +x_406 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_406, 0, x_396); +lean_ctor_set(x_406, 1, x_405); +x_407 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_396); +x_408 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_408, 0, x_396); +lean_ctor_set(x_408, 1, x_407); +x_409 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_410 = l_Lean_addMacroScope(x_398, x_409, x_397); +x_411 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_396); +x_412 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_412, 0, x_396); +lean_ctor_set(x_412, 1, x_411); +lean_ctor_set(x_412, 2, x_410); +lean_ctor_set(x_412, 3, x_401); +x_413 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_396); +x_414 = l_Lean_Syntax_node2(x_396, x_413, x_408, x_412); +x_415 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_396); +x_416 = l_Lean_Syntax_node1(x_396, x_415, x_393); +x_417 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_396); +x_418 = l_Lean_Syntax_node2(x_396, x_417, x_414, x_416); +x_419 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_396); +x_420 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_420, 0, x_396); +lean_ctor_set(x_420, 1, x_419); +x_421 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_396); +x_422 = l_Lean_Syntax_node3(x_396, x_421, x_406, x_418, x_420); +lean_inc(x_396); +x_423 = l_Lean_Syntax_node1(x_396, x_415, x_422); +x_424 = l_Lean_Syntax_node2(x_396, x_417, x_404, x_423); +lean_ctor_set(x_391, 0, x_424); +return x_391; +} +else +{ +lean_object* x_425; lean_object* x_426; lean_object* x_427; uint8_t x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; +x_425 = lean_ctor_get(x_391, 0); +x_426 = lean_ctor_get(x_391, 1); +lean_inc(x_426); +lean_inc(x_425); +lean_dec(x_391); +x_427 = lean_ctor_get(x_2, 5); +lean_inc(x_427); +x_428 = 0; +x_429 = l_Lean_SourceInfo_fromRef(x_427, x_428); +lean_dec(x_427); +x_430 = lean_ctor_get(x_2, 2); +lean_inc(x_430); +x_431 = lean_ctor_get(x_2, 1); +lean_inc(x_431); +lean_dec(x_2); +x_432 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66; +lean_inc(x_430); +lean_inc(x_431); +x_433 = l_Lean_addMacroScope(x_431, x_432, x_430); +x_434 = lean_box(0); +x_435 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64; +x_436 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70; +lean_inc(x_429); +x_437 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_437, 0, x_429); +lean_ctor_set(x_437, 1, x_435); +lean_ctor_set(x_437, 2, x_433); +lean_ctor_set(x_437, 3, x_436); +x_438 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_429); +x_439 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_439, 0, x_429); +lean_ctor_set(x_439, 1, x_438); +x_440 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_429); +x_441 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_441, 0, x_429); +lean_ctor_set(x_441, 1, x_440); +x_442 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_443 = l_Lean_addMacroScope(x_431, x_442, x_430); +x_444 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_429); +x_445 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_445, 0, x_429); +lean_ctor_set(x_445, 1, x_444); +lean_ctor_set(x_445, 2, x_443); +lean_ctor_set(x_445, 3, x_434); +x_446 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_429); +x_447 = l_Lean_Syntax_node2(x_429, x_446, x_441, x_445); +x_448 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_429); +x_449 = l_Lean_Syntax_node1(x_429, x_448, x_425); +x_450 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_429); +x_451 = l_Lean_Syntax_node2(x_429, x_450, x_447, x_449); +x_452 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_429); +x_453 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_453, 0, x_429); +lean_ctor_set(x_453, 1, x_452); +x_454 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_429); +x_455 = l_Lean_Syntax_node3(x_429, x_454, x_439, x_451, x_453); +lean_inc(x_429); +x_456 = l_Lean_Syntax_node1(x_429, x_448, x_455); +x_457 = l_Lean_Syntax_node2(x_429, x_450, x_437, x_456); +x_458 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_458, 0, x_457); +lean_ctor_set(x_458, 1, x_426); +return x_458; +} +} +} +case 7: +{ +lean_object* x_459; lean_object* x_460; uint8_t x_461; +x_459 = lean_ctor_get(x_1, 0); +lean_inc(x_459); +lean_dec(x_1); +lean_inc(x_2); +x_460 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_459, x_2, x_3); +x_461 = !lean_is_exclusive(x_460); +if (x_461 == 0) +{ +lean_object* x_462; lean_object* x_463; uint8_t x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +x_462 = lean_ctor_get(x_460, 0); +x_463 = lean_ctor_get(x_2, 5); +lean_inc(x_463); +x_464 = 0; +x_465 = l_Lean_SourceInfo_fromRef(x_463, x_464); +lean_dec(x_463); +x_466 = lean_ctor_get(x_2, 2); +lean_inc(x_466); +x_467 = lean_ctor_get(x_2, 1); +lean_inc(x_467); +lean_dec(x_2); +x_468 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74; +x_469 = l_Lean_addMacroScope(x_467, x_468, x_466); +x_470 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72; +x_471 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78; +lean_inc(x_465); +x_472 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_472, 0, x_465); +lean_ctor_set(x_472, 1, x_470); +lean_ctor_set(x_472, 2, x_469); +lean_ctor_set(x_472, 3, x_471); +x_473 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_465); +x_474 = l_Lean_Syntax_node1(x_465, x_473, x_462); +x_475 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_476 = l_Lean_Syntax_node2(x_465, x_475, x_472, x_474); +lean_ctor_set(x_460, 0, x_476); +return x_460; +} +else +{ +lean_object* x_477; lean_object* x_478; lean_object* x_479; uint8_t x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +x_477 = lean_ctor_get(x_460, 0); +x_478 = lean_ctor_get(x_460, 1); +lean_inc(x_478); +lean_inc(x_477); +lean_dec(x_460); +x_479 = lean_ctor_get(x_2, 5); +lean_inc(x_479); +x_480 = 0; +x_481 = l_Lean_SourceInfo_fromRef(x_479, x_480); +lean_dec(x_479); +x_482 = lean_ctor_get(x_2, 2); +lean_inc(x_482); +x_483 = lean_ctor_get(x_2, 1); +lean_inc(x_483); +lean_dec(x_2); +x_484 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74; +x_485 = l_Lean_addMacroScope(x_483, x_484, x_482); +x_486 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72; +x_487 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78; +lean_inc(x_481); +x_488 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_488, 0, x_481); +lean_ctor_set(x_488, 1, x_486); +lean_ctor_set(x_488, 2, x_485); +lean_ctor_set(x_488, 3, x_487); +x_489 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_481); +x_490 = l_Lean_Syntax_node1(x_481, x_489, x_477); +x_491 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_492 = l_Lean_Syntax_node2(x_481, x_491, x_488, x_490); +x_493 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_493, 0, x_492); +lean_ctor_set(x_493, 1, x_478); +return x_493; +} +} +case 8: +{ +lean_object* x_494; lean_object* x_495; uint8_t x_496; +x_494 = lean_ctor_get(x_1, 0); +lean_inc(x_494); +lean_dec(x_1); +lean_inc(x_2); +x_495 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_494, x_2, x_3); +x_496 = !lean_is_exclusive(x_495); +if (x_496 == 0) +{ +lean_object* x_497; lean_object* x_498; uint8_t x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; lean_object* x_511; +x_497 = lean_ctor_get(x_495, 0); +x_498 = lean_ctor_get(x_2, 5); +lean_inc(x_498); +x_499 = 0; +x_500 = l_Lean_SourceInfo_fromRef(x_498, x_499); +lean_dec(x_498); +x_501 = lean_ctor_get(x_2, 2); +lean_inc(x_501); +x_502 = lean_ctor_get(x_2, 1); +lean_inc(x_502); +lean_dec(x_2); +x_503 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82; +x_504 = l_Lean_addMacroScope(x_502, x_503, x_501); +x_505 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80; +x_506 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86; +lean_inc(x_500); +x_507 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_507, 0, x_500); +lean_ctor_set(x_507, 1, x_505); +lean_ctor_set(x_507, 2, x_504); +lean_ctor_set(x_507, 3, x_506); +x_508 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_500); +x_509 = l_Lean_Syntax_node1(x_500, x_508, x_497); +x_510 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_511 = l_Lean_Syntax_node2(x_500, x_510, x_507, x_509); +lean_ctor_set(x_495, 0, x_511); +return x_495; +} +else +{ +lean_object* x_512; lean_object* x_513; lean_object* x_514; uint8_t x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; lean_object* x_519; lean_object* x_520; lean_object* x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; lean_object* x_527; lean_object* x_528; +x_512 = lean_ctor_get(x_495, 0); +x_513 = lean_ctor_get(x_495, 1); +lean_inc(x_513); +lean_inc(x_512); +lean_dec(x_495); +x_514 = lean_ctor_get(x_2, 5); +lean_inc(x_514); +x_515 = 0; +x_516 = l_Lean_SourceInfo_fromRef(x_514, x_515); +lean_dec(x_514); +x_517 = lean_ctor_get(x_2, 2); +lean_inc(x_517); +x_518 = lean_ctor_get(x_2, 1); +lean_inc(x_518); +lean_dec(x_2); +x_519 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82; +x_520 = l_Lean_addMacroScope(x_518, x_519, x_517); +x_521 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80; +x_522 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86; +lean_inc(x_516); +x_523 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_523, 0, x_516); +lean_ctor_set(x_523, 1, x_521); +lean_ctor_set(x_523, 2, x_520); +lean_ctor_set(x_523, 3, x_522); +x_524 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_516); +x_525 = l_Lean_Syntax_node1(x_516, x_524, x_512); +x_526 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_527 = l_Lean_Syntax_node2(x_516, x_526, x_523, x_525); +x_528 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_528, 0, x_527); +lean_ctor_set(x_528, 1, x_513); +return x_528; +} +} +case 9: +{ +uint8_t x_529; lean_object* x_530; uint8_t x_531; +x_529 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_530 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_529, x_2, x_3); +x_531 = !lean_is_exclusive(x_530); +if (x_531 == 0) +{ +lean_object* x_532; lean_object* x_533; uint8_t x_534; lean_object* x_535; lean_object* x_536; lean_object* x_537; lean_object* x_538; lean_object* x_539; lean_object* x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; lean_object* x_544; lean_object* x_545; lean_object* x_546; +x_532 = lean_ctor_get(x_530, 0); +x_533 = lean_ctor_get(x_2, 5); +lean_inc(x_533); +x_534 = 0; +x_535 = l_Lean_SourceInfo_fromRef(x_533, x_534); +lean_dec(x_533); +x_536 = lean_ctor_get(x_2, 2); +lean_inc(x_536); +x_537 = lean_ctor_get(x_2, 1); +lean_inc(x_537); +lean_dec(x_2); +x_538 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90; +x_539 = l_Lean_addMacroScope(x_537, x_538, x_536); +x_540 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88; +x_541 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94; +lean_inc(x_535); +x_542 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_542, 0, x_535); +lean_ctor_set(x_542, 1, x_540); +lean_ctor_set(x_542, 2, x_539); +lean_ctor_set(x_542, 3, x_541); +x_543 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_535); +x_544 = l_Lean_Syntax_node1(x_535, x_543, x_532); +x_545 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_546 = l_Lean_Syntax_node2(x_535, x_545, x_542, x_544); +lean_ctor_set(x_530, 0, x_546); +return x_530; +} +else +{ +lean_object* x_547; lean_object* x_548; lean_object* x_549; uint8_t x_550; lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; lean_object* x_555; lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; lean_object* x_561; lean_object* x_562; lean_object* x_563; +x_547 = lean_ctor_get(x_530, 0); +x_548 = lean_ctor_get(x_530, 1); +lean_inc(x_548); +lean_inc(x_547); +lean_dec(x_530); +x_549 = lean_ctor_get(x_2, 5); +lean_inc(x_549); +x_550 = 0; +x_551 = l_Lean_SourceInfo_fromRef(x_549, x_550); +lean_dec(x_549); +x_552 = lean_ctor_get(x_2, 2); +lean_inc(x_552); +x_553 = lean_ctor_get(x_2, 1); +lean_inc(x_553); +lean_dec(x_2); +x_554 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90; +x_555 = l_Lean_addMacroScope(x_553, x_554, x_552); +x_556 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88; +x_557 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94; +lean_inc(x_551); +x_558 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_558, 0, x_551); +lean_ctor_set(x_558, 1, x_556); +lean_ctor_set(x_558, 2, x_555); +lean_ctor_set(x_558, 3, x_557); +x_559 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_551); +x_560 = l_Lean_Syntax_node1(x_551, x_559, x_547); +x_561 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_562 = l_Lean_Syntax_node2(x_551, x_561, x_558, x_560); +x_563 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_563, 0, x_562); +lean_ctor_set(x_563, 1, x_548); +return x_563; +} +} +case 10: +{ +lean_object* x_564; +x_564 = lean_ctor_get(x_1, 0); +lean_inc(x_564); +lean_dec(x_1); +if (lean_obj_tag(x_564) == 0) +{ +lean_object* x_565; lean_object* x_566; uint8_t x_567; +x_565 = lean_ctor_get(x_564, 0); +lean_inc(x_565); +lean_dec(x_564); +lean_inc(x_2); +x_566 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_565, x_2, x_3); +x_567 = !lean_is_exclusive(x_566); +if (x_567 == 0) +{ +lean_object* x_568; lean_object* x_569; uint8_t x_570; lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; lean_object* x_575; lean_object* x_576; lean_object* x_577; lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_586; lean_object* x_587; lean_object* x_588; lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; lean_object* x_593; lean_object* x_594; lean_object* x_595; lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; +x_568 = lean_ctor_get(x_566, 0); +x_569 = lean_ctor_get(x_2, 5); +lean_inc(x_569); +x_570 = 0; +x_571 = l_Lean_SourceInfo_fromRef(x_569, x_570); +lean_dec(x_569); +x_572 = lean_ctor_get(x_2, 2); +lean_inc(x_572); +x_573 = lean_ctor_get(x_2, 1); +lean_inc(x_573); +lean_dec(x_2); +x_574 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +lean_inc(x_572); +lean_inc(x_573); +x_575 = l_Lean_addMacroScope(x_573, x_574, x_572); +x_576 = lean_box(0); +x_577 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96; +x_578 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102; +lean_inc(x_571); +x_579 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_579, 0, x_571); +lean_ctor_set(x_579, 1, x_577); +lean_ctor_set(x_579, 2, x_575); +lean_ctor_set(x_579, 3, x_578); +x_580 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_571); +x_581 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_581, 0, x_571); +lean_ctor_set(x_581, 1, x_580); +x_582 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_571); +x_583 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_583, 0, x_571); +lean_ctor_set(x_583, 1, x_582); +x_584 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_585 = l_Lean_addMacroScope(x_573, x_584, x_572); +x_586 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_571); +x_587 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_587, 0, x_571); +lean_ctor_set(x_587, 1, x_586); +lean_ctor_set(x_587, 2, x_585); +lean_ctor_set(x_587, 3, x_576); +x_588 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_571); +x_589 = l_Lean_Syntax_node2(x_571, x_588, x_583, x_587); +x_590 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_571); +x_591 = l_Lean_Syntax_node1(x_571, x_590, x_568); +x_592 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_571); +x_593 = l_Lean_Syntax_node2(x_571, x_592, x_589, x_591); +x_594 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_571); +x_595 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_595, 0, x_571); +lean_ctor_set(x_595, 1, x_594); +x_596 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_571); +x_597 = l_Lean_Syntax_node3(x_571, x_596, x_581, x_593, x_595); +lean_inc(x_571); +x_598 = l_Lean_Syntax_node1(x_571, x_590, x_597); +x_599 = l_Lean_Syntax_node2(x_571, x_592, x_579, x_598); +lean_ctor_set(x_566, 0, x_599); +return x_566; +} +else +{ +lean_object* x_600; lean_object* x_601; lean_object* x_602; uint8_t x_603; lean_object* x_604; lean_object* x_605; lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_610; lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; lean_object* x_617; lean_object* x_618; lean_object* x_619; lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; lean_object* x_624; lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; lean_object* x_629; lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; +x_600 = lean_ctor_get(x_566, 0); +x_601 = lean_ctor_get(x_566, 1); +lean_inc(x_601); +lean_inc(x_600); +lean_dec(x_566); +x_602 = lean_ctor_get(x_2, 5); +lean_inc(x_602); +x_603 = 0; +x_604 = l_Lean_SourceInfo_fromRef(x_602, x_603); +lean_dec(x_602); +x_605 = lean_ctor_get(x_2, 2); +lean_inc(x_605); +x_606 = lean_ctor_get(x_2, 1); +lean_inc(x_606); +lean_dec(x_2); +x_607 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +lean_inc(x_605); +lean_inc(x_606); +x_608 = l_Lean_addMacroScope(x_606, x_607, x_605); +x_609 = lean_box(0); +x_610 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96; +x_611 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102; +lean_inc(x_604); +x_612 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_612, 0, x_604); +lean_ctor_set(x_612, 1, x_610); +lean_ctor_set(x_612, 2, x_608); +lean_ctor_set(x_612, 3, x_611); +x_613 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_604); +x_614 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_614, 0, x_604); +lean_ctor_set(x_614, 1, x_613); +x_615 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_604); +x_616 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_616, 0, x_604); +lean_ctor_set(x_616, 1, x_615); +x_617 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50; +x_618 = l_Lean_addMacroScope(x_606, x_617, x_605); +x_619 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49; +lean_inc(x_604); +x_620 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_620, 0, x_604); +lean_ctor_set(x_620, 1, x_619); +lean_ctor_set(x_620, 2, x_618); +lean_ctor_set(x_620, 3, x_609); +x_621 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_604); +x_622 = l_Lean_Syntax_node2(x_604, x_621, x_616, x_620); +x_623 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_604); +x_624 = l_Lean_Syntax_node1(x_604, x_623, x_600); +x_625 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_604); +x_626 = l_Lean_Syntax_node2(x_604, x_625, x_622, x_624); +x_627 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_604); +x_628 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_628, 0, x_604); +lean_ctor_set(x_628, 1, x_627); +x_629 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_604); +x_630 = l_Lean_Syntax_node3(x_604, x_629, x_614, x_626, x_628); +lean_inc(x_604); +x_631 = l_Lean_Syntax_node1(x_604, x_623, x_630); +x_632 = l_Lean_Syntax_node2(x_604, x_625, x_612, x_631); +x_633 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_633, 0, x_632); +lean_ctor_set(x_633, 1, x_601); +return x_633; +} +} +else +{ +lean_object* x_634; uint8_t x_635; lean_object* x_636; uint8_t x_637; +x_634 = lean_ctor_get(x_564, 0); +lean_inc(x_634); +lean_dec(x_564); +x_635 = lean_unbox(x_634); +lean_dec(x_634); +lean_inc(x_2); +x_636 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_635, x_2, x_3); +x_637 = !lean_is_exclusive(x_636); +if (x_637 == 0) +{ +lean_object* x_638; lean_object* x_639; uint8_t x_640; lean_object* x_641; lean_object* x_642; lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; lean_object* x_647; lean_object* x_648; lean_object* x_649; lean_object* x_650; lean_object* x_651; lean_object* x_652; lean_object* x_653; lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; lean_object* x_662; lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; lean_object* x_667; lean_object* x_668; lean_object* x_669; +x_638 = lean_ctor_get(x_636, 0); +x_639 = lean_ctor_get(x_2, 5); +lean_inc(x_639); +x_640 = 0; +x_641 = l_Lean_SourceInfo_fromRef(x_639, x_640); +lean_dec(x_639); +x_642 = lean_ctor_get(x_2, 2); +lean_inc(x_642); +x_643 = lean_ctor_get(x_2, 1); +lean_inc(x_643); +lean_dec(x_2); +x_644 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +lean_inc(x_642); +lean_inc(x_643); +x_645 = l_Lean_addMacroScope(x_643, x_644, x_642); +x_646 = lean_box(0); +x_647 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96; +x_648 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102; +lean_inc(x_641); +x_649 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_649, 0, x_641); +lean_ctor_set(x_649, 1, x_647); +lean_ctor_set(x_649, 2, x_645); +lean_ctor_set(x_649, 3, x_648); +x_650 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_641); +x_651 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_651, 0, x_641); +lean_ctor_set(x_651, 1, x_650); +x_652 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_641); +x_653 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_653, 0, x_641); +lean_ctor_set(x_653, 1, x_652); +x_654 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_655 = l_Lean_addMacroScope(x_643, x_654, x_642); +x_656 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_641); +x_657 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_657, 0, x_641); +lean_ctor_set(x_657, 1, x_656); +lean_ctor_set(x_657, 2, x_655); +lean_ctor_set(x_657, 3, x_646); +x_658 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_641); +x_659 = l_Lean_Syntax_node2(x_641, x_658, x_653, x_657); +x_660 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_641); +x_661 = l_Lean_Syntax_node1(x_641, x_660, x_638); +x_662 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_641); +x_663 = l_Lean_Syntax_node2(x_641, x_662, x_659, x_661); +x_664 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_641); +x_665 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_665, 0, x_641); +lean_ctor_set(x_665, 1, x_664); +x_666 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_641); +x_667 = l_Lean_Syntax_node3(x_641, x_666, x_651, x_663, x_665); +lean_inc(x_641); +x_668 = l_Lean_Syntax_node1(x_641, x_660, x_667); +x_669 = l_Lean_Syntax_node2(x_641, x_662, x_649, x_668); +lean_ctor_set(x_636, 0, x_669); +return x_636; +} +else +{ +lean_object* x_670; lean_object* x_671; lean_object* x_672; uint8_t x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; lean_object* x_694; lean_object* x_695; lean_object* x_696; lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; +x_670 = lean_ctor_get(x_636, 0); +x_671 = lean_ctor_get(x_636, 1); +lean_inc(x_671); +lean_inc(x_670); +lean_dec(x_636); +x_672 = lean_ctor_get(x_2, 5); +lean_inc(x_672); +x_673 = 0; +x_674 = l_Lean_SourceInfo_fromRef(x_672, x_673); +lean_dec(x_672); +x_675 = lean_ctor_get(x_2, 2); +lean_inc(x_675); +x_676 = lean_ctor_get(x_2, 1); +lean_inc(x_676); +lean_dec(x_2); +x_677 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98; +lean_inc(x_675); +lean_inc(x_676); +x_678 = l_Lean_addMacroScope(x_676, x_677, x_675); +x_679 = lean_box(0); +x_680 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96; +x_681 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102; +lean_inc(x_674); +x_682 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_682, 0, x_674); +lean_ctor_set(x_682, 1, x_680); +lean_ctor_set(x_682, 2, x_678); +lean_ctor_set(x_682, 3, x_681); +x_683 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44; +lean_inc(x_674); +x_684 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_684, 0, x_674); +lean_ctor_set(x_684, 1, x_683); +x_685 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_674); +x_686 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_686, 0, x_674); +lean_ctor_set(x_686, 1, x_685); +x_687 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54; +x_688 = l_Lean_addMacroScope(x_676, x_687, x_675); +x_689 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53; +lean_inc(x_674); +x_690 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_690, 0, x_674); +lean_ctor_set(x_690, 1, x_689); +lean_ctor_set(x_690, 2, x_688); +lean_ctor_set(x_690, 3, x_679); +x_691 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_674); +x_692 = l_Lean_Syntax_node2(x_674, x_691, x_686, x_690); +x_693 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_674); +x_694 = l_Lean_Syntax_node1(x_674, x_693, x_670); +x_695 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +lean_inc(x_674); +x_696 = l_Lean_Syntax_node2(x_674, x_695, x_692, x_694); +x_697 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +lean_inc(x_674); +x_698 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_698, 0, x_674); +lean_ctor_set(x_698, 1, x_697); +x_699 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43; +lean_inc(x_674); +x_700 = l_Lean_Syntax_node3(x_674, x_699, x_684, x_696, x_698); +lean_inc(x_674); +x_701 = l_Lean_Syntax_node1(x_674, x_693, x_700); +x_702 = l_Lean_Syntax_node2(x_674, x_695, x_682, x_701); +x_703 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_703, 0, x_702); +lean_ctor_set(x_703, 1, x_671); +return x_703; +} +} +} +case 11: +{ +lean_object* x_704; lean_object* x_705; uint8_t x_706; +x_704 = lean_ctor_get(x_1, 0); +lean_inc(x_704); +lean_dec(x_1); +lean_inc(x_2); +x_705 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_704, x_2, x_3); +x_706 = !lean_is_exclusive(x_705); +if (x_706 == 0) +{ +lean_object* x_707; lean_object* x_708; uint8_t x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; lean_object* x_719; lean_object* x_720; lean_object* x_721; +x_707 = lean_ctor_get(x_705, 0); +x_708 = lean_ctor_get(x_2, 5); +lean_inc(x_708); +x_709 = 0; +x_710 = l_Lean_SourceInfo_fromRef(x_708, x_709); +lean_dec(x_708); +x_711 = lean_ctor_get(x_2, 2); +lean_inc(x_711); +x_712 = lean_ctor_get(x_2, 1); +lean_inc(x_712); +lean_dec(x_2); +x_713 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106; +x_714 = l_Lean_addMacroScope(x_712, x_713, x_711); +x_715 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104; +x_716 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110; +lean_inc(x_710); +x_717 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_717, 0, x_710); +lean_ctor_set(x_717, 1, x_715); +lean_ctor_set(x_717, 2, x_714); +lean_ctor_set(x_717, 3, x_716); +x_718 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_710); +x_719 = l_Lean_Syntax_node1(x_710, x_718, x_707); +x_720 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_721 = l_Lean_Syntax_node2(x_710, x_720, x_717, x_719); +lean_ctor_set(x_705, 0, x_721); +return x_705; +} +else +{ +lean_object* x_722; lean_object* x_723; lean_object* x_724; uint8_t x_725; lean_object* x_726; lean_object* x_727; lean_object* x_728; lean_object* x_729; lean_object* x_730; lean_object* x_731; lean_object* x_732; lean_object* x_733; lean_object* x_734; lean_object* x_735; lean_object* x_736; lean_object* x_737; lean_object* x_738; +x_722 = lean_ctor_get(x_705, 0); +x_723 = lean_ctor_get(x_705, 1); +lean_inc(x_723); +lean_inc(x_722); +lean_dec(x_705); +x_724 = lean_ctor_get(x_2, 5); +lean_inc(x_724); +x_725 = 0; +x_726 = l_Lean_SourceInfo_fromRef(x_724, x_725); +lean_dec(x_724); +x_727 = lean_ctor_get(x_2, 2); +lean_inc(x_727); +x_728 = lean_ctor_get(x_2, 1); +lean_inc(x_728); +lean_dec(x_2); +x_729 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106; +x_730 = l_Lean_addMacroScope(x_728, x_729, x_727); +x_731 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104; +x_732 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110; +lean_inc(x_726); +x_733 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_733, 0, x_726); +lean_ctor_set(x_733, 1, x_731); +lean_ctor_set(x_733, 2, x_730); +lean_ctor_set(x_733, 3, x_732); +x_734 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_726); +x_735 = l_Lean_Syntax_node1(x_726, x_734, x_722); +x_736 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_737 = l_Lean_Syntax_node2(x_726, x_736, x_733, x_735); +x_738 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_738, 0, x_737); +lean_ctor_set(x_738, 1, x_723); +return x_738; +} +} +case 12: +{ +uint8_t x_739; lean_object* x_740; uint8_t x_741; +x_739 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_740 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText(x_739, x_2, x_3); +x_741 = !lean_is_exclusive(x_740); +if (x_741 == 0) +{ +lean_object* x_742; lean_object* x_743; uint8_t x_744; lean_object* x_745; lean_object* x_746; lean_object* x_747; lean_object* x_748; lean_object* x_749; lean_object* x_750; lean_object* x_751; lean_object* x_752; lean_object* x_753; lean_object* x_754; lean_object* x_755; lean_object* x_756; +x_742 = lean_ctor_get(x_740, 0); +x_743 = lean_ctor_get(x_2, 5); +lean_inc(x_743); +x_744 = 0; +x_745 = l_Lean_SourceInfo_fromRef(x_743, x_744); +lean_dec(x_743); +x_746 = lean_ctor_get(x_2, 2); +lean_inc(x_746); +x_747 = lean_ctor_get(x_2, 1); +lean_inc(x_747); +lean_dec(x_2); +x_748 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114; +x_749 = l_Lean_addMacroScope(x_747, x_748, x_746); +x_750 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112; +x_751 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118; +lean_inc(x_745); +x_752 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_752, 0, x_745); +lean_ctor_set(x_752, 1, x_750); +lean_ctor_set(x_752, 2, x_749); +lean_ctor_set(x_752, 3, x_751); +x_753 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_745); +x_754 = l_Lean_Syntax_node1(x_745, x_753, x_742); +x_755 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_756 = l_Lean_Syntax_node2(x_745, x_755, x_752, x_754); +lean_ctor_set(x_740, 0, x_756); +return x_740; +} +else +{ +lean_object* x_757; lean_object* x_758; lean_object* x_759; uint8_t x_760; lean_object* x_761; lean_object* x_762; lean_object* x_763; lean_object* x_764; lean_object* x_765; lean_object* x_766; lean_object* x_767; lean_object* x_768; lean_object* x_769; lean_object* x_770; lean_object* x_771; lean_object* x_772; lean_object* x_773; +x_757 = lean_ctor_get(x_740, 0); +x_758 = lean_ctor_get(x_740, 1); +lean_inc(x_758); +lean_inc(x_757); +lean_dec(x_740); +x_759 = lean_ctor_get(x_2, 5); +lean_inc(x_759); +x_760 = 0; +x_761 = l_Lean_SourceInfo_fromRef(x_759, x_760); +lean_dec(x_759); +x_762 = lean_ctor_get(x_2, 2); +lean_inc(x_762); +x_763 = lean_ctor_get(x_2, 1); +lean_inc(x_763); +lean_dec(x_2); +x_764 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114; +x_765 = l_Lean_addMacroScope(x_763, x_764, x_762); +x_766 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112; +x_767 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118; +lean_inc(x_761); +x_768 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_768, 0, x_761); +lean_ctor_set(x_768, 1, x_766); +lean_ctor_set(x_768, 2, x_765); +lean_ctor_set(x_768, 3, x_767); +x_769 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_761); +x_770 = l_Lean_Syntax_node1(x_761, x_769, x_757); +x_771 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_772 = l_Lean_Syntax_node2(x_761, x_771, x_768, x_770); +x_773 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_773, 0, x_772); +lean_ctor_set(x_773, 1, x_758); +return x_773; +} +} +case 13: +{ +lean_object* x_774; lean_object* x_775; uint8_t x_776; +x_774 = lean_ctor_get(x_1, 0); +lean_inc(x_774); +lean_dec(x_1); +lean_inc(x_2); +x_775 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_774, x_2, x_3); +x_776 = !lean_is_exclusive(x_775); +if (x_776 == 0) +{ +lean_object* x_777; lean_object* x_778; uint8_t x_779; lean_object* x_780; lean_object* x_781; lean_object* x_782; lean_object* x_783; lean_object* x_784; lean_object* x_785; lean_object* x_786; lean_object* x_787; lean_object* x_788; lean_object* x_789; lean_object* x_790; lean_object* x_791; +x_777 = lean_ctor_get(x_775, 0); +x_778 = lean_ctor_get(x_2, 5); +lean_inc(x_778); +x_779 = 0; +x_780 = l_Lean_SourceInfo_fromRef(x_778, x_779); +lean_dec(x_778); +x_781 = lean_ctor_get(x_2, 2); +lean_inc(x_781); +x_782 = lean_ctor_get(x_2, 1); +lean_inc(x_782); +lean_dec(x_2); +x_783 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122; +x_784 = l_Lean_addMacroScope(x_782, x_783, x_781); +x_785 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120; +x_786 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126; +lean_inc(x_780); +x_787 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_787, 0, x_780); +lean_ctor_set(x_787, 1, x_785); +lean_ctor_set(x_787, 2, x_784); +lean_ctor_set(x_787, 3, x_786); +x_788 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_780); +x_789 = l_Lean_Syntax_node1(x_780, x_788, x_777); +x_790 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_791 = l_Lean_Syntax_node2(x_780, x_790, x_787, x_789); +lean_ctor_set(x_775, 0, x_791); +return x_775; +} +else +{ +lean_object* x_792; lean_object* x_793; lean_object* x_794; uint8_t x_795; lean_object* x_796; lean_object* x_797; lean_object* x_798; lean_object* x_799; lean_object* x_800; lean_object* x_801; lean_object* x_802; lean_object* x_803; lean_object* x_804; lean_object* x_805; lean_object* x_806; lean_object* x_807; lean_object* x_808; +x_792 = lean_ctor_get(x_775, 0); +x_793 = lean_ctor_get(x_775, 1); +lean_inc(x_793); +lean_inc(x_792); +lean_dec(x_775); +x_794 = lean_ctor_get(x_2, 5); +lean_inc(x_794); +x_795 = 0; +x_796 = l_Lean_SourceInfo_fromRef(x_794, x_795); +lean_dec(x_794); +x_797 = lean_ctor_get(x_2, 2); +lean_inc(x_797); +x_798 = lean_ctor_get(x_2, 1); +lean_inc(x_798); +lean_dec(x_2); +x_799 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122; +x_800 = l_Lean_addMacroScope(x_798, x_799, x_797); +x_801 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120; +x_802 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126; +lean_inc(x_796); +x_803 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_803, 0, x_796); +lean_ctor_set(x_803, 1, x_801); +lean_ctor_set(x_803, 2, x_800); +lean_ctor_set(x_803, 3, x_802); +x_804 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_796); +x_805 = l_Lean_Syntax_node1(x_796, x_804, x_792); +x_806 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_807 = l_Lean_Syntax_node2(x_796, x_806, x_803, x_805); +x_808 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_808, 0, x_807); +lean_ctor_set(x_808, 1, x_793); +return x_808; +} +} +case 14: +{ +lean_object* x_809; lean_object* x_810; uint8_t x_811; +x_809 = lean_ctor_get(x_1, 0); +lean_inc(x_809); +lean_dec(x_1); +lean_inc(x_2); +x_810 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_809, x_2, x_3); +x_811 = !lean_is_exclusive(x_810); +if (x_811 == 0) +{ +lean_object* x_812; lean_object* x_813; uint8_t x_814; lean_object* x_815; lean_object* x_816; lean_object* x_817; lean_object* x_818; lean_object* x_819; lean_object* x_820; lean_object* x_821; lean_object* x_822; lean_object* x_823; lean_object* x_824; lean_object* x_825; lean_object* x_826; +x_812 = lean_ctor_get(x_810, 0); +x_813 = lean_ctor_get(x_2, 5); +lean_inc(x_813); +x_814 = 0; +x_815 = l_Lean_SourceInfo_fromRef(x_813, x_814); +lean_dec(x_813); +x_816 = lean_ctor_get(x_2, 2); +lean_inc(x_816); +x_817 = lean_ctor_get(x_2, 1); +lean_inc(x_817); +lean_dec(x_2); +x_818 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130; +x_819 = l_Lean_addMacroScope(x_817, x_818, x_816); +x_820 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128; +x_821 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134; +lean_inc(x_815); +x_822 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_822, 0, x_815); +lean_ctor_set(x_822, 1, x_820); +lean_ctor_set(x_822, 2, x_819); +lean_ctor_set(x_822, 3, x_821); +x_823 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_815); +x_824 = l_Lean_Syntax_node1(x_815, x_823, x_812); +x_825 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_826 = l_Lean_Syntax_node2(x_815, x_825, x_822, x_824); +lean_ctor_set(x_810, 0, x_826); +return x_810; +} +else +{ +lean_object* x_827; lean_object* x_828; lean_object* x_829; uint8_t x_830; lean_object* x_831; lean_object* x_832; lean_object* x_833; lean_object* x_834; lean_object* x_835; lean_object* x_836; lean_object* x_837; lean_object* x_838; lean_object* x_839; lean_object* x_840; lean_object* x_841; lean_object* x_842; lean_object* x_843; +x_827 = lean_ctor_get(x_810, 0); +x_828 = lean_ctor_get(x_810, 1); +lean_inc(x_828); +lean_inc(x_827); +lean_dec(x_810); +x_829 = lean_ctor_get(x_2, 5); +lean_inc(x_829); +x_830 = 0; +x_831 = l_Lean_SourceInfo_fromRef(x_829, x_830); +lean_dec(x_829); +x_832 = lean_ctor_get(x_2, 2); +lean_inc(x_832); +x_833 = lean_ctor_get(x_2, 1); +lean_inc(x_833); +lean_dec(x_2); +x_834 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130; +x_835 = l_Lean_addMacroScope(x_833, x_834, x_832); +x_836 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128; +x_837 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134; +lean_inc(x_831); +x_838 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_838, 0, x_831); +lean_ctor_set(x_838, 1, x_836); +lean_ctor_set(x_838, 2, x_835); +lean_ctor_set(x_838, 3, x_837); +x_839 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_831); +x_840 = l_Lean_Syntax_node1(x_831, x_839, x_827); +x_841 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_842 = l_Lean_Syntax_node2(x_831, x_841, x_838, x_840); +x_843 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_843, 0, x_842); +lean_ctor_set(x_843, 1, x_828); +return x_843; +} +} +case 15: +{ +lean_object* x_844; lean_object* x_845; uint8_t x_846; +x_844 = lean_ctor_get(x_1, 0); +lean_inc(x_844); +lean_dec(x_1); +lean_inc(x_2); +x_845 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_844, x_2, x_3); +x_846 = !lean_is_exclusive(x_845); +if (x_846 == 0) +{ +lean_object* x_847; lean_object* x_848; uint8_t x_849; lean_object* x_850; lean_object* x_851; lean_object* x_852; lean_object* x_853; lean_object* x_854; lean_object* x_855; lean_object* x_856; lean_object* x_857; lean_object* x_858; lean_object* x_859; lean_object* x_860; lean_object* x_861; +x_847 = lean_ctor_get(x_845, 0); +x_848 = lean_ctor_get(x_2, 5); +lean_inc(x_848); +x_849 = 0; +x_850 = l_Lean_SourceInfo_fromRef(x_848, x_849); +lean_dec(x_848); +x_851 = lean_ctor_get(x_2, 2); +lean_inc(x_851); +x_852 = lean_ctor_get(x_2, 1); +lean_inc(x_852); +lean_dec(x_2); +x_853 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138; +x_854 = l_Lean_addMacroScope(x_852, x_853, x_851); +x_855 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136; +x_856 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142; +lean_inc(x_850); +x_857 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_857, 0, x_850); +lean_ctor_set(x_857, 1, x_855); +lean_ctor_set(x_857, 2, x_854); +lean_ctor_set(x_857, 3, x_856); +x_858 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_850); +x_859 = l_Lean_Syntax_node1(x_850, x_858, x_847); +x_860 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_861 = l_Lean_Syntax_node2(x_850, x_860, x_857, x_859); +lean_ctor_set(x_845, 0, x_861); +return x_845; +} +else +{ +lean_object* x_862; lean_object* x_863; lean_object* x_864; uint8_t x_865; lean_object* x_866; lean_object* x_867; lean_object* x_868; lean_object* x_869; lean_object* x_870; lean_object* x_871; lean_object* x_872; lean_object* x_873; lean_object* x_874; lean_object* x_875; lean_object* x_876; lean_object* x_877; lean_object* x_878; +x_862 = lean_ctor_get(x_845, 0); +x_863 = lean_ctor_get(x_845, 1); +lean_inc(x_863); +lean_inc(x_862); +lean_dec(x_845); +x_864 = lean_ctor_get(x_2, 5); +lean_inc(x_864); +x_865 = 0; +x_866 = l_Lean_SourceInfo_fromRef(x_864, x_865); +lean_dec(x_864); +x_867 = lean_ctor_get(x_2, 2); +lean_inc(x_867); +x_868 = lean_ctor_get(x_2, 1); +lean_inc(x_868); +lean_dec(x_2); +x_869 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138; +x_870 = l_Lean_addMacroScope(x_868, x_869, x_867); +x_871 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136; +x_872 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142; +lean_inc(x_866); +x_873 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_873, 0, x_866); +lean_ctor_set(x_873, 1, x_871); +lean_ctor_set(x_873, 2, x_870); +lean_ctor_set(x_873, 3, x_872); +x_874 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_866); +x_875 = l_Lean_Syntax_node1(x_866, x_874, x_862); +x_876 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_877 = l_Lean_Syntax_node2(x_866, x_876, x_873, x_875); +x_878 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_878, 0, x_877); +lean_ctor_set(x_878, 1, x_863); +return x_878; +} +} +case 16: +{ +lean_object* x_879; lean_object* x_880; uint8_t x_881; +x_879 = lean_ctor_get(x_1, 0); +lean_inc(x_879); +lean_dec(x_1); +lean_inc(x_2); +x_880 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_879, x_2, x_3); +x_881 = !lean_is_exclusive(x_880); +if (x_881 == 0) +{ +lean_object* x_882; lean_object* x_883; uint8_t x_884; lean_object* x_885; lean_object* x_886; lean_object* x_887; lean_object* x_888; lean_object* x_889; lean_object* x_890; lean_object* x_891; lean_object* x_892; lean_object* x_893; lean_object* x_894; lean_object* x_895; lean_object* x_896; +x_882 = lean_ctor_get(x_880, 0); +x_883 = lean_ctor_get(x_2, 5); +lean_inc(x_883); +x_884 = 0; +x_885 = l_Lean_SourceInfo_fromRef(x_883, x_884); +lean_dec(x_883); +x_886 = lean_ctor_get(x_2, 2); +lean_inc(x_886); +x_887 = lean_ctor_get(x_2, 1); +lean_inc(x_887); +lean_dec(x_2); +x_888 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146; +x_889 = l_Lean_addMacroScope(x_887, x_888, x_886); +x_890 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144; +x_891 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150; +lean_inc(x_885); +x_892 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_892, 0, x_885); +lean_ctor_set(x_892, 1, x_890); +lean_ctor_set(x_892, 2, x_889); +lean_ctor_set(x_892, 3, x_891); +x_893 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_885); +x_894 = l_Lean_Syntax_node1(x_885, x_893, x_882); +x_895 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_896 = l_Lean_Syntax_node2(x_885, x_895, x_892, x_894); +lean_ctor_set(x_880, 0, x_896); +return x_880; +} +else +{ +lean_object* x_897; lean_object* x_898; lean_object* x_899; uint8_t x_900; lean_object* x_901; lean_object* x_902; lean_object* x_903; lean_object* x_904; lean_object* x_905; lean_object* x_906; lean_object* x_907; lean_object* x_908; lean_object* x_909; lean_object* x_910; lean_object* x_911; lean_object* x_912; lean_object* x_913; +x_897 = lean_ctor_get(x_880, 0); +x_898 = lean_ctor_get(x_880, 1); +lean_inc(x_898); +lean_inc(x_897); +lean_dec(x_880); +x_899 = lean_ctor_get(x_2, 5); +lean_inc(x_899); +x_900 = 0; +x_901 = l_Lean_SourceInfo_fromRef(x_899, x_900); +lean_dec(x_899); +x_902 = lean_ctor_get(x_2, 2); +lean_inc(x_902); +x_903 = lean_ctor_get(x_2, 1); +lean_inc(x_903); +lean_dec(x_2); +x_904 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146; +x_905 = l_Lean_addMacroScope(x_903, x_904, x_902); +x_906 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144; +x_907 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150; +lean_inc(x_901); +x_908 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_908, 0, x_901); +lean_ctor_set(x_908, 1, x_906); +lean_ctor_set(x_908, 2, x_905); +lean_ctor_set(x_908, 3, x_907); +x_909 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_901); +x_910 = l_Lean_Syntax_node1(x_901, x_909, x_897); +x_911 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_912 = l_Lean_Syntax_node2(x_901, x_911, x_908, x_910); +x_913 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_913, 0, x_912); +lean_ctor_set(x_913, 1, x_898); +return x_913; +} +} +case 17: +{ +lean_object* x_914; lean_object* x_915; uint8_t x_916; +x_914 = lean_ctor_get(x_1, 0); +lean_inc(x_914); +lean_dec(x_1); +lean_inc(x_2); +x_915 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_914, x_2, x_3); +x_916 = !lean_is_exclusive(x_915); +if (x_916 == 0) +{ +lean_object* x_917; lean_object* x_918; uint8_t x_919; lean_object* x_920; lean_object* x_921; lean_object* x_922; lean_object* x_923; lean_object* x_924; lean_object* x_925; lean_object* x_926; lean_object* x_927; lean_object* x_928; lean_object* x_929; lean_object* x_930; lean_object* x_931; +x_917 = lean_ctor_get(x_915, 0); +x_918 = lean_ctor_get(x_2, 5); +lean_inc(x_918); +x_919 = 0; +x_920 = l_Lean_SourceInfo_fromRef(x_918, x_919); +lean_dec(x_918); +x_921 = lean_ctor_get(x_2, 2); +lean_inc(x_921); +x_922 = lean_ctor_get(x_2, 1); +lean_inc(x_922); +lean_dec(x_2); +x_923 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154; +x_924 = l_Lean_addMacroScope(x_922, x_923, x_921); +x_925 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152; +x_926 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158; +lean_inc(x_920); +x_927 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_927, 0, x_920); +lean_ctor_set(x_927, 1, x_925); +lean_ctor_set(x_927, 2, x_924); +lean_ctor_set(x_927, 3, x_926); +x_928 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_920); +x_929 = l_Lean_Syntax_node1(x_920, x_928, x_917); +x_930 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_931 = l_Lean_Syntax_node2(x_920, x_930, x_927, x_929); +lean_ctor_set(x_915, 0, x_931); +return x_915; +} +else +{ +lean_object* x_932; lean_object* x_933; lean_object* x_934; uint8_t x_935; lean_object* x_936; lean_object* x_937; lean_object* x_938; lean_object* x_939; lean_object* x_940; lean_object* x_941; lean_object* x_942; lean_object* x_943; lean_object* x_944; lean_object* x_945; lean_object* x_946; lean_object* x_947; lean_object* x_948; +x_932 = lean_ctor_get(x_915, 0); +x_933 = lean_ctor_get(x_915, 1); +lean_inc(x_933); +lean_inc(x_932); +lean_dec(x_915); +x_934 = lean_ctor_get(x_2, 5); +lean_inc(x_934); +x_935 = 0; +x_936 = l_Lean_SourceInfo_fromRef(x_934, x_935); +lean_dec(x_934); +x_937 = lean_ctor_get(x_2, 2); +lean_inc(x_937); +x_938 = lean_ctor_get(x_2, 1); +lean_inc(x_938); +lean_dec(x_2); +x_939 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154; +x_940 = l_Lean_addMacroScope(x_938, x_939, x_937); +x_941 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152; +x_942 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158; +lean_inc(x_936); +x_943 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_943, 0, x_936); +lean_ctor_set(x_943, 1, x_941); +lean_ctor_set(x_943, 2, x_940); +lean_ctor_set(x_943, 3, x_942); +x_944 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_936); +x_945 = l_Lean_Syntax_node1(x_936, x_944, x_932); +x_946 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_947 = l_Lean_Syntax_node2(x_936, x_946, x_943, x_945); +x_948 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_948, 0, x_947); +lean_ctor_set(x_948, 1, x_933); +return x_948; +} +} +case 18: +{ +lean_object* x_949; lean_object* x_950; uint8_t x_951; +x_949 = lean_ctor_get(x_1, 0); +lean_inc(x_949); +lean_dec(x_1); +lean_inc(x_2); +x_950 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_949, x_2, x_3); +x_951 = !lean_is_exclusive(x_950); +if (x_951 == 0) +{ +lean_object* x_952; lean_object* x_953; uint8_t x_954; lean_object* x_955; lean_object* x_956; lean_object* x_957; lean_object* x_958; lean_object* x_959; lean_object* x_960; lean_object* x_961; lean_object* x_962; lean_object* x_963; lean_object* x_964; lean_object* x_965; lean_object* x_966; +x_952 = lean_ctor_get(x_950, 0); +x_953 = lean_ctor_get(x_2, 5); +lean_inc(x_953); +x_954 = 0; +x_955 = l_Lean_SourceInfo_fromRef(x_953, x_954); +lean_dec(x_953); +x_956 = lean_ctor_get(x_2, 2); +lean_inc(x_956); +x_957 = lean_ctor_get(x_2, 1); +lean_inc(x_957); +lean_dec(x_2); +x_958 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162; +x_959 = l_Lean_addMacroScope(x_957, x_958, x_956); +x_960 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160; +x_961 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166; +lean_inc(x_955); +x_962 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_962, 0, x_955); +lean_ctor_set(x_962, 1, x_960); +lean_ctor_set(x_962, 2, x_959); +lean_ctor_set(x_962, 3, x_961); +x_963 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_955); +x_964 = l_Lean_Syntax_node1(x_955, x_963, x_952); +x_965 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_966 = l_Lean_Syntax_node2(x_955, x_965, x_962, x_964); +lean_ctor_set(x_950, 0, x_966); +return x_950; +} +else +{ +lean_object* x_967; lean_object* x_968; lean_object* x_969; uint8_t x_970; lean_object* x_971; lean_object* x_972; lean_object* x_973; lean_object* x_974; lean_object* x_975; lean_object* x_976; lean_object* x_977; lean_object* x_978; lean_object* x_979; lean_object* x_980; lean_object* x_981; lean_object* x_982; lean_object* x_983; +x_967 = lean_ctor_get(x_950, 0); +x_968 = lean_ctor_get(x_950, 1); +lean_inc(x_968); +lean_inc(x_967); +lean_dec(x_950); +x_969 = lean_ctor_get(x_2, 5); +lean_inc(x_969); +x_970 = 0; +x_971 = l_Lean_SourceInfo_fromRef(x_969, x_970); +lean_dec(x_969); +x_972 = lean_ctor_get(x_2, 2); +lean_inc(x_972); +x_973 = lean_ctor_get(x_2, 1); +lean_inc(x_973); +lean_dec(x_2); +x_974 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162; +x_975 = l_Lean_addMacroScope(x_973, x_974, x_972); +x_976 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160; +x_977 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166; +lean_inc(x_971); +x_978 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_978, 0, x_971); +lean_ctor_set(x_978, 1, x_976); +lean_ctor_set(x_978, 2, x_975); +lean_ctor_set(x_978, 3, x_977); +x_979 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_971); +x_980 = l_Lean_Syntax_node1(x_971, x_979, x_967); +x_981 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_982 = l_Lean_Syntax_node2(x_971, x_981, x_978, x_980); +x_983 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_983, 0, x_982); +lean_ctor_set(x_983, 1, x_968); +return x_983; +} +} +case 19: +{ +lean_object* x_984; lean_object* x_985; uint8_t x_986; +x_984 = lean_ctor_get(x_1, 0); +lean_inc(x_984); +lean_dec(x_1); +lean_inc(x_2); +x_985 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction(x_984, x_2, x_3); +x_986 = !lean_is_exclusive(x_985); +if (x_986 == 0) +{ +lean_object* x_987; lean_object* x_988; uint8_t x_989; lean_object* x_990; lean_object* x_991; lean_object* x_992; lean_object* x_993; lean_object* x_994; lean_object* x_995; lean_object* x_996; lean_object* x_997; lean_object* x_998; lean_object* x_999; lean_object* x_1000; lean_object* x_1001; +x_987 = lean_ctor_get(x_985, 0); +x_988 = lean_ctor_get(x_2, 5); +lean_inc(x_988); +x_989 = 0; +x_990 = l_Lean_SourceInfo_fromRef(x_988, x_989); +lean_dec(x_988); +x_991 = lean_ctor_get(x_2, 2); +lean_inc(x_991); +x_992 = lean_ctor_get(x_2, 1); +lean_inc(x_992); +lean_dec(x_2); +x_993 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170; +x_994 = l_Lean_addMacroScope(x_992, x_993, x_991); +x_995 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168; +x_996 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174; +lean_inc(x_990); +x_997 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_997, 0, x_990); +lean_ctor_set(x_997, 1, x_995); +lean_ctor_set(x_997, 2, x_994); +lean_ctor_set(x_997, 3, x_996); +x_998 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_990); +x_999 = l_Lean_Syntax_node1(x_990, x_998, x_987); +x_1000 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1001 = l_Lean_Syntax_node2(x_990, x_1000, x_997, x_999); +lean_ctor_set(x_985, 0, x_1001); +return x_985; +} +else +{ +lean_object* x_1002; lean_object* x_1003; lean_object* x_1004; uint8_t x_1005; lean_object* x_1006; lean_object* x_1007; lean_object* x_1008; lean_object* x_1009; lean_object* x_1010; lean_object* x_1011; lean_object* x_1012; lean_object* x_1013; lean_object* x_1014; lean_object* x_1015; lean_object* x_1016; lean_object* x_1017; lean_object* x_1018; +x_1002 = lean_ctor_get(x_985, 0); +x_1003 = lean_ctor_get(x_985, 1); +lean_inc(x_1003); +lean_inc(x_1002); +lean_dec(x_985); +x_1004 = lean_ctor_get(x_2, 5); +lean_inc(x_1004); +x_1005 = 0; +x_1006 = l_Lean_SourceInfo_fromRef(x_1004, x_1005); +lean_dec(x_1004); +x_1007 = lean_ctor_get(x_2, 2); +lean_inc(x_1007); +x_1008 = lean_ctor_get(x_2, 1); +lean_inc(x_1008); +lean_dec(x_2); +x_1009 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170; +x_1010 = l_Lean_addMacroScope(x_1008, x_1009, x_1007); +x_1011 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168; +x_1012 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174; +lean_inc(x_1006); +x_1013 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1013, 0, x_1006); +lean_ctor_set(x_1013, 1, x_1011); +lean_ctor_set(x_1013, 2, x_1010); +lean_ctor_set(x_1013, 3, x_1012); +x_1014 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1006); +x_1015 = l_Lean_Syntax_node1(x_1006, x_1014, x_1002); +x_1016 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1017 = l_Lean_Syntax_node2(x_1006, x_1016, x_1013, x_1015); +x_1018 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1018, 0, x_1017); +lean_ctor_set(x_1018, 1, x_1003); +return x_1018; +} +} +case 20: +{ +lean_object* x_1019; lean_object* x_1020; uint8_t x_1021; +x_1019 = lean_ctor_get(x_1, 0); +lean_inc(x_1019); +lean_dec(x_1); +lean_inc(x_2); +x_1020 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_1019, x_2, x_3); +x_1021 = !lean_is_exclusive(x_1020); +if (x_1021 == 0) +{ +lean_object* x_1022; lean_object* x_1023; uint8_t x_1024; lean_object* x_1025; lean_object* x_1026; lean_object* x_1027; lean_object* x_1028; lean_object* x_1029; lean_object* x_1030; lean_object* x_1031; lean_object* x_1032; lean_object* x_1033; lean_object* x_1034; lean_object* x_1035; lean_object* x_1036; +x_1022 = lean_ctor_get(x_1020, 0); +x_1023 = lean_ctor_get(x_2, 5); +lean_inc(x_1023); +x_1024 = 0; +x_1025 = l_Lean_SourceInfo_fromRef(x_1023, x_1024); +lean_dec(x_1023); +x_1026 = lean_ctor_get(x_2, 2); +lean_inc(x_1026); +x_1027 = lean_ctor_get(x_2, 1); +lean_inc(x_1027); +lean_dec(x_2); +x_1028 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178; +x_1029 = l_Lean_addMacroScope(x_1027, x_1028, x_1026); +x_1030 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176; +x_1031 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182; +lean_inc(x_1025); +x_1032 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1032, 0, x_1025); +lean_ctor_set(x_1032, 1, x_1030); +lean_ctor_set(x_1032, 2, x_1029); +lean_ctor_set(x_1032, 3, x_1031); +x_1033 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1025); +x_1034 = l_Lean_Syntax_node1(x_1025, x_1033, x_1022); +x_1035 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1036 = l_Lean_Syntax_node2(x_1025, x_1035, x_1032, x_1034); +lean_ctor_set(x_1020, 0, x_1036); +return x_1020; +} +else +{ +lean_object* x_1037; lean_object* x_1038; lean_object* x_1039; uint8_t x_1040; lean_object* x_1041; lean_object* x_1042; lean_object* x_1043; lean_object* x_1044; lean_object* x_1045; lean_object* x_1046; lean_object* x_1047; lean_object* x_1048; lean_object* x_1049; lean_object* x_1050; lean_object* x_1051; lean_object* x_1052; lean_object* x_1053; +x_1037 = lean_ctor_get(x_1020, 0); +x_1038 = lean_ctor_get(x_1020, 1); +lean_inc(x_1038); +lean_inc(x_1037); +lean_dec(x_1020); +x_1039 = lean_ctor_get(x_2, 5); +lean_inc(x_1039); +x_1040 = 0; +x_1041 = l_Lean_SourceInfo_fromRef(x_1039, x_1040); +lean_dec(x_1039); +x_1042 = lean_ctor_get(x_2, 2); +lean_inc(x_1042); +x_1043 = lean_ctor_get(x_2, 1); +lean_inc(x_1043); +lean_dec(x_2); +x_1044 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178; +x_1045 = l_Lean_addMacroScope(x_1043, x_1044, x_1042); +x_1046 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176; +x_1047 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182; +lean_inc(x_1041); +x_1048 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1048, 0, x_1041); +lean_ctor_set(x_1048, 1, x_1046); +lean_ctor_set(x_1048, 2, x_1045); +lean_ctor_set(x_1048, 3, x_1047); +x_1049 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1041); +x_1050 = l_Lean_Syntax_node1(x_1041, x_1049, x_1037); +x_1051 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1052 = l_Lean_Syntax_node2(x_1041, x_1051, x_1048, x_1050); +x_1053 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1053, 0, x_1052); +lean_ctor_set(x_1053, 1, x_1038); +return x_1053; +} +} +case 21: +{ +lean_object* x_1054; lean_object* x_1055; uint8_t x_1056; +x_1054 = lean_ctor_get(x_1, 0); +lean_inc(x_1054); +lean_dec(x_1); +lean_inc(x_2); +x_1055 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_1054, x_2, x_3); +x_1056 = !lean_is_exclusive(x_1055); +if (x_1056 == 0) +{ +lean_object* x_1057; lean_object* x_1058; uint8_t x_1059; lean_object* x_1060; lean_object* x_1061; lean_object* x_1062; lean_object* x_1063; lean_object* x_1064; lean_object* x_1065; lean_object* x_1066; lean_object* x_1067; lean_object* x_1068; lean_object* x_1069; lean_object* x_1070; lean_object* x_1071; +x_1057 = lean_ctor_get(x_1055, 0); +x_1058 = lean_ctor_get(x_2, 5); +lean_inc(x_1058); +x_1059 = 0; +x_1060 = l_Lean_SourceInfo_fromRef(x_1058, x_1059); +lean_dec(x_1058); +x_1061 = lean_ctor_get(x_2, 2); +lean_inc(x_1061); +x_1062 = lean_ctor_get(x_2, 1); +lean_inc(x_1062); +lean_dec(x_2); +x_1063 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186; +x_1064 = l_Lean_addMacroScope(x_1062, x_1063, x_1061); +x_1065 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184; +x_1066 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190; +lean_inc(x_1060); +x_1067 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1067, 0, x_1060); +lean_ctor_set(x_1067, 1, x_1065); +lean_ctor_set(x_1067, 2, x_1064); +lean_ctor_set(x_1067, 3, x_1066); +x_1068 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1060); +x_1069 = l_Lean_Syntax_node1(x_1060, x_1068, x_1057); +x_1070 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1071 = l_Lean_Syntax_node2(x_1060, x_1070, x_1067, x_1069); +lean_ctor_set(x_1055, 0, x_1071); +return x_1055; +} +else +{ +lean_object* x_1072; lean_object* x_1073; lean_object* x_1074; uint8_t x_1075; lean_object* x_1076; lean_object* x_1077; lean_object* x_1078; lean_object* x_1079; lean_object* x_1080; lean_object* x_1081; lean_object* x_1082; lean_object* x_1083; lean_object* x_1084; lean_object* x_1085; lean_object* x_1086; lean_object* x_1087; lean_object* x_1088; +x_1072 = lean_ctor_get(x_1055, 0); +x_1073 = lean_ctor_get(x_1055, 1); +lean_inc(x_1073); +lean_inc(x_1072); +lean_dec(x_1055); +x_1074 = lean_ctor_get(x_2, 5); +lean_inc(x_1074); +x_1075 = 0; +x_1076 = l_Lean_SourceInfo_fromRef(x_1074, x_1075); +lean_dec(x_1074); +x_1077 = lean_ctor_get(x_2, 2); +lean_inc(x_1077); +x_1078 = lean_ctor_get(x_2, 1); +lean_inc(x_1078); +lean_dec(x_2); +x_1079 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186; +x_1080 = l_Lean_addMacroScope(x_1078, x_1079, x_1077); +x_1081 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184; +x_1082 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190; +lean_inc(x_1076); +x_1083 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1083, 0, x_1076); +lean_ctor_set(x_1083, 1, x_1081); +lean_ctor_set(x_1083, 2, x_1080); +lean_ctor_set(x_1083, 3, x_1082); +x_1084 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1076); +x_1085 = l_Lean_Syntax_node1(x_1076, x_1084, x_1072); +x_1086 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1087 = l_Lean_Syntax_node2(x_1076, x_1086, x_1083, x_1085); +x_1088 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1088, 0, x_1087); +lean_ctor_set(x_1088, 1, x_1073); +return x_1088; +} +} +case 22: +{ +lean_object* x_1089; lean_object* x_1090; uint8_t x_1091; +x_1089 = lean_ctor_get(x_1, 0); +lean_inc(x_1089); +lean_dec(x_1); +lean_inc(x_2); +x_1090 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber(x_1089, x_2, x_3); +x_1091 = !lean_is_exclusive(x_1090); +if (x_1091 == 0) +{ +lean_object* x_1092; lean_object* x_1093; uint8_t x_1094; lean_object* x_1095; lean_object* x_1096; lean_object* x_1097; lean_object* x_1098; lean_object* x_1099; lean_object* x_1100; lean_object* x_1101; lean_object* x_1102; lean_object* x_1103; lean_object* x_1104; lean_object* x_1105; lean_object* x_1106; +x_1092 = lean_ctor_get(x_1090, 0); +x_1093 = lean_ctor_get(x_2, 5); +lean_inc(x_1093); +x_1094 = 0; +x_1095 = l_Lean_SourceInfo_fromRef(x_1093, x_1094); +lean_dec(x_1093); +x_1096 = lean_ctor_get(x_2, 2); +lean_inc(x_1096); +x_1097 = lean_ctor_get(x_2, 1); +lean_inc(x_1097); +lean_dec(x_2); +x_1098 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194; +x_1099 = l_Lean_addMacroScope(x_1097, x_1098, x_1096); +x_1100 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192; +x_1101 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198; +lean_inc(x_1095); +x_1102 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1102, 0, x_1095); +lean_ctor_set(x_1102, 1, x_1100); +lean_ctor_set(x_1102, 2, x_1099); +lean_ctor_set(x_1102, 3, x_1101); +x_1103 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1095); +x_1104 = l_Lean_Syntax_node1(x_1095, x_1103, x_1092); +x_1105 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1106 = l_Lean_Syntax_node2(x_1095, x_1105, x_1102, x_1104); +lean_ctor_set(x_1090, 0, x_1106); +return x_1090; +} +else +{ +lean_object* x_1107; lean_object* x_1108; lean_object* x_1109; uint8_t x_1110; lean_object* x_1111; lean_object* x_1112; lean_object* x_1113; lean_object* x_1114; lean_object* x_1115; lean_object* x_1116; lean_object* x_1117; lean_object* x_1118; lean_object* x_1119; lean_object* x_1120; lean_object* x_1121; lean_object* x_1122; lean_object* x_1123; +x_1107 = lean_ctor_get(x_1090, 0); +x_1108 = lean_ctor_get(x_1090, 1); +lean_inc(x_1108); +lean_inc(x_1107); +lean_dec(x_1090); +x_1109 = lean_ctor_get(x_2, 5); +lean_inc(x_1109); +x_1110 = 0; +x_1111 = l_Lean_SourceInfo_fromRef(x_1109, x_1110); +lean_dec(x_1109); +x_1112 = lean_ctor_get(x_2, 2); +lean_inc(x_1112); +x_1113 = lean_ctor_get(x_2, 1); +lean_inc(x_1113); +lean_dec(x_2); +x_1114 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194; +x_1115 = l_Lean_addMacroScope(x_1113, x_1114, x_1112); +x_1116 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192; +x_1117 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198; +lean_inc(x_1111); +x_1118 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1118, 0, x_1111); +lean_ctor_set(x_1118, 1, x_1116); +lean_ctor_set(x_1118, 2, x_1115); +lean_ctor_set(x_1118, 3, x_1117); +x_1119 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1111); +x_1120 = l_Lean_Syntax_node1(x_1111, x_1119, x_1107); +x_1121 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1122 = l_Lean_Syntax_node2(x_1111, x_1121, x_1118, x_1120); +x_1123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1123, 0, x_1122); +lean_ctor_set(x_1123, 1, x_1108); +return x_1123; +} +} +case 23: +{ +lean_object* x_1124; uint8_t x_1125; lean_object* x_1126; lean_object* x_1127; lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; +x_1124 = lean_ctor_get(x_2, 5); +lean_inc(x_1124); +x_1125 = 0; +x_1126 = l_Lean_SourceInfo_fromRef(x_1124, x_1125); +lean_dec(x_1124); +x_1127 = lean_ctor_get(x_2, 2); +lean_inc(x_1127); +x_1128 = lean_ctor_get(x_2, 1); +lean_inc(x_1128); +lean_dec(x_2); +x_1129 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202; +x_1130 = l_Lean_addMacroScope(x_1128, x_1129, x_1127); +x_1131 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200; +x_1132 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206; +x_1133 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1133, 0, x_1126); +lean_ctor_set(x_1133, 1, x_1131); +lean_ctor_set(x_1133, 2, x_1130); +lean_ctor_set(x_1133, 3, x_1132); +x_1134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1134, 0, x_1133); +lean_ctor_set(x_1134, 1, x_3); +return x_1134; +} +case 24: +{ +uint8_t x_1135; lean_object* x_1136; uint8_t x_1137; +x_1135 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1136 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName(x_1135, x_2, x_3); +x_1137 = !lean_is_exclusive(x_1136); +if (x_1137 == 0) +{ +lean_object* x_1138; lean_object* x_1139; uint8_t x_1140; lean_object* x_1141; lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; +x_1138 = lean_ctor_get(x_1136, 0); +x_1139 = lean_ctor_get(x_2, 5); +lean_inc(x_1139); +x_1140 = 0; +x_1141 = l_Lean_SourceInfo_fromRef(x_1139, x_1140); +lean_dec(x_1139); +x_1142 = lean_ctor_get(x_2, 2); +lean_inc(x_1142); +x_1143 = lean_ctor_get(x_2, 1); +lean_inc(x_1143); +lean_dec(x_2); +x_1144 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210; +x_1145 = l_Lean_addMacroScope(x_1143, x_1144, x_1142); +x_1146 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208; +x_1147 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214; +lean_inc(x_1141); +x_1148 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1148, 0, x_1141); +lean_ctor_set(x_1148, 1, x_1146); +lean_ctor_set(x_1148, 2, x_1145); +lean_ctor_set(x_1148, 3, x_1147); +x_1149 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1141); +x_1150 = l_Lean_Syntax_node1(x_1141, x_1149, x_1138); +x_1151 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1152 = l_Lean_Syntax_node2(x_1141, x_1151, x_1148, x_1150); +lean_ctor_set(x_1136, 0, x_1152); +return x_1136; +} +else +{ +lean_object* x_1153; lean_object* x_1154; lean_object* x_1155; uint8_t x_1156; lean_object* x_1157; lean_object* x_1158; lean_object* x_1159; lean_object* x_1160; lean_object* x_1161; lean_object* x_1162; lean_object* x_1163; lean_object* x_1164; lean_object* x_1165; lean_object* x_1166; lean_object* x_1167; lean_object* x_1168; lean_object* x_1169; +x_1153 = lean_ctor_get(x_1136, 0); +x_1154 = lean_ctor_get(x_1136, 1); +lean_inc(x_1154); +lean_inc(x_1153); +lean_dec(x_1136); +x_1155 = lean_ctor_get(x_2, 5); +lean_inc(x_1155); +x_1156 = 0; +x_1157 = l_Lean_SourceInfo_fromRef(x_1155, x_1156); +lean_dec(x_1155); +x_1158 = lean_ctor_get(x_2, 2); +lean_inc(x_1158); +x_1159 = lean_ctor_get(x_2, 1); +lean_inc(x_1159); +lean_dec(x_2); +x_1160 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210; +x_1161 = l_Lean_addMacroScope(x_1159, x_1160, x_1158); +x_1162 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208; +x_1163 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214; +lean_inc(x_1157); +x_1164 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1164, 0, x_1157); +lean_ctor_set(x_1164, 1, x_1162); +lean_ctor_set(x_1164, 2, x_1161); +lean_ctor_set(x_1164, 3, x_1163); +x_1165 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1157); +x_1166 = l_Lean_Syntax_node1(x_1157, x_1165, x_1153); +x_1167 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1168 = l_Lean_Syntax_node2(x_1157, x_1167, x_1164, x_1166); +x_1169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1169, 0, x_1168); +lean_ctor_set(x_1169, 1, x_1154); +return x_1169; +} +} +case 25: +{ +uint8_t x_1170; lean_object* x_1171; uint8_t x_1172; +x_1170 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1171 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO(x_1170, x_2, x_3); +x_1172 = !lean_is_exclusive(x_1171); +if (x_1172 == 0) +{ +lean_object* x_1173; lean_object* x_1174; uint8_t x_1175; lean_object* x_1176; lean_object* x_1177; lean_object* x_1178; lean_object* x_1179; lean_object* x_1180; lean_object* x_1181; lean_object* x_1182; lean_object* x_1183; lean_object* x_1184; lean_object* x_1185; lean_object* x_1186; lean_object* x_1187; +x_1173 = lean_ctor_get(x_1171, 0); +x_1174 = lean_ctor_get(x_2, 5); +lean_inc(x_1174); +x_1175 = 0; +x_1176 = l_Lean_SourceInfo_fromRef(x_1174, x_1175); +lean_dec(x_1174); +x_1177 = lean_ctor_get(x_2, 2); +lean_inc(x_1177); +x_1178 = lean_ctor_get(x_2, 1); +lean_inc(x_1178); +lean_dec(x_2); +x_1179 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218; +x_1180 = l_Lean_addMacroScope(x_1178, x_1179, x_1177); +x_1181 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216; +x_1182 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222; +lean_inc(x_1176); +x_1183 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1183, 0, x_1176); +lean_ctor_set(x_1183, 1, x_1181); +lean_ctor_set(x_1183, 2, x_1180); +lean_ctor_set(x_1183, 3, x_1182); +x_1184 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1176); +x_1185 = l_Lean_Syntax_node1(x_1176, x_1184, x_1173); +x_1186 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1187 = l_Lean_Syntax_node2(x_1176, x_1186, x_1183, x_1185); +lean_ctor_set(x_1171, 0, x_1187); +return x_1171; +} +else +{ +lean_object* x_1188; lean_object* x_1189; lean_object* x_1190; uint8_t x_1191; lean_object* x_1192; lean_object* x_1193; lean_object* x_1194; lean_object* x_1195; lean_object* x_1196; lean_object* x_1197; lean_object* x_1198; lean_object* x_1199; lean_object* x_1200; lean_object* x_1201; lean_object* x_1202; lean_object* x_1203; lean_object* x_1204; +x_1188 = lean_ctor_get(x_1171, 0); +x_1189 = lean_ctor_get(x_1171, 1); +lean_inc(x_1189); +lean_inc(x_1188); +lean_dec(x_1171); +x_1190 = lean_ctor_get(x_2, 5); +lean_inc(x_1190); +x_1191 = 0; +x_1192 = l_Lean_SourceInfo_fromRef(x_1190, x_1191); +lean_dec(x_1190); +x_1193 = lean_ctor_get(x_2, 2); +lean_inc(x_1193); +x_1194 = lean_ctor_get(x_2, 1); +lean_inc(x_1194); +lean_dec(x_2); +x_1195 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218; +x_1196 = l_Lean_addMacroScope(x_1194, x_1195, x_1193); +x_1197 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216; +x_1198 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222; +lean_inc(x_1192); +x_1199 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1199, 0, x_1192); +lean_ctor_set(x_1199, 1, x_1197); +lean_ctor_set(x_1199, 2, x_1196); +lean_ctor_set(x_1199, 3, x_1198); +x_1200 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1192); +x_1201 = l_Lean_Syntax_node1(x_1192, x_1200, x_1188); +x_1202 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1203 = l_Lean_Syntax_node2(x_1192, x_1202, x_1199, x_1201); +x_1204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1204, 0, x_1203); +lean_ctor_set(x_1204, 1, x_1189); +return x_1204; +} +} +case 26: +{ +uint8_t x_1205; lean_object* x_1206; uint8_t x_1207; +x_1205 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1206 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX(x_1205, x_2, x_3); +x_1207 = !lean_is_exclusive(x_1206); +if (x_1207 == 0) +{ +lean_object* x_1208; lean_object* x_1209; uint8_t x_1210; lean_object* x_1211; lean_object* x_1212; lean_object* x_1213; lean_object* x_1214; lean_object* x_1215; lean_object* x_1216; lean_object* x_1217; lean_object* x_1218; lean_object* x_1219; lean_object* x_1220; lean_object* x_1221; lean_object* x_1222; +x_1208 = lean_ctor_get(x_1206, 0); +x_1209 = lean_ctor_get(x_2, 5); +lean_inc(x_1209); +x_1210 = 0; +x_1211 = l_Lean_SourceInfo_fromRef(x_1209, x_1210); +lean_dec(x_1209); +x_1212 = lean_ctor_get(x_2, 2); +lean_inc(x_1212); +x_1213 = lean_ctor_get(x_2, 1); +lean_inc(x_1213); +lean_dec(x_2); +x_1214 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226; +x_1215 = l_Lean_addMacroScope(x_1213, x_1214, x_1212); +x_1216 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224; +x_1217 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230; +lean_inc(x_1211); +x_1218 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1218, 0, x_1211); +lean_ctor_set(x_1218, 1, x_1216); +lean_ctor_set(x_1218, 2, x_1215); +lean_ctor_set(x_1218, 3, x_1217); +x_1219 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1211); +x_1220 = l_Lean_Syntax_node1(x_1211, x_1219, x_1208); +x_1221 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1222 = l_Lean_Syntax_node2(x_1211, x_1221, x_1218, x_1220); +lean_ctor_set(x_1206, 0, x_1222); +return x_1206; +} +else +{ +lean_object* x_1223; lean_object* x_1224; lean_object* x_1225; uint8_t x_1226; lean_object* x_1227; lean_object* x_1228; lean_object* x_1229; lean_object* x_1230; lean_object* x_1231; lean_object* x_1232; lean_object* x_1233; lean_object* x_1234; lean_object* x_1235; lean_object* x_1236; lean_object* x_1237; lean_object* x_1238; lean_object* x_1239; +x_1223 = lean_ctor_get(x_1206, 0); +x_1224 = lean_ctor_get(x_1206, 1); +lean_inc(x_1224); +lean_inc(x_1223); +lean_dec(x_1206); +x_1225 = lean_ctor_get(x_2, 5); +lean_inc(x_1225); +x_1226 = 0; +x_1227 = l_Lean_SourceInfo_fromRef(x_1225, x_1226); +lean_dec(x_1225); +x_1228 = lean_ctor_get(x_2, 2); +lean_inc(x_1228); +x_1229 = lean_ctor_get(x_2, 1); +lean_inc(x_1229); +lean_dec(x_2); +x_1230 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226; +x_1231 = l_Lean_addMacroScope(x_1229, x_1230, x_1228); +x_1232 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224; +x_1233 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230; +lean_inc(x_1227); +x_1234 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1234, 0, x_1227); +lean_ctor_set(x_1234, 1, x_1232); +lean_ctor_set(x_1234, 2, x_1231); +lean_ctor_set(x_1234, 3, x_1233); +x_1235 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1227); +x_1236 = l_Lean_Syntax_node1(x_1227, x_1235, x_1223); +x_1237 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1238 = l_Lean_Syntax_node2(x_1227, x_1237, x_1234, x_1236); +x_1239 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1239, 0, x_1238); +lean_ctor_set(x_1239, 1, x_1224); +return x_1239; +} +} +case 27: +{ +uint8_t x_1240; lean_object* x_1241; uint8_t x_1242; +x_1240 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1241 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX(x_1240, x_2, x_3); +x_1242 = !lean_is_exclusive(x_1241); +if (x_1242 == 0) +{ +lean_object* x_1243; lean_object* x_1244; uint8_t x_1245; lean_object* x_1246; lean_object* x_1247; lean_object* x_1248; lean_object* x_1249; lean_object* x_1250; lean_object* x_1251; lean_object* x_1252; lean_object* x_1253; lean_object* x_1254; lean_object* x_1255; lean_object* x_1256; lean_object* x_1257; +x_1243 = lean_ctor_get(x_1241, 0); +x_1244 = lean_ctor_get(x_2, 5); +lean_inc(x_1244); +x_1245 = 0; +x_1246 = l_Lean_SourceInfo_fromRef(x_1244, x_1245); +lean_dec(x_1244); +x_1247 = lean_ctor_get(x_2, 2); +lean_inc(x_1247); +x_1248 = lean_ctor_get(x_2, 1); +lean_inc(x_1248); +lean_dec(x_2); +x_1249 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234; +x_1250 = l_Lean_addMacroScope(x_1248, x_1249, x_1247); +x_1251 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232; +x_1252 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238; +lean_inc(x_1246); +x_1253 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1253, 0, x_1246); +lean_ctor_set(x_1253, 1, x_1251); +lean_ctor_set(x_1253, 2, x_1250); +lean_ctor_set(x_1253, 3, x_1252); +x_1254 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1246); +x_1255 = l_Lean_Syntax_node1(x_1246, x_1254, x_1243); +x_1256 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1257 = l_Lean_Syntax_node2(x_1246, x_1256, x_1253, x_1255); +lean_ctor_set(x_1241, 0, x_1257); +return x_1241; +} +else +{ +lean_object* x_1258; lean_object* x_1259; lean_object* x_1260; uint8_t x_1261; lean_object* x_1262; lean_object* x_1263; lean_object* x_1264; lean_object* x_1265; lean_object* x_1266; lean_object* x_1267; lean_object* x_1268; lean_object* x_1269; lean_object* x_1270; lean_object* x_1271; lean_object* x_1272; lean_object* x_1273; lean_object* x_1274; +x_1258 = lean_ctor_get(x_1241, 0); +x_1259 = lean_ctor_get(x_1241, 1); +lean_inc(x_1259); +lean_inc(x_1258); +lean_dec(x_1241); +x_1260 = lean_ctor_get(x_2, 5); +lean_inc(x_1260); +x_1261 = 0; +x_1262 = l_Lean_SourceInfo_fromRef(x_1260, x_1261); +lean_dec(x_1260); +x_1263 = lean_ctor_get(x_2, 2); +lean_inc(x_1263); +x_1264 = lean_ctor_get(x_2, 1); +lean_inc(x_1264); +lean_dec(x_2); +x_1265 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234; +x_1266 = l_Lean_addMacroScope(x_1264, x_1265, x_1263); +x_1267 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232; +x_1268 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238; +lean_inc(x_1262); +x_1269 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1269, 0, x_1262); +lean_ctor_set(x_1269, 1, x_1267); +lean_ctor_set(x_1269, 2, x_1266); +lean_ctor_set(x_1269, 3, x_1268); +x_1270 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1262); +x_1271 = l_Lean_Syntax_node1(x_1262, x_1270, x_1258); +x_1272 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1273 = l_Lean_Syntax_node2(x_1262, x_1272, x_1269, x_1271); +x_1274 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1274, 0, x_1273); +lean_ctor_set(x_1274, 1, x_1259); +return x_1274; +} +} +default: +{ +uint8_t x_1275; lean_object* x_1276; uint8_t x_1277; +x_1275 = lean_ctor_get_uint8(x_1, 0); +lean_dec(x_1); +lean_inc(x_2); +x_1276 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ(x_1275, x_2, x_3); +x_1277 = !lean_is_exclusive(x_1276); +if (x_1277 == 0) +{ +lean_object* x_1278; lean_object* x_1279; uint8_t x_1280; lean_object* x_1281; lean_object* x_1282; lean_object* x_1283; lean_object* x_1284; lean_object* x_1285; lean_object* x_1286; lean_object* x_1287; lean_object* x_1288; lean_object* x_1289; lean_object* x_1290; lean_object* x_1291; lean_object* x_1292; +x_1278 = lean_ctor_get(x_1276, 0); +x_1279 = lean_ctor_get(x_2, 5); +lean_inc(x_1279); +x_1280 = 0; +x_1281 = l_Lean_SourceInfo_fromRef(x_1279, x_1280); +lean_dec(x_1279); +x_1282 = lean_ctor_get(x_2, 2); +lean_inc(x_1282); +x_1283 = lean_ctor_get(x_2, 1); +lean_inc(x_1283); +lean_dec(x_2); +x_1284 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242; +x_1285 = l_Lean_addMacroScope(x_1283, x_1284, x_1282); +x_1286 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240; +x_1287 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246; +lean_inc(x_1281); +x_1288 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1288, 0, x_1281); +lean_ctor_set(x_1288, 1, x_1286); +lean_ctor_set(x_1288, 2, x_1285); +lean_ctor_set(x_1288, 3, x_1287); +x_1289 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1281); +x_1290 = l_Lean_Syntax_node1(x_1281, x_1289, x_1278); +x_1291 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1292 = l_Lean_Syntax_node2(x_1281, x_1291, x_1288, x_1290); +lean_ctor_set(x_1276, 0, x_1292); +return x_1276; +} +else +{ +lean_object* x_1293; lean_object* x_1294; lean_object* x_1295; uint8_t x_1296; lean_object* x_1297; lean_object* x_1298; lean_object* x_1299; lean_object* x_1300; lean_object* x_1301; lean_object* x_1302; lean_object* x_1303; lean_object* x_1304; lean_object* x_1305; lean_object* x_1306; lean_object* x_1307; lean_object* x_1308; lean_object* x_1309; +x_1293 = lean_ctor_get(x_1276, 0); +x_1294 = lean_ctor_get(x_1276, 1); +lean_inc(x_1294); +lean_inc(x_1293); +lean_dec(x_1276); +x_1295 = lean_ctor_get(x_2, 5); +lean_inc(x_1295); +x_1296 = 0; +x_1297 = l_Lean_SourceInfo_fromRef(x_1295, x_1296); +lean_dec(x_1295); +x_1298 = lean_ctor_get(x_2, 2); +lean_inc(x_1298); +x_1299 = lean_ctor_get(x_2, 1); +lean_inc(x_1299); +lean_dec(x_2); +x_1300 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242; +x_1301 = l_Lean_addMacroScope(x_1299, x_1300, x_1298); +x_1302 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240; +x_1303 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246; +lean_inc(x_1297); +x_1304 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_1304, 0, x_1297); +lean_ctor_set(x_1304, 1, x_1302); +lean_ctor_set(x_1304, 2, x_1301); +lean_ctor_set(x_1304, 3, x_1303); +x_1305 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_1297); +x_1306 = l_Lean_Syntax_node1(x_1297, x_1305, x_1293); +x_1307 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_1308 = l_Lean_Syntax_node2(x_1297, x_1307, x_1304, x_1306); +x_1309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1309, 0, x_1308); +lean_ctor_set(x_1309, 1, x_1294); +return x_1309; +} +} +} +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("string", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("modifier", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_ctor_get(x_2, 5); +lean_inc(x_5); +x_6 = 0; +x_7 = l_Lean_SourceInfo_fromRef(x_5, x_6); +lean_dec(x_5); +x_8 = lean_ctor_get(x_2, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_dec(x_2); +x_10 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_7); +x_11 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +x_12 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3; +x_13 = l_Lean_addMacroScope(x_9, x_12, x_8); +x_14 = lean_box(0); +x_15 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2; +lean_inc(x_7); +x_16 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_16, 0, x_7); +lean_ctor_set(x_16, 1, x_15); +lean_ctor_set(x_16, 2, x_13); +lean_ctor_set(x_16, 3, x_14); +x_17 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_7); +x_18 = l_Lean_Syntax_node2(x_7, x_17, x_11, x_16); +x_19 = lean_box(2); +x_20 = l_Lean_Syntax_mkStrLit(x_4, x_19); +lean_dec(x_4); +x_21 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_7); +x_22 = l_Lean_Syntax_node1(x_7, x_21, x_20); +x_23 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_24 = l_Lean_Syntax_node2(x_7, x_23, x_18, x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_3); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +lean_inc(x_2); +x_27 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier(x_26, x_2, x_3); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_29 = lean_ctor_get(x_27, 0); +x_30 = lean_ctor_get(x_2, 5); +lean_inc(x_30); +x_31 = 0; +x_32 = l_Lean_SourceInfo_fromRef(x_30, x_31); +lean_dec(x_30); +x_33 = lean_ctor_get(x_2, 2); +lean_inc(x_33); +x_34 = lean_ctor_get(x_2, 1); +lean_inc(x_34); +lean_dec(x_2); +x_35 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_32); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_35); +x_37 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6; +x_38 = l_Lean_addMacroScope(x_34, x_37, x_33); +x_39 = lean_box(0); +x_40 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5; +lean_inc(x_32); +x_41 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_41, 0, x_32); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_38); +lean_ctor_set(x_41, 3, x_39); +x_42 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_32); +x_43 = l_Lean_Syntax_node2(x_32, x_42, x_36, x_41); +x_44 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_32); +x_45 = l_Lean_Syntax_node1(x_32, x_44, x_29); +x_46 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_47 = l_Lean_Syntax_node2(x_32, x_46, x_43, x_45); +lean_ctor_set(x_27, 0, x_47); +return x_27; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_48 = lean_ctor_get(x_27, 0); +x_49 = lean_ctor_get(x_27, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_27); +x_50 = lean_ctor_get(x_2, 5); +lean_inc(x_50); +x_51 = 0; +x_52 = l_Lean_SourceInfo_fromRef(x_50, x_51); +lean_dec(x_50); +x_53 = lean_ctor_get(x_2, 2); +lean_inc(x_53); +x_54 = lean_ctor_get(x_2, 1); +lean_inc(x_54); +lean_dec(x_2); +x_55 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47; +lean_inc(x_52); +x_56 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_56, 0, x_52); +lean_ctor_set(x_56, 1, x_55); +x_57 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6; +x_58 = l_Lean_addMacroScope(x_54, x_57, x_53); +x_59 = lean_box(0); +x_60 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5; +lean_inc(x_52); +x_61 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_61, 0, x_52); +lean_ctor_set(x_61, 1, x_60); +lean_ctor_set(x_61, 2, x_58); +lean_ctor_set(x_61, 3, x_59); +x_62 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46; +lean_inc(x_52); +x_63 = l_Lean_Syntax_node2(x_52, x_62, x_56, x_61); +x_64 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_52); +x_65 = l_Lean_Syntax_node1(x_52, x_64, x_48); +x_66 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5; +x_67 = l_Lean_Syntax_node2(x_52, x_66, x_63, x_65); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_49); +return x_68; +} +} +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("termDatespec(_)", 15, 15); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4; +x_3 = l_Std_Time_termDatespec_x28___x29___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("andthen", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_termDatespec_x28___x29___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("datespec(", 9, 9); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__5; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("str", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_termDatespec_x28___x29___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__8; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__4; +x_2 = l_Std_Time_termDatespec_x28___x29___closed__6; +x_3 = l_Std_Time_termDatespec_x28___x29___closed__9; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__4; +x_2 = l_Std_Time_termDatespec_x28___x29___closed__10; +x_3 = l_Std_Time_termDatespec_x28___x29___closed__11; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_Std_Time_termDatespec_x28___x29___closed__12; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_termDatespec_x28___x29() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_termDatespec_x28___x29___closed__13; +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_3); +x_5 = l_List_reverse___rarg(x_2); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_1); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_10 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart(x_8, x_3, x_4); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_11); +{ +lean_object* _tmp_0 = x_9; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_3 = x_12; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_4 = _tmp_3; +} +goto _start; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_1, 0); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_1); +lean_inc(x_3); +x_16 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart(x_14, x_3, x_4); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_2); +x_1 = x_15; +x_2 = x_19; +x_4 = x_18; +goto _start; +} +} +} +} +static lean_object* _init_l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +return x_2; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1; +x_7 = l_Lean_Syntax_TSepArray_push___rarg(x_6, x_2, x_4); +x_2 = x_7; +x_3 = x_5; +goto _start; +} +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot compile spec: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("anonymousCtor", 13, 13); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1; +x_2 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2; +x_3 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3; +x_4 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟨", 3, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("term[_]", 7, 7); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("[", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟩", 3, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_termDatespec_x28___x29___closed__2; +lean_inc(x_1); +x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +lean_dec(x_1); +x_6 = lean_box(1); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l_Std_Time_termDatespec_x28___x29___closed__8; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +lean_dec(x_2); +x_12 = lean_box(1); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_3); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_box(0); +x_15 = l_Lean_TSyntax_getString(x_9); +x_16 = l_Std_Time_GenericFormat_spec___rarg(x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1; +x_19 = lean_string_append(x_18, x_17); +lean_dec(x_17); +x_20 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2; +x_21 = lean_string_append(x_19, x_20); +x_22 = l_Lean_Macro_throwErrorAt___rarg(x_9, x_21, x_2, x_3); +lean_dec(x_9); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_9); +x_23 = lean_ctor_get(x_16, 0); +lean_inc(x_23); +lean_dec(x_16); +lean_inc(x_2); +x_24 = l_List_mapM_loop___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__1(x_23, x_14, x_2, x_3); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_26 = lean_ctor_get(x_24, 0); +x_27 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5; +x_28 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6; +x_29 = l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2(x_27, x_28, x_26); +x_30 = lean_ctor_get(x_2, 5); +lean_inc(x_30); +lean_dec(x_2); +x_31 = 0; +x_32 = l_Lean_SourceInfo_fromRef(x_30, x_31); +lean_dec(x_30); +x_33 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9; +lean_inc(x_32); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12; +lean_inc(x_32); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13; +x_38 = l_Array_append___rarg(x_37, x_29); +lean_dec(x_29); +x_39 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_32); +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_32); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_38); +x_41 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14; +lean_inc(x_32); +x_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_32); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11; +lean_inc(x_32); +x_44 = l_Lean_Syntax_node3(x_32, x_43, x_36, x_40, x_42); +lean_inc(x_32); +x_45 = l_Lean_Syntax_node1(x_32, x_39, x_44); +x_46 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15; +lean_inc(x_32); +x_47 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_47, 0, x_32); +lean_ctor_set(x_47, 1, x_46); +x_48 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8; +x_49 = l_Lean_Syntax_node3(x_32, x_48, x_34, x_45, x_47); +lean_ctor_set(x_24, 0, x_49); +return x_24; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_50 = lean_ctor_get(x_24, 0); +x_51 = lean_ctor_get(x_24, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_24); +x_52 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5; +x_53 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6; +x_54 = l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2(x_52, x_53, x_50); +x_55 = lean_ctor_get(x_2, 5); +lean_inc(x_55); +lean_dec(x_2); +x_56 = 0; +x_57 = l_Lean_SourceInfo_fromRef(x_55, x_56); +lean_dec(x_55); +x_58 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9; +lean_inc(x_57); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12; +lean_inc(x_57); +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_57); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13; +x_63 = l_Array_append___rarg(x_62, x_54); +lean_dec(x_54); +x_64 = l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16; +lean_inc(x_57); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_57); +lean_ctor_set(x_65, 1, x_64); +lean_ctor_set(x_65, 2, x_63); +x_66 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14; +lean_inc(x_57); +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_57); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11; +lean_inc(x_57); +x_69 = l_Lean_Syntax_node3(x_57, x_68, x_61, x_65, x_67); +lean_inc(x_57); +x_70 = l_Lean_Syntax_node1(x_57, x_64, x_69); +x_71 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15; +lean_inc(x_57); +x_72 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_72, 0, x_57); +lean_ctor_set(x_72, 1, x_71); +x_73 = l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8; +x_74 = l_Lean_Syntax_node3(x_57, x_73, x_59, x_70, x_72); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_51); +return x_75; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +lean_object* initialize_Std_Time_Date(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Format_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Notation_Spec(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Date(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Format_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__18); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__19); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__20); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__21); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__22); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__23); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__24); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__25); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__26); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertText___closed__27); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertNumber___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFraction___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__18); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__19); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__20); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__21); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__22); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__23); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__24); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertYear___closed__25); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertZoneName___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__18); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__19); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__20); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__21); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__22); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__23); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__24); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__25); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__26); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__27); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__28); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__29); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__30); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__31); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__32); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__33); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__34); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__35); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__36); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__37); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__38); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__39); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__40); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetX___closed__41); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetO___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__18); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__19); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__20); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__21); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertOffsetZ___closed__22); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__6); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__7); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__8); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__9); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__10); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__11); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__12); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__13); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__14); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__15); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__16); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__17); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__18); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__19); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__20); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__21); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__22); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__23); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__24); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__25); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__26); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__27); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__28); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__29); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__30); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__31); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__32); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__33); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__34); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__35); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__36); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__37); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__38); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__39); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__40); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__41); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__42); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__43); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__44); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__45); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__46); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__47); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__48); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__49); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__50); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__51); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__52); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__53); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__54); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__55); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__56); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__57); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__58); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__59); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__60); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__61); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__62); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__63); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__64); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__65); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__66); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__67); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__68); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__69); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__70); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__71); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__72); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__73); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__74); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__75); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__76); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__77); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__78); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__79); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__80); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__81); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__82); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__83); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__84); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__85); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__86); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__87); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__88); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__89); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__90); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__91); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__92); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__93); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__94); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__95); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__96); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__97); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__98); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__99); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__100); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__101); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__102); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__103); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__104); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__105); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__106); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__107); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__108); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__109); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__110); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__111); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__112); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__113); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__114); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__115); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__116); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__117); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__118); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__119); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__120); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__121); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__122); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__123); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__124); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__125); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__126); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__127); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__128); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__129); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__130); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__131); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__132); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__133); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__134); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__135); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__136); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__137); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__138); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__139); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__140); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__141); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__142); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__143); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__144); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__145); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__146); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__147); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__148); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__149); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__150); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__151); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__152); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__153); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__154); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__155); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__156); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__157); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__158); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__159); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__160); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__161); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__162); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__163); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__164); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__165); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__166); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__167); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__168); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__169); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__170); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__171); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__172); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__173); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__174); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__175); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__176); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__177); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__178); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__179); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__180); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__181); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__182); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__183); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__184); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__185); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__186); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__187); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__188); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__189); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__190); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__191); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__192); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__193); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__194); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__195); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__196); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__197); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__198); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__199); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__200); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__201); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__202); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__203); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__204); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__205); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__206); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__207); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__208); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__209); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__210); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__211); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__212); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__213); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__214); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__215); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__216); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__217); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__218); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__219); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__220); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__221); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__222); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__223); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__224); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__225); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__226); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__227); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__228); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__229); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__230); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__231); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__232); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__233); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__234); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__235); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__236); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__237); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__238); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__239); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__240); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__241); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__242); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__243); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__244); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__245); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertModifier___closed__246); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__1); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__2); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__3); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__4); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__5); +l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6 = _init_l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6(); +lean_mark_persistent(l___private_Std_Time_Notation_Spec_0__Std_Time_convertFormatPart___closed__6); +l_Std_Time_termDatespec_x28___x29___closed__1 = _init_l_Std_Time_termDatespec_x28___x29___closed__1(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__1); +l_Std_Time_termDatespec_x28___x29___closed__2 = _init_l_Std_Time_termDatespec_x28___x29___closed__2(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__2); +l_Std_Time_termDatespec_x28___x29___closed__3 = _init_l_Std_Time_termDatespec_x28___x29___closed__3(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__3); +l_Std_Time_termDatespec_x28___x29___closed__4 = _init_l_Std_Time_termDatespec_x28___x29___closed__4(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__4); +l_Std_Time_termDatespec_x28___x29___closed__5 = _init_l_Std_Time_termDatespec_x28___x29___closed__5(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__5); +l_Std_Time_termDatespec_x28___x29___closed__6 = _init_l_Std_Time_termDatespec_x28___x29___closed__6(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__6); +l_Std_Time_termDatespec_x28___x29___closed__7 = _init_l_Std_Time_termDatespec_x28___x29___closed__7(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__7); +l_Std_Time_termDatespec_x28___x29___closed__8 = _init_l_Std_Time_termDatespec_x28___x29___closed__8(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__8); +l_Std_Time_termDatespec_x28___x29___closed__9 = _init_l_Std_Time_termDatespec_x28___x29___closed__9(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__9); +l_Std_Time_termDatespec_x28___x29___closed__10 = _init_l_Std_Time_termDatespec_x28___x29___closed__10(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__10); +l_Std_Time_termDatespec_x28___x29___closed__11 = _init_l_Std_Time_termDatespec_x28___x29___closed__11(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__11); +l_Std_Time_termDatespec_x28___x29___closed__12 = _init_l_Std_Time_termDatespec_x28___x29___closed__12(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__12); +l_Std_Time_termDatespec_x28___x29___closed__13 = _init_l_Std_Time_termDatespec_x28___x29___closed__13(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29___closed__13); +l_Std_Time_termDatespec_x28___x29 = _init_l_Std_Time_termDatespec_x28___x29(); +lean_mark_persistent(l_Std_Time_termDatespec_x28___x29); +l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1 = _init_l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1(); +lean_mark_persistent(l_List_foldl___at_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___spec__2___closed__1); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__1); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__2); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__3); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__4); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__5); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__6); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__7); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__8); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__9); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__10); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__11); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__12); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__13); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__14); +l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15 = _init_l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15(); +lean_mark_persistent(l_Std_Time___aux__Std__Time__Notation__Spec______macroRules__Std__Time__termDatespec_x28___x29__1___closed__15); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time.c b/stage0/stdlib/Std/Time/Time.c new file mode 100644 index 000000000000..ed17210489d7 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time.c @@ -0,0 +1,37 @@ +// Lean compiler output +// Module: Std.Time.Time +// Imports: Std.Time.Time.Basic Std.Time.Time.HourMarker Std.Time.Time.PlainTime +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Time_Time_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_HourMarker(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_PlainTime(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_HourMarker(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_PlainTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Basic.c b/stage0/stdlib/Std/Time/Time/Basic.c new file mode 100644 index 000000000000..4e9417337acb --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Basic.c @@ -0,0 +1,29 @@ +// Lean compiler output +// Module: Std.Time.Time.Basic +// Imports: Std.Time.Time.Unit.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +lean_object* initialize_Std_Time_Time_Unit_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time_Unit_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/HourMarker.c b/stage0/stdlib/Std/Time/Time/HourMarker.c new file mode 100644 index 000000000000..ddf03df46a59 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/HourMarker.c @@ -0,0 +1,920 @@ +// Lean compiler output +// Module: Std.Time.Time.HourMarker +// Imports: Std.Time.Time.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_HourMarker_toRelative___closed__1; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__3; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__6; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8; +static lean_object* l_Std_Time_instReprHourMarker___closed__1; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9; +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__15; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__14; +LEAN_EXPORT uint8_t l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__7; +LEAN_EXPORT lean_object* l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toRelative___closed__3; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__4; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqHourMarker; +LEAN_EXPORT lean_object* l_Std_Time_instReprHourMarker; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toAbsolute(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4; +static lean_object* l_Std_Time_HourMarker_toRelative___closed__2; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toAbsolute___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__9; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion(lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative___lambda__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__5; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__8; +static lean_object* l_Std_Time_HourMarker_noConfusion___rarg___closed__1; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_ofOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toCtorIdx(uint8_t); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__11; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11; +static lean_object* l_Std_Time_instBEqHourMarker___closed__1; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__2; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1; +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_HourMarker_ofOrdinal(lean_object*); +static lean_object* l_Std_Time_HourMarker_ofOrdinal___closed__1; +lean_object* lean_int_add(lean_object*, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___lambda__1(lean_object*); +lean_object* lean_int_neg(lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___lambda__1___boxed(lean_object*); +static lean_object* l_Std_Time_HourMarker_toRelative___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative(lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13; +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11_(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14; +static lean_object* l_Std_Time_HourMarker_toAbsolute___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_HourMarker_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_HourMarker_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_HourMarker_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_HourMarker_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_HourMarker_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_HourMarker_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_HourMarker_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.HourMarker.am", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3; +x_2 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6; +x_2 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.HourMarker.pm", 22, 22); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3; +x_2 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6; +x_2 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instReprHourMarker___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprHourMarker() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprHourMarker___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95_(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Std_Time_HourMarker_toCtorIdx(x_1); +x_4 = l_Std_Time_HourMarker_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_instBEqHourMarker___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Time_HourMarker_0__Std_Time_beqHourMarker____x40_Std_Time_Time_HourMarker___hyg_95____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instBEqHourMarker() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instBEqHourMarker___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_HourMarker_ofOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_HourMarker_ofOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_3 = lean_int_dec_le(x_2, x_1); +if (x_3 == 0) +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +else +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_ofOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_HourMarker_ofOrdinal(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__3; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__4; +x_2 = l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__6; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__7; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__8; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__9; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__12; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__13; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__14; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toAbsolute___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_HourMarker_toAbsolute___closed__15; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toAbsolute(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_4 = lean_int_dec_eq(x_2, x_3); +if (x_4 == 0) +{ +lean_inc(x_2); +return x_2; +} +else +{ +lean_object* x_5; +x_5 = l_Std_Time_HourMarker_toAbsolute___closed__10; +return x_5; +} +} +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_7 = lean_int_dec_eq(x_2, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_int_add(x_2, x_6); +x_9 = l_Std_Time_HourMarker_toAbsolute___closed__11; +x_10 = lean_int_emod(x_8, x_9); +lean_dec(x_8); +return x_10; +} +else +{ +lean_object* x_11; +x_11 = l_Std_Time_HourMarker_toAbsolute___closed__16; +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toAbsolute___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_HourMarker_toAbsolute(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; uint8_t x_7; +x_6 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_7 = lean_int_dec_lt(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_4); +x_8 = lean_apply_1(x_5, lean_box(0)); +return x_8; +} +else +{ +lean_object* x_9; +lean_dec(x_5); +x_9 = lean_apply_1(x_4, lean_box(0)); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative___lambda__1(lean_object* x_1) { +_start: +{ +lean_internal_panic_unreachable(); +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative___lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_3 = 0; +x_4 = lean_box(x_3); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toRelative___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toRelative___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_HourMarker_toRelative___lambda__1), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toRelative___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_2 = 1; +x_3 = lean_box(x_2); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_HourMarker_toRelative___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_2 = 0; +x_3 = lean_box(x_2); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_HourMarker_toRelative(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_HourMarker_toAbsolute___closed__1; +x_3 = lean_int_dec_eq(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Std_Time_HourMarker_ofOrdinal___closed__1; +x_5 = lean_int_dec_le(x_1, x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_6 = l_Std_Time_HourMarker_toRelative___closed__1; +x_7 = lean_int_add(x_1, x_6); +lean_dec(x_1); +x_8 = 1; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = lean_int_dec_eq(x_1, x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_inc(x_1); +x_12 = lean_alloc_closure((void*)(l_Std_Time_HourMarker_toRelative___lambda__2), 2, 1); +lean_closure_set(x_12, 0, x_1); +x_13 = l_Std_Time_HourMarker_toRelative___closed__2; +x_14 = l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1(x_1, lean_box(0), lean_box(0), x_13, x_12); +lean_dec(x_1); +return x_14; +} +else +{ +lean_object* x_15; +lean_dec(x_1); +x_15 = l_Std_Time_HourMarker_toRelative___closed__3; +return x_15; +} +} +} +else +{ +lean_object* x_16; +lean_dec(x_1); +x_16 = l_Std_Time_HourMarker_toRelative___closed__4; +return x_16; +} +} +} +LEAN_EXPORT lean_object* l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Or_by__cases___at_Std_Time_HourMarker_toRelative___spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +lean_object* initialize_Std_Time_Time_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_HourMarker(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_HourMarker_noConfusion___rarg___closed__1 = _init_l_Std_Time_HourMarker_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_HourMarker_noConfusion___rarg___closed__1); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__1); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__2); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__3); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__4); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__5); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__6); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__7); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__8); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__9); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__10); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__11); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__12); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__13); +l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14 = _init_l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14(); +lean_mark_persistent(l___private_Std_Time_Time_HourMarker_0__Std_Time_reprHourMarker____x40_Std_Time_Time_HourMarker___hyg_11____closed__14); +l_Std_Time_instReprHourMarker___closed__1 = _init_l_Std_Time_instReprHourMarker___closed__1(); +lean_mark_persistent(l_Std_Time_instReprHourMarker___closed__1); +l_Std_Time_instReprHourMarker = _init_l_Std_Time_instReprHourMarker(); +lean_mark_persistent(l_Std_Time_instReprHourMarker); +l_Std_Time_instBEqHourMarker___closed__1 = _init_l_Std_Time_instBEqHourMarker___closed__1(); +lean_mark_persistent(l_Std_Time_instBEqHourMarker___closed__1); +l_Std_Time_instBEqHourMarker = _init_l_Std_Time_instBEqHourMarker(); +lean_mark_persistent(l_Std_Time_instBEqHourMarker); +l_Std_Time_HourMarker_ofOrdinal___closed__1 = _init_l_Std_Time_HourMarker_ofOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_HourMarker_ofOrdinal___closed__1); +l_Std_Time_HourMarker_toAbsolute___closed__1 = _init_l_Std_Time_HourMarker_toAbsolute___closed__1(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__1); +l_Std_Time_HourMarker_toAbsolute___closed__2 = _init_l_Std_Time_HourMarker_toAbsolute___closed__2(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__2); +l_Std_Time_HourMarker_toAbsolute___closed__3 = _init_l_Std_Time_HourMarker_toAbsolute___closed__3(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__3); +l_Std_Time_HourMarker_toAbsolute___closed__4 = _init_l_Std_Time_HourMarker_toAbsolute___closed__4(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__4); +l_Std_Time_HourMarker_toAbsolute___closed__5 = _init_l_Std_Time_HourMarker_toAbsolute___closed__5(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__5); +l_Std_Time_HourMarker_toAbsolute___closed__6 = _init_l_Std_Time_HourMarker_toAbsolute___closed__6(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__6); +l_Std_Time_HourMarker_toAbsolute___closed__7 = _init_l_Std_Time_HourMarker_toAbsolute___closed__7(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__7); +l_Std_Time_HourMarker_toAbsolute___closed__8 = _init_l_Std_Time_HourMarker_toAbsolute___closed__8(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__8); +l_Std_Time_HourMarker_toAbsolute___closed__9 = _init_l_Std_Time_HourMarker_toAbsolute___closed__9(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__9); +l_Std_Time_HourMarker_toAbsolute___closed__10 = _init_l_Std_Time_HourMarker_toAbsolute___closed__10(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__10); +l_Std_Time_HourMarker_toAbsolute___closed__11 = _init_l_Std_Time_HourMarker_toAbsolute___closed__11(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__11); +l_Std_Time_HourMarker_toAbsolute___closed__12 = _init_l_Std_Time_HourMarker_toAbsolute___closed__12(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__12); +l_Std_Time_HourMarker_toAbsolute___closed__13 = _init_l_Std_Time_HourMarker_toAbsolute___closed__13(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__13); +l_Std_Time_HourMarker_toAbsolute___closed__14 = _init_l_Std_Time_HourMarker_toAbsolute___closed__14(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__14); +l_Std_Time_HourMarker_toAbsolute___closed__15 = _init_l_Std_Time_HourMarker_toAbsolute___closed__15(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__15); +l_Std_Time_HourMarker_toAbsolute___closed__16 = _init_l_Std_Time_HourMarker_toAbsolute___closed__16(); +lean_mark_persistent(l_Std_Time_HourMarker_toAbsolute___closed__16); +l_Std_Time_HourMarker_toRelative___closed__1 = _init_l_Std_Time_HourMarker_toRelative___closed__1(); +lean_mark_persistent(l_Std_Time_HourMarker_toRelative___closed__1); +l_Std_Time_HourMarker_toRelative___closed__2 = _init_l_Std_Time_HourMarker_toRelative___closed__2(); +lean_mark_persistent(l_Std_Time_HourMarker_toRelative___closed__2); +l_Std_Time_HourMarker_toRelative___closed__3 = _init_l_Std_Time_HourMarker_toRelative___closed__3(); +lean_mark_persistent(l_Std_Time_HourMarker_toRelative___closed__3); +l_Std_Time_HourMarker_toRelative___closed__4 = _init_l_Std_Time_HourMarker_toRelative___closed__4(); +lean_mark_persistent(l_Std_Time_HourMarker_toRelative___closed__4); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/PlainTime.c b/stage0/stdlib/Std/Time/Time/PlainTime.c new file mode 100644 index 000000000000..4f6001dfddf0 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/PlainTime.c @@ -0,0 +1,2608 @@ +// Lean compiler output +// Module: Std.Time.Time.PlainTime +// Imports: Std.Time.Time.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toHours(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22; +static lean_object* l_Std_Time_PlainTime_toNanoseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withSeconds(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofSeconds___boxed(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHAddOffset__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSeconds___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subSeconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_millisecond___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMinutes(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16; +static lean_object* l_Std_Time_PlainTime_instHSubOffset___closed__1; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34; +static lean_object* l_Std_Time_PlainTime_instHAddOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHSubOffset; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25; +static lean_object* l_Std_Time_PlainTime_instHSubOffset__3___closed__1; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subHours___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +static lean_object* l_Std_Time_PlainTime_instHAddOffset___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__11; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5; +static lean_object* l_Std_Time_PlainTime_toMilliseconds___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addHours___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instReprPlainTime___closed__1; +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHAddOffset__1; +static lean_object* l_Std_Time_PlainTime_midnight___closed__4; +static lean_object* l_Std_Time_PlainTime_instHAddOffset__1___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__3; +static lean_object* l_Std_Time_PlainTime_midnight___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMilliseconds(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addSeconds___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +lean_object* l_Fin_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqPlainTime___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHSubOffset__1; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHAddOffset; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHAddOffset__3; +static lean_object* l_Std_Time_PlainTime_toSeconds___closed__2; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1; +static lean_object* l_Std_Time_PlainTime_toNanoseconds___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subSeconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSeconds(uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toSeconds___boxed(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMilliseconds___boxed(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofNanoseconds___boxed(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withHours(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_midnight; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMinutes___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__18; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMilliseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHSubOffset__2; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSecondsNano___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addNanoseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addSeconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMinutes___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__19; +lean_object* l_Int_repr(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__13; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMinutes(lean_object*); +lean_object* lean_int_div(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSecondsNano(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHSubOffset__3; +static lean_object* l_Std_Time_PlainTime_midnight___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHAddOffset__4; +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__14; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMilliseconds___boxed(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_instHSubOffset__4; +static lean_object* l_Std_Time_PlainTime_toMilliseconds___closed__2; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__5; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13; +static lean_object* l_Std_Time_PlainTime_instHAddOffset__3___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__15; +static lean_object* l_Std_Time_PlainTime_instHSubOffset__1___closed__1; +static lean_object* l_Std_Time_PlainTime_toSeconds___closed__1; +static lean_object* l_Std_Time_PlainTime_midnight___closed__2; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4; +lean_object* lean_int_mul(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_millisecond(lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__17; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainTime_instHAddOffset__4___closed__1; +static lean_object* l_Std_Time_PlainTime_ofNanoseconds___closed__1; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__1; +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subNanoseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainTime_instHSubOffset__4___closed__1; +static lean_object* l_Std_Time_PlainTime_toMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instBEqPlainTime(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMilliseconds(lean_object*, lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__20; +static lean_object* l_Std_Time_PlainTime_toMilliseconds___closed__4; +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedPlainTime; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprPlainTime; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withNanoseconds(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMinutes___boxed(lean_object*); +static lean_object* l_Std_Time_PlainTime_toNanoseconds___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMilliseconds(lean_object*); +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__7; +static lean_object* l_Std_Time_PlainTime_instHSubOffset__2___closed__1; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__9; +static lean_object* l_Std_Time_instInhabitedPlainTime___closed__16; +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("hour", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("minute", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("second", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("nanosecond", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(14u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(", ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟨", 3, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("⟩", 3, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_int_dec_lt(x_6, x_4); +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 3); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_int_dec_lt(x_9, x_4); +if (x_5 == 0) +{ +lean_object* x_183; lean_object* x_184; +x_183 = l_Int_repr(x_3); +lean_dec(x_3); +x_184 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_184, 0, x_183); +x_11 = x_184; +goto block_182; +} +else +{ +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_185 = l_Int_repr(x_3); +lean_dec(x_3); +x_186 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_186, 0, x_185); +x_187 = lean_unsigned_to_nat(0u); +x_188 = l_Repr_addAppParen(x_186, x_187); +x_11 = x_188; +goto block_182; +} +block_182: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_12 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7; +x_13 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = 0; +x_15 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set_uint8(x_15, sizeof(void*)*1, x_14); +x_16 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_box(1); +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12; +x_23 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5; +x_25 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +if (x_7 == 0) +{ +lean_object* x_176; lean_object* x_177; +x_176 = l_Int_repr(x_6); +lean_dec(x_6); +x_177 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_177, 0, x_176); +x_26 = x_177; +goto block_175; +} +else +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_178 = l_Int_repr(x_6); +lean_dec(x_6); +x_179 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_179, 0, x_178); +x_180 = lean_unsigned_to_nat(0u); +x_181 = l_Repr_addAppParen(x_179, x_180); +x_26 = x_181; +goto block_175; +} +block_175: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_74; +x_27 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13; +x_28 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +x_29 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set_uint8(x_29, sizeof(void*)*1, x_14); +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_25); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_18); +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_20); +x_33 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15; +x_34 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_24); +x_74 = !lean_is_exclusive(x_8); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; uint8_t x_77; uint8_t x_78; +x_75 = lean_ctor_get(x_8, 0); +x_76 = lean_ctor_get(x_8, 1); +x_77 = lean_int_dec_lt(x_76, x_4); +x_78 = lean_unbox(x_75); +lean_dec(x_75); +if (x_78 == 0) +{ +if (x_77 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_79 = l_Int_repr(x_76); +lean_dec(x_76); +x_80 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_80, 0, x_79); +x_81 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35; +lean_ctor_set_tag(x_8, 5); +lean_ctor_set(x_8, 1, x_80); +lean_ctor_set(x_8, 0, x_81); +x_82 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_83 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_8); +x_84 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_85 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_87 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_85); +x_88 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set_uint8(x_88, sizeof(void*)*1, x_14); +x_36 = x_88; +goto block_73; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_89 = l_Int_repr(x_76); +lean_dec(x_76); +x_90 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_90, 0, x_89); +x_91 = lean_unsigned_to_nat(0u); +x_92 = l_Repr_addAppParen(x_90, x_91); +x_93 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35; +lean_ctor_set_tag(x_8, 5); +lean_ctor_set(x_8, 1, x_92); +lean_ctor_set(x_8, 0, x_93); +x_94 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_95 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_8); +x_96 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_97 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_97, 0, x_95); +lean_ctor_set(x_97, 1, x_96); +x_98 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_99 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set_uint8(x_100, sizeof(void*)*1, x_14); +x_36 = x_100; +goto block_73; +} +} +else +{ +if (x_77 == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_101 = l_Int_repr(x_76); +lean_dec(x_76); +x_102 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_102, 0, x_101); +x_103 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38; +lean_ctor_set_tag(x_8, 5); +lean_ctor_set(x_8, 1, x_102); +lean_ctor_set(x_8, 0, x_103); +x_104 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_105 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_105, 0, x_104); +lean_ctor_set(x_105, 1, x_8); +x_106 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_107 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +x_108 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_109 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +x_110 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set_uint8(x_110, sizeof(void*)*1, x_14); +x_36 = x_110; +goto block_73; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_111 = l_Int_repr(x_76); +lean_dec(x_76); +x_112 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = lean_unsigned_to_nat(0u); +x_114 = l_Repr_addAppParen(x_112, x_113); +x_115 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38; +lean_ctor_set_tag(x_8, 5); +lean_ctor_set(x_8, 1, x_114); +lean_ctor_set(x_8, 0, x_115); +x_116 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_117 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_8); +x_118 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_119 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +x_120 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_121 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_119); +x_122 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set_uint8(x_122, sizeof(void*)*1, x_14); +x_36 = x_122; +goto block_73; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; uint8_t x_125; uint8_t x_126; +x_123 = lean_ctor_get(x_8, 0); +x_124 = lean_ctor_get(x_8, 1); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_8); +x_125 = lean_int_dec_lt(x_124, x_4); +x_126 = lean_unbox(x_123); +lean_dec(x_123); +if (x_126 == 0) +{ +if (x_125 == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_127 = l_Int_repr(x_124); +lean_dec(x_124); +x_128 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_128, 0, x_127); +x_129 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35; +x_130 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_128); +x_131 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_132 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_130); +x_133 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_134 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +x_135 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_136 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_134); +x_137 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set_uint8(x_137, sizeof(void*)*1, x_14); +x_36 = x_137; +goto block_73; +} +else +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; +x_138 = l_Int_repr(x_124); +lean_dec(x_124); +x_139 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_139, 0, x_138); +x_140 = lean_unsigned_to_nat(0u); +x_141 = l_Repr_addAppParen(x_139, x_140); +x_142 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35; +x_143 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_141); +x_144 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_145 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_143); +x_146 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_147 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_146); +x_148 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_149 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_147); +x_150 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set_uint8(x_150, sizeof(void*)*1, x_14); +x_36 = x_150; +goto block_73; +} +} +else +{ +if (x_125 == 0) +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_151 = l_Int_repr(x_124); +lean_dec(x_124); +x_152 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_152, 0, x_151); +x_153 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38; +x_154 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_152); +x_155 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_156 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_154); +x_157 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_158 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_160 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_158); +x_161 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set_uint8(x_161, sizeof(void*)*1, x_14); +x_36 = x_161; +goto block_73; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_162 = l_Int_repr(x_124); +lean_dec(x_124); +x_163 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_163, 0, x_162); +x_164 = lean_unsigned_to_nat(0u); +x_165 = l_Repr_addAppParen(x_163, x_164); +x_166 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38; +x_167 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_165); +x_168 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30; +x_169 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_167); +x_170 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32; +x_171 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_171, 0, x_169); +lean_ctor_set(x_171, 1, x_170); +x_172 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29; +x_173 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_171); +x_174 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set_uint8(x_174, sizeof(void*)*1, x_14); +x_36 = x_174; +goto block_73; +} +} +} +block_73: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_37 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_37, 0, x_27); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set_uint8(x_38, sizeof(void*)*1, x_14); +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_35); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_18); +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_20); +x_42 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17; +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_24); +if (x_10 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_45 = l_Int_repr(x_9); +lean_dec(x_9); +x_46 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18; +x_48 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set_uint8(x_49, sizeof(void*)*1, x_14); +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_44); +lean_ctor_set(x_50, 1, x_49); +x_51 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22; +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +x_53 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24; +x_54 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +x_55 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21; +x_56 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +x_57 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set_uint8(x_57, sizeof(void*)*1, x_14); +return x_57; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_58 = l_Int_repr(x_9); +lean_dec(x_9); +x_59 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l_Repr_addAppParen(x_59, x_60); +x_62 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18; +x_63 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set_uint8(x_64, sizeof(void*)*1, x_14); +x_65 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_65, 0, x_44); +lean_ctor_set(x_65, 1, x_64); +x_66 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22; +x_67 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_65); +x_68 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24; +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21; +x_71 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_69); +x_72 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set_uint8(x_72, sizeof(void*)*1, x_14); +return x_72; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprPlainTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprPlainTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprPlainTime___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__2; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__3; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__6; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__7; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__8; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__9; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__11; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__12; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__13; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__6; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__15; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__14; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__16; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__14; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__17; +x_2 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__19() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__18; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__10; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__18; +x_3 = l_Std_Time_instInhabitedPlainTime___closed__19; +x_4 = l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_instInhabitedPlainTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__20; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_instBEqPlainTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_int_dec_eq(x_3, x_4); +if (x_5 == 0) +{ +uint8_t x_6; +x_6 = 0; +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 1); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_int_dec_eq(x_7, x_8); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_1, 2); +x_12 = lean_ctor_get(x_11, 1); +x_13 = lean_ctor_get(x_2, 2); +x_14 = lean_ctor_get(x_13, 1); +x_15 = lean_int_dec_eq(x_12, x_14); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_1, 3); +x_18 = lean_ctor_get(x_2, 3); +x_19 = lean_int_dec_eq(x_17, x_18); +return x_19; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqPlainTime___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_instBEqPlainTime(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_PlainTime_midnight___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__18; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_PlainTime_midnight___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(999999999u); +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Fin_ofNat(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_PlainTime_midnight___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_PlainTime_midnight___closed__2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_midnight___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_instInhabitedPlainTime___closed__10; +x_2 = l_Std_Time_instInhabitedPlainTime___closed__18; +x_3 = l_Std_Time_PlainTime_midnight___closed__1; +x_4 = l_Std_Time_PlainTime_midnight___closed__3; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_PlainTime_midnight() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_midnight___closed__4; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSecondsNano(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_box(x_1); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_3); +lean_ctor_set(x_8, 2, x_7); +lean_ctor_set(x_8, 3, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSecondsNano___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint8_t x_6; lean_object* x_7; +x_6 = lean_unbox(x_1); +lean_dec(x_1); +x_7 = l_Std_Time_PlainTime_ofHourMinuteSecondsNano(x_6, x_2, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSeconds(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_box(x_1); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = l_Std_Time_PlainTime_midnight___closed__3; +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_3); +lean_ctor_set(x_8, 2, x_6); +lean_ctor_set(x_8, 3, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHourMinuteSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_1); +lean_dec(x_1); +x_6 = l_Std_Time_PlainTime_ofHourMinuteSeconds(x_5, x_2, x_3, x_4); +return x_6; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toMilliseconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toMilliseconds___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toMilliseconds___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainTime_toMilliseconds___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Std_Time_PlainTime_toMilliseconds___closed__2; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +lean_dec(x_4); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_9, 1); +x_11 = l_Std_Time_PlainTime_toMilliseconds___closed__3; +x_12 = lean_int_mul(x_10, x_11); +x_13 = lean_int_add(x_8, x_12); +lean_dec(x_12); +lean_dec(x_8); +x_14 = lean_ctor_get(x_1, 3); +x_15 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_16 = lean_int_div(x_14, x_15); +x_17 = lean_int_add(x_13, x_16); +lean_dec(x_16); +lean_dec(x_13); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toNanoseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("3600000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toNanoseconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("60000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toNanoseconds___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainTime_toNanoseconds___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Std_Time_PlainTime_toNanoseconds___closed__2; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +lean_dec(x_4); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_9, 1); +x_11 = l_Std_Time_PlainTime_toNanoseconds___closed__3; +x_12 = lean_int_mul(x_10, x_11); +x_13 = lean_int_add(x_8, x_12); +lean_dec(x_12); +lean_dec(x_8); +x_14 = lean_ctor_get(x_1, 3); +x_15 = lean_int_add(x_13, x_14); +lean_dec(x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_toSeconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainTime_toSeconds___closed__1; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = l_Std_Time_PlainTime_toSeconds___closed__2; +x_7 = lean_int_mul(x_5, x_6); +x_8 = lean_int_add(x_4, x_7); +lean_dec(x_7); +lean_dec(x_4); +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_9, 1); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Std_Time_PlainTime_toSeconds___closed__2; +x_4 = lean_int_mul(x_2, x_3); +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_int_add(x_4, x_5); +lean_dec(x_4); +x_7 = lean_ctor_get(x_1, 2); +x_8 = lean_ctor_get(x_7, 1); +x_9 = lean_int_div(x_8, x_3); +x_10 = lean_int_add(x_6, x_9); +lean_dec(x_9); +lean_dec(x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_ofNanoseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_2 = l_Std_Time_PlainTime_toNanoseconds___closed__3; +x_3 = lean_int_ediv(x_1, x_2); +x_4 = lean_int_emod(x_1, x_2); +x_5 = l_Std_Time_PlainTime_toSeconds___closed__1; +x_6 = lean_int_ediv(x_3, x_5); +x_7 = l_Std_Time_PlainTime_ofNanoseconds___closed__1; +x_8 = lean_int_emod(x_6, x_7); +lean_dec(x_6); +x_9 = lean_int_emod(x_3, x_5); +x_10 = l_Std_Time_PlainTime_toSeconds___closed__2; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = lean_int_emod(x_3, x_10); +lean_dec(x_3); +x_13 = 0; +x_14 = lean_box(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_12); +x_16 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_16, 0, x_8); +lean_ctor_set(x_16, 1, x_11); +lean_ctor_set(x_16, 2, x_15); +lean_ctor_set(x_16, 3, x_4); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_PlainTime_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainTime_toNanoseconds___closed__3; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_PlainTime_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainTime_toNanoseconds___closed__2; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_PlainTime_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_PlainTime_toNanoseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +x_4 = l_Std_Time_PlainTime_ofNanoseconds(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_4 = l_Std_Time_PlainTime_toNanoseconds___closed__3; +x_5 = lean_int_mul(x_2, x_4); +x_6 = lean_int_add(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_7 = l_Std_Time_PlainTime_ofNanoseconds(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_addSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_5 = l_Std_Time_PlainTime_toNanoseconds___closed__3; +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainTime_ofNanoseconds(x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_subSeconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_4 = l_Std_Time_PlainTime_toNanoseconds___closed__2; +x_5 = lean_int_mul(x_2, x_4); +x_6 = lean_int_add(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_7 = l_Std_Time_PlainTime_ofNanoseconds(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_addMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_5 = l_Std_Time_PlainTime_toNanoseconds___closed__2; +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainTime_ofNanoseconds(x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_subMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_4 = l_Std_Time_PlainTime_toNanoseconds___closed__1; +x_5 = lean_int_mul(x_2, x_4); +x_6 = lean_int_add(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_7 = l_Std_Time_PlainTime_ofNanoseconds(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_addHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_5 = l_Std_Time_PlainTime_toNanoseconds___closed__1; +x_6 = lean_int_mul(x_3, x_5); +lean_dec(x_3); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +x_8 = l_Std_Time_PlainTime_ofNanoseconds(x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_subHours(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_PlainTime_toNanoseconds(x_1); +x_4 = lean_int_add(x_3, x_2); +lean_dec(x_3); +x_5 = l_Std_Time_PlainTime_ofNanoseconds(x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_addNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainTime_addNanoseconds(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_subNanoseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_PlainTime_toMilliseconds(x_1); +x_4 = lean_int_add(x_3, x_2); +lean_dec(x_3); +x_5 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_6 = lean_int_mul(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_PlainTime_ofNanoseconds(x_6); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_addMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_int_neg(x_2); +x_4 = l_Std_Time_PlainTime_addMilliseconds(x_1, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_subMilliseconds(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 2); +lean_dec(x_4); +lean_ctor_set(x_1, 2, x_2); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +lean_ctor_set(x_8, 2, x_2); +lean_ctor_set(x_8, 3, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_dec(x_4); +lean_ctor_set(x_1, 1, x_2); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_6); +lean_ctor_set(x_8, 3, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_1, 3); +x_5 = l_Std_Time_PlainTime_toMilliseconds___closed__3; +x_6 = lean_int_emod(x_4, x_5); +lean_dec(x_4); +x_7 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_8 = lean_int_mul(x_2, x_7); +x_9 = lean_int_add(x_8, x_6); +lean_dec(x_6); +lean_dec(x_8); +lean_ctor_set(x_1, 3, x_9); +return x_1; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 2); +x_13 = lean_ctor_get(x_1, 3); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_1); +x_14 = l_Std_Time_PlainTime_toMilliseconds___closed__3; +x_15 = lean_int_emod(x_13, x_14); +lean_dec(x_13); +x_16 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_17 = lean_int_mul(x_2, x_16); +x_18 = lean_int_add(x_17, x_15); +lean_dec(x_15); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_19, 0, x_10); +lean_ctor_set(x_19, 1, x_11); +lean_ctor_set(x_19, 2, x_12); +lean_ctor_set(x_19, 3, x_18); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainTime_withMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 3); +lean_dec(x_4); +lean_ctor_set(x_1, 3, x_2); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_ctor_get(x_1, 2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_6); +lean_ctor_set(x_8, 2, x_7); +lean_ctor_set(x_8, 3, x_2); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_withHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 0); +lean_dec(x_4); +lean_ctor_set(x_1, 0, x_2); +return x_1; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_1, 1); +x_6 = lean_ctor_get(x_1, 2); +x_7 = lean_ctor_get(x_1, 3); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_dec(x_1); +x_8 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_5); +lean_ctor_set(x_8, 2, x_6); +lean_ctor_set(x_8, 3, x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_millisecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 3); +x_3 = l_Std_Time_PlainTime_toMilliseconds___closed__4; +x_4 = lean_int_ediv(x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_millisecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_PlainTime_millisecond(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_addNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_subNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_addMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_subMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHSubOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_addSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHAddOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_subSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHSubOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_addMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHAddOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_subMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHSubOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_addHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHAddOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_PlainTime_subHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_PlainTime_instHSubOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_PlainTime_instHSubOffset__4___closed__1; +return x_1; +} +} +lean_object* initialize_Std_Time_Time_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_PlainTime(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__1); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__2); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__3); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__4); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__5); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__6); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__7); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__8); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__9); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__10); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__11); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__12); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__13); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__14); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__15); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__16); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__17); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__18); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__19); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__20); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__21); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__22); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__23); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__24); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__25); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__26); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__27); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__28); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__29); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__30); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__31); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__32); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__33); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__34); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__35); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__36); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__37); +l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38 = _init_l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38(); +lean_mark_persistent(l___private_Std_Time_Time_PlainTime_0__Std_Time_reprPlainTime____x40_Std_Time_Time_PlainTime___hyg_43____closed__38); +l_Std_Time_instReprPlainTime___closed__1 = _init_l_Std_Time_instReprPlainTime___closed__1(); +lean_mark_persistent(l_Std_Time_instReprPlainTime___closed__1); +l_Std_Time_instReprPlainTime = _init_l_Std_Time_instReprPlainTime(); +lean_mark_persistent(l_Std_Time_instReprPlainTime); +l_Std_Time_instInhabitedPlainTime___closed__1 = _init_l_Std_Time_instInhabitedPlainTime___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__1); +l_Std_Time_instInhabitedPlainTime___closed__2 = _init_l_Std_Time_instInhabitedPlainTime___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__2); +l_Std_Time_instInhabitedPlainTime___closed__3 = _init_l_Std_Time_instInhabitedPlainTime___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__3); +l_Std_Time_instInhabitedPlainTime___closed__4 = _init_l_Std_Time_instInhabitedPlainTime___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__4); +l_Std_Time_instInhabitedPlainTime___closed__5 = _init_l_Std_Time_instInhabitedPlainTime___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__5); +l_Std_Time_instInhabitedPlainTime___closed__6 = _init_l_Std_Time_instInhabitedPlainTime___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__6); +l_Std_Time_instInhabitedPlainTime___closed__7 = _init_l_Std_Time_instInhabitedPlainTime___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__7); +l_Std_Time_instInhabitedPlainTime___closed__8 = _init_l_Std_Time_instInhabitedPlainTime___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__8); +l_Std_Time_instInhabitedPlainTime___closed__9 = _init_l_Std_Time_instInhabitedPlainTime___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__9); +l_Std_Time_instInhabitedPlainTime___closed__10 = _init_l_Std_Time_instInhabitedPlainTime___closed__10(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__10); +l_Std_Time_instInhabitedPlainTime___closed__11 = _init_l_Std_Time_instInhabitedPlainTime___closed__11(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__11); +l_Std_Time_instInhabitedPlainTime___closed__12 = _init_l_Std_Time_instInhabitedPlainTime___closed__12(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__12); +l_Std_Time_instInhabitedPlainTime___closed__13 = _init_l_Std_Time_instInhabitedPlainTime___closed__13(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__13); +l_Std_Time_instInhabitedPlainTime___closed__14 = _init_l_Std_Time_instInhabitedPlainTime___closed__14(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__14); +l_Std_Time_instInhabitedPlainTime___closed__15 = _init_l_Std_Time_instInhabitedPlainTime___closed__15(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__15); +l_Std_Time_instInhabitedPlainTime___closed__16 = _init_l_Std_Time_instInhabitedPlainTime___closed__16(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__16); +l_Std_Time_instInhabitedPlainTime___closed__17 = _init_l_Std_Time_instInhabitedPlainTime___closed__17(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__17); +l_Std_Time_instInhabitedPlainTime___closed__18 = _init_l_Std_Time_instInhabitedPlainTime___closed__18(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__18); +l_Std_Time_instInhabitedPlainTime___closed__19 = _init_l_Std_Time_instInhabitedPlainTime___closed__19(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__19); +l_Std_Time_instInhabitedPlainTime___closed__20 = _init_l_Std_Time_instInhabitedPlainTime___closed__20(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime___closed__20); +l_Std_Time_instInhabitedPlainTime = _init_l_Std_Time_instInhabitedPlainTime(); +lean_mark_persistent(l_Std_Time_instInhabitedPlainTime); +l_Std_Time_PlainTime_midnight___closed__1 = _init_l_Std_Time_PlainTime_midnight___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_midnight___closed__1); +l_Std_Time_PlainTime_midnight___closed__2 = _init_l_Std_Time_PlainTime_midnight___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_midnight___closed__2); +l_Std_Time_PlainTime_midnight___closed__3 = _init_l_Std_Time_PlainTime_midnight___closed__3(); +lean_mark_persistent(l_Std_Time_PlainTime_midnight___closed__3); +l_Std_Time_PlainTime_midnight___closed__4 = _init_l_Std_Time_PlainTime_midnight___closed__4(); +lean_mark_persistent(l_Std_Time_PlainTime_midnight___closed__4); +l_Std_Time_PlainTime_midnight = _init_l_Std_Time_PlainTime_midnight(); +lean_mark_persistent(l_Std_Time_PlainTime_midnight); +l_Std_Time_PlainTime_toMilliseconds___closed__1 = _init_l_Std_Time_PlainTime_toMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_toMilliseconds___closed__1); +l_Std_Time_PlainTime_toMilliseconds___closed__2 = _init_l_Std_Time_PlainTime_toMilliseconds___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_toMilliseconds___closed__2); +l_Std_Time_PlainTime_toMilliseconds___closed__3 = _init_l_Std_Time_PlainTime_toMilliseconds___closed__3(); +lean_mark_persistent(l_Std_Time_PlainTime_toMilliseconds___closed__3); +l_Std_Time_PlainTime_toMilliseconds___closed__4 = _init_l_Std_Time_PlainTime_toMilliseconds___closed__4(); +lean_mark_persistent(l_Std_Time_PlainTime_toMilliseconds___closed__4); +l_Std_Time_PlainTime_toNanoseconds___closed__1 = _init_l_Std_Time_PlainTime_toNanoseconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_toNanoseconds___closed__1); +l_Std_Time_PlainTime_toNanoseconds___closed__2 = _init_l_Std_Time_PlainTime_toNanoseconds___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_toNanoseconds___closed__2); +l_Std_Time_PlainTime_toNanoseconds___closed__3 = _init_l_Std_Time_PlainTime_toNanoseconds___closed__3(); +lean_mark_persistent(l_Std_Time_PlainTime_toNanoseconds___closed__3); +l_Std_Time_PlainTime_toSeconds___closed__1 = _init_l_Std_Time_PlainTime_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_toSeconds___closed__1); +l_Std_Time_PlainTime_toSeconds___closed__2 = _init_l_Std_Time_PlainTime_toSeconds___closed__2(); +lean_mark_persistent(l_Std_Time_PlainTime_toSeconds___closed__2); +l_Std_Time_PlainTime_ofNanoseconds___closed__1 = _init_l_Std_Time_PlainTime_ofNanoseconds___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_ofNanoseconds___closed__1); +l_Std_Time_PlainTime_instHAddOffset___closed__1 = _init_l_Std_Time_PlainTime_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset___closed__1); +l_Std_Time_PlainTime_instHAddOffset = _init_l_Std_Time_PlainTime_instHAddOffset(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset); +l_Std_Time_PlainTime_instHSubOffset___closed__1 = _init_l_Std_Time_PlainTime_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset___closed__1); +l_Std_Time_PlainTime_instHSubOffset = _init_l_Std_Time_PlainTime_instHSubOffset(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset); +l_Std_Time_PlainTime_instHAddOffset__1___closed__1 = _init_l_Std_Time_PlainTime_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__1___closed__1); +l_Std_Time_PlainTime_instHAddOffset__1 = _init_l_Std_Time_PlainTime_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__1); +l_Std_Time_PlainTime_instHSubOffset__1___closed__1 = _init_l_Std_Time_PlainTime_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__1___closed__1); +l_Std_Time_PlainTime_instHSubOffset__1 = _init_l_Std_Time_PlainTime_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__1); +l_Std_Time_PlainTime_instHAddOffset__2___closed__1 = _init_l_Std_Time_PlainTime_instHAddOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__2___closed__1); +l_Std_Time_PlainTime_instHAddOffset__2 = _init_l_Std_Time_PlainTime_instHAddOffset__2(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__2); +l_Std_Time_PlainTime_instHSubOffset__2___closed__1 = _init_l_Std_Time_PlainTime_instHSubOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__2___closed__1); +l_Std_Time_PlainTime_instHSubOffset__2 = _init_l_Std_Time_PlainTime_instHSubOffset__2(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__2); +l_Std_Time_PlainTime_instHAddOffset__3___closed__1 = _init_l_Std_Time_PlainTime_instHAddOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__3___closed__1); +l_Std_Time_PlainTime_instHAddOffset__3 = _init_l_Std_Time_PlainTime_instHAddOffset__3(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__3); +l_Std_Time_PlainTime_instHSubOffset__3___closed__1 = _init_l_Std_Time_PlainTime_instHSubOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__3___closed__1); +l_Std_Time_PlainTime_instHSubOffset__3 = _init_l_Std_Time_PlainTime_instHSubOffset__3(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__3); +l_Std_Time_PlainTime_instHAddOffset__4___closed__1 = _init_l_Std_Time_PlainTime_instHAddOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__4___closed__1); +l_Std_Time_PlainTime_instHAddOffset__4 = _init_l_Std_Time_PlainTime_instHAddOffset__4(); +lean_mark_persistent(l_Std_Time_PlainTime_instHAddOffset__4); +l_Std_Time_PlainTime_instHSubOffset__4___closed__1 = _init_l_Std_Time_PlainTime_instHSubOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__4___closed__1); +l_Std_Time_PlainTime_instHSubOffset__4 = _init_l_Std_Time_PlainTime_instHSubOffset__4(); +lean_mark_persistent(l_Std_Time_PlainTime_instHSubOffset__4); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Basic.c b/stage0/stdlib/Std/Time/Time/Unit/Basic.c new file mode 100644 index 000000000000..8fbb1adaa95d --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Basic.c @@ -0,0 +1,960 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Basic +// Imports: Std.Time.Internal Std.Time.Time.Unit.Hour Std.Time.Time.Unit.Minute Std.Time.Time.Unit.Second Std.Time.Time.Unit.Nanosecond Std.Time.Time.Unit.Millisecond +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMilliseconds___boxed(lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofSeconds___boxed(lean_object*); +static lean_object* l_Std_Time_Millisecond_Offset_toHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofSeconds(lean_object*); +static lean_object* l_Std_Time_Millisecond_Offset_toSeconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMinutes___boxed(lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toMinutes(lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toSeconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMinutes___boxed(lean_object*); +lean_object* lean_int_div(lean_object*, lean_object*); +static lean_object* l_Std_Time_Millisecond_Offset_toMinutes___closed__1; +static lean_object* l_Std_Time_Second_Offset_toHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofMilliseconds___boxed(lean_object*); +static lean_object* l_Std_Time_Second_Offset_toMinutes___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toHours___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNanoseconds(lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMinutes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMilliseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMinutes___boxed(lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toMinutes___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofHours(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toHours___boxed(lean_object*); +static lean_object* l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMinutes___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNanoseconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofHours(lean_object*); +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toSeconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("60000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Offset_toHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_cstr_to_nat("3600000000000"); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_Offset_toSeconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toSeconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_Offset_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_Offset_toHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toSeconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toSeconds___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toSeconds___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_Offset_toMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_Offset_toHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_toHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_toHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNanoseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Nanosecond_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNanoseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofNanoseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMilliseconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Millisecond_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMilliseconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofMilliseconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toHours___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toHours___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_toMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_toMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMinutes(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_Second_Offset_toMinutes___closed__1; +x_3 = lean_int_div(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofMinutes___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofMinutes(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Hour(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Minute(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Second(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Nanosecond(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Millisecond(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Hour(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Minute(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Second(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Nanosecond(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Millisecond(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toMilliseconds___closed__1); +l_Std_Time_Nanosecond_Offset_toSeconds___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toSeconds___closed__1); +l_Std_Time_Nanosecond_Offset_toMinutes___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toMinutes___closed__1); +l_Std_Time_Nanosecond_Offset_toHours___closed__1 = _init_l_Std_Time_Nanosecond_Offset_toHours___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_Offset_toHours___closed__1); +l_Std_Time_Millisecond_Offset_toSeconds___closed__1 = _init_l_Std_Time_Millisecond_Offset_toSeconds___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_Offset_toSeconds___closed__1); +l_Std_Time_Millisecond_Offset_toMinutes___closed__1 = _init_l_Std_Time_Millisecond_Offset_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_Offset_toMinutes___closed__1); +l_Std_Time_Millisecond_Offset_toHours___closed__1 = _init_l_Std_Time_Millisecond_Offset_toHours___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_Offset_toHours___closed__1); +l_Std_Time_Second_Offset_toMinutes___closed__1 = _init_l_Std_Time_Second_Offset_toMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_Second_Offset_toMinutes___closed__1); +l_Std_Time_Second_Offset_toHours___closed__1 = _init_l_Std_Time_Second_Offset_toHours___closed__1(); +lean_mark_persistent(l_Std_Time_Second_Offset_toHours___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Hour.c b/stage0/stdlib/Std/Time/Time/Unit/Hour.c new file mode 100644 index 000000000000..6827c8cef051 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Hour.c @@ -0,0 +1,728 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Hour +// Imports: Std.Internal.Rat Std.Time.Internal Std.Time.Time.Unit.Minute Std.Time.Time.Unit.Second +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___boxed(lean_object*); +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOrdinalBEq; +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__8; +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOrdinalLE; +lean_object* lean_int_emod(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__10; +static lean_object* l_Std_Time_Hour_instOrdinalRepr___closed__1; +static lean_object* l_Std_Time_Hour_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofInt___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__3; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetRepr; +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetBEq; +static lean_object* l_Std_Time_Hour_instOfNatOrdinal___closed__1; +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2; +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instOffsetNeg___closed__1; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Std_Time_Hour_Ordinal_toRelative___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNat(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetNeg; +static lean_object* l_Std_Time_Hour_instOrdinalBEq___closed__1; +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__7; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLtOrdinal(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__1; +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofInt___boxed(lean_object*); +static lean_object* l_Std_Time_Hour_instOffsetBEq___closed__1; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOrdinalLT; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofInt(lean_object*); +static lean_object* l_Std_Time_Hour_instOffsetRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetAdd; +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLeOrdinal(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOrdinalRepr; +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1; +static lean_object* l_Std_Time_Hour_instOffsetSub___closed__1; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOrdinal(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetInhabited; +static lean_object* l_Std_Time_Hour_Ordinal_toRelative___closed__2; +static lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetSub; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instInhabitedOrdinal; +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOffsetToString; +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative___boxed(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofFin(lean_object*); +static lean_object* l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Hour_instOffsetAdd___closed__1; +static lean_object* _init_l_Std_Time_Hour_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(23u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__9; +x_2 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instInhabitedOrdinal___closed__10; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Hour_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Hour_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Hour_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOfNatOrdinal___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Hour_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Hour_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Hour_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_toRelative___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_toRelative___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l_Std_Time_Hour_Ordinal_toRelative___closed__2; +x_3 = lean_int_add(x_1, x_2); +x_4 = l_Std_Time_Hour_Ordinal_toRelative___closed__1; +x_5 = lean_int_emod(x_3, x_4); +lean_dec(x_3); +x_6 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toRelative___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Ordinal_toRelative(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2; +x_2 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4; +x_2 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5; +x_2 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l_Std_Time_Hour_instInhabitedOrdinal___closed__4; +x_3 = lean_int_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_inc(x_1); +return x_1; +} +else +{ +lean_object* x_4; +x_4 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7; +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Ordinal_shiftTo1BasedHour(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Hour_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Hour_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Minute(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Second(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Hour(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Minute(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Second(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Hour_instOrdinalRepr___closed__1 = _init_l_Std_Time_Hour_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalRepr___closed__1); +l_Std_Time_Hour_instOrdinalRepr = _init_l_Std_Time_Hour_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalRepr); +l_Std_Time_Hour_instOrdinalBEq___closed__1 = _init_l_Std_Time_Hour_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalBEq___closed__1); +l_Std_Time_Hour_instOrdinalBEq = _init_l_Std_Time_Hour_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalBEq); +l_Std_Time_Hour_instOrdinalLE = _init_l_Std_Time_Hour_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalLE); +l_Std_Time_Hour_instOrdinalLT = _init_l_Std_Time_Hour_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Hour_instOrdinalLT); +l_Std_Time_Hour_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Hour_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOfNatOrdinal___closed__1); +l_Std_Time_Hour_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__1); +l_Std_Time_Hour_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__2); +l_Std_Time_Hour_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__3); +l_Std_Time_Hour_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__4); +l_Std_Time_Hour_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__5); +l_Std_Time_Hour_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__6); +l_Std_Time_Hour_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__7); +l_Std_Time_Hour_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__8); +l_Std_Time_Hour_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__9); +l_Std_Time_Hour_instInhabitedOrdinal___closed__10 = _init_l_Std_Time_Hour_instInhabitedOrdinal___closed__10(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal___closed__10); +l_Std_Time_Hour_instInhabitedOrdinal = _init_l_Std_Time_Hour_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Hour_instInhabitedOrdinal); +l_Std_Time_Hour_instOffsetRepr___closed__1 = _init_l_Std_Time_Hour_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetRepr___closed__1); +l_Std_Time_Hour_instOffsetRepr = _init_l_Std_Time_Hour_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetRepr); +l_Std_Time_Hour_instOffsetBEq___closed__1 = _init_l_Std_Time_Hour_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetBEq___closed__1); +l_Std_Time_Hour_instOffsetBEq = _init_l_Std_Time_Hour_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetBEq); +l_Std_Time_Hour_instOffsetInhabited = _init_l_Std_Time_Hour_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetInhabited); +l_Std_Time_Hour_instOffsetAdd___closed__1 = _init_l_Std_Time_Hour_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetAdd___closed__1); +l_Std_Time_Hour_instOffsetAdd = _init_l_Std_Time_Hour_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetAdd); +l_Std_Time_Hour_instOffsetSub___closed__1 = _init_l_Std_Time_Hour_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetSub___closed__1); +l_Std_Time_Hour_instOffsetSub = _init_l_Std_Time_Hour_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetSub); +l_Std_Time_Hour_instOffsetNeg___closed__1 = _init_l_Std_Time_Hour_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetNeg___closed__1); +l_Std_Time_Hour_instOffsetNeg = _init_l_Std_Time_Hour_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetNeg); +l_Std_Time_Hour_instOffsetToString___closed__1 = _init_l_Std_Time_Hour_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetToString___closed__1); +l_Std_Time_Hour_instOffsetToString = _init_l_Std_Time_Hour_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Hour_instOffsetToString); +l_Std_Time_Hour_Ordinal_toRelative___closed__1 = _init_l_Std_Time_Hour_Ordinal_toRelative___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_toRelative___closed__1); +l_Std_Time_Hour_Ordinal_toRelative___closed__2 = _init_l_Std_Time_Hour_Ordinal_toRelative___closed__2(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_toRelative___closed__2); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__1); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__2); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__3); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__4); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__5); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__6); +l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7 = _init_l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7(); +lean_mark_persistent(l_Std_Time_Hour_Ordinal_shiftTo1BasedHour___closed__7); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c b/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c new file mode 100644 index 000000000000..eff3dae7cf2f --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Millisecond.c @@ -0,0 +1,576 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Millisecond +// Imports: Std.Internal.Rat Std.Time.Internal Std.Time.Time.Unit.Nanosecond +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Millisecond_instOffsetNeg___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetLT; +static lean_object* l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_toOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofInt___boxed(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofInt(lean_object*); +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofInt___boxed(lean_object*); +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10; +static lean_object* l_Std_Time_Millisecond_instOffsetSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNat(lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOrdinalRepr; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetLE; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofInt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Millisecond_instOrdinalRepr___closed__1; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOrdinalLT; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetAdd; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOfNatOrdinal(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLeOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_toOffset___boxed(lean_object*); +static lean_object* l_Std_Time_Millisecond_instOffsetRepr___closed__1; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetInhabited; +static lean_object* l_Std_Time_Millisecond_instOrdinalBEq___closed__1; +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOrdinalLE; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetNeg; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetSub; +static lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetBEq; +static lean_object* l_Std_Time_Millisecond_instOffsetAdd___closed__1; +lean_object* lean_int_sub(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofFin(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOrdinalBEq; +static lean_object* l_Std_Time_Millisecond_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOfNatOffset(lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetRepr; +static lean_object* l_Std_Time_Millisecond_instOffsetBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instInhabitedOrdinal; +lean_object* lean_int_add(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOffsetToString; +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLtOrdinal(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(999u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(999u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9; +x_2 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Millisecond_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Millisecond_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Millisecond_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOfNatOrdinal___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Millisecond_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Millisecond_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Millisecond_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Millisecond_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Millisecond_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Nanosecond(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Millisecond(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Nanosecond(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Millisecond_instOrdinalRepr___closed__1 = _init_l_Std_Time_Millisecond_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalRepr___closed__1); +l_Std_Time_Millisecond_instOrdinalRepr = _init_l_Std_Time_Millisecond_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalRepr); +l_Std_Time_Millisecond_instOrdinalBEq___closed__1 = _init_l_Std_Time_Millisecond_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalBEq___closed__1); +l_Std_Time_Millisecond_instOrdinalBEq = _init_l_Std_Time_Millisecond_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalBEq); +l_Std_Time_Millisecond_instOrdinalLE = _init_l_Std_Time_Millisecond_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalLE); +l_Std_Time_Millisecond_instOrdinalLT = _init_l_Std_Time_Millisecond_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Millisecond_instOrdinalLT); +l_Std_Time_Millisecond_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Millisecond_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOfNatOrdinal___closed__1); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__1); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__2); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__3); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__4); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__5); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__6); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__7); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__8); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__9); +l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10 = _init_l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal___closed__10); +l_Std_Time_Millisecond_instInhabitedOrdinal = _init_l_Std_Time_Millisecond_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Millisecond_instInhabitedOrdinal); +l_Std_Time_Millisecond_instOffsetRepr___closed__1 = _init_l_Std_Time_Millisecond_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetRepr___closed__1); +l_Std_Time_Millisecond_instOffsetRepr = _init_l_Std_Time_Millisecond_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetRepr); +l_Std_Time_Millisecond_instOffsetBEq___closed__1 = _init_l_Std_Time_Millisecond_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetBEq___closed__1); +l_Std_Time_Millisecond_instOffsetBEq = _init_l_Std_Time_Millisecond_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetBEq); +l_Std_Time_Millisecond_instOffsetInhabited = _init_l_Std_Time_Millisecond_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetInhabited); +l_Std_Time_Millisecond_instOffsetAdd___closed__1 = _init_l_Std_Time_Millisecond_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetAdd___closed__1); +l_Std_Time_Millisecond_instOffsetAdd = _init_l_Std_Time_Millisecond_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetAdd); +l_Std_Time_Millisecond_instOffsetSub___closed__1 = _init_l_Std_Time_Millisecond_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetSub___closed__1); +l_Std_Time_Millisecond_instOffsetSub = _init_l_Std_Time_Millisecond_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetSub); +l_Std_Time_Millisecond_instOffsetNeg___closed__1 = _init_l_Std_Time_Millisecond_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetNeg___closed__1); +l_Std_Time_Millisecond_instOffsetNeg = _init_l_Std_Time_Millisecond_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetNeg); +l_Std_Time_Millisecond_instOffsetLE = _init_l_Std_Time_Millisecond_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetLE); +l_Std_Time_Millisecond_instOffsetLT = _init_l_Std_Time_Millisecond_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetLT); +l_Std_Time_Millisecond_instOffsetToString___closed__1 = _init_l_Std_Time_Millisecond_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetToString___closed__1); +l_Std_Time_Millisecond_instOffsetToString = _init_l_Std_Time_Millisecond_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Millisecond_instOffsetToString); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Minute.c b/stage0/stdlib/Std/Time/Time/Unit/Minute.c new file mode 100644 index 000000000000..6197984455e6 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Minute.c @@ -0,0 +1,554 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Minute +// Imports: Std.Internal.Rat Std.Time.Internal Std.Time.Time.Unit.Second +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofInt___boxed(lean_object*); +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__2; +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__6; +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instOfNatOrdinal___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_toOffset___boxed(lean_object*); +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__9; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetRepr; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetAdd; +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instInhabitedOrdinal; +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofFin(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalBEq; +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalLT; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalRepr; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instOrdinalBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_toOffset(lean_object*); +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__3; +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__1; +lean_object* lean_int_sub(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetToString; +static lean_object* l_Std_Time_Minute_instOffsetRepr___closed__1; +static lean_object* l_Std_Time_Minute_instOrdinalRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofInt(lean_object*); +static lean_object* l_Std_Time_Minute_instOffsetNeg___closed__1; +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOrdinalLE; +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__5; +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLtOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOfNatOrdinal(lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instOffsetBEq___closed__1; +static lean_object* l_Std_Time_Minute_instInhabitedOrdinal___closed__8; +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLeOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOfNatOffset(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetNeg; +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetBEq; +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOffsetSub; +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Minute_instOffsetAdd___closed__1; +static lean_object* l_Std_Time_Minute_instOffsetSub___closed__1; +static lean_object* _init_l_Std_Time_Minute_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +x_3 = lean_unsigned_to_nat(59u); +x_4 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_2, x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +x_2 = l_Std_Time_Minute_instInhabitedOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__2; +x_2 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__3; +x_2 = l_Std_Time_Minute_instInhabitedOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__6; +x_2 = l_Std_Time_Minute_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__7; +x_2 = l_Std_Time_Minute_instInhabitedOrdinal___closed__5; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__8; +x_2 = l_Std_Time_Minute_instInhabitedOrdinal___closed__5; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__9; +x_2 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Minute_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instInhabitedOrdinal___closed__10; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Minute_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Minute_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Minute_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOfNatOrdinal___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Minute_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Minute_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Minute_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Minute_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Minute_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Second(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Minute(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Second(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Minute_instOrdinalRepr___closed__1 = _init_l_Std_Time_Minute_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalRepr___closed__1); +l_Std_Time_Minute_instOrdinalRepr = _init_l_Std_Time_Minute_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalRepr); +l_Std_Time_Minute_instOrdinalBEq___closed__1 = _init_l_Std_Time_Minute_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalBEq___closed__1); +l_Std_Time_Minute_instOrdinalBEq = _init_l_Std_Time_Minute_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalBEq); +l_Std_Time_Minute_instOrdinalLE = _init_l_Std_Time_Minute_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalLE); +l_Std_Time_Minute_instOrdinalLT = _init_l_Std_Time_Minute_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Minute_instOrdinalLT); +l_Std_Time_Minute_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Minute_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOfNatOrdinal___closed__1); +l_Std_Time_Minute_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__1); +l_Std_Time_Minute_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__2); +l_Std_Time_Minute_instInhabitedOrdinal___closed__3 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__3); +l_Std_Time_Minute_instInhabitedOrdinal___closed__4 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__4); +l_Std_Time_Minute_instInhabitedOrdinal___closed__5 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__5); +l_Std_Time_Minute_instInhabitedOrdinal___closed__6 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__6(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__6); +l_Std_Time_Minute_instInhabitedOrdinal___closed__7 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__7(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__7); +l_Std_Time_Minute_instInhabitedOrdinal___closed__8 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__8(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__8); +l_Std_Time_Minute_instInhabitedOrdinal___closed__9 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__9(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__9); +l_Std_Time_Minute_instInhabitedOrdinal___closed__10 = _init_l_Std_Time_Minute_instInhabitedOrdinal___closed__10(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal___closed__10); +l_Std_Time_Minute_instInhabitedOrdinal = _init_l_Std_Time_Minute_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Minute_instInhabitedOrdinal); +l_Std_Time_Minute_instOffsetRepr___closed__1 = _init_l_Std_Time_Minute_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetRepr___closed__1); +l_Std_Time_Minute_instOffsetRepr = _init_l_Std_Time_Minute_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetRepr); +l_Std_Time_Minute_instOffsetBEq___closed__1 = _init_l_Std_Time_Minute_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetBEq___closed__1); +l_Std_Time_Minute_instOffsetBEq = _init_l_Std_Time_Minute_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetBEq); +l_Std_Time_Minute_instOffsetInhabited = _init_l_Std_Time_Minute_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetInhabited); +l_Std_Time_Minute_instOffsetAdd___closed__1 = _init_l_Std_Time_Minute_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetAdd___closed__1); +l_Std_Time_Minute_instOffsetAdd = _init_l_Std_Time_Minute_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetAdd); +l_Std_Time_Minute_instOffsetSub___closed__1 = _init_l_Std_Time_Minute_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetSub___closed__1); +l_Std_Time_Minute_instOffsetSub = _init_l_Std_Time_Minute_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetSub); +l_Std_Time_Minute_instOffsetNeg___closed__1 = _init_l_Std_Time_Minute_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetNeg___closed__1); +l_Std_Time_Minute_instOffsetNeg = _init_l_Std_Time_Minute_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetNeg); +l_Std_Time_Minute_instOffsetToString___closed__1 = _init_l_Std_Time_Minute_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetToString___closed__1); +l_Std_Time_Minute_instOffsetToString = _init_l_Std_Time_Minute_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Minute_instOffsetToString); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c b/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c new file mode 100644 index 000000000000..00cc03dc0bc3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Nanosecond.c @@ -0,0 +1,615 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Nanosecond +// Imports: Std.Internal.Rat Std.Time.Internal +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Nanosecond_instOffsetSub___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalLE; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofFin(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetLT; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalBEq; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLeOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofInt___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetRepr; +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOrdinal___boxed(lean_object*); +static lean_object* l_Std_Time_Nanosecond_instOrdinalBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_instOfDayBEq; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instInhabitedSpan; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLeOffset___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetLE; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetNeg; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Span_toOffset(lean_object*); +lean_object* l_Fin_ofNat(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofInt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instSpanLE; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_instOfDayRepr; +lean_object* l_Std_Time_Internal_Bounded_instRepr___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Nanosecond_instOffsetToString___closed__1; +lean_object* lean_nat_to_int(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLtOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_toOffset(lean_object*); +static lean_object* l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1; +static lean_object* l_Std_Time_Nanosecond_instOrdinalRepr___closed__1; +static lean_object* l_Std_Time_Nanosecond_instOffsetInhabited___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalLT; +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOrdinal(lean_object*); +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOrdinalRepr; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Span_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetToString; +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLeOffset(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLeOrdinal(lean_object*, lean_object*); +static lean_object* l_Std_Time_Nanosecond_instOffsetBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_instOfDayLT; +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLtOrdinal(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instSpanBEq; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instSpanLT; +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_instOfDayLE; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instSpanRepr; +static lean_object* l_Std_Time_Nanosecond_instOffsetAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofInt___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetAdd; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instInhabitedOrdinal; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetSub; +static lean_object* l_Std_Time_Nanosecond_instOffsetRepr___closed__1; +lean_object* l_Std_Time_Internal_Bounded_instBEq___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOffsetBEq; +static lean_object* l_Std_Time_Nanosecond_instOffsetNeg___closed__1; +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_Bounded_instBEq___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOrdinalLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOrdinal(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_unsigned_to_nat(999999999u); +x_3 = l_Fin_ofNat(x_2, x_1); +x_4 = lean_nat_to_int(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOrdinal___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_instOfNatOrdinal(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(999999999u); +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Fin_ofNat(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instInhabitedOrdinal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLeOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLeOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Nanosecond_instDecidableLeOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLtOrdinal(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLtOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Nanosecond_instDecidableLtOrdinal(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetInhabited___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetInhabited___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Nanosecond_instDecidableLeOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instDecidableLeOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Nanosecond_instDecidableLeOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instSpanRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instSpanBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instSpanLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instSpanLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_instInhabitedSpan() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOffsetInhabited___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Span_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Span_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Span_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Ordinal_instOfDayRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Ordinal_instOfDayBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Nanosecond_instOrdinalBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Ordinal_instOfDayLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Nanosecond_Ordinal_instOfDayLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofInt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofInt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Nanosecond_Ordinal_ofInt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofNat(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_ofFin(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_toOffset(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Nanosecond_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Nanosecond_Ordinal_toOffset(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Internal(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Nanosecond(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Nanosecond_instOrdinalRepr___closed__1 = _init_l_Std_Time_Nanosecond_instOrdinalRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalRepr___closed__1); +l_Std_Time_Nanosecond_instOrdinalRepr = _init_l_Std_Time_Nanosecond_instOrdinalRepr(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalRepr); +l_Std_Time_Nanosecond_instOrdinalBEq___closed__1 = _init_l_Std_Time_Nanosecond_instOrdinalBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalBEq___closed__1); +l_Std_Time_Nanosecond_instOrdinalBEq = _init_l_Std_Time_Nanosecond_instOrdinalBEq(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalBEq); +l_Std_Time_Nanosecond_instOrdinalLE = _init_l_Std_Time_Nanosecond_instOrdinalLE(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalLE); +l_Std_Time_Nanosecond_instOrdinalLT = _init_l_Std_Time_Nanosecond_instOrdinalLT(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOrdinalLT); +l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1 = _init_l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__1); +l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2 = _init_l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Nanosecond_instInhabitedOrdinal___closed__2); +l_Std_Time_Nanosecond_instInhabitedOrdinal = _init_l_Std_Time_Nanosecond_instInhabitedOrdinal(); +lean_mark_persistent(l_Std_Time_Nanosecond_instInhabitedOrdinal); +l_Std_Time_Nanosecond_instOffsetRepr___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetRepr___closed__1); +l_Std_Time_Nanosecond_instOffsetRepr = _init_l_Std_Time_Nanosecond_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetRepr); +l_Std_Time_Nanosecond_instOffsetBEq___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetBEq___closed__1); +l_Std_Time_Nanosecond_instOffsetBEq = _init_l_Std_Time_Nanosecond_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetBEq); +l_Std_Time_Nanosecond_instOffsetInhabited___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetInhabited___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetInhabited___closed__1); +l_Std_Time_Nanosecond_instOffsetInhabited = _init_l_Std_Time_Nanosecond_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetInhabited); +l_Std_Time_Nanosecond_instOffsetAdd___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetAdd___closed__1); +l_Std_Time_Nanosecond_instOffsetAdd = _init_l_Std_Time_Nanosecond_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetAdd); +l_Std_Time_Nanosecond_instOffsetSub___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetSub___closed__1); +l_Std_Time_Nanosecond_instOffsetSub = _init_l_Std_Time_Nanosecond_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetSub); +l_Std_Time_Nanosecond_instOffsetNeg___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetNeg___closed__1); +l_Std_Time_Nanosecond_instOffsetNeg = _init_l_Std_Time_Nanosecond_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetNeg); +l_Std_Time_Nanosecond_instOffsetLE = _init_l_Std_Time_Nanosecond_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetLE); +l_Std_Time_Nanosecond_instOffsetLT = _init_l_Std_Time_Nanosecond_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetLT); +l_Std_Time_Nanosecond_instOffsetToString___closed__1 = _init_l_Std_Time_Nanosecond_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetToString___closed__1); +l_Std_Time_Nanosecond_instOffsetToString = _init_l_Std_Time_Nanosecond_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Nanosecond_instOffsetToString); +l_Std_Time_Nanosecond_instSpanRepr = _init_l_Std_Time_Nanosecond_instSpanRepr(); +lean_mark_persistent(l_Std_Time_Nanosecond_instSpanRepr); +l_Std_Time_Nanosecond_instSpanBEq = _init_l_Std_Time_Nanosecond_instSpanBEq(); +lean_mark_persistent(l_Std_Time_Nanosecond_instSpanBEq); +l_Std_Time_Nanosecond_instSpanLE = _init_l_Std_Time_Nanosecond_instSpanLE(); +lean_mark_persistent(l_Std_Time_Nanosecond_instSpanLE); +l_Std_Time_Nanosecond_instSpanLT = _init_l_Std_Time_Nanosecond_instSpanLT(); +lean_mark_persistent(l_Std_Time_Nanosecond_instSpanLT); +l_Std_Time_Nanosecond_instInhabitedSpan = _init_l_Std_Time_Nanosecond_instInhabitedSpan(); +lean_mark_persistent(l_Std_Time_Nanosecond_instInhabitedSpan); +l_Std_Time_Nanosecond_Ordinal_instOfDayRepr = _init_l_Std_Time_Nanosecond_Ordinal_instOfDayRepr(); +lean_mark_persistent(l_Std_Time_Nanosecond_Ordinal_instOfDayRepr); +l_Std_Time_Nanosecond_Ordinal_instOfDayBEq = _init_l_Std_Time_Nanosecond_Ordinal_instOfDayBEq(); +lean_mark_persistent(l_Std_Time_Nanosecond_Ordinal_instOfDayBEq); +l_Std_Time_Nanosecond_Ordinal_instOfDayLE = _init_l_Std_Time_Nanosecond_Ordinal_instOfDayLE(); +lean_mark_persistent(l_Std_Time_Nanosecond_Ordinal_instOfDayLE); +l_Std_Time_Nanosecond_Ordinal_instOfDayLT = _init_l_Std_Time_Nanosecond_Ordinal_instOfDayLT(); +lean_mark_persistent(l_Std_Time_Nanosecond_Ordinal_instOfDayLT); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Time/Unit/Second.c b/stage0/stdlib/Std/Time/Time/Unit/Second.c new file mode 100644 index 000000000000..7a364fa63304 --- /dev/null +++ b/stage0/stdlib/Std/Time/Time/Unit/Second.c @@ -0,0 +1,722 @@ +// Lean compiler output +// Module: Std.Time.Time.Unit.Second +// Imports: Std.Internal.Rat Std.Time.Time.Unit.Nanosecond +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetNeg; +LEAN_EXPORT uint8_t l_Std_Time_Second_instBEqOrdinal___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetLE; +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat___rarg(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instToString___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOrdinal___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofInt(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___rarg___boxed(lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_instOffsetRepr___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___rarg(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___rarg(lean_object*); +lean_object* l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instLEOrdinal(uint8_t); +lean_object* l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___boxed(lean_object*); +static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetAdd; +lean_object* l_Int_repr(lean_object*); +static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__3; +lean_object* l_Std_Time_Internal_UnitVal_add___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_instOffsetToString___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetRepr; +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetLT; +lean_object* l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin(uint8_t); +static lean_object* l_Std_Time_Second_instOffsetBEq___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofInt___boxed(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOrdinal(uint8_t, lean_object*); +static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__1; +lean_object* l_Std_Time_Internal_UnitVal_sub___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_instOfNatOrdinal___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Second_instLEOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOffset(lean_object*); +static lean_object* l_Std_Time_Second_instOffsetSub___closed__1; +lean_object* l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___rarg___boxed(lean_object*, lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_instOffsetNeg___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetSub; +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal___boxed(lean_object*); +static lean_object* l_Std_Time_Second_instOffsetAdd___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetBEq; +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLtOrdinal___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin___boxed(lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLeOrdinal___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instLTOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNat(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetInhabited; +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_Second_instOffsetToString; +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Second_instLTOrdinal(uint8_t); +LEAN_EXPORT uint8_t l_Std_Time_Second_instBEqOrdinal___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_instBEqOrdinal___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Second_instBEqOrdinal___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instBEqOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instBEqOrdinal(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instLEOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instLEOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instLEOrdinal(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instLTOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instLTOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instLTOrdinal(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Second_instReprOrdinal___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +x_4 = lean_int_dec_lt(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Int_repr(x_1); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l_Int_repr(x_1); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_instReprOrdinal___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Second_instReprOrdinal___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instReprOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instReprOrdinal(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Second_instOfNatOrdinal___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_instOfNatOrdinal___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +x_2 = l_Std_Time_Second_instOfNatOrdinal___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Second_instOfNatOrdinal___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Second_instOfNatOrdinal___closed__2; +x_2 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Second_instOfNatOrdinal___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Second_instOfNatOrdinal___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Second_instOfNatOrdinal___closed__3; +x_2 = l_Std_Time_Second_instOfNatOrdinal___closed__4; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOrdinal(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +x_4 = lean_unsigned_to_nat(59u); +x_5 = l_Std_Time_Internal_Bounded_LE_instOfNatHAddIntCast(x_3, x_2, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_6 = lean_nat_to_int(x_2); +x_7 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +x_8 = lean_int_sub(x_6, x_7); +lean_dec(x_6); +x_9 = l_Std_Time_Second_instOfNatOrdinal___closed__5; +x_10 = lean_int_emod(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_add(x_10, x_9); +lean_dec(x_10); +x_12 = lean_int_emod(x_11, x_9); +lean_dec(x_11); +x_13 = lean_int_add(x_12, x_7); +lean_dec(x_12); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOrdinal___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l_Std_Time_Second_instOfNatOrdinal(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLeOrdinal___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_le(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_instDecidableLeOrdinal___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Second_instDecidableLeOrdinal___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLeOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instDecidableLeOrdinal(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_Second_instDecidableLtOrdinal___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_lt(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_instDecidableLtOrdinal___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_Second_instDecidableLtOrdinal___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instDecidableLtOrdinal___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_instDecidableLtOrdinal(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetRepr___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instRepr___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetRepr() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetRepr___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetBEq___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Internal_UnitVal_0__Std_Time_Internal_beqUnitVal____x40_Std_Time_Internal_UnitVal___hyg_35____rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetBEq() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetBEq___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetInhabited() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instReprOrdinal___rarg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetAdd___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_add___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetAdd() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetAdd___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetSub___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_sub___rarg___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetSub() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetSub___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetNeg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instNeg___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetNeg() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetNeg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetLE() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetLT() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetToString___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Internal_UnitVal_instToString___rarg___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Second_instOffsetToString() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Second_instOffsetToString___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_instOfNatOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofNat(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofInt(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Offset_ofInt___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Offset_ofInt(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_Ordinal_ofInt___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Second_Ordinal_ofInt___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofInt___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_Ordinal_ofInt(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_nat_to_int(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_Ordinal_ofNat___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofNat___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_Ordinal_ofNat(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_Ordinal_ofFin___rarg), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_ofFin___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_Ordinal_ofFin(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset(uint8_t x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Second_Ordinal_toOffset___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Second_Ordinal_toOffset___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Second_Ordinal_toOffset___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_Second_Ordinal_toOffset(x_2); +return x_3; +} +} +lean_object* initialize_Std_Internal_Rat(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Time_Unit_Nanosecond(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Time_Unit_Second(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Internal_Rat(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Time_Unit_Nanosecond(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Second_instReprOrdinal___rarg___closed__1 = _init_l_Std_Time_Second_instReprOrdinal___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instReprOrdinal___rarg___closed__1); +l_Std_Time_Second_instOfNatOrdinal___closed__1 = _init_l_Std_Time_Second_instOfNatOrdinal___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOfNatOrdinal___closed__1); +l_Std_Time_Second_instOfNatOrdinal___closed__2 = _init_l_Std_Time_Second_instOfNatOrdinal___closed__2(); +lean_mark_persistent(l_Std_Time_Second_instOfNatOrdinal___closed__2); +l_Std_Time_Second_instOfNatOrdinal___closed__3 = _init_l_Std_Time_Second_instOfNatOrdinal___closed__3(); +lean_mark_persistent(l_Std_Time_Second_instOfNatOrdinal___closed__3); +l_Std_Time_Second_instOfNatOrdinal___closed__4 = _init_l_Std_Time_Second_instOfNatOrdinal___closed__4(); +lean_mark_persistent(l_Std_Time_Second_instOfNatOrdinal___closed__4); +l_Std_Time_Second_instOfNatOrdinal___closed__5 = _init_l_Std_Time_Second_instOfNatOrdinal___closed__5(); +lean_mark_persistent(l_Std_Time_Second_instOfNatOrdinal___closed__5); +l_Std_Time_Second_instOffsetRepr___closed__1 = _init_l_Std_Time_Second_instOffsetRepr___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetRepr___closed__1); +l_Std_Time_Second_instOffsetRepr = _init_l_Std_Time_Second_instOffsetRepr(); +lean_mark_persistent(l_Std_Time_Second_instOffsetRepr); +l_Std_Time_Second_instOffsetBEq___closed__1 = _init_l_Std_Time_Second_instOffsetBEq___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetBEq___closed__1); +l_Std_Time_Second_instOffsetBEq = _init_l_Std_Time_Second_instOffsetBEq(); +lean_mark_persistent(l_Std_Time_Second_instOffsetBEq); +l_Std_Time_Second_instOffsetInhabited = _init_l_Std_Time_Second_instOffsetInhabited(); +lean_mark_persistent(l_Std_Time_Second_instOffsetInhabited); +l_Std_Time_Second_instOffsetAdd___closed__1 = _init_l_Std_Time_Second_instOffsetAdd___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetAdd___closed__1); +l_Std_Time_Second_instOffsetAdd = _init_l_Std_Time_Second_instOffsetAdd(); +lean_mark_persistent(l_Std_Time_Second_instOffsetAdd); +l_Std_Time_Second_instOffsetSub___closed__1 = _init_l_Std_Time_Second_instOffsetSub___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetSub___closed__1); +l_Std_Time_Second_instOffsetSub = _init_l_Std_Time_Second_instOffsetSub(); +lean_mark_persistent(l_Std_Time_Second_instOffsetSub); +l_Std_Time_Second_instOffsetNeg___closed__1 = _init_l_Std_Time_Second_instOffsetNeg___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetNeg___closed__1); +l_Std_Time_Second_instOffsetNeg = _init_l_Std_Time_Second_instOffsetNeg(); +lean_mark_persistent(l_Std_Time_Second_instOffsetNeg); +l_Std_Time_Second_instOffsetLE = _init_l_Std_Time_Second_instOffsetLE(); +lean_mark_persistent(l_Std_Time_Second_instOffsetLE); +l_Std_Time_Second_instOffsetLT = _init_l_Std_Time_Second_instOffsetLT(); +lean_mark_persistent(l_Std_Time_Second_instOffsetLT); +l_Std_Time_Second_instOffsetToString___closed__1 = _init_l_Std_Time_Second_instOffsetToString___closed__1(); +lean_mark_persistent(l_Std_Time_Second_instOffsetToString___closed__1); +l_Std_Time_Second_instOffsetToString = _init_l_Std_Time_Second_instOffsetToString(); +lean_mark_persistent(l_Std_Time_Second_instOffsetToString); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned.c b/stage0/stdlib/Std/Time/Zoned.c new file mode 100644 index 000000000000..f69aad12dabe --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned.c @@ -0,0 +1,3042 @@ +// Lean compiler output +// Module: Std.Time.Zoned +// Imports: Std.Time.Zoned.DateTime Std.Time.Zoned.ZoneRules Std.Time.Zoned.ZonedDateTime Std.Time.Zoned.Database +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainTime(lean_object*); +lean_object* l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(lean_object*, lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampWithZone(lean_object*, lean_object*); +lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nowAt(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampWithZone___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now___lambda__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestamp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampWithZone(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_now(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_now(lean_object*); +lean_object* l_Std_Time_Database_defaultGetLocalZoneRules(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampWithZone___boxed(lean_object*, lean_object*); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_now___closed__2; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +extern lean_object* l_Std_Time_PlainTime_midnight; +lean_object* l_Array_back_x3f___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_now(lean_object*); +static lean_object* l_Std_Time_DateTime_ofPlainDate___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate(lean_object*); +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___boxed(lean_object*); +lean_object* lean_thunk_get_own(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDate___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_thunk(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDate(lean_object*); +lean_object* l_Std_Time_TimeZone_Transition_timezoneAt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestamp(lean_object*, lean_object*); +lean_object* lean_nat_abs(lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_now(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___rarg(lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___boxed(lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___rarg(lean_object*); +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1; +lean_object* lean_get_current_time(lean_object*); +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainTime___boxed(lean_object*); +lean_object* l_Std_Time_Database_defaultGetZoneRules(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_of(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_PlainDateTime_now___closed__1; +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_PlainDateTime_now___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_PlainDateTime_now___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_now(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_get_current_time(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Std_Time_Database_defaultGetLocalZoneRules(x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_inc(x_3); +x_9 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_8, x_3); +lean_dec(x_8); +x_10 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +x_12 = l_Std_Time_PlainTime_toSeconds(x_11); +x_13 = l_Std_Time_PlainTime_toNanoseconds(x_11); +lean_dec(x_11); +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_16 = lean_ctor_get(x_7, 0); +lean_inc(x_16); +lean_dec(x_7); +x_17 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_int_add(x_12, x_18); +lean_dec(x_12); +x_20 = l_Std_Time_PlainDateTime_now___closed__1; +x_21 = lean_int_ediv(x_19, x_20); +lean_dec(x_19); +x_22 = l_Std_Time_PlainDateTime_now___closed__2; +x_23 = lean_int_mul(x_18, x_22); +lean_dec(x_18); +x_24 = lean_int_add(x_13, x_23); +lean_dec(x_23); +lean_dec(x_13); +x_25 = l_Std_Time_PlainTime_ofNanoseconds(x_24); +lean_dec(x_24); +x_26 = lean_int_add(x_15, x_21); +lean_dec(x_21); +lean_dec(x_15); +x_27 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +lean_ctor_set(x_5, 0, x_28); +return x_5; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_7); +x_29 = lean_ctor_get(x_9, 0); +lean_inc(x_29); +lean_dec(x_9); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_30); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_int_add(x_12, x_32); +lean_dec(x_12); +x_34 = l_Std_Time_PlainDateTime_now___closed__1; +x_35 = lean_int_ediv(x_33, x_34); +lean_dec(x_33); +x_36 = l_Std_Time_PlainDateTime_now___closed__2; +x_37 = lean_int_mul(x_32, x_36); +lean_dec(x_32); +x_38 = lean_int_add(x_13, x_37); +lean_dec(x_37); +lean_dec(x_13); +x_39 = l_Std_Time_PlainTime_ofNanoseconds(x_38); +lean_dec(x_38); +x_40 = lean_int_add(x_15, x_35); +lean_dec(x_35); +lean_dec(x_15); +x_41 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_40); +lean_dec(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_5, 0, x_42); +return x_5; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_43 = lean_ctor_get(x_5, 0); +x_44 = lean_ctor_get(x_5, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_5); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_inc(x_3); +x_46 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_45, x_3); +lean_dec(x_45); +x_47 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +x_49 = l_Std_Time_PlainTime_toSeconds(x_48); +x_50 = l_Std_Time_PlainTime_toNanoseconds(x_48); +lean_dec(x_48); +x_51 = lean_ctor_get(x_47, 0); +lean_inc(x_51); +lean_dec(x_47); +x_52 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_51); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_53 = lean_ctor_get(x_43, 0); +lean_inc(x_53); +lean_dec(x_43); +x_54 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_53); +lean_dec(x_53); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_int_add(x_49, x_55); +lean_dec(x_49); +x_57 = l_Std_Time_PlainDateTime_now___closed__1; +x_58 = lean_int_ediv(x_56, x_57); +lean_dec(x_56); +x_59 = l_Std_Time_PlainDateTime_now___closed__2; +x_60 = lean_int_mul(x_55, x_59); +lean_dec(x_55); +x_61 = lean_int_add(x_50, x_60); +lean_dec(x_60); +lean_dec(x_50); +x_62 = l_Std_Time_PlainTime_ofNanoseconds(x_61); +lean_dec(x_61); +x_63 = lean_int_add(x_52, x_58); +lean_dec(x_58); +lean_dec(x_52); +x_64 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_63); +lean_dec(x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_44); +return x_66; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_43); +x_67 = lean_ctor_get(x_46, 0); +lean_inc(x_67); +lean_dec(x_46); +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +lean_dec(x_67); +x_69 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_68); +lean_dec(x_68); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +lean_dec(x_69); +x_71 = lean_int_add(x_49, x_70); +lean_dec(x_49); +x_72 = l_Std_Time_PlainDateTime_now___closed__1; +x_73 = lean_int_ediv(x_71, x_72); +lean_dec(x_71); +x_74 = l_Std_Time_PlainDateTime_now___closed__2; +x_75 = lean_int_mul(x_70, x_74); +lean_dec(x_70); +x_76 = lean_int_add(x_50, x_75); +lean_dec(x_75); +lean_dec(x_50); +x_77 = l_Std_Time_PlainTime_ofNanoseconds(x_76); +lean_dec(x_76); +x_78 = lean_int_add(x_52, x_73); +lean_dec(x_73); +lean_dec(x_52); +x_79 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_78); +lean_dec(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_77); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_44); +return x_81; +} +} +} +else +{ +uint8_t x_82; +lean_dec(x_3); +x_82 = !lean_is_exclusive(x_5); +if (x_82 == 0) +{ +return x_5; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_5, 0); +x_84 = lean_ctor_get(x_5, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_5); +x_85 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +return x_85; +} +} +} +else +{ +uint8_t x_86; +x_86 = !lean_is_exclusive(x_2); +if (x_86 == 0) +{ +return x_2; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_2, 0); +x_88 = lean_ctor_get(x_2, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_2); +x_89 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +return x_89; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_now(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_get_current_time(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Std_Time_Database_defaultGetLocalZoneRules(x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_inc(x_3); +x_9 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_8, x_3); +lean_dec(x_8); +x_10 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +x_12 = l_Std_Time_PlainTime_toSeconds(x_11); +lean_dec(x_11); +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_15 = lean_ctor_get(x_7, 0); +lean_inc(x_15); +lean_dec(x_7); +x_16 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_15); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_int_add(x_12, x_17); +lean_dec(x_17); +lean_dec(x_12); +x_19 = l_Std_Time_PlainDateTime_now___closed__1; +x_20 = lean_int_ediv(x_18, x_19); +lean_dec(x_18); +x_21 = lean_int_add(x_14, x_20); +lean_dec(x_20); +lean_dec(x_14); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +lean_ctor_set(x_5, 0, x_22); +return x_5; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_7); +x_23 = lean_ctor_get(x_9, 0); +lean_inc(x_23); +lean_dec(x_9); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +x_25 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_24); +lean_dec(x_24); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_int_add(x_12, x_26); +lean_dec(x_26); +lean_dec(x_12); +x_28 = l_Std_Time_PlainDateTime_now___closed__1; +x_29 = lean_int_ediv(x_27, x_28); +lean_dec(x_27); +x_30 = lean_int_add(x_14, x_29); +lean_dec(x_29); +lean_dec(x_14); +x_31 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_30); +lean_dec(x_30); +lean_ctor_set(x_5, 0, x_31); +return x_5; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_inc(x_3); +x_35 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_34, x_3); +lean_dec(x_34); +x_36 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +x_38 = l_Std_Time_PlainTime_toSeconds(x_37); +lean_dec(x_37); +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_40 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_39); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_41 = lean_ctor_get(x_32, 0); +lean_inc(x_41); +lean_dec(x_32); +x_42 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_41); +lean_dec(x_41); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_int_add(x_38, x_43); +lean_dec(x_43); +lean_dec(x_38); +x_45 = l_Std_Time_PlainDateTime_now___closed__1; +x_46 = lean_int_ediv(x_44, x_45); +lean_dec(x_44); +x_47 = lean_int_add(x_40, x_46); +lean_dec(x_46); +lean_dec(x_40); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_33); +return x_49; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_32); +x_50 = lean_ctor_get(x_35, 0); +lean_inc(x_50); +lean_dec(x_35); +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +lean_dec(x_50); +x_52 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_51); +lean_dec(x_51); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_int_add(x_38, x_53); +lean_dec(x_53); +lean_dec(x_38); +x_55 = l_Std_Time_PlainDateTime_now___closed__1; +x_56 = lean_int_ediv(x_54, x_55); +lean_dec(x_54); +x_57 = lean_int_add(x_40, x_56); +lean_dec(x_56); +lean_dec(x_40); +x_58 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_57); +lean_dec(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_33); +return x_59; +} +} +} +else +{ +uint8_t x_60; +lean_dec(x_3); +x_60 = !lean_is_exclusive(x_5); +if (x_60 == 0) +{ +return x_5; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_5, 0); +x_62 = lean_ctor_get(x_5, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_5); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +uint8_t x_64; +x_64 = !lean_is_exclusive(x_2); +if (x_64 == 0) +{ +return x_2; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_2, 0); +x_66 = lean_ctor_get(x_2, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_2); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainTime_now(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_get_current_time(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Std_Time_Database_defaultGetLocalZoneRules(x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_7, 1); +lean_inc(x_8); +lean_inc(x_3); +x_9 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_8, x_3); +lean_dec(x_8); +x_10 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_11); +lean_dec(x_11); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_13 = lean_ctor_get(x_7, 0); +lean_inc(x_13); +lean_dec(x_7); +x_14 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Std_Time_PlainDateTime_now___closed__2; +x_17 = lean_int_mul(x_15, x_16); +lean_dec(x_15); +x_18 = lean_int_add(x_12, x_17); +lean_dec(x_17); +lean_dec(x_12); +x_19 = l_Std_Time_PlainTime_ofNanoseconds(x_18); +lean_dec(x_18); +lean_ctor_set(x_5, 0, x_19); +return x_5; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_7); +x_20 = lean_ctor_get(x_9, 0); +lean_inc(x_20); +lean_dec(x_9); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDateTime_now___closed__2; +x_25 = lean_int_mul(x_23, x_24); +lean_dec(x_23); +x_26 = lean_int_add(x_12, x_25); +lean_dec(x_25); +lean_dec(x_12); +x_27 = l_Std_Time_PlainTime_ofNanoseconds(x_26); +lean_dec(x_26); +lean_ctor_set(x_5, 0, x_27); +return x_5; +} +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_28 = lean_ctor_get(x_5, 0); +x_29 = lean_ctor_get(x_5, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_5); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_inc(x_3); +x_31 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_30, x_3); +lean_dec(x_30); +x_32 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +lean_dec(x_32); +x_34 = l_Std_Time_PlainTime_toNanoseconds(x_33); +lean_dec(x_33); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_35 = lean_ctor_get(x_28, 0); +lean_inc(x_35); +lean_dec(x_28); +x_36 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_35); +lean_dec(x_35); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Std_Time_PlainDateTime_now___closed__2; +x_39 = lean_int_mul(x_37, x_38); +lean_dec(x_37); +x_40 = lean_int_add(x_34, x_39); +lean_dec(x_39); +lean_dec(x_34); +x_41 = l_Std_Time_PlainTime_ofNanoseconds(x_40); +lean_dec(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_29); +return x_42; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +lean_dec(x_28); +x_43 = lean_ctor_get(x_31, 0); +lean_inc(x_43); +lean_dec(x_31); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_44); +lean_dec(x_44); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +lean_dec(x_45); +x_47 = l_Std_Time_PlainDateTime_now___closed__2; +x_48 = lean_int_mul(x_46, x_47); +lean_dec(x_46); +x_49 = lean_int_add(x_34, x_48); +lean_dec(x_48); +lean_dec(x_34); +x_50 = l_Std_Time_PlainTime_ofNanoseconds(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_29); +return x_51; +} +} +} +else +{ +uint8_t x_52; +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_5); +if (x_52 == 0) +{ +return x_5; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_5, 0); +x_54 = lean_ctor_get(x_5, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_5); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +uint8_t x_56; +x_56 = !lean_is_exclusive(x_2); +if (x_56 == 0) +{ +return x_2; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_2, 0); +x_58 = lean_ctor_get(x_2, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_2); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = lean_int_add(x_8, x_6); +lean_dec(x_8); +x_10 = lean_int_ediv(x_9, x_3); +lean_dec(x_9); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_12 = l_Std_Time_PlainDateTime_now___closed__2; +x_13 = lean_int_mul(x_6, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_5, 0); +lean_inc(x_16); +lean_dec(x_5); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +static lean_object* _init_l_Std_Time_DateTime_ofPlainDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_1); +x_4 = l_Std_Time_PlainDateTime_now___closed__1; +x_5 = lean_int_mul(x_3, x_4); +lean_dec(x_3); +x_6 = l_Std_Time_DateTime_ofPlainDate___closed__1; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +lean_inc(x_7); +x_8 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDate___lambda__1___boxed), 4, 3); +lean_closure_set(x_8, 0, x_7); +lean_closure_set(x_8, 1, x_2); +lean_closure_set(x_8, 2, x_4); +x_9 = lean_mk_thunk(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDate___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_DateTime_ofPlainDate___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_PlainDateTime_now___closed__1; +x_5 = lean_int_ediv(x_3, x_4); +x_6 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toPlainDate___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainDate___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDate___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainDate(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toPlainTime___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainTime___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_5); +lean_dec(x_7); +x_9 = l_Std_Time_PlainDateTime_now___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l_Std_Time_PlainDateTime_now___closed__2; +x_13 = lean_int_mul(x_5, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_get_current_time(x_2); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +x_6 = lean_alloc_closure((void*)(l_Std_Time_DateTime_now___lambda__1___boxed), 3, 2); +lean_closure_set(x_6, 0, x_5); +lean_closure_set(x_6, 1, x_1); +x_7 = lean_mk_thunk(x_6); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_5); +lean_ctor_set(x_8, 1, x_7); +lean_ctor_set(x_3, 0, x_8); +return x_3; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_3); +lean_inc(x_9); +x_11 = lean_alloc_closure((void*)(l_Std_Time_DateTime_now___lambda__1___boxed), 3, 2); +lean_closure_set(x_11, 0, x_9); +lean_closure_set(x_11, 1, x_1); +x_12 = lean_mk_thunk(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +} +else +{ +uint8_t x_15; +lean_dec(x_1); +x_15 = !lean_is_exclusive(x_3); +if (x_15 == 0) +{ +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_3); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_now___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_now___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_now(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_get_current_time(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l_Std_Time_Database_defaultGetLocalZoneRules(x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_15; lean_object* x_16; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + x_8 = x_5; +} else { + lean_dec_ref(x_5); + x_8 = lean_box(0); +} +x_15 = lean_ctor_get(x_6, 1); +lean_inc(x_15); +lean_inc(x_3); +x_16 = l_Std_Time_TimeZone_Transition_timezoneAt(x_15, x_3); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_16); +x_17 = lean_ctor_get(x_6, 0); +lean_inc(x_17); +x_18 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_17); +lean_dec(x_17); +x_9 = x_18; +goto block_14; +} +else +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_16, 0); +lean_inc(x_19); +lean_dec(x_16); +x_9 = x_19; +goto block_14; +} +block_14: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_9); +lean_inc(x_3); +x_10 = lean_alloc_closure((void*)(l_Std_Time_DateTime_now___lambda__1___boxed), 3, 2); +lean_closure_set(x_10, 0, x_3); +lean_closure_set(x_10, 1, x_9); +x_11 = lean_mk_thunk(x_10); +x_12 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_3); +lean_ctor_set(x_12, 2, x_6); +lean_ctor_set(x_12, 3, x_9); +if (lean_is_scalar(x_8)) { + x_13 = lean_alloc_ctor(0, 2, 0); +} else { + x_13 = x_8; +} +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_7); +return x_13; +} +} +else +{ +uint8_t x_20; +lean_dec(x_3); +x_20 = !lean_is_exclusive(x_5); +if (x_20 == 0) +{ +return x_5; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_5, 0); +x_22 = lean_ctor_get(x_5, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_5); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_2); +if (x_24 == 0) +{ +return x_2; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_2, 0); +x_26 = lean_ctor_get(x_2, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_2); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nowAt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_get_current_time(x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l_Std_Time_Database_defaultGetZoneRules(x_1, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_16; lean_object* x_17; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +if (lean_is_exclusive(x_6)) { + lean_ctor_release(x_6, 0); + lean_ctor_release(x_6, 1); + x_9 = x_6; +} else { + lean_dec_ref(x_6); + x_9 = lean_box(0); +} +x_16 = lean_ctor_get(x_7, 1); +lean_inc(x_16); +lean_inc(x_4); +x_17 = l_Std_Time_TimeZone_Transition_timezoneAt(x_16, x_4); +lean_dec(x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_17); +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +x_19 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_18); +lean_dec(x_18); +x_10 = x_19; +goto block_15; +} +else +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +lean_dec(x_17); +x_10 = x_20; +goto block_15; +} +block_15: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_inc(x_10); +lean_inc(x_4); +x_11 = lean_alloc_closure((void*)(l_Std_Time_DateTime_now___lambda__1___boxed), 3, 2); +lean_closure_set(x_11, 0, x_4); +lean_closure_set(x_11, 1, x_10); +x_12 = lean_mk_thunk(x_11); +x_13 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_4); +lean_ctor_set(x_13, 2, x_7); +lean_ctor_set(x_13, 3, x_10); +if (lean_is_scalar(x_9)) { + x_14 = lean_alloc_ctor(0, 2, 0); +} else { + x_14 = x_9; +} +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; +} +} +else +{ +uint8_t x_21; +lean_dec(x_4); +x_21 = !lean_is_exclusive(x_6); +if (x_21 == 0) +{ +return x_6; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_6, 0); +x_23 = lean_ctor_get(x_6, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_6); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +else +{ +uint8_t x_25; +lean_dec(x_1); +x_25 = !lean_is_exclusive(x_3); +if (x_25 == 0) +{ +return x_3; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_3, 0); +x_27 = lean_ctor_get(x_3, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_3); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; +} +} +} +} +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_dec_le(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_2); +lean_dec(x_7); +x_9 = l_Std_Time_PlainDateTime_now___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = lean_int_mul(x_2, x_3); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +lean_dec(x_5); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_10); +lean_dec(x_10); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofPlainDate___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_DateTime_ofPlainDate___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_PlainTime_midnight; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +x_5 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_inc(x_6); +x_7 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_7, 0, x_6); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_Array_findIdx_x3f_loop___rarg(x_7, x_8, x_9); +x_11 = l_Std_Time_PlainDateTime_now___closed__2; +x_12 = lean_int_mul(x_6, x_11); +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_28; +lean_dec(x_6); +x_28 = l_Array_back_x3f___rarg(x_8); +lean_dec(x_8); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_2, 0); +lean_inc(x_29); +x_15 = x_29; +goto block_27; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_15 = x_31; +goto block_27; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_10, 0); +lean_inc(x_32); +lean_dec(x_10); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_sub(x_32, x_33); +x_35 = l_Array_get_x3f___rarg(x_8, x_34); +lean_dec(x_34); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +lean_dec(x_32); +lean_dec(x_8); +lean_dec(x_6); +x_36 = lean_ctor_get(x_2, 0); +lean_inc(x_36); +x_15 = x_36; +goto block_27; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Array_get_x3f___rarg(x_8, x_32); +lean_dec(x_32); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; +x_39 = l_Array_back_x3f___rarg(x_8); +lean_dec(x_8); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +lean_dec(x_37); +lean_dec(x_6); +x_40 = lean_ctor_get(x_2, 0); +lean_inc(x_40); +x_15 = x_40; +goto block_27; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_37, 1); +lean_inc(x_43); +lean_dec(x_37); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_nat_abs(x_44); +lean_dec(x_44); +x_46 = lean_nat_to_int(x_45); +x_47 = lean_int_sub(x_42, x_46); +lean_dec(x_46); +lean_dec(x_42); +x_48 = lean_int_dec_lt(x_6, x_47); +lean_dec(x_47); +lean_dec(x_6); +if (x_48 == 0) +{ +lean_object* x_49; +lean_dec(x_43); +x_49 = lean_ctor_get(x_41, 1); +lean_inc(x_49); +lean_dec(x_41); +x_15 = x_49; +goto block_27; +} +else +{ +lean_dec(x_41); +x_15 = x_43; +goto block_27; +} +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +lean_dec(x_8); +x_50 = lean_ctor_get(x_38, 0); +lean_inc(x_50); +lean_dec(x_38); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_ctor_get(x_37, 1); +lean_inc(x_52); +lean_dec(x_37); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_nat_abs(x_53); +lean_dec(x_53); +x_55 = lean_nat_to_int(x_54); +x_56 = lean_int_sub(x_51, x_55); +lean_dec(x_55); +lean_dec(x_51); +x_57 = lean_int_dec_lt(x_6, x_56); +lean_dec(x_56); +lean_dec(x_6); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_52); +x_58 = lean_ctor_get(x_50, 1); +lean_inc(x_58); +lean_dec(x_50); +x_15 = x_58; +goto block_27; +} +else +{ +lean_dec(x_50); +x_15 = x_52; +goto block_27; +} +} +} +} +block_27: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_16 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_15); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_int_neg(x_17); +x_19 = lean_int_mul(x_18, x_11); +lean_dec(x_18); +x_20 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_19); +x_22 = lean_int_add(x_14, x_21); +lean_dec(x_21); +lean_dec(x_14); +x_23 = l_Std_Time_Duration_ofNanoseconds(x_22); +lean_dec(x_22); +lean_inc(x_23); +x_24 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2___boxed), 4, 3); +lean_closure_set(x_24, 0, x_23); +lean_closure_set(x_24, 1, x_17); +lean_closure_set(x_24, 2, x_11); +x_25 = lean_mk_thunk(x_24); +x_26 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +lean_ctor_set(x_26, 2, x_2); +lean_ctor_set(x_26, 3, x_16); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1; +x_2 = l_Array_back_x3f___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_3 = l_Std_Time_PlainTime_midnight; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_9 = 0; +x_10 = 1; +lean_inc(x_6); +lean_inc(x_7); +lean_inc(x_5); +x_11 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_11, 0, x_5); +lean_ctor_set(x_11, 1, x_7); +lean_ctor_set(x_11, 2, x_6); +lean_ctor_set_uint8(x_11, sizeof(void*)*3, x_8); +lean_ctor_set_uint8(x_11, sizeof(void*)*3 + 1, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*3 + 2, x_10); +x_12 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1; +lean_inc(x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_4); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_inc(x_15); +x_16 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_16, 0, x_15); +x_17 = lean_unsigned_to_nat(0u); +x_18 = l_Array_findIdx_x3f_loop___rarg(x_16, x_12, x_17); +x_19 = l_Std_Time_PlainDateTime_now___closed__2; +x_20 = lean_int_mul(x_15, x_19); +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_dec(x_14); +x_22 = lean_int_add(x_20, x_21); +lean_dec(x_21); +lean_dec(x_20); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_36; +lean_dec(x_15); +x_36 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_36) == 0) +{ +x_23 = x_11; +goto block_35; +} +else +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_11); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_23 = x_38; +goto block_35; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_18, 0); +lean_inc(x_39); +lean_dec(x_18); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_sub(x_39, x_40); +x_42 = l_Array_get_x3f___rarg(x_12, x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_dec(x_39); +lean_dec(x_15); +x_23 = x_11; +goto block_35; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +lean_dec(x_42); +x_44 = l_Array_get_x3f___rarg(x_12, x_39); +lean_dec(x_39); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_45) == 0) +{ +lean_dec(x_43); +lean_dec(x_15); +x_23 = x_11; +goto block_35; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +lean_dec(x_11); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_43, 1); +lean_inc(x_48); +lean_dec(x_43); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_nat_abs(x_49); +lean_dec(x_49); +x_51 = lean_nat_to_int(x_50); +x_52 = lean_int_sub(x_47, x_51); +lean_dec(x_51); +lean_dec(x_47); +x_53 = lean_int_dec_lt(x_15, x_52); +lean_dec(x_52); +lean_dec(x_15); +if (x_53 == 0) +{ +lean_object* x_54; +lean_dec(x_48); +x_54 = lean_ctor_get(x_46, 1); +lean_inc(x_54); +lean_dec(x_46); +x_23 = x_54; +goto block_35; +} +else +{ +lean_dec(x_46); +x_23 = x_48; +goto block_35; +} +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +lean_dec(x_11); +x_55 = lean_ctor_get(x_44, 0); +lean_inc(x_55); +lean_dec(x_44); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_43, 1); +lean_inc(x_57); +lean_dec(x_43); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_nat_abs(x_58); +lean_dec(x_58); +x_60 = lean_nat_to_int(x_59); +x_61 = lean_int_sub(x_56, x_60); +lean_dec(x_60); +lean_dec(x_56); +x_62 = lean_int_dec_lt(x_15, x_61); +lean_dec(x_61); +lean_dec(x_15); +if (x_62 == 0) +{ +lean_object* x_63; +lean_dec(x_57); +x_63 = lean_ctor_get(x_55, 1); +lean_inc(x_63); +lean_dec(x_55); +x_23 = x_63; +goto block_35; +} +else +{ +lean_dec(x_55); +x_23 = x_57; +goto block_35; +} +} +} +} +block_35: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_24 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_23); +lean_dec(x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_int_neg(x_25); +x_27 = lean_int_mul(x_26, x_19); +lean_dec(x_26); +x_28 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_29 = lean_int_add(x_27, x_28); +lean_dec(x_27); +x_30 = lean_int_add(x_22, x_29); +lean_dec(x_29); +lean_dec(x_22); +x_31 = l_Std_Time_Duration_ofNanoseconds(x_30); +lean_dec(x_30); +lean_inc(x_31); +x_32 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2___boxed), 4, 3); +lean_closure_set(x_32, 0, x_31); +lean_closure_set(x_32, 1, x_25); +lean_closure_set(x_32, 2, x_19); +x_33 = lean_mk_thunk(x_32); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_31); +lean_ctor_set(x_34, 2, x_13); +lean_ctor_set(x_34, 3, x_24); +return x_34; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateWithZone___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDate(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDate___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_toPlainDate(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_toPlainTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_of(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_defaultGetZoneRules(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +if (lean_is_exclusive(x_4)) { + lean_ctor_release(x_4, 0); + lean_ctor_release(x_4, 1); + x_7 = x_4; +} else { + lean_dec_ref(x_4); + x_7 = lean_box(0); +} +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_inc(x_9); +x_10 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_10, 0, x_9); +x_11 = lean_ctor_get(x_5, 1); +lean_inc(x_11); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_findIdx_x3f_loop___rarg(x_10, x_11, x_12); +x_14 = l_Std_Time_PlainDateTime_now___closed__2; +x_15 = lean_int_mul(x_9, x_14); +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_int_add(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_32; +lean_dec(x_9); +x_32 = l_Array_back_x3f___rarg(x_11); +lean_dec(x_11); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_5, 0); +lean_inc(x_33); +x_18 = x_33; +goto block_31; +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_18 = x_35; +goto block_31; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_13, 0); +lean_inc(x_36); +lean_dec(x_13); +x_37 = lean_unsigned_to_nat(1u); +x_38 = lean_nat_sub(x_36, x_37); +x_39 = l_Array_get_x3f___rarg(x_11, x_38); +lean_dec(x_38); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +lean_dec(x_36); +lean_dec(x_11); +lean_dec(x_9); +x_40 = lean_ctor_get(x_5, 0); +lean_inc(x_40); +x_18 = x_40; +goto block_31; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +lean_dec(x_39); +x_42 = l_Array_get_x3f___rarg(x_11, x_36); +lean_dec(x_36); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +x_43 = l_Array_back_x3f___rarg(x_11); +lean_dec(x_11); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; +lean_dec(x_41); +lean_dec(x_9); +x_44 = lean_ctor_get(x_5, 0); +lean_inc(x_44); +x_18 = x_44; +goto block_31; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_41, 1); +lean_inc(x_47); +lean_dec(x_41); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_nat_abs(x_48); +lean_dec(x_48); +x_50 = lean_nat_to_int(x_49); +x_51 = lean_int_sub(x_46, x_50); +lean_dec(x_50); +lean_dec(x_46); +x_52 = lean_int_dec_lt(x_9, x_51); +lean_dec(x_51); +lean_dec(x_9); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_47); +x_53 = lean_ctor_get(x_45, 1); +lean_inc(x_53); +lean_dec(x_45); +x_18 = x_53; +goto block_31; +} +else +{ +lean_dec(x_45); +x_18 = x_47; +goto block_31; +} +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +lean_dec(x_11); +x_54 = lean_ctor_get(x_42, 0); +lean_inc(x_54); +lean_dec(x_42); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_41, 1); +lean_inc(x_56); +lean_dec(x_41); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_nat_abs(x_57); +lean_dec(x_57); +x_59 = lean_nat_to_int(x_58); +x_60 = lean_int_sub(x_55, x_59); +lean_dec(x_59); +lean_dec(x_55); +x_61 = lean_int_dec_lt(x_9, x_60); +lean_dec(x_60); +lean_dec(x_9); +if (x_61 == 0) +{ +lean_object* x_62; +lean_dec(x_56); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_18 = x_62; +goto block_31; +} +else +{ +lean_dec(x_54); +x_18 = x_56; +goto block_31; +} +} +} +} +block_31: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_19 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_18); +lean_dec(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_int_neg(x_20); +x_22 = lean_int_mul(x_21, x_14); +lean_dec(x_21); +x_23 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_24 = lean_int_add(x_22, x_23); +lean_dec(x_22); +x_25 = lean_int_add(x_17, x_24); +lean_dec(x_24); +lean_dec(x_17); +x_26 = l_Std_Time_Duration_ofNanoseconds(x_25); +lean_dec(x_25); +lean_inc(x_26); +x_27 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__2___boxed), 4, 3); +lean_closure_set(x_27, 0, x_26); +lean_closure_set(x_27, 1, x_20); +lean_closure_set(x_27, 2, x_14); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +lean_ctor_set(x_29, 2, x_5); +lean_ctor_set(x_29, 3, x_19); +if (lean_is_scalar(x_7)) { + x_30 = lean_alloc_ctor(0, 2, 0); +} else { + x_30 = x_7; +} +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_6); +return x_30; +} +} +else +{ +uint8_t x_63; +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_4); +if (x_63 == 0) +{ +return x_4; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_4, 0); +x_65 = lean_ctor_get(x_4, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_4); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_inc(x_4); +x_5 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_5, 0, x_4); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Array_findIdx_x3f_loop___rarg(x_5, x_6, x_7); +x_9 = l_Std_Time_PlainDateTime_now___closed__2; +x_10 = lean_int_mul(x_4, x_9); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_dec(x_3); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_23; +lean_dec(x_4); +x_23 = l_Array_back_x3f___rarg(x_6); +lean_dec(x_6); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_2, 0); +lean_inc(x_24); +lean_dec(x_2); +x_13 = x_24; +goto block_22; +} +else +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_2); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_13 = x_26; +goto block_22; +} +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_8, 0); +lean_inc(x_27); +lean_dec(x_8); +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_sub(x_27, x_28); +x_30 = l_Array_get_x3f___rarg(x_6, x_29); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; +lean_dec(x_27); +lean_dec(x_6); +lean_dec(x_4); +x_31 = lean_ctor_get(x_2, 0); +lean_inc(x_31); +lean_dec(x_2); +x_13 = x_31; +goto block_22; +} +else +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 0); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Array_get_x3f___rarg(x_6, x_27); +lean_dec(x_27); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; +x_34 = l_Array_back_x3f___rarg(x_6); +lean_dec(x_6); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +lean_dec(x_32); +lean_dec(x_4); +x_35 = lean_ctor_get(x_2, 0); +lean_inc(x_35); +lean_dec(x_2); +x_13 = x_35; +goto block_22; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +lean_dec(x_2); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_32, 1); +lean_inc(x_38); +lean_dec(x_32); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_nat_abs(x_39); +lean_dec(x_39); +x_41 = lean_nat_to_int(x_40); +x_42 = lean_int_sub(x_37, x_41); +lean_dec(x_41); +lean_dec(x_37); +x_43 = lean_int_dec_lt(x_4, x_42); +lean_dec(x_42); +lean_dec(x_4); +if (x_43 == 0) +{ +lean_object* x_44; +lean_dec(x_38); +x_44 = lean_ctor_get(x_36, 1); +lean_inc(x_44); +lean_dec(x_36); +x_13 = x_44; +goto block_22; +} +else +{ +lean_dec(x_36); +x_13 = x_38; +goto block_22; +} +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +lean_dec(x_6); +lean_dec(x_2); +x_45 = lean_ctor_get(x_33, 0); +lean_inc(x_45); +lean_dec(x_33); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_32, 1); +lean_inc(x_47); +lean_dec(x_32); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_nat_abs(x_48); +lean_dec(x_48); +x_50 = lean_nat_to_int(x_49); +x_51 = lean_int_sub(x_46, x_50); +lean_dec(x_50); +lean_dec(x_46); +x_52 = lean_int_dec_lt(x_4, x_51); +lean_dec(x_51); +lean_dec(x_4); +if (x_52 == 0) +{ +lean_object* x_53; +lean_dec(x_47); +x_53 = lean_ctor_get(x_45, 1); +lean_inc(x_53); +lean_dec(x_45); +x_13 = x_53; +goto block_22; +} +else +{ +lean_dec(x_45); +x_13 = x_47; +goto block_22; +} +} +} +} +block_22: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_14 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_int_neg(x_15); +lean_dec(x_15); +x_17 = lean_int_mul(x_16, x_9); +lean_dec(x_16); +x_18 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_19 = lean_int_add(x_17, x_18); +lean_dec(x_17); +x_20 = lean_int_add(x_12, x_19); +lean_dec(x_19); +lean_dec(x_12); +x_21 = l_Std_Time_Duration_ofNanoseconds(x_20); +lean_dec(x_20); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampWithZone(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_7 = 0; +x_8 = 1; +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_3); +x_9 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_5); +lean_ctor_set(x_9, 2, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*3, x_6); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 1, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 2, x_8); +x_10 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_12, 0, x_11); +x_13 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1; +x_14 = lean_unsigned_to_nat(0u); +x_15 = l_Array_findIdx_x3f_loop___rarg(x_12, x_13, x_14); +x_16 = l_Std_Time_PlainDateTime_now___closed__2; +x_17 = lean_int_mul(x_11, x_16); +x_18 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_int_add(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_30; +lean_dec(x_11); +x_30 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_30) == 0) +{ +x_20 = x_9; +goto block_29; +} +else +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_9); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_20 = x_32; +goto block_29; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_15, 0); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +x_36 = l_Array_get_x3f___rarg(x_13, x_35); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_dec(x_33); +lean_dec(x_11); +x_20 = x_9; +goto block_29; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Array_get_x3f___rarg(x_13, x_33); +lean_dec(x_33); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; +x_39 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_39) == 0) +{ +lean_dec(x_37); +lean_dec(x_11); +x_20 = x_9; +goto block_29; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +lean_dec(x_9); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_dec(x_37); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_nat_abs(x_43); +lean_dec(x_43); +x_45 = lean_nat_to_int(x_44); +x_46 = lean_int_sub(x_41, x_45); +lean_dec(x_45); +lean_dec(x_41); +x_47 = lean_int_dec_lt(x_11, x_46); +lean_dec(x_46); +lean_dec(x_11); +if (x_47 == 0) +{ +lean_object* x_48; +lean_dec(x_42); +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_20 = x_48; +goto block_29; +} +else +{ +lean_dec(x_40); +x_20 = x_42; +goto block_29; +} +} +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +lean_dec(x_9); +x_49 = lean_ctor_get(x_38, 0); +lean_inc(x_49); +lean_dec(x_38); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_37, 1); +lean_inc(x_51); +lean_dec(x_37); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_nat_abs(x_52); +lean_dec(x_52); +x_54 = lean_nat_to_int(x_53); +x_55 = lean_int_sub(x_50, x_54); +lean_dec(x_54); +lean_dec(x_50); +x_56 = lean_int_dec_lt(x_11, x_55); +lean_dec(x_55); +lean_dec(x_11); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_51); +x_57 = lean_ctor_get(x_49, 1); +lean_inc(x_57); +lean_dec(x_49); +x_20 = x_57; +goto block_29; +} +else +{ +lean_dec(x_49); +x_20 = x_51; +goto block_29; +} +} +} +} +block_29: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_21 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_20); +lean_dec(x_20); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_int_neg(x_22); +lean_dec(x_22); +x_24 = lean_int_mul(x_23, x_16); +lean_dec(x_23); +x_25 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_26 = lean_int_add(x_24, x_25); +lean_dec(x_24); +x_27 = lean_int_add(x_19, x_26); +lean_dec(x_26); +lean_dec(x_19); +x_28 = l_Std_Time_Duration_ofNanoseconds(x_27); +lean_dec(x_27); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDateTime_toTimestampWithZone___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDateTime_toTimestampWithZone(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = l_Std_Time_PlainTime_midnight; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +x_5 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_inc(x_6); +x_7 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_7, 0, x_6); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l_Array_findIdx_x3f_loop___rarg(x_7, x_8, x_9); +x_11 = l_Std_Time_PlainDateTime_now___closed__2; +x_12 = lean_int_mul(x_6, x_11); +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_25; +lean_dec(x_6); +x_25 = l_Array_back_x3f___rarg(x_8); +lean_dec(x_8); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; +x_26 = lean_ctor_get(x_2, 0); +lean_inc(x_26); +lean_dec(x_2); +x_15 = x_26; +goto block_24; +} +else +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_2); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_15 = x_28; +goto block_24; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_10, 0); +lean_inc(x_29); +lean_dec(x_10); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_sub(x_29, x_30); +x_32 = l_Array_get_x3f___rarg(x_8, x_31); +lean_dec(x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_dec(x_29); +lean_dec(x_8); +lean_dec(x_6); +x_33 = lean_ctor_get(x_2, 0); +lean_inc(x_33); +lean_dec(x_2); +x_15 = x_33; +goto block_24; +} +else +{ +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Array_get_x3f___rarg(x_8, x_29); +lean_dec(x_29); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = l_Array_back_x3f___rarg(x_8); +lean_dec(x_8); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +lean_dec(x_34); +lean_dec(x_6); +x_37 = lean_ctor_get(x_2, 0); +lean_inc(x_37); +lean_dec(x_2); +x_15 = x_37; +goto block_24; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_dec(x_2); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_34, 1); +lean_inc(x_40); +lean_dec(x_34); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_nat_abs(x_41); +lean_dec(x_41); +x_43 = lean_nat_to_int(x_42); +x_44 = lean_int_sub(x_39, x_43); +lean_dec(x_43); +lean_dec(x_39); +x_45 = lean_int_dec_lt(x_6, x_44); +lean_dec(x_44); +lean_dec(x_6); +if (x_45 == 0) +{ +lean_object* x_46; +lean_dec(x_40); +x_46 = lean_ctor_get(x_38, 1); +lean_inc(x_46); +lean_dec(x_38); +x_15 = x_46; +goto block_24; +} +else +{ +lean_dec(x_38); +x_15 = x_40; +goto block_24; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +lean_dec(x_8); +lean_dec(x_2); +x_47 = lean_ctor_get(x_35, 0); +lean_inc(x_47); +lean_dec(x_35); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_34, 1); +lean_inc(x_49); +lean_dec(x_34); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_nat_abs(x_50); +lean_dec(x_50); +x_52 = lean_nat_to_int(x_51); +x_53 = lean_int_sub(x_48, x_52); +lean_dec(x_52); +lean_dec(x_48); +x_54 = lean_int_dec_lt(x_6, x_53); +lean_dec(x_53); +lean_dec(x_6); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_49); +x_55 = lean_ctor_get(x_47, 1); +lean_inc(x_55); +lean_dec(x_47); +x_15 = x_55; +goto block_24; +} +else +{ +lean_dec(x_47); +x_15 = x_49; +goto block_24; +} +} +} +} +block_24: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_16 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_15); +lean_dec(x_15); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_int_neg(x_17); +lean_dec(x_17); +x_19 = lean_int_mul(x_18, x_11); +lean_dec(x_18); +x_20 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_19); +x_22 = lean_int_add(x_14, x_21); +lean_dec(x_21); +lean_dec(x_14); +x_23 = l_Std_Time_Duration_ofNanoseconds(x_22); +lean_dec(x_22); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampWithZone(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; uint8_t x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_3 = l_Std_Time_PlainTime_midnight; +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_9 = 0; +x_10 = 1; +lean_inc(x_6); +lean_inc(x_7); +lean_inc(x_5); +x_11 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_11, 0, x_5); +lean_ctor_set(x_11, 1, x_7); +lean_ctor_set(x_11, 2, x_6); +lean_ctor_set_uint8(x_11, sizeof(void*)*3, x_8); +lean_ctor_set_uint8(x_11, sizeof(void*)*3 + 1, x_9); +lean_ctor_set_uint8(x_11, sizeof(void*)*3 + 2, x_10); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_4); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDate___lambda__1___boxed), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_findIdx_x3f_loop___rarg(x_14, x_15, x_16); +x_18 = l_Std_Time_PlainDateTime_now___closed__2; +x_19 = lean_int_mul(x_13, x_18); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_32; +lean_dec(x_13); +x_32 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_32) == 0) +{ +x_22 = x_11; +goto block_31; +} +else +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_11); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_22 = x_34; +goto block_31; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_17, 0); +lean_inc(x_35); +lean_dec(x_17); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_sub(x_35, x_36); +x_38 = l_Array_get_x3f___rarg(x_15, x_37); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_dec(x_35); +lean_dec(x_13); +x_22 = x_11; +goto block_31; +} +else +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +lean_dec(x_38); +x_40 = l_Array_get_x3f___rarg(x_15, x_35); +lean_dec(x_35); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +x_41 = l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2; +if (lean_obj_tag(x_41) == 0) +{ +lean_dec(x_39); +lean_dec(x_13); +x_22 = x_11; +goto block_31; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +lean_dec(x_11); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_39, 1); +lean_inc(x_44); +lean_dec(x_39); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_nat_abs(x_45); +lean_dec(x_45); +x_47 = lean_nat_to_int(x_46); +x_48 = lean_int_sub(x_43, x_47); +lean_dec(x_47); +lean_dec(x_43); +x_49 = lean_int_dec_lt(x_13, x_48); +lean_dec(x_48); +lean_dec(x_13); +if (x_49 == 0) +{ +lean_object* x_50; +lean_dec(x_44); +x_50 = lean_ctor_get(x_42, 1); +lean_inc(x_50); +lean_dec(x_42); +x_22 = x_50; +goto block_31; +} +else +{ +lean_dec(x_42); +x_22 = x_44; +goto block_31; +} +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_11); +x_51 = lean_ctor_get(x_40, 0); +lean_inc(x_51); +lean_dec(x_40); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_39, 1); +lean_inc(x_53); +lean_dec(x_39); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_nat_abs(x_54); +lean_dec(x_54); +x_56 = lean_nat_to_int(x_55); +x_57 = lean_int_sub(x_52, x_56); +lean_dec(x_56); +lean_dec(x_52); +x_58 = lean_int_dec_lt(x_13, x_57); +lean_dec(x_57); +lean_dec(x_13); +if (x_58 == 0) +{ +lean_object* x_59; +lean_dec(x_53); +x_59 = lean_ctor_get(x_51, 1); +lean_inc(x_59); +lean_dec(x_51); +x_22 = x_59; +goto block_31; +} +else +{ +lean_dec(x_51); +x_22 = x_53; +goto block_31; +} +} +} +} +block_31: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_22); +lean_dec(x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +x_25 = lean_int_neg(x_24); +lean_dec(x_24); +x_26 = lean_int_mul(x_25, x_18); +lean_dec(x_25); +x_27 = l_Std_Time_ZonedDateTime_ofPlainDate___closed__1; +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_26); +x_29 = lean_int_add(x_21, x_28); +lean_dec(x_28); +lean_dec(x_21); +x_30 = l_Std_Time_Duration_ofNanoseconds(x_29); +lean_dec(x_29); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_PlainDate_toTimestampWithZone___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_PlainDate_toTimestampWithZone(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +lean_object* initialize_Std_Time_Zoned_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_ZonedDateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Zoned_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_ZoneRules(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_ZonedDateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_PlainDateTime_now___closed__1 = _init_l_Std_Time_PlainDateTime_now___closed__1(); +lean_mark_persistent(l_Std_Time_PlainDateTime_now___closed__1); +l_Std_Time_PlainDateTime_now___closed__2 = _init_l_Std_Time_PlainDateTime_now___closed__2(); +lean_mark_persistent(l_Std_Time_PlainDateTime_now___closed__2); +l_Std_Time_DateTime_ofPlainDate___closed__1 = _init_l_Std_Time_DateTime_ofPlainDate___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_ofPlainDate___closed__1); +l_Std_Time_ZonedDateTime_ofPlainDate___closed__1 = _init_l_Std_Time_ZonedDateTime_ofPlainDate___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofPlainDate___closed__1); +l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1 = _init_l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__1); +l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2 = _init_l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofPlainDateWithZone___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Database.c b/stage0/stdlib/Std/Time/Zoned/Database.c new file mode 100644 index 000000000000..e35ab2f1ff02 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Database.c @@ -0,0 +1,179 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Database +// Imports: Std.Time.Zoned.ZonedDateTime Std.Time.Zoned.Database.Basic Std.Time.Zoned.Database.TZdb Std.Time.Zoned.Database.Windows Init.System.Platform +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static uint64_t l_Std_Time_Database_defaultGetLocalZoneRules___closed__2; +lean_object* l_Std_Time_Database_TZdb_localRules(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_defaultGetLocalZoneRules(lean_object*); +static uint64_t l_Std_Time_Database_defaultGetLocalZoneRules___closed__3; +uint64_t lean_int64_neg(uint64_t); +static lean_object* l_Std_Time_Database_defaultGetLocalZoneRules___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1; +static lean_object* l_Std_Time_Database_defaultGetZoneRules___closed__1; +static lean_object* l_Std_Time_Database_defaultGetLocalZoneRules___closed__5; +lean_object* l_Std_Time_Database_Windows_getLocalTimeZoneIdentifierAt___boxed(lean_object*, lean_object*); +uint64_t lean_int64_of_nat(lean_object*); +lean_object* l_Bind_bindLeft___at_Std_Time_Database_WindowsDb_inst___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_Database_Windows_getZoneRules___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Database_TZdb_readRulesFromDisk(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_defaultGetZoneRules(lean_object*, lean_object*); +extern uint8_t l_System_Platform_isWindows; +static lean_object* l_Std_Time_Database_defaultGetLocalZoneRules___closed__1; +lean_object* l_Std_Time_Database_Windows_getZoneRules(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Database_defaultGetZoneRules___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("/usr/share/zoneinfo/", 20, 20); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_defaultGetZoneRules(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = l_System_Platform_isWindows; +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Std_Time_Database_defaultGetZoneRules___closed__1; +x_5 = l_Std_Time_Database_TZdb_readRulesFromDisk(x_4, x_1, x_2); +return x_5; +} +else +{ +lean_object* x_6; +x_6 = l_Std_Time_Database_Windows_getZoneRules(x_1, x_2); +lean_dec(x_1); +return x_6; +} +} +} +static lean_object* _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("/etc/localtime", 14, 14); +return x_1; +} +} +static uint64_t _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__2() { +_start: +{ +lean_object* x_1; uint64_t x_2; +x_1 = lean_unsigned_to_nat(2147483648u); +x_2 = lean_int64_of_nat(x_1); +return x_2; +} +} +static uint64_t _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__3() { +_start: +{ +uint64_t x_1; uint64_t x_2; +x_1 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__2; +x_2 = lean_int64_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1() { +_start: +{ +uint64_t x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__3; +x_2 = lean_box_uint64(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Database_Windows_getLocalTimeZoneIdentifierAt___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_Windows_getZoneRules___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_defaultGetLocalZoneRules(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = l_System_Platform_isWindows; +if (x_2 == 0) +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__1; +x_4 = l_Std_Time_Database_TZdb_localRules(x_3, x_1); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__5; +x_6 = l_Std_Time_Database_defaultGetLocalZoneRules___closed__4; +x_7 = l_Bind_bindLeft___at_Std_Time_Database_WindowsDb_inst___spec__1(x_5, x_6, x_1); +return x_7; +} +} +} +lean_object* initialize_Std_Time_Zoned_ZonedDateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_TZdb(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_Windows(uint8_t builtin, lean_object*); +lean_object* initialize_Init_System_Platform(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Database(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Zoned_ZonedDateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_TZdb(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_Windows(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Init_System_Platform(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Database_defaultGetZoneRules___closed__1 = _init_l_Std_Time_Database_defaultGetZoneRules___closed__1(); +lean_mark_persistent(l_Std_Time_Database_defaultGetZoneRules___closed__1); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__1 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__1(); +lean_mark_persistent(l_Std_Time_Database_defaultGetLocalZoneRules___closed__1); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__2 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__2(); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__3 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__3(); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1(); +lean_mark_persistent(l_Std_Time_Database_defaultGetLocalZoneRules___closed__4___boxed__const__1); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__4 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__4(); +lean_mark_persistent(l_Std_Time_Database_defaultGetLocalZoneRules___closed__4); +l_Std_Time_Database_defaultGetLocalZoneRules___closed__5 = _init_l_Std_Time_Database_defaultGetLocalZoneRules___closed__5(); +lean_mark_persistent(l_Std_Time_Database_defaultGetLocalZoneRules___closed__5); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Database/Basic.c b/stage0/stdlib/Std/Time/Zoned/Database/Basic.c new file mode 100644 index 000000000000..b92f56a5a2f3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Database/Basic.c @@ -0,0 +1,1001 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Database.Basic +// Imports: Std.Time.Zoned.ZoneRules Std.Time.Zoned.Database.TzIf +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertLocalTimeType(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2; +extern lean_object* l_Std_Time_TimeZone_instInhabitedLocalTimeType; +lean_object* lean_uint32_to_nat(uint32_t); +extern lean_object* l_Int_instInhabited; +lean_object* lean_array_push(lean_object*, lean_object*); +static uint8_t l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_convertTZifV1___closed__3; +static lean_object* l_Std_Time_TimeZone_convertTZifV1___closed__2; +static lean_object* l_Std_Time_TimeZone_convertTZifV1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV2(lean_object*, lean_object*); +static uint8_t l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZif(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_convertWall(uint8_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTransition(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_convertUt(uint8_t); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV2___boxed(lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +lean_object* lean_uint8_to_nat(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertWall___boxed(lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertLocalTimeType___boxed(lean_object*, lean_object*, lean_object*); +extern uint8_t l_UInt8_instInhabited; +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertUt___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZif___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTransition___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_convertWall(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertWall___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_TimeZone_convertWall(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_convertUt(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertUt___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_TimeZone_convertUt(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +static uint8_t _init_l_Std_Time_TimeZone_convertLocalTimeType___closed__1() { +_start: +{ +uint8_t x_1; uint8_t x_2; +x_1 = 1; +x_2 = l_Std_Time_TimeZone_convertWall(x_1); +return x_2; +} +} +static uint8_t _init_l_Std_Time_TimeZone_convertLocalTimeType___closed__2() { +_start: +{ +uint8_t x_1; uint8_t x_2; +x_1 = 1; +x_2 = l_Std_Time_TimeZone_convertUt(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertLocalTimeType(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_2, 3); +x_5 = l_Array_get_x3f___rarg(x_4, x_1); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; +lean_dec(x_3); +x_6 = lean_box(0); +return x_6; +} +else +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_8 = lean_ctor_get(x_5, 0); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get_uint8(x_8, sizeof(void*)*1); +lean_dec(x_8); +x_11 = lean_ctor_get(x_2, 4); +x_12 = 1; +lean_inc(x_9); +x_13 = l_Std_Time_TimeZone_Offset_toIsoString(x_9, x_12); +x_14 = lean_array_get_size(x_11); +x_15 = lean_nat_dec_lt(x_1, x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_2, 6); +x_17 = lean_array_get_size(x_16); +x_18 = lean_nat_dec_lt(x_1, x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_2, 7); +x_20 = lean_array_get_size(x_19); +x_21 = lean_nat_dec_lt(x_1, x_20); +lean_dec(x_20); +if (x_15 == 0) +{ +if (x_18 == 0) +{ +if (x_21 == 0) +{ +uint8_t x_22; uint8_t x_23; lean_object* x_24; +x_22 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_23 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_24 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_24, 0, x_9); +lean_ctor_set(x_24, 1, x_13); +lean_ctor_set(x_24, 2, x_3); +lean_ctor_set_uint8(x_24, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_24, sizeof(void*)*3 + 1, x_22); +lean_ctor_set_uint8(x_24, sizeof(void*)*3 + 2, x_23); +lean_ctor_set(x_5, 0, x_24); +return x_5; +} +else +{ +lean_object* x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; +x_25 = lean_array_fget(x_19, x_1); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +x_27 = l_Std_Time_TimeZone_convertUt(x_26); +x_28 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_29 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_29, 0, x_9); +lean_ctor_set(x_29, 1, x_13); +lean_ctor_set(x_29, 2, x_3); +lean_ctor_set_uint8(x_29, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_29, sizeof(void*)*3 + 1, x_28); +lean_ctor_set_uint8(x_29, sizeof(void*)*3 + 2, x_27); +lean_ctor_set(x_5, 0, x_29); +return x_5; +} +} +else +{ +lean_object* x_30; uint8_t x_31; uint8_t x_32; +x_30 = lean_array_fget(x_16, x_1); +x_31 = lean_unbox(x_30); +lean_dec(x_30); +x_32 = l_Std_Time_TimeZone_convertWall(x_31); +if (x_21 == 0) +{ +uint8_t x_33; lean_object* x_34; +x_33 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_34 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_34, 0, x_9); +lean_ctor_set(x_34, 1, x_13); +lean_ctor_set(x_34, 2, x_3); +lean_ctor_set_uint8(x_34, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_34, sizeof(void*)*3 + 1, x_32); +lean_ctor_set_uint8(x_34, sizeof(void*)*3 + 2, x_33); +lean_ctor_set(x_5, 0, x_34); +return x_5; +} +else +{ +lean_object* x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; +x_35 = lean_array_fget(x_19, x_1); +x_36 = lean_unbox(x_35); +lean_dec(x_35); +x_37 = l_Std_Time_TimeZone_convertUt(x_36); +x_38 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_38, 0, x_9); +lean_ctor_set(x_38, 1, x_13); +lean_ctor_set(x_38, 2, x_3); +lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_38, sizeof(void*)*3 + 1, x_32); +lean_ctor_set_uint8(x_38, sizeof(void*)*3 + 2, x_37); +lean_ctor_set(x_5, 0, x_38); +return x_5; +} +} +} +else +{ +lean_object* x_39; +lean_dec(x_13); +x_39 = lean_array_fget(x_11, x_1); +if (x_18 == 0) +{ +if (x_21 == 0) +{ +uint8_t x_40; uint8_t x_41; lean_object* x_42; +x_40 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_41 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_42 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_42, 0, x_9); +lean_ctor_set(x_42, 1, x_39); +lean_ctor_set(x_42, 2, x_3); +lean_ctor_set_uint8(x_42, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_42, sizeof(void*)*3 + 1, x_40); +lean_ctor_set_uint8(x_42, sizeof(void*)*3 + 2, x_41); +lean_ctor_set(x_5, 0, x_42); +return x_5; +} +else +{ +lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; lean_object* x_47; +x_43 = lean_array_fget(x_19, x_1); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +x_45 = l_Std_Time_TimeZone_convertUt(x_44); +x_46 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_47 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_47, 0, x_9); +lean_ctor_set(x_47, 1, x_39); +lean_ctor_set(x_47, 2, x_3); +lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_47, sizeof(void*)*3 + 1, x_46); +lean_ctor_set_uint8(x_47, sizeof(void*)*3 + 2, x_45); +lean_ctor_set(x_5, 0, x_47); +return x_5; +} +} +else +{ +lean_object* x_48; uint8_t x_49; uint8_t x_50; +x_48 = lean_array_fget(x_16, x_1); +x_49 = lean_unbox(x_48); +lean_dec(x_48); +x_50 = l_Std_Time_TimeZone_convertWall(x_49); +if (x_21 == 0) +{ +uint8_t x_51; lean_object* x_52; +x_51 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_52 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_52, 0, x_9); +lean_ctor_set(x_52, 1, x_39); +lean_ctor_set(x_52, 2, x_3); +lean_ctor_set_uint8(x_52, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_52, sizeof(void*)*3 + 1, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*3 + 2, x_51); +lean_ctor_set(x_5, 0, x_52); +return x_5; +} +else +{ +lean_object* x_53; uint8_t x_54; uint8_t x_55; lean_object* x_56; +x_53 = lean_array_fget(x_19, x_1); +x_54 = lean_unbox(x_53); +lean_dec(x_53); +x_55 = l_Std_Time_TimeZone_convertUt(x_54); +x_56 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_56, 0, x_9); +lean_ctor_set(x_56, 1, x_39); +lean_ctor_set(x_56, 2, x_3); +lean_ctor_set_uint8(x_56, sizeof(void*)*3, x_10); +lean_ctor_set_uint8(x_56, sizeof(void*)*3 + 1, x_50); +lean_ctor_set_uint8(x_56, sizeof(void*)*3 + 2, x_55); +lean_ctor_set(x_5, 0, x_56); +return x_5; +} +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; uint8_t x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_57 = lean_ctor_get(x_5, 0); +lean_inc(x_57); +lean_dec(x_5); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get_uint8(x_57, sizeof(void*)*1); +lean_dec(x_57); +x_60 = lean_ctor_get(x_2, 4); +x_61 = 1; +lean_inc(x_58); +x_62 = l_Std_Time_TimeZone_Offset_toIsoString(x_58, x_61); +x_63 = lean_array_get_size(x_60); +x_64 = lean_nat_dec_lt(x_1, x_63); +lean_dec(x_63); +x_65 = lean_ctor_get(x_2, 6); +x_66 = lean_array_get_size(x_65); +x_67 = lean_nat_dec_lt(x_1, x_66); +lean_dec(x_66); +x_68 = lean_ctor_get(x_2, 7); +x_69 = lean_array_get_size(x_68); +x_70 = lean_nat_dec_lt(x_1, x_69); +lean_dec(x_69); +if (x_64 == 0) +{ +if (x_67 == 0) +{ +if (x_70 == 0) +{ +uint8_t x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; +x_71 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_72 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_73 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_73, 0, x_58); +lean_ctor_set(x_73, 1, x_62); +lean_ctor_set(x_73, 2, x_3); +lean_ctor_set_uint8(x_73, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_73, sizeof(void*)*3 + 1, x_71); +lean_ctor_set_uint8(x_73, sizeof(void*)*3 + 2, x_72); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +return x_74; +} +else +{ +lean_object* x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; +x_75 = lean_array_fget(x_68, x_1); +x_76 = lean_unbox(x_75); +lean_dec(x_75); +x_77 = l_Std_Time_TimeZone_convertUt(x_76); +x_78 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_79 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_79, 0, x_58); +lean_ctor_set(x_79, 1, x_62); +lean_ctor_set(x_79, 2, x_3); +lean_ctor_set_uint8(x_79, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_79, sizeof(void*)*3 + 1, x_78); +lean_ctor_set_uint8(x_79, sizeof(void*)*3 + 2, x_77); +x_80 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_80, 0, x_79); +return x_80; +} +} +else +{ +lean_object* x_81; uint8_t x_82; uint8_t x_83; +x_81 = lean_array_fget(x_65, x_1); +x_82 = lean_unbox(x_81); +lean_dec(x_81); +x_83 = l_Std_Time_TimeZone_convertWall(x_82); +if (x_70 == 0) +{ +uint8_t x_84; lean_object* x_85; lean_object* x_86; +x_84 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_85 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_85, 0, x_58); +lean_ctor_set(x_85, 1, x_62); +lean_ctor_set(x_85, 2, x_3); +lean_ctor_set_uint8(x_85, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_85, sizeof(void*)*3 + 1, x_83); +lean_ctor_set_uint8(x_85, sizeof(void*)*3 + 2, x_84); +x_86 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_86, 0, x_85); +return x_86; +} +else +{ +lean_object* x_87; uint8_t x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; +x_87 = lean_array_fget(x_68, x_1); +x_88 = lean_unbox(x_87); +lean_dec(x_87); +x_89 = l_Std_Time_TimeZone_convertUt(x_88); +x_90 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_90, 0, x_58); +lean_ctor_set(x_90, 1, x_62); +lean_ctor_set(x_90, 2, x_3); +lean_ctor_set_uint8(x_90, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_90, sizeof(void*)*3 + 1, x_83); +lean_ctor_set_uint8(x_90, sizeof(void*)*3 + 2, x_89); +x_91 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_90); +return x_91; +} +} +} +else +{ +lean_object* x_92; +lean_dec(x_62); +x_92 = lean_array_fget(x_60, x_1); +if (x_67 == 0) +{ +if (x_70 == 0) +{ +uint8_t x_93; uint8_t x_94; lean_object* x_95; lean_object* x_96; +x_93 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_94 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_95 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_95, 0, x_58); +lean_ctor_set(x_95, 1, x_92); +lean_ctor_set(x_95, 2, x_3); +lean_ctor_set_uint8(x_95, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_95, sizeof(void*)*3 + 1, x_93); +lean_ctor_set_uint8(x_95, sizeof(void*)*3 + 2, x_94); +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_95); +return x_96; +} +else +{ +lean_object* x_97; uint8_t x_98; uint8_t x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; +x_97 = lean_array_fget(x_68, x_1); +x_98 = lean_unbox(x_97); +lean_dec(x_97); +x_99 = l_Std_Time_TimeZone_convertUt(x_98); +x_100 = l_Std_Time_TimeZone_convertLocalTimeType___closed__1; +x_101 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_101, 0, x_58); +lean_ctor_set(x_101, 1, x_92); +lean_ctor_set(x_101, 2, x_3); +lean_ctor_set_uint8(x_101, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_101, sizeof(void*)*3 + 1, x_100); +lean_ctor_set_uint8(x_101, sizeof(void*)*3 + 2, x_99); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_101); +return x_102; +} +} +else +{ +lean_object* x_103; uint8_t x_104; uint8_t x_105; +x_103 = lean_array_fget(x_65, x_1); +x_104 = lean_unbox(x_103); +lean_dec(x_103); +x_105 = l_Std_Time_TimeZone_convertWall(x_104); +if (x_70 == 0) +{ +uint8_t x_106; lean_object* x_107; lean_object* x_108; +x_106 = l_Std_Time_TimeZone_convertLocalTimeType___closed__2; +x_107 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_107, 0, x_58); +lean_ctor_set(x_107, 1, x_92); +lean_ctor_set(x_107, 2, x_3); +lean_ctor_set_uint8(x_107, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_107, sizeof(void*)*3 + 1, x_105); +lean_ctor_set_uint8(x_107, sizeof(void*)*3 + 2, x_106); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +return x_108; +} +else +{ +lean_object* x_109; uint8_t x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; +x_109 = lean_array_fget(x_68, x_1); +x_110 = lean_unbox(x_109); +lean_dec(x_109); +x_111 = l_Std_Time_TimeZone_convertUt(x_110); +x_112 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_112, 0, x_58); +lean_ctor_set(x_112, 1, x_92); +lean_ctor_set(x_112, 2, x_3); +lean_ctor_set_uint8(x_112, sizeof(void*)*3, x_59); +lean_ctor_set_uint8(x_112, sizeof(void*)*3 + 1, x_105); +lean_ctor_set_uint8(x_112, sizeof(void*)*3 + 2, x_111); +x_113 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_113, 0, x_112); +return x_113; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertLocalTimeType___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_convertLocalTimeType(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTransition(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_4 = lean_ctor_get(x_3, 1); +x_5 = l_Int_instInhabited; +x_6 = lean_array_get(x_5, x_4, x_2); +x_7 = lean_ctor_get(x_3, 2); +x_8 = l_UInt8_instInhabited; +x_9 = lean_box(x_8); +x_10 = lean_array_get(x_9, x_7, x_2); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +x_12 = lean_uint8_to_nat(x_11); +x_13 = l_Std_Time_TimeZone_instInhabitedLocalTimeType; +x_14 = lean_array_get(x_13, x_1, x_12); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_6); +lean_ctor_set(x_15, 1, x_14); +x_16 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_16, 0, x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTransition___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_convertTransition(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot convert local time ", 26, 26); +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" of the file", 12, 12); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_nat_dec_lt(x_8, x_5); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_10); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_7, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_7, x_15); +lean_dec(x_7); +lean_inc(x_2); +x_17 = l_Std_Time_TimeZone_convertLocalTimeType(x_8, x_1, x_2); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_2); +x_18 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_19 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1; +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2; +x_22 = lean_string_append(x_20, x_21); +x_23 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_23, 0, x_22); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 0); +lean_inc(x_24); +lean_dec(x_17); +x_25 = lean_array_push(x_10, x_24); +x_26 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_16; +x_8 = x_26; +x_9 = lean_box(0); +x_10 = x_25; +goto _start; +} +} +else +{ +lean_object* x_28; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +x_28 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_28, 0, x_10); +return x_28; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot convert transition ", 26, 26); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_nat_dec_lt(x_8, x_5); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_8); +lean_dec(x_7); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_10); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_7, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_7, x_15); +lean_dec(x_7); +x_17 = l_Std_Time_TimeZone_convertTransition(x_2, x_8, x_1); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_array_push(x_10, x_18); +x_20 = lean_nat_add(x_8, x_6); +lean_dec(x_8); +x_7 = x_16; +x_8 = x_20; +x_9 = lean_box(0); +x_10 = x_19; +goto _start; +} +else +{ +lean_object* x_22; +lean_dec(x_8); +lean_dec(x_7); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_10); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +x_4 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_4, 0, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_convertTZifV1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_convertTZifV1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("empty transitions for ", 22, 22); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_convertTZifV1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get_uint32(x_3, 16); +x_5 = lean_uint32_to_nat(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_unsigned_to_nat(1u); +lean_inc(x_5); +x_8 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_8, 0, x_6); +lean_ctor_set(x_8, 1, x_5); +lean_ctor_set(x_8, 2, x_7); +x_9 = l_Std_Time_TimeZone_convertTZifV1___closed__1; +lean_inc(x_5); +lean_inc(x_2); +x_10 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1(x_1, x_2, x_8, x_6, x_5, x_7, x_5, x_6, lean_box(0), x_9); +lean_dec(x_5); +lean_dec(x_8); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_2); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_ctor_get(x_1, 1); +x_16 = lean_array_get_size(x_15); +lean_inc(x_16); +x_17 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set(x_17, 2, x_7); +lean_inc(x_16); +x_18 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2(x_1, x_14, x_17, x_6, x_16, x_7, x_16, x_6, lean_box(0), x_9); +lean_dec(x_16); +lean_dec(x_17); +lean_dec(x_14); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_2); +x_21 = l_Std_Time_TimeZone_convertLocalTimeType(x_6, x_1, x_2); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_20); +x_22 = l_Std_Time_TimeZone_convertTZifV1___closed__2; +x_23 = lean_string_append(x_22, x_2); +lean_dec(x_2); +x_24 = l_Std_Time_TimeZone_convertTZifV1___closed__3; +x_25 = lean_string_append(x_23, x_24); +lean_ctor_set_tag(x_18, 0); +lean_ctor_set(x_18, 0, x_25); +return x_18; +} +else +{ +lean_object* x_26; lean_object* x_27; +lean_free_object(x_18); +lean_dec(x_2); +x_26 = lean_ctor_get(x_21, 0); +lean_inc(x_26); +lean_dec(x_21); +x_27 = l_Std_Time_TimeZone_convertTZifV1___lambda__1(x_20, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_18, 0); +lean_inc(x_28); +lean_dec(x_18); +lean_inc(x_2); +x_29 = l_Std_Time_TimeZone_convertLocalTimeType(x_6, x_1, x_2); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_28); +x_30 = l_Std_Time_TimeZone_convertTZifV1___closed__2; +x_31 = lean_string_append(x_30, x_2); +lean_dec(x_2); +x_32 = l_Std_Time_TimeZone_convertTZifV1___closed__3; +x_33 = lean_string_append(x_31, x_32); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_33); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_2); +x_35 = lean_ctor_get(x_29, 0); +lean_inc(x_35); +lean_dec(x_29); +x_36 = l_Std_Time_TimeZone_convertTZifV1___lambda__1(x_28, x_35); +return x_36; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_convertTZifV1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_TimeZone_convertTZifV1(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZifV2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_convertTZifV2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZif(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 1); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +x_5 = l_Std_Time_TimeZone_convertTZifV1(x_4, x_2); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 0); +x_7 = l_Std_Time_TimeZone_convertTZifV2(x_6, x_2); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_convertTZif___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_convertTZif(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_TzIf(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Database_Basic(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Zoned_ZoneRules(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_TzIf(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_TimeZone_convertLocalTimeType___closed__1 = _init_l_Std_Time_TimeZone_convertLocalTimeType___closed__1(); +l_Std_Time_TimeZone_convertLocalTimeType___closed__2 = _init_l_Std_Time_TimeZone_convertLocalTimeType___closed__2(); +l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__1); +l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2 = _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__1___closed__2); +l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___closed__1 = _init_l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at_Std_Time_TimeZone_convertTZifV1___spec__2___closed__1); +l_Std_Time_TimeZone_convertTZifV1___closed__1 = _init_l_Std_Time_TimeZone_convertTZifV1___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_convertTZifV1___closed__1); +l_Std_Time_TimeZone_convertTZifV1___closed__2 = _init_l_Std_Time_TimeZone_convertTZifV1___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_convertTZifV1___closed__2); +l_Std_Time_TimeZone_convertTZifV1___closed__3 = _init_l_Std_Time_TimeZone_convertTZifV1___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_convertTZifV1___closed__3); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Database/TZdb.c b/stage0/stdlib/Std/Time/Zoned/Database/TZdb.c new file mode 100644 index 000000000000..223f541796b3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Database/TZdb.c @@ -0,0 +1,830 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Database.TZdb +// Imports: Std.Time.DateTime Std.Time.Zoned.TimeZone Std.Time.Zoned.ZoneRules Std.Time.Zoned.Database.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__2; +lean_object* l_Std_Time_TimeZone_TZif_parse(lean_object*); +static lean_object* l_Std_Time_Database_TZdb_idFromPath___closed__2; +lean_object* l_System_FilePath_join(lean_object*, lean_object*); +lean_object* l_Std_Internal_Parsec_ByteArray_Parser_run___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZif(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_inst___closed__3; +uint8_t lean_string_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_inst___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_inst; +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1; +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__3; +lean_object* l_Std_Time_TimeZone_convertTZif(lean_object*, lean_object*); +lean_object* l_System_FilePath_components(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__5; +static lean_object* l_Std_Time_Database_TZdb_parseTZif___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_default; +static lean_object* l_Std_Time_Database_TZdb_default___closed__1; +static lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1; +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__1; +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Std_Time_Database_TZdb_parseTZIfFromDisk___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_inst___lambda__2(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_idFromPath(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_inst___closed__1; +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2; +static lean_object* l_Std_Time_Database_TZdb_default___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_readRulesFromDisk(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_default___closed__3; +static lean_object* l_Std_Time_Database_TZdb_localRules___closed__6; +lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_IO_FS_readBinFile(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_idFromPath___closed__3; +lean_object* l_IO_Process_run(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_inst___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_TZdb_idFromPath___closed__1; +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_Database_TZdb_default___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("/etc/localtime", 14, 14); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_default___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("/usr/share/zoneinfo/", 20, 20); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_default___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Database_TZdb_default___closed__1; +x_2 = l_Std_Time_Database_TZdb_default___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_default() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Database_TZdb_default___closed__3; +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_parseTZif___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_TZif_parse), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZif(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Time_Database_TZdb_parseTZif___closed__1; +x_4 = l_Std_Internal_Parsec_ByteArray_Parser_run___rarg(x_3, x_1); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +lean_dec(x_2); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +lean_dec(x_4); +x_9 = l_Std_Time_TimeZone_convertTZif(x_8, x_2); +lean_dec(x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Std_Time_Database_TZdb_parseTZIfFromDisk___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; +lean_ctor_set_tag(x_1, 18); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_2); +return x_7; +} +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_2); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Std_Time_Database_TZdb_parseTZif(x_2, x_1); +x_5 = l_IO_ofExcept___at_Std_Time_Database_TZdb_parseTZIfFromDisk___spec__1(x_4, x_3); +return x_5; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot find ", 12, 12); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" in the local timezone database", 31, 31); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_IO_FS_readBinFile(x_1, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = l_Std_Time_Database_TZdb_parseTZIfFromDisk___lambda__1(x_2, x_5, x_6); +return x_7; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_4, 0); +lean_dec(x_9); +x_10 = l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1; +x_11 = lean_string_append(x_10, x_2); +lean_dec(x_2); +x_12 = l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2; +x_13 = lean_string_append(x_11, x_12); +x_14 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_4, 0, x_14); +return x_4; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +lean_dec(x_4); +x_16 = l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1; +x_17 = lean_string_append(x_16, x_2); +lean_dec(x_2); +x_18 = l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2; +x_19 = lean_string_append(x_17, x_18); +x_20 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_20, 0, x_19); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_15); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_parseTZIfFromDisk___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_TZdb_parseTZIfFromDisk(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = 0; +return x_4; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_5; +x_5 = 0; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_1, 0); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_string_dec_eq(x_6, x_7); +return x_8; +} +} +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_idFromPath___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("zoneinfo", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_idFromPath___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_TZdb_idFromPath___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_idFromPath___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("/", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_idFromPath(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l_System_FilePath_components(x_1); +x_3 = lean_array_mk(x_2); +x_4 = lean_array_get_size(x_3); +x_5 = lean_unsigned_to_nat(1u); +x_6 = lean_nat_sub(x_4, x_5); +x_7 = l_Array_get_x3f___rarg(x_3, x_6); +lean_dec(x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +x_8 = lean_box(0); +return x_8; +} +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_7); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_7, 0); +x_11 = lean_unsigned_to_nat(2u); +x_12 = lean_nat_sub(x_4, x_11); +lean_dec(x_4); +x_13 = l_Array_get_x3f___rarg(x_3, x_12); +lean_dec(x_12); +lean_dec(x_3); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; +lean_free_object(x_7); +lean_dec(x_10); +x_14 = lean_box(0); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_13); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); +x_17 = l_Std_Time_Database_TZdb_idFromPath___closed__2; +x_18 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(x_13, x_17); +lean_dec(x_13); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = l_Std_Time_Database_TZdb_idFromPath___closed__3; +x_20 = lean_string_append(x_16, x_19); +x_21 = lean_string_append(x_20, x_10); +lean_dec(x_10); +lean_ctor_set(x_7, 0, x_21); +return x_7; +} +else +{ +lean_dec(x_16); +return x_7; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_13, 0); +lean_inc(x_22); +lean_dec(x_13); +lean_inc(x_22); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = l_Std_Time_Database_TZdb_idFromPath___closed__2; +x_25 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(x_23, x_24); +lean_dec(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = l_Std_Time_Database_TZdb_idFromPath___closed__3; +x_27 = lean_string_append(x_22, x_26); +x_28 = lean_string_append(x_27, x_10); +lean_dec(x_10); +lean_ctor_set(x_7, 0, x_28); +return x_7; +} +else +{ +lean_dec(x_22); +return x_7; +} +} +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_7, 0); +lean_inc(x_29); +lean_dec(x_7); +x_30 = lean_unsigned_to_nat(2u); +x_31 = lean_nat_sub(x_4, x_30); +lean_dec(x_4); +x_32 = l_Array_get_x3f___rarg(x_3, x_31); +lean_dec(x_31); +lean_dec(x_3); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; +lean_dec(x_29); +x_33 = lean_box(0); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_34 = lean_ctor_get(x_32, 0); +lean_inc(x_34); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + x_35 = x_32; +} else { + lean_dec_ref(x_32); + x_35 = lean_box(0); +} +lean_inc(x_34); +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(1, 1, 0); +} else { + x_36 = x_35; +} +lean_ctor_set(x_36, 0, x_34); +x_37 = l_Std_Time_Database_TZdb_idFromPath___closed__2; +x_38 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(x_36, x_37); +lean_dec(x_36); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = l_Std_Time_Database_TZdb_idFromPath___closed__3; +x_40 = lean_string_append(x_34, x_39); +x_41 = lean_string_append(x_40, x_29); +lean_dec(x_29); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +return x_42; +} +else +{ +lean_object* x_43; +lean_dec(x_34); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_29); +return x_43; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Std_Time_Database_TZdb_idFromPath___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot read the id of the path.", 31, 31); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_TZdb_idFromPath(x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2; +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_3); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +lean_dec(x_4); +x_8 = l_Std_Time_Database_TZdb_parseTZIfFromDisk(x_1, x_7, x_3); +return x_8; +} +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; +x_1 = 1; +x_2 = lean_alloc_ctor(0, 0, 3); +lean_ctor_set_uint8(x_2, 0, x_1); +lean_ctor_set_uint8(x_2, 1, x_1); +lean_ctor_set_uint8(x_2, 2, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-f", 2, 2); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("readlink", 8, 8); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot find the local timezone database", 39, 39); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_localRules___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_TZdb_localRules___closed__5; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; +x_3 = lean_box(0); +lean_inc(x_1); +x_4 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +x_5 = l_Std_Time_Database_TZdb_localRules___closed__2; +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +x_7 = lean_array_mk(x_6); +x_8 = lean_box(0); +x_9 = l_Std_Time_Database_TZdb_localRules___closed__1; +x_10 = l_Std_Time_Database_TZdb_localRules___closed__4; +x_11 = l_Std_Time_Database_TZdb_localRules___closed__3; +x_12 = 0; +x_13 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_13, 0, x_9); +lean_ctor_set(x_13, 1, x_10); +lean_ctor_set(x_13, 2, x_7); +lean_ctor_set(x_13, 3, x_8); +lean_ctor_set(x_13, 4, x_11); +lean_ctor_set_uint8(x_13, sizeof(void*)*5, x_12); +x_14 = l_IO_Process_run(x_13, x_2); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Std_Time_Database_TZdb_localRules___lambda__1(x_1, x_15, x_16); +lean_dec(x_1); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_14); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_14, 0); +lean_dec(x_19); +x_20 = l_Std_Time_Database_TZdb_localRules___closed__6; +lean_ctor_set(x_14, 0, x_20); +return x_14; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_14, 1); +lean_inc(x_21); +lean_dec(x_14); +x_22 = l_Std_Time_Database_TZdb_localRules___closed__6; +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_localRules___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_TZdb_localRules___lambda__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_readRulesFromDisk(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_System_FilePath_join(x_1, x_2); +x_5 = l_Std_Time_Database_TZdb_parseTZIfFromDisk(x_4, x_2, x_3); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_inst___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Std_Time_Database_TZdb_readRulesFromDisk(x_4, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_TZdb_inst___lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = l_Std_Time_Database_TZdb_localRules(x_3, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_inst___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_TZdb_inst___lambda__1), 3, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_inst___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_TZdb_inst___lambda__2), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_inst___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Database_TZdb_inst___closed__1; +x_2 = l_Std_Time_Database_TZdb_inst___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Database_TZdb_inst() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Database_TZdb_inst___closed__3; +return x_1; +} +} +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_TimeZone(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Database_TZdb(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_TimeZone(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_ZoneRules(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_Database_TZdb_default___closed__1 = _init_l_Std_Time_Database_TZdb_default___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_default___closed__1); +l_Std_Time_Database_TZdb_default___closed__2 = _init_l_Std_Time_Database_TZdb_default___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_default___closed__2); +l_Std_Time_Database_TZdb_default___closed__3 = _init_l_Std_Time_Database_TZdb_default___closed__3(); +lean_mark_persistent(l_Std_Time_Database_TZdb_default___closed__3); +l_Std_Time_Database_TZdb_default = _init_l_Std_Time_Database_TZdb_default(); +lean_mark_persistent(l_Std_Time_Database_TZdb_default); +l_Std_Time_Database_TZdb_parseTZif___closed__1 = _init_l_Std_Time_Database_TZdb_parseTZif___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_parseTZif___closed__1); +l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1 = _init_l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__1); +l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2 = _init_l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_parseTZIfFromDisk___closed__2); +l_Std_Time_Database_TZdb_idFromPath___closed__1 = _init_l_Std_Time_Database_TZdb_idFromPath___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_idFromPath___closed__1); +l_Std_Time_Database_TZdb_idFromPath___closed__2 = _init_l_Std_Time_Database_TZdb_idFromPath___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_idFromPath___closed__2); +l_Std_Time_Database_TZdb_idFromPath___closed__3 = _init_l_Std_Time_Database_TZdb_idFromPath___closed__3(); +lean_mark_persistent(l_Std_Time_Database_TZdb_idFromPath___closed__3); +l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1 = _init_l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___lambda__1___closed__1); +l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2 = _init_l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___lambda__1___closed__2); +l_Std_Time_Database_TZdb_localRules___closed__1 = _init_l_Std_Time_Database_TZdb_localRules___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__1); +l_Std_Time_Database_TZdb_localRules___closed__2 = _init_l_Std_Time_Database_TZdb_localRules___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__2); +l_Std_Time_Database_TZdb_localRules___closed__3 = _init_l_Std_Time_Database_TZdb_localRules___closed__3(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__3); +l_Std_Time_Database_TZdb_localRules___closed__4 = _init_l_Std_Time_Database_TZdb_localRules___closed__4(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__4); +l_Std_Time_Database_TZdb_localRules___closed__5 = _init_l_Std_Time_Database_TZdb_localRules___closed__5(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__5); +l_Std_Time_Database_TZdb_localRules___closed__6 = _init_l_Std_Time_Database_TZdb_localRules___closed__6(); +lean_mark_persistent(l_Std_Time_Database_TZdb_localRules___closed__6); +l_Std_Time_Database_TZdb_inst___closed__1 = _init_l_Std_Time_Database_TZdb_inst___closed__1(); +lean_mark_persistent(l_Std_Time_Database_TZdb_inst___closed__1); +l_Std_Time_Database_TZdb_inst___closed__2 = _init_l_Std_Time_Database_TZdb_inst___closed__2(); +lean_mark_persistent(l_Std_Time_Database_TZdb_inst___closed__2); +l_Std_Time_Database_TZdb_inst___closed__3 = _init_l_Std_Time_Database_TZdb_inst___closed__3(); +lean_mark_persistent(l_Std_Time_Database_TZdb_inst___closed__3); +l_Std_Time_Database_TZdb_inst = _init_l_Std_Time_Database_TZdb_inst(); +lean_mark_persistent(l_Std_Time_Database_TZdb_inst); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Database/TzIf.c b/stage0/stdlib/Std/Time/Zoned/Database/TzIf.c new file mode 100644 index 000000000000..0dd039d48e6e --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Database/TzIf.c @@ -0,0 +1,6987 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Database.TzIf +// Imports: Init.Data.Range Std.Internal.Parsec Std.Internal.Parsec.ByteArray +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_parse(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZif; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(uint8_t); +uint8_t lean_byte_array_fget(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3; +lean_object* lean_uint32_to_nat(uint32_t); +static lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1; +static lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8; +static lean_object* l_Std_Time_TimeZone_TZif_instReprTZif___closed__1; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__12(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes(lean_object*, uint32_t, lean_object*); +lean_object* l_String_quote(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(uint8_t); +uint32_t lean_uint8_to_uint32(uint8_t); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20; +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1; +lean_object* l_Std_Internal_Parsec_ByteArray_skipBytes(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi64(lean_object*); +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__8(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu64(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____boxed(lean_object*, lean_object*); +lean_object* lean_string_push(lean_object*, uint32_t); +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprTZifV1; +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2; +static lean_object* l_Std_Time_TimeZone_TZif_instReprHeader___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices(uint32_t, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__4___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__6(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21; +uint8_t l_instDecidableNot___rarg(uint8_t); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter(lean_object*); +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7; +static lean_object* l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1; +LEAN_EXPORT lean_object* l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2; +LEAN_EXPORT lean_object* l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13; +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11; +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__10(lean_object*, lean_object*); +lean_object* l_Std_Internal_Parsec_ByteArray_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(uint32_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds(lean_object*, uint32_t, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___boxed(lean_object*); +lean_object* lean_uint64_to_nat(uint64_t); +extern lean_object* l_Std_Internal_Parsec_unexpectedEndOfInput; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSecond(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes(uint32_t, lean_object*); +lean_object* lean_array_to_list(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +static lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2; +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1; +lean_object* l_Int_repr(lean_object*); +lean_object* l_ByteArray_toUInt64BE_x21(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pbool(lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(lean_object*); +static lean_object* l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19; +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__9(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__15(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprLocalTimeType; +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1; +static lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprTZifV2; +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1(uint8_t); +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprLeapSecond; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__3___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedHeader; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6; +lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__4(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprHeader; +lean_object* lean_string_length(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2(lean_object*); +lean_object* l_ByteArray_toUInt64LE_x21(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instReprTZif; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14; +lean_object* lean_string_to_utf8(lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint32_t lean_uint32_lor(uint32_t, uint32_t); +lean_object* lean_panic_fn(lean_object*, lean_object*); +uint32_t lean_uint32_shift_left(uint32_t, uint32_t); +lean_object* lean_nat_shiftl(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4; +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1; +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi32(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24; +static lean_object* l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11; +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__3(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30; +static lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3; +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17; +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__14(lean_object*, lean_object*); +extern uint32_t l_UInt32_instInhabited; +uint8_t lean_byte_array_get(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11; +lean_object* lean_uint8_to_nat(uint8_t); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____boxed(lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__11(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations(lean_object*, uint32_t, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22; +static lean_object* l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__13(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_array_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZifV2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8; +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3; +lean_object* lean_array_get_size(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18; +static lean_object* l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +extern uint8_t l_UInt8_instInhabited; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1; +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_instInhabitedTZifV1; +lean_object* lean_byte_array_size(lean_object*); +lean_object* l_Int_negOfNat(lean_object*); +uint8_t lean_uint8_dec_eq(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2; +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5; +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeType(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5; +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1; +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("version", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isutcnt", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isstdcnt", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("leapcnt", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("timecnt", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("typecnt", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("charcnt", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76_(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint32_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint32_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint32_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint32_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint32_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint32_t x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_3 = lean_ctor_get_uint8(x_1, 24); +x_4 = lean_uint8_to_nat(x_3); +x_5 = l___private_Init_Data_Repr_0__Nat_reprFast(x_4); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7; +x_8 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +x_9 = 0; +x_10 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set_uint8(x_10, sizeof(void*)*1, x_9); +x_11 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_15 = lean_box(1); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11; +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_ctor_get_uint32(x_1, 0); +x_22 = lean_uint32_to_nat(x_21); +x_23 = l___private_Init_Data_Repr_0__Nat_reprFast(x_22); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_7); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set_uint8(x_26, sizeof(void*)*1, x_9); +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_13); +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_15); +x_30 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13; +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_19); +x_33 = lean_ctor_get_uint32(x_1, 4); +x_34 = lean_uint32_to_nat(x_33); +x_35 = l___private_Init_Data_Repr_0__Nat_reprFast(x_34); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14; +x_38 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +x_39 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set_uint8(x_39, sizeof(void*)*1, x_9); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_32); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_13); +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_15); +x_43 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16; +x_44 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_19); +x_46 = lean_ctor_get_uint32(x_1, 8); +x_47 = lean_uint32_to_nat(x_46); +x_48 = l___private_Init_Data_Repr_0__Nat_reprFast(x_47); +x_49 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_49, 0, x_48); +x_50 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_50, 0, x_7); +lean_ctor_set(x_50, 1, x_49); +x_51 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set_uint8(x_51, sizeof(void*)*1, x_9); +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_45); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_13); +x_54 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_15); +x_55 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18; +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_19); +x_58 = lean_ctor_get_uint32(x_1, 12); +x_59 = lean_uint32_to_nat(x_58); +x_60 = l___private_Init_Data_Repr_0__Nat_reprFast(x_59); +x_61 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_62, 0, x_7); +lean_ctor_set(x_62, 1, x_61); +x_63 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set_uint8(x_63, sizeof(void*)*1, x_9); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_57); +lean_ctor_set(x_64, 1, x_63); +x_65 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_13); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_15); +x_67 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20; +x_68 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_19); +x_70 = lean_ctor_get_uint32(x_1, 16); +x_71 = lean_uint32_to_nat(x_70); +x_72 = l___private_Init_Data_Repr_0__Nat_reprFast(x_71); +x_73 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_73, 0, x_72); +x_74 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_74, 0, x_7); +lean_ctor_set(x_74, 1, x_73); +x_75 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set_uint8(x_75, sizeof(void*)*1, x_9); +x_76 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_76, 0, x_69); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_13); +x_78 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_15); +x_79 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22; +x_80 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_19); +x_82 = lean_ctor_get_uint32(x_1, 20); +x_83 = lean_uint32_to_nat(x_82); +x_84 = l___private_Init_Data_Repr_0__Nat_reprFast(x_83); +x_85 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_86 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_86, 0, x_7); +lean_ctor_set(x_86, 1, x_85); +x_87 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set_uint8(x_87, sizeof(void*)*1, x_9); +x_88 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_88, 0, x_81); +lean_ctor_set(x_88, 1, x_87); +x_89 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_90 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_88); +x_91 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_92 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +x_93 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_94 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_92); +x_95 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set_uint8(x_95, sizeof(void*)*1, x_9); +return x_95; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprHeader___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprHeader() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprHeader___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1() { +_start: +{ +uint8_t x_1; uint32_t x_2; lean_object* x_3; +x_1 = 0; +x_2 = 0; +x_3 = lean_alloc_ctor(0, 0, 25); +lean_ctor_set_uint8(x_3, 24, x_1); +lean_ctor_set_uint32(x_3, 0, x_2); +lean_ctor_set_uint32(x_3, 4, x_2); +lean_ctor_set_uint32(x_3, 8, x_2); +lean_ctor_set_uint32(x_3, 12, x_2); +lean_ctor_set_uint32(x_3, 16, x_2); +lean_ctor_set_uint32(x_3, 20, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedHeader() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("gmtOffset", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(13u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isDst", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("abbreviationIndex", 17, 17); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(21u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); +x_7 = lean_ctor_get_uint8(x_1, sizeof(void*)*1 + 1); +x_8 = lean_uint8_to_nat(x_7); +x_9 = l___private_Init_Data_Repr_0__Nat_reprFast(x_8); +x_10 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12; +x_12 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = 0; +x_14 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set_uint8(x_14, sizeof(void*)*1, x_13); +if (x_5 == 0) +{ +lean_object* x_60; lean_object* x_61; +x_60 = l_Int_repr(x_3); +x_61 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_15 = x_61; +goto block_59; +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = l_Int_repr(x_3); +x_63 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_63, 0, x_62); +x_64 = lean_unsigned_to_nat(0u); +x_65 = l_Repr_addAppParen(x_63, x_64); +x_15 = x_65; +goto block_59; +} +block_59: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_16 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5; +x_17 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_15); +x_18 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_13); +x_19 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +x_21 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_22 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_box(1); +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8; +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +if (x_6 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_29 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16; +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_21); +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_23); +x_33 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11; +x_34 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_27); +x_36 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_14); +x_37 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_38 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_36); +x_39 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_42 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*1, x_13); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_44 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20; +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_28); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_21); +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_23); +x_48 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11; +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_27); +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_14); +x_52 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_53 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_57 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +x_58 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set_uint8(x_58, sizeof(void*)*1, x_13); +return x_58; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprLocalTimeType() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = 0; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 1, 2); +lean_ctor_set(x_4, 0, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*1, x_3); +lean_ctor_set_uint8(x_4, sizeof(void*)*1 + 1, x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("transitionTime", 14, 14); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(18u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("correction", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(14u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_int_dec_lt(x_6, x_4); +if (x_5 == 0) +{ +lean_object* x_52; lean_object* x_53; +x_52 = l_Int_repr(x_3); +x_53 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_53, 0, x_52); +x_8 = x_53; +goto block_51; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = l_Int_repr(x_3); +x_55 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_unsigned_to_nat(0u); +x_57 = l_Repr_addAppParen(x_55, x_56); +x_8 = x_57; +goto block_51; +} +block_51: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_9 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5; +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +x_13 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = lean_box(1); +x_18 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7; +x_20 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_22 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_23 = l_Int_repr(x_6); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8; +x_26 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_11); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_22); +lean_ctor_set(x_28, 1, x_27); +x_29 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_34 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_11); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_36 = l_Int_repr(x_6); +x_37 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_unsigned_to_nat(0u); +x_39 = l_Repr_addAppParen(x_37, x_38); +x_40 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8; +x_41 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set_uint8(x_42, sizeof(void*)*1, x_11); +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_22); +lean_ctor_set(x_43, 1, x_42); +x_44 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_49 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set_uint8(x_50, sizeof(void*)*1, x_11); +return x_50; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprLeapSecond() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6; +x_3 = lean_int_dec_lt(x_1, x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Int_repr(x_1); +x_5 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_5, 0, x_4); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = l_Int_repr(x_1); +x_7 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_7, 0, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Repr_addAppParen(x_7, x_8); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(uint8_t x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_uint8_to_nat(x_1); +x_3 = l___private_Init_Data_Repr_0__Nat_reprFast(x_2); +x_4 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_4, 0, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14; +return x_2; +} +else +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unbox(x_5); +lean_dec(x_5); +x_8 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unbox(x_11); +lean_dec(x_11); +x_15 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__6(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unbox(x_5); +lean_dec(x_5); +x_7 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(x_6); +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +x_10 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(x_9); +x_11 = l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__7(x_2, x_10, x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_5, x_7); +lean_dec(x_5); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_11, x_14); +lean_dec(x_11); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_5, x_6); +lean_dec(x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388_(x_8, x_9); +lean_dec(x_8); +x_11 = l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__9(x_2, x_10, x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_5, x_7); +lean_dec(x_5); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_11, x_14); +lean_dec(x_11); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__10(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_5, x_6); +lean_dec(x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278_(x_8, x_9); +lean_dec(x_8); +x_11 = l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__11(x_2, x_10, x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unbox(x_5); +lean_dec(x_5); +x_8 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(x_7); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unbox(x_11); +lean_dec(x_11); +x_15 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(x_14); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__12(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; uint8_t x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unbox(x_5); +lean_dec(x_5); +x_7 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(x_6); +return x_7; +} +else +{ +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +x_10 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(x_9); +x_11 = l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__13(x_2, x_10, x_4); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(x_5); +lean_dec(x_5); +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_3); +lean_ctor_set(x_8, 1, x_7); +x_2 = x_8; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +lean_inc(x_1); +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_1); +x_13 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(x_10); +lean_dec(x_10); +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +x_2 = x_14; +x_3 = x_11; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__14(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(x_5); +lean_dec(x_5); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(x_7); +lean_dec(x_7); +x_9 = l_List_foldl___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__15(x_2, x_8, x_4); +return x_9; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("header", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("transitionTimes", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(19u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("transitionIndices", 17, 17); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("localTimeTypes", 14, 14); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("abbreviations", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(17u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("leapSeconds", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(15u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("stdWallIndicators", 17, 17); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("utLocalIndicators", 17, 17); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[]", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76_(x_3, x_4); +lean_dec(x_3); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +x_21 = lean_array_get_size(x_20); +x_22 = lean_nat_dec_eq(x_21, x_4); +lean_dec(x_21); +x_23 = lean_ctor_get(x_1, 2); +lean_inc(x_23); +x_24 = lean_array_get_size(x_23); +x_25 = lean_nat_dec_eq(x_24, x_4); +lean_dec(x_24); +x_26 = lean_ctor_get(x_1, 3); +lean_inc(x_26); +x_27 = lean_array_get_size(x_26); +x_28 = lean_nat_dec_eq(x_27, x_4); +lean_dec(x_27); +x_29 = lean_ctor_get(x_1, 4); +lean_inc(x_29); +x_30 = lean_array_get_size(x_29); +x_31 = lean_nat_dec_eq(x_30, x_4); +lean_dec(x_30); +x_32 = lean_ctor_get(x_1, 5); +lean_inc(x_32); +x_33 = lean_array_get_size(x_32); +x_34 = lean_nat_dec_eq(x_33, x_4); +lean_dec(x_33); +x_35 = lean_ctor_get(x_1, 6); +lean_inc(x_35); +x_36 = lean_array_get_size(x_35); +x_37 = lean_nat_dec_eq(x_36, x_4); +lean_dec(x_36); +x_38 = lean_ctor_get(x_1, 7); +lean_inc(x_38); +lean_dec(x_1); +x_39 = lean_array_get_size(x_38); +x_40 = lean_nat_dec_eq(x_39, x_4); +lean_dec(x_39); +if (x_22 == 0) +{ +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; uint8_t x_205; lean_object* x_206; +x_196 = lean_array_to_list(x_20); +x_197 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_198 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__14(x_196, x_197); +x_199 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_200 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_200, 0, x_199); +lean_ctor_set(x_200, 1, x_198); +x_201 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_202 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_202, 0, x_200); +lean_ctor_set(x_202, 1, x_201); +x_203 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_204 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_204, 0, x_203); +lean_ctor_set(x_204, 1, x_202); +x_205 = 1; +x_206 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_206, 0, x_204); +lean_ctor_set_uint8(x_206, sizeof(void*)*1, x_205); +x_41 = x_206; +goto block_195; +} +else +{ +lean_object* x_207; +lean_dec(x_20); +x_207 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_41 = x_207; +goto block_195; +} +block_195: +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_42 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8; +x_43 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set_uint8(x_44, sizeof(void*)*1, x_8); +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_19); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_12); +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_14); +x_48 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10; +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_18); +if (x_25 == 0) +{ +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; lean_object* x_193; +x_183 = lean_array_to_list(x_23); +x_184 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_185 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__12(x_183, x_184); +x_186 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_187 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_185); +x_188 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_189 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_189, 0, x_187); +lean_ctor_set(x_189, 1, x_188); +x_190 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_191 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_189); +x_192 = 1; +x_193 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_193, 0, x_191); +lean_ctor_set_uint8(x_193, sizeof(void*)*1, x_192); +x_51 = x_193; +goto block_182; +} +else +{ +lean_object* x_194; +lean_dec(x_23); +x_194 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_51 = x_194; +goto block_182; +} +block_182: +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_52 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12; +x_53 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_8); +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_50); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_12); +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_14); +x_58 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12; +x_59 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_18); +if (x_28 == 0) +{ +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; uint8_t x_179; lean_object* x_180; +x_170 = lean_array_to_list(x_26); +x_171 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_172 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__10(x_170, x_171); +x_173 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_174 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set(x_174, 1, x_172); +x_175 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_176 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +x_177 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_178 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_176); +x_179 = 1; +x_180 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_180, 0, x_178); +lean_ctor_set_uint8(x_180, sizeof(void*)*1, x_179); +x_61 = x_180; +goto block_169; +} +else +{ +lean_object* x_181; +lean_dec(x_26); +x_181 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_61 = x_181; +goto block_169; +} +block_169: +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_62 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5; +x_63 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +x_64 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set_uint8(x_64, sizeof(void*)*1, x_8); +x_65 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_65, 0, x_60); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_12); +x_67 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_14); +x_68 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14; +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_18); +if (x_31 == 0) +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; +x_157 = lean_array_to_list(x_29); +x_158 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_159 = l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2054____spec__3(x_157, x_158); +x_160 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_161 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_159); +x_162 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_163 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_163, 0, x_161); +lean_ctor_set(x_163, 1, x_162); +x_164 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_165 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_163); +x_166 = 1; +x_167 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set_uint8(x_167, sizeof(void*)*1, x_166); +x_71 = x_167; +goto block_156; +} +else +{ +lean_object* x_168; +lean_dec(x_29); +x_168 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_71 = x_168; +goto block_156; +} +block_156: +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_72 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15; +x_73 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set_uint8(x_74, sizeof(void*)*1, x_8); +x_75 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_74); +x_76 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_12); +x_77 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_14); +x_78 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17; +x_79 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_18); +if (x_34 == 0) +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; lean_object* x_154; +x_144 = lean_array_to_list(x_32); +x_145 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_146 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__8(x_144, x_145); +x_147 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_148 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_146); +x_149 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_150 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_150, 0, x_148); +lean_ctor_set(x_150, 1, x_149); +x_151 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_152 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_150); +x_153 = 1; +x_154 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set_uint8(x_154, sizeof(void*)*1, x_153); +x_81 = x_154; +goto block_143; +} +else +{ +lean_object* x_155; +lean_dec(x_32); +x_155 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_81 = x_155; +goto block_143; +} +block_143: +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_82 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18; +x_83 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_83, 0, x_82); +lean_ctor_set(x_83, 1, x_81); +x_84 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set_uint8(x_84, sizeof(void*)*1, x_8); +x_85 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_85, 0, x_80); +lean_ctor_set(x_85, 1, x_84); +x_86 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_12); +x_87 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_14); +x_88 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20; +x_89 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_18); +if (x_37 == 0) +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; +x_131 = lean_array_to_list(x_35); +x_132 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_133 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__6(x_131, x_132); +x_134 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_135 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_133); +x_136 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_137 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_138 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_139 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_137); +x_140 = 1; +x_141 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set_uint8(x_141, sizeof(void*)*1, x_140); +x_91 = x_141; +goto block_130; +} +else +{ +lean_object* x_142; +lean_dec(x_35); +x_142 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31; +x_91 = x_142; +goto block_130; +} +block_130: +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_92 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_92, 0, x_52); +lean_ctor_set(x_92, 1, x_91); +x_93 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set_uint8(x_93, sizeof(void*)*1, x_8); +x_94 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_94, 0, x_90); +lean_ctor_set(x_94, 1, x_93); +x_95 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_95, 1, x_12); +x_96 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_14); +x_97 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22; +x_98 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_18); +if (x_40 == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_100 = lean_array_to_list(x_38); +x_101 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23; +x_102 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__6(x_100, x_101); +x_103 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27; +x_104 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_102); +x_105 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29; +x_106 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +x_107 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26; +x_108 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_106); +x_109 = 1; +x_110 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set_uint8(x_110, sizeof(void*)*1, x_109); +x_111 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_111, 0, x_52); +lean_ctor_set(x_111, 1, x_110); +x_112 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set_uint8(x_112, sizeof(void*)*1, x_8); +x_113 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_113, 0, x_99); +lean_ctor_set(x_113, 1, x_112); +x_114 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_115 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_113); +x_116 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_117 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +x_118 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_119 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set_uint8(x_120, sizeof(void*)*1, x_8); +return x_120; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +lean_dec(x_38); +x_121 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33; +x_122 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_122, 0, x_99); +lean_ctor_set(x_122, 1, x_121); +x_123 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_124 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_122); +x_125 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_126 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_126, 0, x_124); +lean_ctor_set(x_126, 1, x_125); +x_127 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_128 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_126); +x_129 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set_uint8(x_129, sizeof(void*)*1, x_8); +return x_129; +} +} +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__2(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__3___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__3(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__4___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__4(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____spec__5(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZifV1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1; +x_2 = l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1; +x_3 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 2, x_2); +lean_ctor_set(x_3, 3, x_2); +lean_ctor_set(x_3, 4, x_2); +lean_ctor_set(x_3, 5, x_2); +lean_ctor_set(x_3, 6, x_2); +lean_ctor_set(x_3, 7, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2; +return x_1; +} +} +static lean_object* _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("none", 4, 4); +return x_1; +} +} +static lean_object* _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("some ", 5, 5); +return x_1; +} +} +static lean_object* _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2; +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_String_quote(x_5); +lean_dec(x_5); +lean_ctor_set_tag(x_1, 3); +lean_ctor_set(x_1, 0, x_6); +x_7 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_1); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = l_String_quote(x_10); +lean_dec(x_10); +x_12 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l_Repr_addAppParen(x_14, x_2); +return x_15; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("toTZifV1", 8, 8); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("footer", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532_(x_3, x_4); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1(x_20, x_4); +x_22 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5; +x_23 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*1, x_8); +x_25 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set_uint8(x_32, sizeof(void*)*1, x_8); +return x_32; +} +} +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZifV2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2; +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_unsigned_to_nat(1024u); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748_(x_4, x_5); +x_7 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4; +x_8 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +x_9 = l_Repr_addAppParen(x_8, x_2); +return x_9; +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("v1", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(6u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("v2", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532_(x_3, x_4); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1(x_20, x_4); +x_22 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_22, 0, x_6); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set_uint8(x_23, sizeof(void*)*1, x_8); +x_24 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_24, 0, x_19); +lean_ctor_set(x_24, 1, x_23); +x_25 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26; +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28; +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25; +x_30 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_8); +return x_31; +} +} +LEAN_EXPORT lean_object* l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____spec__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZif___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instReprTZif() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instReprTZif___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_TZif_instInhabitedTZif() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1; +return x_1; +} +} +static lean_object* _init_l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1() { +_start: +{ +uint32_t x_1; lean_object* x_2; +x_1 = l_UInt32_instInhabited; +x_2 = lean_box_uint32(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("assertion violation: ", 21, 21); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("bs.size == 4\n ", 15, 15); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.Zoned.Database.TzIf", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("_private.Std.Time.Zoned.Database.TzIf.0.Std.Time.TimeZone.TZif.toUInt32", 71, 71); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5; +x_3 = lean_unsigned_to_nat(177u); +x_4 = lean_unsigned_to_nat(2u); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_byte_array_size(x_1); +x_3 = lean_unsigned_to_nat(4u); +x_4 = lean_nat_dec_eq(x_2, x_3); +lean_dec(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6; +x_6 = l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1(x_5); +return x_6; +} +else +{ +lean_object* x_7; uint8_t x_8; uint32_t x_9; uint32_t x_10; uint32_t x_11; lean_object* x_12; uint8_t x_13; uint32_t x_14; uint32_t x_15; uint32_t x_16; uint32_t x_17; lean_object* x_18; uint8_t x_19; uint32_t x_20; uint32_t x_21; uint32_t x_22; uint32_t x_23; lean_object* x_24; uint8_t x_25; uint32_t x_26; uint32_t x_27; lean_object* x_28; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_byte_array_get(x_1, x_7); +x_9 = lean_uint8_to_uint32(x_8); +x_10 = 24; +x_11 = lean_uint32_shift_left(x_9, x_10); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_byte_array_get(x_1, x_12); +x_14 = lean_uint8_to_uint32(x_13); +x_15 = 16; +x_16 = lean_uint32_shift_left(x_14, x_15); +x_17 = lean_uint32_lor(x_11, x_16); +x_18 = lean_unsigned_to_nat(2u); +x_19 = lean_byte_array_get(x_1, x_18); +x_20 = lean_uint8_to_uint32(x_19); +x_21 = 8; +x_22 = lean_uint32_shift_left(x_20, x_21); +x_23 = lean_uint32_lor(x_17, x_22); +x_24 = lean_unsigned_to_nat(3u); +x_25 = lean_byte_array_get(x_1, x_24); +x_26 = lean_uint8_to_uint32(x_25); +x_27 = lean_uint32_lor(x_23, x_26); +x_28 = lean_box_uint32(x_27); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_unsigned_to_nat(31u); +x_3 = lean_nat_shiftl(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint32_t x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(x_1); +x_3 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_4 = lean_uint32_to_nat(x_3); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1; +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_cstr_to_nat("4294967296"); +x_8 = lean_nat_sub(x_7, x_4); +lean_dec(x_4); +x_9 = l_Int_negOfNat(x_8); +lean_dec(x_8); +return x_9; +} +else +{ +lean_object* x_10; +x_10 = lean_nat_to_int(x_4); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_unsigned_to_nat(63u); +x_3 = lean_nat_shiftl(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_cstr_to_nat("18446744073709551616"); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint64_t x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = l_ByteArray_toUInt64BE_x21(x_1); +x_3 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_4 = lean_uint64_to_nat(x_3); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1; +x_6 = lean_nat_dec_lt(x_4, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2; +x_8 = lean_nat_sub(x_7, x_4); +lean_dec(x_4); +x_9 = l_Int_negOfNat(x_8); +lean_dec(x_8); +return x_9; +} +else +{ +lean_object* x_10; +x_10 = lean_nat_to_int(x_4); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_nat_dec_lt(x_7, x_4); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_9); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_6, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_6, x_15); +lean_dec(x_6); +lean_inc(x_1); +x_17 = lean_apply_1(x_1, x_10); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_array_push(x_9, x_19); +x_21 = lean_nat_add(x_7, x_5); +lean_dec(x_7); +x_6 = x_16; +x_7 = x_21; +x_8 = lean_box(0); +x_9 = x_20; +x_10 = x_18; +goto _start; +} +else +{ +uint8_t x_23; +lean_dec(x_16); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_17); +if (x_23 == 0) +{ +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_17, 0); +x_25 = lean_ctor_get(x_17, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_17); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +else +{ +lean_object* x_27; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_1); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_9); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg___boxed), 10, 0); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_unsigned_to_nat(1u); +lean_inc(x_1); +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_5); +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1; +lean_inc(x_1); +x_8 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg(x_2, x_6, x_4, x_1, x_5, x_1, x_4, lean_box(0), x_7, x_3); +lean_dec(x_1); +lean_dec(x_6); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +return x_8; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___spec__1___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu64(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(8u); +x_3 = l_Std_Internal_Parsec_ByteArray_take(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint64_t x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 1); +x_6 = l_ByteArray_toUInt64LE_x21(x_5); +lean_dec(x_5); +x_7 = lean_unbox_uint64(x_6); +lean_dec(x_6); +x_8 = lean_box_uint64(x_7); +lean_ctor_set(x_3, 1, x_8); +return x_3; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_3); +x_11 = l_ByteArray_toUInt64LE_x21(x_10); +lean_dec(x_10); +x_12 = lean_unbox_uint64(x_11); +lean_dec(x_11); +x_13 = lean_box_uint64(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_9); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_3); +if (x_15 == 0) +{ +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_3); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi64(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(8u); +x_3 = l_Std_Internal_Parsec_ByteArray_take(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64(x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_3); +x_9 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64(x_8); +lean_dec(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_3); +if (x_11 == 0) +{ +return x_3; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(4u); +x_3 = l_Std_Internal_Parsec_ByteArray_take(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 1); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(x_5); +lean_dec(x_5); +x_7 = lean_unbox_uint32(x_6); +lean_dec(x_6); +x_8 = lean_box_uint32(x_7); +lean_ctor_set(x_3, 1, x_8); +return x_3; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint32_t x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_3, 0); +x_10 = lean_ctor_get(x_3, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_3); +x_11 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32(x_10); +lean_dec(x_10); +x_12 = lean_unbox_uint32(x_11); +lean_dec(x_11); +x_13 = lean_box_uint32(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_9); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_3); +if (x_15 == 0) +{ +return x_3; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_3, 0); +x_17 = lean_ctor_get(x_3, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_3); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi32(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(4u); +x_3 = l_Std_Internal_Parsec_ByteArray_take(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_3, 1); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32(x_5); +lean_dec(x_5); +lean_ctor_set(x_3, 1, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_3); +x_9 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32(x_8); +lean_dec(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_3); +if (x_11 == 0) +{ +return x_3; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_3, 0); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_3); +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_byte_array_size(x_2); +x_5 = lean_nat_dec_lt(x_3, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +lean_dec(x_3); +lean_dec(x_2); +x_6 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set(x_7, 1, x_6); +return x_7; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 1); +lean_dec(x_9); +x_10 = lean_ctor_get(x_1, 0); +lean_dec(x_10); +x_11 = lean_byte_array_fget(x_2, x_3); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_3, x_12); +lean_dec(x_3); +lean_ctor_set(x_1, 1, x_13); +x_14 = lean_box(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +else +{ +uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_1); +x_16 = lean_byte_array_fget(x_2, x_3); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_3, x_17); +lean_dec(x_3); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_box(x_16); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pbool(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(x_1); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 1); +x_5 = 0; +x_6 = lean_unbox(x_4); +lean_dec(x_4); +x_7 = lean_uint8_dec_eq(x_6, x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; +x_8 = 1; +x_9 = lean_box(x_8); +lean_ctor_set(x_2, 1, x_9); +return x_2; +} +else +{ +uint8_t x_10; lean_object* x_11; +x_10 = 0; +x_11 = lean_box(x_10); +lean_ctor_set(x_2, 1, x_11); +return x_2; +} +} +else +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_2); +x_14 = 0; +x_15 = lean_unbox(x_13); +lean_dec(x_13); +x_16 = lean_uint8_dec_eq(x_15, x_14); +if (x_16 == 0) +{ +uint8_t x_17; lean_object* x_18; lean_object* x_19; +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 1, x_18); +return x_19; +} +else +{ +uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_20 = 0; +x_21 = lean_box(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_12); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +return x_2; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_2, 0); +x_25 = lean_ctor_get(x_2, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_2); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("TZif", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1; +x_2 = lean_string_to_utf8(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2; +x_3 = l_Std_Internal_Parsec_ByteArray_skipBytes(x_2, x_1); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_unsigned_to_nat(15u); +x_9 = l_Std_Internal_Parsec_ByteArray_take(x_8, x_6); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +x_11 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_18); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu32(x_24); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; uint32_t x_31; uint32_t x_32; uint32_t x_33; uint32_t x_34; uint32_t x_35; uint32_t x_36; +x_28 = lean_ctor_get(x_26, 1); +x_29 = lean_alloc_ctor(0, 0, 25); +x_30 = lean_unbox(x_7); +lean_dec(x_7); +lean_ctor_set_uint8(x_29, 24, x_30); +x_31 = lean_unbox_uint32(x_13); +lean_dec(x_13); +lean_ctor_set_uint32(x_29, 0, x_31); +x_32 = lean_unbox_uint32(x_16); +lean_dec(x_16); +lean_ctor_set_uint32(x_29, 4, x_32); +x_33 = lean_unbox_uint32(x_19); +lean_dec(x_19); +lean_ctor_set_uint32(x_29, 8, x_33); +x_34 = lean_unbox_uint32(x_22); +lean_dec(x_22); +lean_ctor_set_uint32(x_29, 12, x_34); +x_35 = lean_unbox_uint32(x_25); +lean_dec(x_25); +lean_ctor_set_uint32(x_29, 16, x_35); +x_36 = lean_unbox_uint32(x_28); +lean_dec(x_28); +lean_ctor_set_uint32(x_29, 20, x_36); +lean_ctor_set(x_26, 1, x_29); +return x_26; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; uint32_t x_41; uint32_t x_42; uint32_t x_43; uint32_t x_44; uint32_t x_45; uint32_t x_46; lean_object* x_47; +x_37 = lean_ctor_get(x_26, 0); +x_38 = lean_ctor_get(x_26, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_26); +x_39 = lean_alloc_ctor(0, 0, 25); +x_40 = lean_unbox(x_7); +lean_dec(x_7); +lean_ctor_set_uint8(x_39, 24, x_40); +x_41 = lean_unbox_uint32(x_13); +lean_dec(x_13); +lean_ctor_set_uint32(x_39, 0, x_41); +x_42 = lean_unbox_uint32(x_16); +lean_dec(x_16); +lean_ctor_set_uint32(x_39, 4, x_42); +x_43 = lean_unbox_uint32(x_19); +lean_dec(x_19); +lean_ctor_set_uint32(x_39, 8, x_43); +x_44 = lean_unbox_uint32(x_22); +lean_dec(x_22); +lean_ctor_set_uint32(x_39, 12, x_44); +x_45 = lean_unbox_uint32(x_25); +lean_dec(x_25); +lean_ctor_set_uint32(x_39, 16, x_45); +x_46 = lean_unbox_uint32(x_38); +lean_dec(x_38); +lean_ctor_set_uint32(x_39, 20, x_46); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_37); +lean_ctor_set(x_47, 1, x_39); +return x_47; +} +} +else +{ +uint8_t x_48; +lean_dec(x_25); +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_7); +x_48 = !lean_is_exclusive(x_26); +if (x_48 == 0) +{ +return x_26; +} +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_26, 0); +x_50 = lean_ctor_get(x_26, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_26); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +else +{ +uint8_t x_52; +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_7); +x_52 = !lean_is_exclusive(x_23); +if (x_52 == 0) +{ +return x_23; +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_23, 0); +x_54 = lean_ctor_get(x_23, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_23); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +else +{ +uint8_t x_56; +lean_dec(x_19); +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_7); +x_56 = !lean_is_exclusive(x_20); +if (x_56 == 0) +{ +return x_20; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_20, 0); +x_58 = lean_ctor_get(x_20, 1); +lean_inc(x_58); +lean_inc(x_57); +lean_dec(x_20); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +else +{ +uint8_t x_60; +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_7); +x_60 = !lean_is_exclusive(x_17); +if (x_60 == 0) +{ +return x_17; +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_17, 0); +x_62 = lean_ctor_get(x_17, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_17); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +} +else +{ +uint8_t x_64; +lean_dec(x_13); +lean_dec(x_7); +x_64 = !lean_is_exclusive(x_14); +if (x_64 == 0) +{ +return x_14; +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_14, 0); +x_66 = lean_ctor_get(x_14, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_14); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +uint8_t x_68; +lean_dec(x_7); +x_68 = !lean_is_exclusive(x_11); +if (x_68 == 0) +{ +return x_11; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_11, 0); +x_70 = lean_ctor_get(x_11, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_11); +x_71 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +return x_71; +} +} +} +else +{ +uint8_t x_72; +lean_dec(x_7); +x_72 = !lean_is_exclusive(x_9); +if (x_72 == 0) +{ +return x_9; +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_9, 0); +x_74 = lean_ctor_get(x_9, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_9); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +return x_75; +} +} +} +else +{ +uint8_t x_76; +x_76 = !lean_is_exclusive(x_5); +if (x_76 == 0) +{ +return x_5; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_5, 0); +x_78 = lean_ctor_get(x_5, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_5); +x_79 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +return x_79; +} +} +} +else +{ +uint8_t x_80; +x_80 = !lean_is_exclusive(x_3); +if (x_80 == 0) +{ +return x_3; +} +else +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_3, 0); +x_82 = lean_ctor_get(x_3, 1); +lean_inc(x_82); +lean_inc(x_81); +lean_dec(x_3); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_81); +lean_ctor_set(x_83, 1, x_82); +return x_83; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeType(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi32(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pbool(x_3); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(x_6); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; uint8_t x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_alloc_ctor(0, 1, 2); +lean_ctor_set(x_11, 0, x_4); +x_12 = lean_unbox(x_7); +lean_dec(x_7); +lean_ctor_set_uint8(x_11, sizeof(void*)*1, x_12); +x_13 = lean_unbox(x_10); +lean_dec(x_10); +lean_ctor_set_uint8(x_11, sizeof(void*)*1 + 1, x_13); +lean_ctor_set(x_8, 1, x_11); +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +x_16 = lean_alloc_ctor(0, 1, 2); +lean_ctor_set(x_16, 0, x_4); +x_17 = lean_unbox(x_7); +lean_dec(x_7); +lean_ctor_set_uint8(x_16, sizeof(void*)*1, x_17); +x_18 = lean_unbox(x_15); +lean_dec(x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*1 + 1, x_18); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_14); +lean_ctor_set(x_19, 1, x_16); +return x_19; +} +} +else +{ +uint8_t x_20; +lean_dec(x_7); +lean_dec(x_4); +x_20 = !lean_is_exclusive(x_8); +if (x_20 == 0) +{ +return x_8; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_8, 0); +x_22 = lean_ctor_get(x_8, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_8); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +else +{ +uint8_t x_24; +lean_dec(x_4); +x_24 = !lean_is_exclusive(x_5); +if (x_24 == 0) +{ +return x_5; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_5, 0); +x_26 = lean_ctor_get(x_5, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_5); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +else +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_2); +if (x_28 == 0) +{ +return x_2; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_2, 0); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_2); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSecond(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_apply_1(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi32(x_4); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_6, 1, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_10); +lean_ctor_set(x_13, 1, x_12); +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_5); +x_14 = !lean_is_exclusive(x_6); +if (x_14 == 0) +{ +return x_6; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_6, 0); +x_16 = lean_ctor_get(x_6, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_6); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +} +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_3); +if (x_18 == 0) +{ +return x_3; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_3, 0); +x_20 = lean_ctor_get(x_3, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_3); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes(lean_object* x_1, uint32_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_uint32_to_nat(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_4, x_1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes(x_1, x_4, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices(uint32_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_uint32_to_nat(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1; +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices(x_3, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeType), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes(uint32_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_uint32_to_nat(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1; +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes(x_3, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_nat_dec_lt(x_7, x_4); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_9); +return x_12; +} +else +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_unsigned_to_nat(0u); +x_14 = lean_nat_dec_eq(x_6, x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_6, x_15); +lean_dec(x_6); +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; uint8_t x_25; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +x_20 = l_UInt8_instInhabited; +x_21 = lean_box(x_20); +x_22 = lean_array_get(x_21, x_1, x_7); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +x_24 = 0; +x_25 = lean_uint8_dec_eq(x_23, x_24); +if (x_25 == 0) +{ +uint32_t x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_uint8_to_uint32(x_23); +x_27 = lean_string_push(x_18, x_26); +lean_ctor_set(x_9, 0, x_27); +x_28 = lean_nat_add(x_7, x_5); +lean_dec(x_7); +x_6 = x_16; +x_7 = x_28; +x_8 = lean_box(0); +goto _start; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_dec(x_16); +lean_dec(x_7); +x_30 = lean_array_push(x_19, x_18); +x_31 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1; +lean_ctor_set(x_9, 1, x_30); +lean_ctor_set(x_9, 0, x_31); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_10); +lean_ctor_set(x_32, 1, x_9); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; +x_33 = lean_ctor_get(x_9, 0); +x_34 = lean_ctor_get(x_9, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_9); +x_35 = l_UInt8_instInhabited; +x_36 = lean_box(x_35); +x_37 = lean_array_get(x_36, x_1, x_7); +x_38 = lean_unbox(x_37); +lean_dec(x_37); +x_39 = 0; +x_40 = lean_uint8_dec_eq(x_38, x_39); +if (x_40 == 0) +{ +uint32_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_uint8_to_uint32(x_38); +x_42 = lean_string_push(x_33, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_34); +x_44 = lean_nat_add(x_7, x_5); +lean_dec(x_7); +x_6 = x_16; +x_7 = x_44; +x_8 = lean_box(0); +x_9 = x_43; +goto _start; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_16); +lean_dec(x_7); +x_46 = lean_array_push(x_34, x_33); +x_47 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_10); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else +{ +lean_object* x_50; +lean_dec(x_7); +lean_dec(x_6); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_10); +lean_ctor_set(x_50, 1, x_9); +return x_50; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, size_t x_6, size_t x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_7, x_6); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_2); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +else +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_uget(x_5, x_7); +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_14 = lean_ctor_get_uint8(x_12, sizeof(void*)*1 + 1); +lean_dec(x_12); +x_15 = lean_uint8_to_nat(x_14); +x_16 = lean_unsigned_to_nat(1u); +lean_inc(x_2); +lean_inc(x_15); +x_17 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_16); +lean_inc(x_2); +lean_inc(x_15); +x_18 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1(x_3, x_17, x_15, x_2, x_16, x_2, x_15, lean_box(0), x_8, x_9); +lean_dec(x_15); +lean_dec(x_17); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = !lean_is_exclusive(x_19); +if (x_21 == 0) +{ +size_t x_22; size_t x_23; +x_22 = 1; +x_23 = lean_usize_add(x_7, x_22); +x_7 = x_23; +x_8 = x_19; +x_9 = x_20; +goto _start; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = 1; +x_29 = lean_usize_add(x_7, x_28); +x_7 = x_29; +x_8 = x_27; +x_9 = x_20; +goto _start; +} +} +else +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; +x_31 = lean_ctor_get(x_8, 0); +x_32 = lean_ctor_get(x_8, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_8); +x_33 = lean_ctor_get_uint8(x_12, sizeof(void*)*1 + 1); +lean_dec(x_12); +x_34 = lean_uint8_to_nat(x_33); +x_35 = lean_unsigned_to_nat(1u); +lean_inc(x_2); +lean_inc(x_34); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_2); +lean_ctor_set(x_36, 2, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_31); +lean_ctor_set(x_37, 1, x_32); +lean_inc(x_2); +lean_inc(x_34); +x_38 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1(x_3, x_36, x_34, x_2, x_35, x_2, x_34, lean_box(0), x_37, x_9); +lean_dec(x_34); +lean_dec(x_36); +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_39)) { + lean_ctor_release(x_39, 0); + lean_ctor_release(x_39, 1); + x_43 = x_39; +} else { + lean_dec_ref(x_39); + x_43 = lean_box(0); +} +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(0, 2, 0); +} else { + x_44 = x_43; +} +lean_ctor_set(x_44, 0, x_41); +lean_ctor_set(x_44, 1, x_42); +x_45 = 1; +x_46 = lean_usize_add(x_7, x_45); +x_7 = x_46; +x_8 = x_44; +x_9 = x_40; +goto _start; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations(lean_object* x_1, uint32_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_uint32_to_nat(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1; +lean_inc(x_4); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_4, x_5, x_3); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_box(0); +x_10 = lean_array_size(x_1); +x_11 = 0; +x_12 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1; +x_13 = l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2(x_1, x_4, x_8, x_9, x_1, x_10, x_11, x_12, x_7); +lean_dec(x_8); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 1); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_ctor_set(x_13, 1, x_16); +return x_13; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_13, 1); +x_18 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_inc(x_18); +lean_dec(x_13); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_4); +x_21 = !lean_is_exclusive(x_6); +if (x_21 == 0) +{ +return x_6; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_6, 0); +x_23 = lean_ctor_get(x_6, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_6); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_11 = lean_unbox_usize(x_7); +lean_dec(x_7); +x_12 = l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__2(x_1, x_2, x_3, x_4, x_5, x_10, x_11, x_8, x_9); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations(x_1, x_4, x_3); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds(lean_object* x_1, uint32_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_uint32_to_nat(x_2); +x_5 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSecond), 2, 1); +lean_closure_set(x_5, 0, x_1); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_4, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint32_t x_4; lean_object* x_5; +x_4 = lean_unbox_uint32(x_2); +lean_dec(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds(x_1, x_4, x_3); +return x_5; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pbool), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(uint32_t x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_uint32_to_nat(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1; +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg(x_3, x_4, x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint32_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(x_3, x_2); +return x_4; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi32), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; uint32_t x_5; lean_object* x_6; lean_object* x_7; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_ctor_get_uint32(x_4, 12); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1; +x_7 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes(x_6, x_5, x_3); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices(x_5, x_8); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; uint32_t x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get_uint32(x_4, 16); +x_14 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes(x_13, x_11); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; uint32_t x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_ctor_get_uint32(x_4, 20); +x_18 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations(x_16, x_17, x_15); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; uint32_t x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get_uint32(x_4, 8); +x_22 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds(x_6, x_21, x_19); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; uint32_t x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_ctor_get_uint32(x_4, 4); +x_26 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(x_25, x_23); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; uint32_t x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get_uint32(x_4, 0); +x_30 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(x_29, x_27); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_30, 1); +x_33 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_33, 0, x_4); +lean_ctor_set(x_33, 1, x_9); +lean_ctor_set(x_33, 2, x_12); +lean_ctor_set(x_33, 3, x_16); +lean_ctor_set(x_33, 4, x_20); +lean_ctor_set(x_33, 5, x_24); +lean_ctor_set(x_33, 6, x_28); +lean_ctor_set(x_33, 7, x_32); +lean_ctor_set(x_30, 1, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_30); +x_36 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_36, 0, x_4); +lean_ctor_set(x_36, 1, x_9); +lean_ctor_set(x_36, 2, x_12); +lean_ctor_set(x_36, 3, x_16); +lean_ctor_set(x_36, 4, x_20); +lean_ctor_set(x_36, 5, x_24); +lean_ctor_set(x_36, 6, x_28); +lean_ctor_set(x_36, 7, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +else +{ +uint8_t x_38; +lean_dec(x_28); +lean_dec(x_24); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_4); +x_38 = !lean_is_exclusive(x_30); +if (x_38 == 0) +{ +return x_30; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_30, 0); +x_40 = lean_ctor_get(x_30, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_30); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +} +} +else +{ +uint8_t x_42; +lean_dec(x_24); +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_4); +x_42 = !lean_is_exclusive(x_26); +if (x_42 == 0) +{ +return x_26; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_26, 0); +x_44 = lean_ctor_get(x_26, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_26); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +} +else +{ +uint8_t x_46; +lean_dec(x_20); +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_4); +x_46 = !lean_is_exclusive(x_22); +if (x_46 == 0) +{ +return x_22; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_22, 0); +x_48 = lean_ctor_get(x_22, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_22); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else +{ +uint8_t x_50; +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_4); +x_50 = !lean_is_exclusive(x_18); +if (x_50 == 0) +{ +return x_18; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_18, 0); +x_52 = lean_ctor_get(x_18, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_18); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +else +{ +uint8_t x_54; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_4); +x_54 = !lean_is_exclusive(x_14); +if (x_54 == 0) +{ +return x_14; +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_14, 0); +x_56 = lean_ctor_get(x_14, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_14); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +else +{ +uint8_t x_58; +lean_dec(x_9); +lean_dec(x_4); +x_58 = !lean_is_exclusive(x_10); +if (x_58 == 0) +{ +return x_10; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_10, 0); +x_60 = lean_ctor_get(x_10, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_10); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +else +{ +uint8_t x_62; +lean_dec(x_4); +x_62 = !lean_is_exclusive(x_7); +if (x_62 == 0) +{ +return x_7; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_7, 0); +x_64 = lean_ctor_get(x_7, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_7); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} +} +} +else +{ +uint8_t x_66; +x_66 = !lean_is_exclusive(x_2); +if (x_66 == 0) +{ +return x_2; +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_2, 0); +x_68 = lean_ctor_get(x_2, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_2); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +return x_69; +} +} +} +} +static lean_object* _init_l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("condition not satisfied", 23, 23); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +x_5 = lean_byte_array_size(x_3); +x_6 = lean_nat_dec_lt(x_4, x_5); +lean_dec(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_7 = l_Std_Internal_Parsec_unexpectedEndOfInput; +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +uint8_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_byte_array_fget(x_3, x_4); +x_10 = lean_box(x_9); +x_11 = lean_apply_1(x_1, x_10); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_4); +lean_dec(x_3); +x_13 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_2, 1); +lean_dec(x_16); +x_17 = lean_ctor_get(x_2, 0); +lean_dec(x_17); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_4, x_18); +lean_dec(x_4); +lean_ctor_set(x_2, 1, x_19); +x_20 = lean_box(x_9); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_2); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_2); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_4, x_22); +lean_dec(x_4); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_3); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_box(x_9); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +} +LEAN_EXPORT uint8_t l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1(uint8_t x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; uint8_t x_4; +x_2 = 10; +x_3 = lean_uint8_dec_eq(x_1, x_2); +x_4 = l_instDecidableNot___rarg(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1; +lean_inc(x_2); +x_4 = l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_array_push(x_1, x_6); +x_1 = x_7; +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_4); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_4, 0); +x_11 = lean_ctor_get(x_4, 1); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +x_14 = lean_nat_dec_eq(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_dec(x_1); +return x_4; +} +else +{ +lean_dec(x_11); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 1, x_1); +return x_4; +} +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_4, 0); +x_16 = lean_ctor_get(x_4, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_4); +x_17 = lean_ctor_get(x_2, 1); +lean_inc(x_17); +lean_dec(x_2); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +x_19 = lean_nat_dec_eq(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_1); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_15); +lean_ctor_set(x_20, 1, x_16); +return x_20; +} +else +{ +lean_object* x_21; +lean_dec(x_16); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_15); +lean_ctor_set(x_21, 1, x_1); +return x_21; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_5, x_4); +if (x_8 == 0) +{ +lean_object* x_9; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_6); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; uint32_t x_12; lean_object* x_13; size_t x_14; size_t x_15; +x_10 = lean_array_uget(x_3, x_5); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +x_12 = lean_uint8_to_uint32(x_11); +x_13 = lean_string_push(x_6, x_12); +x_14 = 1; +x_15 = lean_usize_add(x_5, x_14); +x_5 = x_15; +x_6 = x_13; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1; +x_4 = l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_box(0); +x_8 = lean_array_size(x_6); +x_9 = 0; +x_10 = l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1; +x_11 = l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3(x_6, x_7, x_6, x_8, x_9, x_10, x_5); +lean_dec(x_6); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_11, 1); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_11, 1, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_15); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_4); +if (x_19 == 0) +{ +return x_4; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_4, 0); +x_21 = lean_ctor_get(x_4, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_4); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pu8(x_1); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; uint8_t x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1; +x_7 = 10; +x_8 = lean_unbox(x_5); +lean_dec(x_5); +x_9 = lean_uint8_dec_eq(x_8, x_7); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_box(0); +lean_ctor_set(x_2, 1, x_10); +return x_2; +} +else +{ +lean_object* x_11; lean_object* x_12; +lean_free_object(x_2); +x_11 = lean_box(0); +x_12 = lean_apply_2(x_6, x_11, x_4); +return x_12; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; +x_13 = lean_ctor_get(x_2, 0); +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_2); +x_15 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1; +x_16 = 10; +x_17 = lean_unbox(x_14); +lean_dec(x_14); +x_18 = lean_uint8_dec_eq(x_17, x_16); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_13); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = lean_apply_2(x_15, x_21, x_13); +return x_22; +} +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_2); +if (x_23 == 0) +{ +return x_2; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_2, 0); +x_25 = lean_ctor_get(x_2, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_2); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___lambda__1(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +size_t x_8; size_t x_9; lean_object* x_10; +x_8 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_9 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_10 = l_Array_forIn_x27Unsafe_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__3(x_1, x_2, x_3, x_8, x_9, x_6, x_7); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___lambda__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_pi64), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_11; +lean_inc(x_1); +x_11 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader(x_1); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; uint32_t x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get_uint32(x_13, 12); +x_15 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1; +x_16 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionTimes(x_15, x_14, x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices(x_14, x_17); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; uint32_t x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_ctor_get_uint32(x_13, 16); +x_23 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes(x_22, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; uint32_t x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = lean_ctor_get_uint32(x_13, 20); +x_27 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations(x_25, x_26, x_24); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; uint32_t x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_ctor_get_uint32(x_13, 8); +x_31 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLeapSeconds(x_15, x_30, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; uint32_t x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_ctor_get_uint32(x_13, 4); +x_35 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(x_34, x_32); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; lean_object* x_37; uint32_t x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get_uint32(x_13, 0); +x_39 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators(x_38, x_36); +if (lean_obj_tag(x_39) == 0) +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_39, 0); +x_42 = lean_ctor_get(x_39, 1); +x_43 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter(x_41); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +lean_dec(x_1); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_43, 1); +x_46 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_46, 0, x_13); +lean_ctor_set(x_46, 1, x_18); +lean_ctor_set(x_46, 2, x_21); +lean_ctor_set(x_46, 3, x_25); +lean_ctor_set(x_46, 4, x_29); +lean_ctor_set(x_46, 5, x_33); +lean_ctor_set(x_46, 6, x_37); +lean_ctor_set(x_46, 7, x_42); +lean_ctor_set(x_39, 1, x_45); +lean_ctor_set(x_39, 0, x_46); +x_47 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_47, 0, x_39); +lean_ctor_set(x_43, 1, x_47); +return x_43; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_48 = lean_ctor_get(x_43, 0); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_43); +x_50 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_50, 0, x_13); +lean_ctor_set(x_50, 1, x_18); +lean_ctor_set(x_50, 2, x_21); +lean_ctor_set(x_50, 3, x_25); +lean_ctor_set(x_50, 4, x_29); +lean_ctor_set(x_50, 5, x_33); +lean_ctor_set(x_50, 6, x_37); +lean_ctor_set(x_50, 7, x_42); +lean_ctor_set(x_39, 1, x_49); +lean_ctor_set(x_39, 0, x_50); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_39); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_48); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +else +{ +lean_object* x_53; lean_object* x_54; +lean_free_object(x_39); +lean_dec(x_42); +lean_dec(x_37); +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_53 = lean_ctor_get(x_43, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_43, 1); +lean_inc(x_54); +lean_dec(x_43); +x_2 = x_53; +x_3 = x_54; +goto block_10; +} +} +else +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_39, 0); +x_56 = lean_ctor_get(x_39, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_39); +x_57 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter(x_55); +if (lean_obj_tag(x_57) == 0) +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_1); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_60 = x_57; +} else { + lean_dec_ref(x_57); + x_60 = lean_box(0); +} +x_61 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_61, 0, x_13); +lean_ctor_set(x_61, 1, x_18); +lean_ctor_set(x_61, 2, x_21); +lean_ctor_set(x_61, 3, x_25); +lean_ctor_set(x_61, 4, x_29); +lean_ctor_set(x_61, 5, x_33); +lean_ctor_set(x_61, 6, x_37); +lean_ctor_set(x_61, 7, x_56); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_59); +x_63 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_63, 0, x_62); +if (lean_is_scalar(x_60)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_60; +} +lean_ctor_set(x_64, 0, x_58); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_56); +lean_dec(x_37); +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_65 = lean_ctor_get(x_57, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_57, 1); +lean_inc(x_66); +lean_dec(x_57); +x_2 = x_65; +x_3 = x_66; +goto block_10; +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; +lean_dec(x_37); +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_67 = lean_ctor_get(x_39, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_39, 1); +lean_inc(x_68); +lean_dec(x_39); +x_2 = x_67; +x_3 = x_68; +goto block_10; +} +} +else +{ +lean_object* x_69; lean_object* x_70; +lean_dec(x_33); +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_69 = lean_ctor_get(x_35, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_35, 1); +lean_inc(x_70); +lean_dec(x_35); +x_2 = x_69; +x_3 = x_70; +goto block_10; +} +} +else +{ +lean_object* x_71; lean_object* x_72; +lean_dec(x_29); +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_71 = lean_ctor_get(x_31, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_31, 1); +lean_inc(x_72); +lean_dec(x_31); +x_2 = x_71; +x_3 = x_72; +goto block_10; +} +} +else +{ +lean_object* x_73; lean_object* x_74; +lean_dec(x_25); +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_73 = lean_ctor_get(x_27, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_27, 1); +lean_inc(x_74); +lean_dec(x_27); +x_2 = x_73; +x_3 = x_74; +goto block_10; +} +} +else +{ +lean_object* x_75; lean_object* x_76; +lean_dec(x_21); +lean_dec(x_18); +lean_dec(x_13); +x_75 = lean_ctor_get(x_23, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_23, 1); +lean_inc(x_76); +lean_dec(x_23); +x_2 = x_75; +x_3 = x_76; +goto block_10; +} +} +else +{ +lean_object* x_77; lean_object* x_78; +lean_dec(x_18); +lean_dec(x_13); +x_77 = lean_ctor_get(x_19, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_19, 1); +lean_inc(x_78); +lean_dec(x_19); +x_2 = x_77; +x_3 = x_78; +goto block_10; +} +} +else +{ +lean_object* x_79; lean_object* x_80; +lean_dec(x_13); +x_79 = lean_ctor_get(x_16, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_16, 1); +lean_inc(x_80); +lean_dec(x_16); +x_2 = x_79; +x_3 = x_80; +goto block_10; +} +} +else +{ +lean_object* x_81; lean_object* x_82; +x_81 = lean_ctor_get(x_11, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_11, 1); +lean_inc(x_82); +lean_dec(x_11); +x_2 = x_81; +x_3 = x_82; +goto block_10; +} +block_10: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = lean_nat_dec_eq(x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_2); +lean_ctor_set(x_7, 1, x_3); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_8 = lean_box(0); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_2); +lean_ctor_set(x_9, 1, x_8); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_TZif_parse(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1(x_1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_2, 1); +lean_inc(x_4); +lean_dec(x_2); +x_5 = l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2(x_3); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_4); +lean_ctor_set(x_8, 1, x_7); +lean_ctor_set(x_5, 1, x_8); +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_4); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_9); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +lean_dec(x_4); +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +return x_5; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_5); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_2); +if (x_17 == 0) +{ +return x_2; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_2, 0); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_2); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +lean_object* initialize_Init_Data_Range(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Parsec(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal_Parsec_ByteArray(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Database_TzIf(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Init_Data_Range(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal_Parsec(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal_Parsec_ByteArray(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__7); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__8); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__9); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__10); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__11); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__12); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__13); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__14); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__15); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__16); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__17); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__18); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__19); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__20); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__21); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__22); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__23); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__24); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__25); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__26); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__27); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprHeader____x40_Std_Time_Zoned_Database_TzIf___hyg_76____closed__28); +l_Std_Time_TimeZone_TZif_instReprHeader___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprHeader___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprHeader___closed__1); +l_Std_Time_TimeZone_TZif_instReprHeader = _init_l_Std_Time_TimeZone_TZif_instReprHeader(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprHeader); +l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedHeader___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedHeader = _init_l_Std_Time_TimeZone_TZif_instInhabitedHeader(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedHeader); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__7); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__8); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__9); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__10); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__11); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__12); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__13); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__14); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__15); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__16); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__17); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__18); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__19); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLocalTimeType____x40_Std_Time_Zoned_Database_TzIf___hyg_278____closed__20); +l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprLocalTimeType___closed__1); +l_Std_Time_TimeZone_TZif_instReprLocalTimeType = _init_l_Std_Time_TimeZone_TZif_instReprLocalTimeType(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprLocalTimeType); +l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType = _init_l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedLocalTimeType); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__7); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprLeapSecond____x40_Std_Time_Zoned_Database_TzIf___hyg_388____closed__8); +l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprLeapSecond___closed__1); +l_Std_Time_TimeZone_TZif_instReprLeapSecond = _init_l_Std_Time_TimeZone_TZif_instReprLeapSecond(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprLeapSecond); +l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond = _init_l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedLeapSecond); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__7); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__8); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__9); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__10); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__11); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__12); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__13); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__14); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__15); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__16); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__17); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__18); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__19); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__20); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__21); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__22); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__23); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__24); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__25); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__26); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__27); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__28); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__29); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__30); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__31); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__32); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV1____x40_Std_Time_Zoned_Database_TzIf___hyg_532____closed__33); +l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZifV1___closed__1); +l_Std_Time_TimeZone_TZif_instReprTZifV1 = _init_l_Std_Time_TimeZone_TZif_instReprTZifV1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZifV1); +l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZifV1___closed__2); +l_Std_Time_TimeZone_TZif_instInhabitedTZifV1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZifV1); +l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1 = _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1(); +lean_mark_persistent(l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__1); +l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2 = _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2(); +lean_mark_persistent(l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__2); +l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3 = _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3(); +lean_mark_persistent(l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__3); +l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4 = _init_l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4(); +lean_mark_persistent(l_Option_repr___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____spec__1___closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZifV2____x40_Std_Time_Zoned_Database_TzIf___hyg_748____closed__6); +l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZifV2___closed__1); +l_Std_Time_TimeZone_TZif_instReprTZifV2 = _init_l_Std_Time_TimeZone_TZif_instReprTZifV2(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZifV2); +l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZifV2___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedTZifV2 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZifV2(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZifV2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_reprTZif____x40_Std_Time_Zoned_Database_TzIf___hyg_838____closed__7); +l_Std_Time_TimeZone_TZif_instReprTZif___closed__1 = _init_l_Std_Time_TimeZone_TZif_instReprTZif___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZif___closed__1); +l_Std_Time_TimeZone_TZif_instReprTZif = _init_l_Std_Time_TimeZone_TZif_instReprTZif(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instReprTZif); +l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1 = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZif___closed__1); +l_Std_Time_TimeZone_TZif_instInhabitedTZif = _init_l_Std_Time_TimeZone_TZif_instInhabitedTZif(); +lean_mark_persistent(l_Std_Time_TimeZone_TZif_instInhabitedTZif); +l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1 = _init_l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1(); +lean_mark_persistent(l_panic___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___spec__1___boxed__const__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__3); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__4); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__5); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toUInt32___closed__6); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt32___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_toInt64___closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_manyN___rarg___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseHeader___closed__2); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTransitionIndices___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseLocalTimeTypes___closed__1); +l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1 = _init_l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_x27_loop___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___spec__1___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseAbbreviations___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseIndicators___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV1___closed__1); +l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1 = _init_l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1(); +lean_mark_persistent(l_Std_Internal_Parsec_satisfy___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__1___closed__1); +l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1 = _init_l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1(); +lean_mark_persistent(l_Std_Internal_Parsec_manyCore___at___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___spec__2___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseFooter___closed__1); +l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1 = _init_l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Database_TzIf_0__Std_Time_TimeZone_TZif_parseTZifV2___closed__1); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Database/Windows.c b/stage0/stdlib/Std/Time/Zoned/Database/Windows.c new file mode 100644 index 000000000000..bfe7290106d5 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Database/Windows.c @@ -0,0 +1,1039 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Database.Windows +// Imports: Std.Time.DateTime Std.Time.Zoned.TimeZone Std.Time.Zoned.ZoneRules Std.Time.Zoned.Database.Basic +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst; +lean_object* lean_int64_to_int_sint(uint64_t); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(lean_object*, lean_object*, uint64_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_Windows_getZoneRules___closed__3; +static lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1; +static lean_object* l_Std_Time_Database_WindowsDb_inst___closed__2; +static lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2; +static lean_object* l_Std_Time_Database_Windows_getZoneRules___closed__4; +uint64_t lean_int64_neg(uint64_t); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___lambda__1(lean_object*, uint64_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules_toLocalTime(lean_object*); +uint8_t lean_int64_dec_le(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2(lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_WindowsDb_inst___closed__3; +static uint64_t l_Std_Time_Database_Windows_getZoneRules___closed__1; +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_windows_get_next_transition(lean_object*, uint64_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_toCtorIdx(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getNextTransition___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_WindowsDb_inst___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getLocalTimeZoneIdentifierAt___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_Database_Windows_getZoneRules___closed__5; +uint64_t lean_int64_of_nat(lean_object*); +static uint64_t l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules_toLocalTime___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_toCtorIdx___boxed(lean_object*); +lean_object* lean_get_windows_local_timezone_id_at(uint64_t, lean_object*); +LEAN_EXPORT lean_object* l_Bind_bindLeft___at_Std_Time_Database_WindowsDb_inst___spec__1(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1; +static uint64_t l_Std_Time_Database_Windows_getZoneRules___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_default; +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getNextTransition___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint64_t x_5; uint8_t x_6; lean_object* x_7; +x_5 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_6 = lean_unbox(x_3); +lean_dec(x_3); +x_7 = lean_windows_get_next_transition(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getLocalTimeZoneIdentifierAt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; lean_object* x_4; +x_3 = lean_unbox_uint64(x_1); +lean_dec(x_1); +x_4 = lean_get_windows_local_timezone_id_at(x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules_toLocalTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_1, 2); +x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_6 = 0; +x_7 = 1; +lean_inc(x_3); +lean_inc(x_4); +lean_inc(x_2); +x_8 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_4); +lean_ctor_set(x_8, 2, x_3); +lean_ctor_set_uint8(x_8, sizeof(void*)*3, x_5); +lean_ctor_set_uint8(x_8, sizeof(void*)*3 + 1, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*3 + 2, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules_toLocalTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Database_Windows_getZoneRules_toLocalTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, uint64_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_2); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_5); +return x_9; +} +} +static uint64_t _init_l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1() { +_start: +{ +lean_object* x_1; uint64_t x_2; +x_1 = lean_cstr_to_nat("32503690800"); +x_2 = lean_int64_of_nat(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; uint64_t x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = 0; +x_8 = lean_unbox_uint64(x_5); +x_9 = lean_windows_get_next_transition(x_1, x_8, x_7, x_3); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +lean_ctor_set(x_9, 0, x_2); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; uint64_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint64_t x_26; uint64_t x_27; uint8_t x_28; +x_16 = lean_ctor_get(x_9, 1); +x_17 = lean_ctor_get(x_9, 0); +lean_dec(x_17); +x_18 = lean_ctor_get(x_10, 0); +lean_inc(x_18); +lean_dec(x_10); +x_19 = lean_unbox_uint64(x_5); +x_20 = lean_int64_to_int_sint(x_19); +x_21 = lean_ctor_get(x_18, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +x_23 = l_Std_Time_Database_Windows_getZoneRules_toLocalTime(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_20); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_array_push(x_6, x_24); +x_26 = lean_unbox_uint64(x_21); +x_27 = lean_unbox_uint64(x_5); +x_28 = lean_int64_dec_le(x_26, x_27); +if (x_28 == 0) +{ +uint64_t x_29; uint64_t x_30; uint8_t x_31; +x_29 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1; +x_30 = lean_unbox_uint64(x_21); +lean_dec(x_21); +x_31 = lean_int64_dec_le(x_29, x_30); +if (x_31 == 0) +{ +lean_object* x_32; uint64_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_free_object(x_9); +lean_free_object(x_2); +x_32 = lean_box(0); +x_33 = lean_unbox_uint64(x_5); +lean_dec(x_5); +x_34 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(x_18, x_25, x_33, x_32, x_16); +lean_dec(x_18); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_2 = x_37; +x_3 = x_36; +goto _start; +} +else +{ +lean_dec(x_18); +lean_ctor_set(x_2, 1, x_25); +lean_ctor_set(x_9, 0, x_2); +return x_9; +} +} +else +{ +lean_dec(x_21); +lean_dec(x_18); +lean_ctor_set(x_2, 1, x_25); +lean_ctor_set(x_9, 0, x_2); +return x_9; +} +} +else +{ +lean_object* x_39; lean_object* x_40; uint64_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint64_t x_48; uint64_t x_49; uint8_t x_50; +x_39 = lean_ctor_get(x_9, 1); +lean_inc(x_39); +lean_dec(x_9); +x_40 = lean_ctor_get(x_10, 0); +lean_inc(x_40); +lean_dec(x_10); +x_41 = lean_unbox_uint64(x_5); +x_42 = lean_int64_to_int_sint(x_41); +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_40, 1); +lean_inc(x_44); +x_45 = l_Std_Time_Database_Windows_getZoneRules_toLocalTime(x_44); +lean_dec(x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_42); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_array_push(x_6, x_46); +x_48 = lean_unbox_uint64(x_43); +x_49 = lean_unbox_uint64(x_5); +x_50 = lean_int64_dec_le(x_48, x_49); +if (x_50 == 0) +{ +uint64_t x_51; uint64_t x_52; uint8_t x_53; +x_51 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1; +x_52 = lean_unbox_uint64(x_43); +lean_dec(x_43); +x_53 = lean_int64_dec_le(x_51, x_52); +if (x_53 == 0) +{ +lean_object* x_54; uint64_t x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_free_object(x_2); +x_54 = lean_box(0); +x_55 = lean_unbox_uint64(x_5); +lean_dec(x_5); +x_56 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(x_40, x_47, x_55, x_54, x_39); +lean_dec(x_40); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_ctor_get(x_57, 0); +lean_inc(x_59); +lean_dec(x_57); +x_2 = x_59; +x_3 = x_58; +goto _start; +} +else +{ +lean_object* x_61; +lean_dec(x_40); +lean_ctor_set(x_2, 1, x_47); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_2); +lean_ctor_set(x_61, 1, x_39); +return x_61; +} +} +else +{ +lean_object* x_62; +lean_dec(x_43); +lean_dec(x_40); +lean_ctor_set(x_2, 1, x_47); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_2); +lean_ctor_set(x_62, 1, x_39); +return x_62; +} +} +} +} +else +{ +uint8_t x_63; +lean_free_object(x_2); +lean_dec(x_6); +lean_dec(x_5); +x_63 = !lean_is_exclusive(x_9); +if (x_63 == 0) +{ +return x_9; +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_9, 0); +x_65 = lean_ctor_get(x_9, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_9); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +return x_66; +} +} +} +else +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; uint64_t x_70; lean_object* x_71; +x_67 = lean_ctor_get(x_2, 0); +x_68 = lean_ctor_get(x_2, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_2); +x_69 = 0; +x_70 = lean_unbox_uint64(x_67); +x_71 = lean_windows_get_next_transition(x_1, x_70, x_69, x_3); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_74 = x_71; +} else { + lean_dec_ref(x_71); + x_74 = lean_box(0); +} +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_67); +lean_ctor_set(x_75, 1, x_68); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_74; +} +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; uint64_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint64_t x_87; uint64_t x_88; uint8_t x_89; +x_77 = lean_ctor_get(x_71, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_78 = x_71; +} else { + lean_dec_ref(x_71); + x_78 = lean_box(0); +} +x_79 = lean_ctor_get(x_72, 0); +lean_inc(x_79); +lean_dec(x_72); +x_80 = lean_unbox_uint64(x_67); +x_81 = lean_int64_to_int_sint(x_80); +x_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_79, 1); +lean_inc(x_83); +x_84 = l_Std_Time_Database_Windows_getZoneRules_toLocalTime(x_83); +lean_dec(x_83); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_81); +lean_ctor_set(x_85, 1, x_84); +x_86 = lean_array_push(x_68, x_85); +x_87 = lean_unbox_uint64(x_82); +x_88 = lean_unbox_uint64(x_67); +x_89 = lean_int64_dec_le(x_87, x_88); +if (x_89 == 0) +{ +uint64_t x_90; uint64_t x_91; uint8_t x_92; +x_90 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1; +x_91 = lean_unbox_uint64(x_82); +lean_dec(x_82); +x_92 = lean_int64_dec_le(x_90, x_91); +if (x_92 == 0) +{ +lean_object* x_93; uint64_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_78); +x_93 = lean_box(0); +x_94 = lean_unbox_uint64(x_67); +lean_dec(x_67); +x_95 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(x_79, x_86, x_94, x_93, x_77); +lean_dec(x_79); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_ctor_get(x_96, 0); +lean_inc(x_98); +lean_dec(x_96); +x_2 = x_98; +x_3 = x_97; +goto _start; +} +else +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_79); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_67); +lean_ctor_set(x_100, 1, x_86); +if (lean_is_scalar(x_78)) { + x_101 = lean_alloc_ctor(0, 2, 0); +} else { + x_101 = x_78; +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_77); +return x_101; +} +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_dec(x_82); +lean_dec(x_79); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_67); +lean_ctor_set(x_102, 1, x_86); +if (lean_is_scalar(x_78)) { + x_103 = lean_alloc_ctor(0, 2, 0); +} else { + x_103 = x_78; +} +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_77); +return x_103; +} +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_68); +lean_dec(x_67); +x_104 = lean_ctor_get(x_71, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_71, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_106 = x_71; +} else { + lean_dec_ref(x_71); + x_106 = lean_box(0); +} +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(1, 2, 0); +} else { + x_107 = x_106; +} +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_105); +return x_107; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___lambda__1(lean_object* x_1, uint64_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_box_uint64(x_2); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +x_8 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1(x_1, x_7, x_5); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_8, 0); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +lean_ctor_set(x_10, 0, x_4); +return x_8; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_4); +lean_ctor_set(x_14, 1, x_13); +lean_ctor_set(x_8, 0, x_14); +return x_8; +} +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_15 = lean_ctor_get(x_8, 0); +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_8); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +if (lean_is_exclusive(x_15)) { + lean_ctor_release(x_15, 0); + lean_ctor_release(x_15, 1); + x_18 = x_15; +} else { + lean_dec_ref(x_15); + x_18 = lean_box(0); +} +if (lean_is_scalar(x_18)) { + x_19 = lean_alloc_ctor(0, 2, 0); +} else { + x_19 = x_18; +} +lean_ctor_set(x_19, 0, x_4); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_16); +return x_20; +} +} +else +{ +uint8_t x_21; +lean_dec(x_4); +x_21 = !lean_is_exclusive(x_8); +if (x_21 == 0) +{ +return x_8; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_8, 0); +x_23 = lean_ctor_get(x_8, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_8); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +} +static uint64_t _init_l_Std_Time_Database_Windows_getZoneRules___closed__1() { +_start: +{ +lean_object* x_1; uint64_t x_2; +x_1 = lean_unsigned_to_nat(2147483648u); +x_2 = lean_int64_of_nat(x_1); +return x_2; +} +} +static uint64_t _init_l_Std_Time_Database_Windows_getZoneRules___closed__2() { +_start: +{ +uint64_t x_1; uint64_t x_2; +x_1 = l_Std_Time_Database_Windows_getZoneRules___closed__1; +x_2 = lean_int64_neg(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_Windows_getZoneRules___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_Windows_getZoneRules___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot find first transition in zone rules", 42, 42); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_Windows_getZoneRules___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_Windows_getZoneRules___closed__4; +x_2 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint64_t x_3; uint8_t x_4; lean_object* x_5; +x_3 = l_Std_Time_Database_Windows_getZoneRules___closed__2; +x_4 = 1; +x_5 = lean_windows_get_next_transition(x_1, x_3, x_4, x_2); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +if (lean_obj_tag(x_6) == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_5, 0); +lean_dec(x_8); +x_9 = l_Std_Time_Database_Windows_getZoneRules___closed__5; +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 0, x_9); +return x_5; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_dec(x_5); +x_11 = l_Std_Time_Database_Windows_getZoneRules___closed__5; +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = lean_ctor_get(x_6, 0); +lean_inc(x_14); +lean_dec(x_6); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l_Std_Time_Database_Windows_getZoneRules_toLocalTime(x_15); +lean_dec(x_15); +x_17 = l_Std_Time_Database_Windows_getZoneRules___closed__3; +x_18 = l_Std_Time_Database_Windows_getZoneRules___lambda__1(x_1, x_3, x_17, x_16, x_13); +return x_18; +} +} +else +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_5); +if (x_19 == 0) +{ +return x_5; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_5, 0); +x_21 = lean_ctor_get(x_5, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_5); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +return x_22; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint64_t x_6; lean_object* x_7; +x_6 = lean_unbox_uint64(x_3); +lean_dec(x_3); +x_7 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___lambda__1(x_1, x_2, x_6, x_4, x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +uint64_t x_6; lean_object* x_7; +x_6 = lean_unbox_uint64(x_2); +lean_dec(x_2); +x_7 = l_Std_Time_Database_Windows_getZoneRules___lambda__1(x_1, x_6, x_3, x_4, x_5); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_Windows_getZoneRules___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Database_Windows_getZoneRules(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_toCtorIdx(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Database_WindowsDb_toCtorIdx(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___rarg(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Std_Time_Database_WindowsDb_noConfusion___rarg___boxed), 1, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_Database_WindowsDb_noConfusion___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_noConfusion___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_Database_WindowsDb_noConfusion(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Bind_bindLeft___at_Std_Time_Database_WindowsDb_inst___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_apply_1(x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_apply_2(x_1, x_5, x_6); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_4); +if (x_8 == 0) +{ +return x_4; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_4); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_Windows_getZoneRules(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1() { +_start: +{ +uint64_t x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_Windows_getZoneRules___closed__2; +x_2 = lean_box_uint64(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1; +x_2 = lean_alloc_closure((void*)(l_Std_Time_Database_Windows_getLocalTimeZoneIdentifierAt___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_Windows_getZoneRules___boxed), 2, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2; +x_4 = l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1; +x_5 = l_Bind_bindLeft___at_Std_Time_Database_WindowsDb_inst___spec__1(x_3, x_4, x_2); +return x_5; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_WindowsDb_inst___lambda__1___boxed), 3, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_Database_WindowsDb_inst___lambda__2___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_Database_WindowsDb_inst___closed__1; +x_2 = l_Std_Time_Database_WindowsDb_inst___closed__2; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_Database_WindowsDb_inst() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_Database_WindowsDb_inst___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_Database_WindowsDb_inst___lambda__1(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_Database_WindowsDb_inst___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_Database_WindowsDb_inst___lambda__2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_TimeZone(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_Database_Basic(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Database_Windows(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_TimeZone(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_ZoneRules(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_Database_Basic(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1 = _init_l_Lean_Loop_forIn_loop___at_Std_Time_Database_Windows_getZoneRules___spec__1___closed__1(); +l_Std_Time_Database_Windows_getZoneRules___closed__1 = _init_l_Std_Time_Database_Windows_getZoneRules___closed__1(); +l_Std_Time_Database_Windows_getZoneRules___closed__2 = _init_l_Std_Time_Database_Windows_getZoneRules___closed__2(); +l_Std_Time_Database_Windows_getZoneRules___closed__3 = _init_l_Std_Time_Database_Windows_getZoneRules___closed__3(); +lean_mark_persistent(l_Std_Time_Database_Windows_getZoneRules___closed__3); +l_Std_Time_Database_Windows_getZoneRules___closed__4 = _init_l_Std_Time_Database_Windows_getZoneRules___closed__4(); +lean_mark_persistent(l_Std_Time_Database_Windows_getZoneRules___closed__4); +l_Std_Time_Database_Windows_getZoneRules___closed__5 = _init_l_Std_Time_Database_Windows_getZoneRules___closed__5(); +lean_mark_persistent(l_Std_Time_Database_Windows_getZoneRules___closed__5); +l_Std_Time_Database_WindowsDb_default = _init_l_Std_Time_Database_WindowsDb_default(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_default); +l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1 = _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1___boxed__const__1); +l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1 = _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__1); +l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2 = _init_l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___lambda__2___closed__2); +l_Std_Time_Database_WindowsDb_inst___closed__1 = _init_l_Std_Time_Database_WindowsDb_inst___closed__1(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___closed__1); +l_Std_Time_Database_WindowsDb_inst___closed__2 = _init_l_Std_Time_Database_WindowsDb_inst___closed__2(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___closed__2); +l_Std_Time_Database_WindowsDb_inst___closed__3 = _init_l_Std_Time_Database_WindowsDb_inst___closed__3(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst___closed__3); +l_Std_Time_Database_WindowsDb_inst = _init_l_Std_Time_Database_WindowsDb_inst(); +lean_mark_persistent(l_Std_Time_Database_WindowsDb_inst); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/DateTime.c b/stage0/stdlib/Std/Time/Zoned/DateTime.c new file mode 100644 index 000000000000..576795946634 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/DateTime.c @@ -0,0 +1,7828 @@ +// Lean compiler output +// Module: Std.Time.Zoned.DateTime +// Imports: Std.Time.DateTime Std.Time.Zoned.TimeZone Std.Internal +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_DateTime_withDaysClip___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subDays(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMinutes(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday(lean_object*); +lean_object* lean_int_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__4(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second(lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsClip___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsClip___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMinutes___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___boxed(lean_object*); +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__5(lean_object*); +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsRollOver___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withSeconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysRollOver___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___boxed(lean_object*); +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMinutes___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime(lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsRollOver___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___boxed(lean_object*); +lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addNanoseconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMinutes(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_addWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___rarg(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_DateTime_era___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofDaysSinceUNIXEpoch___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__2(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subWeeks___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthRollOver___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_withMilliseconds___closed__2; +static lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMilliseconds(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_addMinutes___closed__1; +static lean_object* l_Std_Time_DateTime_withDaysClip___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withSeconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___boxed(lean_object*); +lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subSeconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsClip___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___closed__2; +LEAN_EXPORT uint8_t l_Std_Time_DateTime_inLeapYear___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___rarg___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instBEqDateTime___rarg(lean_object*, lean_object*); +uint8_t l_instDecidableNot___rarg(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddDuration___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDateTime_weekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withWeekday___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withHours(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__1(lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subWeeks(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearRollOver___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsClip___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___boxed(lean_object*); +static lean_object* l_Std_Time_DateTime_withMilliseconds___closed__1; +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMilliseconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__6(lean_object*); +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__4(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___rarg(lean_object*); +static lean_object* l_Std_Time_DateTime_addHours___closed__1; +lean_object* lean_thunk_get_own(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subDays___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddDuration(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_addYearsRollOver___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__2(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMilliseconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearClip___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withNanoseconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subNanoseconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addDays___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_mk_thunk(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day(lean_object*); +lean_object* l_Std_Time_PlainDate_rollOver(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDateTime_withWeekday(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__3(lean_object*); +lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withWeekday(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addWeeks___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withNanoseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear___boxed(lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMinutes___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_DateTime_withDaysClip___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsRollOver___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg(lean_object*); +lean_object* l_Std_Time_PlainTime_addMilliseconds(lean_object*, lean_object*); +uint8_t l_Std_Time_Year_Offset_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday___boxed(lean_object*); +static lean_object* l_Std_Time_DateTime_addMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsClip(lean_object*, lean_object*, lean_object*); +uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +lean_object* l_Std_Time_PlainDateTime_addMonthsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subHours(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__3(lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addWeeks(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMilliseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMilliseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subSeconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysClip___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___rarg___boxed(lean_object*); +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearClip(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___rarg___boxed(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofDaysSinceUNIXEpoch(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__6(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addSeconds(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___rarg(lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMinutes(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subNanoseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthClip___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withHours___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMilliseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addDays(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addSeconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsRollOver___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addHours___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addHours(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__5(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsRollOver(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedDateTime___closed__2; +static lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +LEAN_EXPORT uint8_t l_Std_Time_DateTime_weekday___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subHours___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addNanoseconds___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_instBEqDateTime___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = l___private_Std_Time_DateTime_Timestamp_0__Std_Time_beqTimestamp____x40_Std_Time_DateTime_Timestamp___hyg_56_(x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_instBEqDateTime___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_instBEqDateTime___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instBEqDateTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instBEqDateTime(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_4 = l_Std_Time_instInhabitedDateTime___lambda__1___closed__1; +x_5 = lean_int_add(x_1, x_4); +x_6 = lean_int_sub(x_5, x_1); +lean_dec(x_5); +x_7 = lean_int_add(x_6, x_1); +lean_dec(x_6); +x_8 = lean_int_sub(x_1, x_1); +x_9 = lean_int_emod(x_8, x_7); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = lean_int_emod(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = lean_int_add(x_11, x_1); +lean_dec(x_11); +x_13 = l_Std_Time_instInhabitedDateTime___lambda__1___closed__2; +x_14 = lean_int_add(x_1, x_13); +x_15 = lean_int_sub(x_14, x_1); +lean_dec(x_14); +x_16 = lean_int_add(x_15, x_1); +lean_dec(x_15); +x_17 = lean_int_emod(x_8, x_16); +lean_dec(x_8); +x_18 = lean_int_add(x_17, x_16); +lean_dec(x_17); +x_19 = lean_int_emod(x_18, x_16); +lean_dec(x_16); +lean_dec(x_18); +x_20 = lean_int_add(x_19, x_1); +lean_dec(x_19); +lean_inc(x_2); +x_21 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_21, 0, x_2); +lean_ctor_set(x_21, 1, x_12); +lean_ctor_set(x_21, 2, x_20); +x_22 = l_Std_Time_instInhabitedDateTime___lambda__1___closed__3; +x_23 = lean_int_add(x_2, x_22); +x_24 = lean_int_sub(x_23, x_2); +lean_dec(x_23); +x_25 = lean_int_add(x_24, x_1); +lean_dec(x_24); +x_26 = lean_int_sub(x_2, x_2); +x_27 = lean_int_emod(x_26, x_25); +x_28 = lean_int_add(x_27, x_25); +lean_dec(x_27); +x_29 = lean_int_emod(x_28, x_25); +lean_dec(x_25); +lean_dec(x_28); +x_30 = lean_int_add(x_29, x_2); +lean_dec(x_29); +x_31 = l_Std_Time_instInhabitedDateTime___lambda__1___closed__4; +x_32 = lean_int_add(x_2, x_31); +x_33 = lean_int_sub(x_32, x_2); +lean_dec(x_32); +x_34 = lean_int_add(x_33, x_1); +lean_dec(x_33); +x_35 = lean_int_emod(x_26, x_34); +lean_dec(x_26); +x_36 = lean_int_add(x_35, x_34); +lean_dec(x_35); +x_37 = lean_int_emod(x_36, x_34); +lean_dec(x_34); +lean_dec(x_36); +x_38 = lean_int_add(x_37, x_2); +lean_dec(x_37); +x_39 = 0; +x_40 = lean_box(x_39); +lean_inc(x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +x_42 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_42, 0, x_30); +lean_ctor_set(x_42, 1, x_38); +lean_ctor_set(x_42, 2, x_41); +lean_ctor_set(x_42, 3, x_2); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_21); +lean_ctor_set(x_43, 1, x_42); +return x_43; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedDateTime___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedDateTime___closed__2; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = l_Std_Time_instInhabitedDateTime___closed__1; +x_3 = l_Std_Time_instInhabitedDateTime___closed__2; +x_4 = lean_alloc_closure((void*)(l_Std_Time_instInhabitedDateTime___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_3); +x_5 = lean_mk_thunk(x_4); +x_6 = l_Std_Time_instInhabitedDateTime___closed__3; +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_instInhabitedDateTime___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedDateTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instInhabitedDateTime(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_5); +lean_dec(x_7); +x_9 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_mul(x_5, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_inc(x_1); +x_3 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +x_4 = lean_mk_thunk(x_3); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofTimestamp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_ofTimestamp___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toDaysSinceUNIXEpoch___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toDaysSinceUNIXEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toDaysSinceUNIXEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toTimestamp___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toTimestamp___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toTimestamp___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toTimestamp(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_1); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_1, 1); +lean_dec(x_5); +lean_inc(x_4); +x_6 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_6, 0, x_4); +lean_closure_set(x_6, 1, x_2); +x_7 = lean_mk_thunk(x_6); +lean_ctor_set(x_1, 1, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_2); +x_10 = lean_mk_thunk(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_8); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_convertTimeZone___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_convertTimeZone___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_convertTimeZone(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_4 = lean_ctor_get(x_1, 0); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = l_Std_Time_PlainTime_toSeconds(x_5); +x_7 = lean_int_add(x_6, x_4); +lean_dec(x_6); +x_8 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_5); +lean_dec(x_5); +x_11 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_12 = lean_int_mul(x_4, x_11); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_2, 0); +lean_inc(x_15); +lean_dec(x_2); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_9); +lean_dec(x_9); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_inc(x_1); +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_4 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_1); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_3); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_ofPlainDateTimeAssumingUTC___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_neg(x_3); +x_5 = lean_ctor_get(x_1, 1); +lean_inc(x_5); +x_6 = l_Std_Time_PlainTime_toSeconds(x_5); +x_7 = lean_int_add(x_6, x_4); +lean_dec(x_6); +x_8 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_10 = l_Std_Time_PlainTime_toNanoseconds(x_5); +lean_dec(x_5); +x_11 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_12 = lean_int_mul(x_4, x_11); +lean_dec(x_4); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_1, 0); +lean_inc(x_15); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_9); +lean_dec(x_9); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +x_20 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_19); +x_21 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_21, 0, x_1); +x_22 = lean_mk_thunk(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_ofPlainDateTime___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofPlainDateTime___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_ofPlainDateTime(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_DateTime_addHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addHours(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = l_Std_Time_DateTime_addHours___closed__1; +x_9 = lean_int_mul(x_3, x_8); +x_10 = lean_int_add(x_7, x_9); +lean_dec(x_7); +x_11 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_12 = lean_int_ediv(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_14 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_15 = lean_int_mul(x_9, x_14); +lean_dec(x_9); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_5, 0); +lean_inc(x_18); +lean_dec(x_5); +x_19 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_18); +x_20 = lean_int_add(x_19, x_12); +lean_dec(x_12); +lean_dec(x_19); +x_21 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_20); +lean_dec(x_20); +lean_inc(x_17); +lean_inc(x_21); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_17); +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_int_neg(x_23); +x_25 = l_Std_Time_PlainTime_toSeconds(x_17); +x_26 = lean_int_add(x_25, x_24); +lean_dec(x_25); +x_27 = lean_int_ediv(x_26, x_11); +lean_dec(x_26); +x_28 = l_Std_Time_PlainTime_toNanoseconds(x_17); +lean_dec(x_17); +x_29 = lean_int_mul(x_24, x_14); +lean_dec(x_24); +x_30 = lean_int_add(x_28, x_29); +lean_dec(x_29); +lean_dec(x_28); +x_31 = l_Std_Time_PlainTime_ofNanoseconds(x_30); +lean_dec(x_30); +x_32 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_21); +x_33 = lean_int_add(x_32, x_27); +lean_dec(x_27); +lean_dec(x_32); +x_34 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_33); +lean_dec(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +x_36 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_35); +x_37 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_37, 0, x_22); +x_38 = lean_mk_thunk(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addHours___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addHours(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subHours(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_int_neg(x_3); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = l_Std_Time_DateTime_addHours___closed__1; +x_10 = lean_int_mul(x_6, x_9); +lean_dec(x_6); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_8); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_10, x_15); +lean_dec(x_10); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_5, 0); +lean_inc(x_19); +lean_dec(x_5); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +lean_inc(x_18); +lean_inc(x_22); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_int_neg(x_24); +x_26 = l_Std_Time_PlainTime_toSeconds(x_18); +x_27 = lean_int_add(x_26, x_25); +lean_dec(x_26); +x_28 = lean_int_ediv(x_27, x_12); +lean_dec(x_27); +x_29 = l_Std_Time_PlainTime_toNanoseconds(x_18); +lean_dec(x_18); +x_30 = lean_int_mul(x_25, x_15); +lean_dec(x_25); +x_31 = lean_int_add(x_29, x_30); +lean_dec(x_30); +lean_dec(x_29); +x_32 = l_Std_Time_PlainTime_ofNanoseconds(x_31); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_22); +x_34 = lean_int_add(x_33, x_28); +lean_dec(x_28); +lean_dec(x_33); +x_35 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_34); +lean_dec(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_36); +x_38 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_38, 0, x_23); +x_39 = lean_mk_thunk(x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subHours___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subHours(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_addMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMinutes(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = l_Std_Time_DateTime_addMinutes___closed__1; +x_9 = lean_int_mul(x_3, x_8); +x_10 = lean_int_add(x_7, x_9); +lean_dec(x_7); +x_11 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_12 = lean_int_ediv(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_14 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_15 = lean_int_mul(x_9, x_14); +lean_dec(x_9); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_5, 0); +lean_inc(x_18); +lean_dec(x_5); +x_19 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_18); +x_20 = lean_int_add(x_19, x_12); +lean_dec(x_12); +lean_dec(x_19); +x_21 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_20); +lean_dec(x_20); +lean_inc(x_17); +lean_inc(x_21); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_17); +x_23 = lean_ctor_get(x_1, 0); +x_24 = lean_int_neg(x_23); +x_25 = l_Std_Time_PlainTime_toSeconds(x_17); +x_26 = lean_int_add(x_25, x_24); +lean_dec(x_25); +x_27 = lean_int_ediv(x_26, x_11); +lean_dec(x_26); +x_28 = l_Std_Time_PlainTime_toNanoseconds(x_17); +lean_dec(x_17); +x_29 = lean_int_mul(x_24, x_14); +lean_dec(x_24); +x_30 = lean_int_add(x_28, x_29); +lean_dec(x_29); +lean_dec(x_28); +x_31 = l_Std_Time_PlainTime_ofNanoseconds(x_30); +lean_dec(x_30); +x_32 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_21); +x_33 = lean_int_add(x_32, x_27); +lean_dec(x_27); +lean_dec(x_32); +x_34 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_33); +lean_dec(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +x_36 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_35); +x_37 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_37, 0, x_22); +x_38 = lean_mk_thunk(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMinutes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addMinutes(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMinutes(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_int_neg(x_3); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = l_Std_Time_DateTime_addMinutes___closed__1; +x_10 = lean_int_mul(x_6, x_9); +lean_dec(x_6); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_8); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_10, x_15); +lean_dec(x_10); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_5, 0); +lean_inc(x_19); +lean_dec(x_5); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +lean_inc(x_18); +lean_inc(x_22); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_int_neg(x_24); +x_26 = l_Std_Time_PlainTime_toSeconds(x_18); +x_27 = lean_int_add(x_26, x_25); +lean_dec(x_26); +x_28 = lean_int_ediv(x_27, x_12); +lean_dec(x_27); +x_29 = l_Std_Time_PlainTime_toNanoseconds(x_18); +lean_dec(x_18); +x_30 = lean_int_mul(x_25, x_15); +lean_dec(x_25); +x_31 = lean_int_add(x_29, x_30); +lean_dec(x_30); +lean_dec(x_29); +x_32 = l_Std_Time_PlainTime_ofNanoseconds(x_31); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_22); +x_34 = lean_int_add(x_33, x_28); +lean_dec(x_28); +lean_dec(x_33); +x_35 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_34); +lean_dec(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_36); +x_38 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_38, 0, x_23); +x_39 = lean_mk_thunk(x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMinutes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subMinutes(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addSeconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_3); +lean_dec(x_7); +x_9 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_mul(x_3, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_5, 0); +lean_inc(x_16); +lean_dec(x_5); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +lean_inc(x_15); +lean_inc(x_19); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_int_neg(x_21); +x_23 = l_Std_Time_PlainTime_toSeconds(x_15); +x_24 = lean_int_add(x_23, x_22); +lean_dec(x_23); +x_25 = lean_int_ediv(x_24, x_9); +lean_dec(x_24); +x_26 = l_Std_Time_PlainTime_toNanoseconds(x_15); +lean_dec(x_15); +x_27 = lean_int_mul(x_22, x_12); +lean_dec(x_22); +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_27); +lean_dec(x_26); +x_29 = l_Std_Time_PlainTime_ofNanoseconds(x_28); +lean_dec(x_28); +x_30 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_31 = lean_int_add(x_30, x_25); +lean_dec(x_25); +lean_dec(x_30); +x_32 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_31); +lean_dec(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +x_34 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_33); +x_35 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_35, 0, x_20); +x_36 = lean_mk_thunk(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addSeconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subSeconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_int_neg(x_3); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = lean_int_add(x_8, x_6); +lean_dec(x_8); +x_10 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_13 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_14 = lean_int_mul(x_6, x_13); +lean_dec(x_6); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_PlainTime_ofNanoseconds(x_15); +lean_dec(x_15); +x_17 = lean_ctor_get(x_5, 0); +lean_inc(x_17); +lean_dec(x_5); +x_18 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_17); +x_19 = lean_int_add(x_18, x_11); +lean_dec(x_11); +lean_dec(x_18); +x_20 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_19); +lean_dec(x_19); +lean_inc(x_16); +lean_inc(x_20); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_16); +x_22 = lean_ctor_get(x_1, 0); +x_23 = lean_int_neg(x_22); +x_24 = l_Std_Time_PlainTime_toSeconds(x_16); +x_25 = lean_int_add(x_24, x_23); +lean_dec(x_24); +x_26 = lean_int_ediv(x_25, x_10); +lean_dec(x_25); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_16); +lean_dec(x_16); +x_28 = lean_int_mul(x_23, x_13); +lean_dec(x_23); +x_29 = lean_int_add(x_27, x_28); +lean_dec(x_28); +lean_dec(x_27); +x_30 = l_Std_Time_PlainTime_ofNanoseconds(x_29); +lean_dec(x_29); +x_31 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_20); +x_32 = lean_int_add(x_31, x_26); +lean_dec(x_26); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_36 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_36, 0, x_21); +x_37 = lean_mk_thunk(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subSeconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_addMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMilliseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toMilliseconds(x_6); +x_8 = lean_int_add(x_7, x_3); +lean_dec(x_7); +x_9 = l_Std_Time_DateTime_addMilliseconds___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_addMilliseconds(x_6, x_3); +lean_dec(x_6); +x_12 = lean_ctor_get(x_5, 0); +lean_inc(x_12); +lean_dec(x_5); +x_13 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_14 = lean_int_add(x_13, x_10); +lean_dec(x_10); +lean_dec(x_13); +x_15 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_14); +lean_dec(x_14); +lean_inc(x_11); +lean_inc(x_15); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_11); +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_int_neg(x_17); +x_19 = l_Std_Time_PlainTime_toSeconds(x_11); +x_20 = lean_int_add(x_19, x_18); +lean_dec(x_19); +x_21 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_22 = lean_int_ediv(x_20, x_21); +lean_dec(x_20); +x_23 = l_Std_Time_PlainTime_toNanoseconds(x_11); +lean_dec(x_11); +x_24 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_25 = lean_int_mul(x_18, x_24); +lean_dec(x_18); +x_26 = lean_int_add(x_23, x_25); +lean_dec(x_25); +lean_dec(x_23); +x_27 = l_Std_Time_PlainTime_ofNanoseconds(x_26); +lean_dec(x_26); +x_28 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_29 = lean_int_add(x_28, x_22); +lean_dec(x_22); +lean_dec(x_28); +x_30 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_29); +lean_dec(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +x_32 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_31); +x_33 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_33, 0, x_16); +x_34 = lean_mk_thunk(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addMilliseconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMilliseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_int_neg(x_3); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toMilliseconds(x_7); +x_9 = lean_int_add(x_8, x_6); +lean_dec(x_8); +x_10 = l_Std_Time_DateTime_addMilliseconds___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_addMilliseconds(x_7, x_6); +lean_dec(x_6); +lean_dec(x_7); +x_13 = lean_ctor_get(x_5, 0); +lean_inc(x_13); +lean_dec(x_5); +x_14 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_15 = lean_int_add(x_14, x_11); +lean_dec(x_11); +lean_dec(x_14); +x_16 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_15); +lean_dec(x_15); +lean_inc(x_12); +lean_inc(x_16); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_12); +x_18 = lean_ctor_get(x_1, 0); +x_19 = lean_int_neg(x_18); +x_20 = l_Std_Time_PlainTime_toSeconds(x_12); +x_21 = lean_int_add(x_20, x_19); +lean_dec(x_20); +x_22 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_23 = lean_int_ediv(x_21, x_22); +lean_dec(x_21); +x_24 = l_Std_Time_PlainTime_toNanoseconds(x_12); +lean_dec(x_12); +x_25 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_26 = lean_int_mul(x_19, x_25); +lean_dec(x_19); +x_27 = lean_int_add(x_24, x_26); +lean_dec(x_26); +lean_dec(x_24); +x_28 = l_Std_Time_PlainTime_ofNanoseconds(x_27); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_30 = lean_int_add(x_29, x_23); +lean_dec(x_23); +lean_dec(x_29); +x_31 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_30); +lean_dec(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_28); +x_33 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_32); +x_34 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_34, 0, x_17); +x_35 = lean_mk_thunk(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subMilliseconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addNanoseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_ctor_get(x_8, 3); +lean_inc(x_9); +x_10 = lean_int_add(x_9, x_3); +lean_dec(x_9); +x_11 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_12 = lean_int_ediv(x_10, x_11); +x_13 = lean_int_emod(x_10, x_11); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_15 = lean_int_mul(x_12, x_11); +lean_dec(x_12); +x_16 = lean_int_add(x_14, x_15); +lean_dec(x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_19 = lean_ctor_get(x_17, 3); +lean_dec(x_19); +lean_ctor_set(x_17, 3, x_13); +lean_inc(x_17); +lean_inc(x_7); +lean_ctor_set(x_5, 1, x_17); +x_20 = lean_ctor_get(x_1, 0); +x_21 = lean_int_neg(x_20); +x_22 = l_Std_Time_PlainTime_toSeconds(x_17); +x_23 = lean_int_add(x_22, x_21); +lean_dec(x_22); +x_24 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_25 = lean_int_ediv(x_23, x_24); +lean_dec(x_23); +x_26 = l_Std_Time_PlainTime_toNanoseconds(x_17); +lean_dec(x_17); +x_27 = lean_int_mul(x_21, x_11); +lean_dec(x_21); +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_27); +lean_dec(x_26); +x_29 = l_Std_Time_PlainTime_ofNanoseconds(x_28); +lean_dec(x_28); +x_30 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_31 = lean_int_add(x_30, x_25); +lean_dec(x_25); +lean_dec(x_30); +x_32 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_31); +lean_dec(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +x_34 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_33); +x_35 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_35, 0, x_5); +x_36 = lean_mk_thunk(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_38 = lean_ctor_get(x_17, 0); +x_39 = lean_ctor_get(x_17, 1); +x_40 = lean_ctor_get(x_17, 2); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_17); +x_41 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_39); +lean_ctor_set(x_41, 2, x_40); +lean_ctor_set(x_41, 3, x_13); +lean_inc(x_41); +lean_inc(x_7); +lean_ctor_set(x_5, 1, x_41); +x_42 = lean_ctor_get(x_1, 0); +x_43 = lean_int_neg(x_42); +x_44 = l_Std_Time_PlainTime_toSeconds(x_41); +x_45 = lean_int_add(x_44, x_43); +lean_dec(x_44); +x_46 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_47 = lean_int_ediv(x_45, x_46); +lean_dec(x_45); +x_48 = l_Std_Time_PlainTime_toNanoseconds(x_41); +lean_dec(x_41); +x_49 = lean_int_mul(x_43, x_11); +lean_dec(x_43); +x_50 = lean_int_add(x_48, x_49); +lean_dec(x_49); +lean_dec(x_48); +x_51 = l_Std_Time_PlainTime_ofNanoseconds(x_50); +lean_dec(x_50); +x_52 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_53 = lean_int_add(x_52, x_47); +lean_dec(x_47); +lean_dec(x_52); +x_54 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_53); +lean_dec(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_51); +x_56 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_55); +x_57 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_57, 0, x_5); +x_58 = lean_mk_thunk(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_60 = lean_ctor_get(x_5, 0); +x_61 = lean_ctor_get(x_5, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_5); +x_62 = lean_ctor_get(x_61, 3); +lean_inc(x_62); +x_63 = lean_int_add(x_62, x_3); +lean_dec(x_62); +x_64 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_65 = lean_int_ediv(x_63, x_64); +x_66 = lean_int_emod(x_63, x_64); +lean_dec(x_63); +x_67 = l_Std_Time_PlainTime_toNanoseconds(x_61); +lean_dec(x_61); +x_68 = lean_int_mul(x_65, x_64); +lean_dec(x_65); +x_69 = lean_int_add(x_67, x_68); +lean_dec(x_68); +lean_dec(x_67); +x_70 = l_Std_Time_PlainTime_ofNanoseconds(x_69); +lean_dec(x_69); +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +x_73 = lean_ctor_get(x_70, 2); +lean_inc(x_73); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + lean_ctor_release(x_70, 2); + lean_ctor_release(x_70, 3); + x_74 = x_70; +} else { + lean_dec_ref(x_70); + x_74 = lean_box(0); +} +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(0, 4, 0); +} else { + x_75 = x_74; +} +lean_ctor_set(x_75, 0, x_71); +lean_ctor_set(x_75, 1, x_72); +lean_ctor_set(x_75, 2, x_73); +lean_ctor_set(x_75, 3, x_66); +lean_inc(x_75); +lean_inc(x_60); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_60); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_ctor_get(x_1, 0); +x_78 = lean_int_neg(x_77); +x_79 = l_Std_Time_PlainTime_toSeconds(x_75); +x_80 = lean_int_add(x_79, x_78); +lean_dec(x_79); +x_81 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_82 = lean_int_ediv(x_80, x_81); +lean_dec(x_80); +x_83 = l_Std_Time_PlainTime_toNanoseconds(x_75); +lean_dec(x_75); +x_84 = lean_int_mul(x_78, x_64); +lean_dec(x_78); +x_85 = lean_int_add(x_83, x_84); +lean_dec(x_84); +lean_dec(x_83); +x_86 = l_Std_Time_PlainTime_ofNanoseconds(x_85); +lean_dec(x_85); +x_87 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_60); +x_88 = lean_int_add(x_87, x_82); +lean_dec(x_82); +lean_dec(x_87); +x_89 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_88); +lean_dec(x_88); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_86); +x_91 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_90); +x_92 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_92, 0, x_76); +x_93 = lean_mk_thunk(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addNanoseconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subNanoseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_int_neg(x_3); +x_7 = !lean_is_exclusive(x_5); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_8 = lean_ctor_get(x_5, 0); +x_9 = lean_ctor_get(x_5, 1); +x_10 = lean_ctor_get(x_9, 3); +lean_inc(x_10); +x_11 = lean_int_add(x_10, x_6); +lean_dec(x_6); +lean_dec(x_10); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_ediv(x_11, x_12); +x_14 = lean_int_emod(x_11, x_12); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_16 = lean_int_mul(x_13, x_12); +lean_dec(x_13); +x_17 = lean_int_add(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_20 = lean_ctor_get(x_18, 3); +lean_dec(x_20); +lean_ctor_set(x_18, 3, x_14); +lean_inc(x_18); +lean_inc(x_8); +lean_ctor_set(x_5, 1, x_18); +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_int_neg(x_21); +x_23 = l_Std_Time_PlainTime_toSeconds(x_18); +x_24 = lean_int_add(x_23, x_22); +lean_dec(x_23); +x_25 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_26 = lean_int_ediv(x_24, x_25); +lean_dec(x_24); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_18); +lean_dec(x_18); +x_28 = lean_int_mul(x_22, x_12); +lean_dec(x_22); +x_29 = lean_int_add(x_27, x_28); +lean_dec(x_28); +lean_dec(x_27); +x_30 = l_Std_Time_PlainTime_ofNanoseconds(x_29); +lean_dec(x_29); +x_31 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_32 = lean_int_add(x_31, x_26); +lean_dec(x_26); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_36 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_36, 0, x_5); +x_37 = lean_mk_thunk(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_39 = lean_ctor_get(x_18, 0); +x_40 = lean_ctor_get(x_18, 1); +x_41 = lean_ctor_get(x_18, 2); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_18); +x_42 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_41); +lean_ctor_set(x_42, 3, x_14); +lean_inc(x_42); +lean_inc(x_8); +lean_ctor_set(x_5, 1, x_42); +x_43 = lean_ctor_get(x_1, 0); +x_44 = lean_int_neg(x_43); +x_45 = l_Std_Time_PlainTime_toSeconds(x_42); +x_46 = lean_int_add(x_45, x_44); +lean_dec(x_45); +x_47 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_48 = lean_int_ediv(x_46, x_47); +lean_dec(x_46); +x_49 = l_Std_Time_PlainTime_toNanoseconds(x_42); +lean_dec(x_42); +x_50 = lean_int_mul(x_44, x_12); +lean_dec(x_44); +x_51 = lean_int_add(x_49, x_50); +lean_dec(x_50); +lean_dec(x_49); +x_52 = l_Std_Time_PlainTime_ofNanoseconds(x_51); +lean_dec(x_51); +x_53 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_54 = lean_int_add(x_53, x_48); +lean_dec(x_48); +lean_dec(x_53); +x_55 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_54); +lean_dec(x_54); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_52); +x_57 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_56); +x_58 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_58, 0, x_5); +x_59 = lean_mk_thunk(x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_61 = lean_ctor_get(x_5, 0); +x_62 = lean_ctor_get(x_5, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_5); +x_63 = lean_ctor_get(x_62, 3); +lean_inc(x_63); +x_64 = lean_int_add(x_63, x_6); +lean_dec(x_6); +lean_dec(x_63); +x_65 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_66 = lean_int_ediv(x_64, x_65); +x_67 = lean_int_emod(x_64, x_65); +lean_dec(x_64); +x_68 = l_Std_Time_PlainTime_toNanoseconds(x_62); +lean_dec(x_62); +x_69 = lean_int_mul(x_66, x_65); +lean_dec(x_66); +x_70 = lean_int_add(x_68, x_69); +lean_dec(x_69); +lean_dec(x_68); +x_71 = l_Std_Time_PlainTime_ofNanoseconds(x_70); +lean_dec(x_70); +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_71, 2); +lean_inc(x_74); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + lean_ctor_release(x_71, 2); + lean_ctor_release(x_71, 3); + x_75 = x_71; +} else { + lean_dec_ref(x_71); + x_75 = lean_box(0); +} +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(0, 4, 0); +} else { + x_76 = x_75; +} +lean_ctor_set(x_76, 0, x_72); +lean_ctor_set(x_76, 1, x_73); +lean_ctor_set(x_76, 2, x_74); +lean_ctor_set(x_76, 3, x_67); +lean_inc(x_76); +lean_inc(x_61); +x_77 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_77, 0, x_61); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_ctor_get(x_1, 0); +x_79 = lean_int_neg(x_78); +x_80 = l_Std_Time_PlainTime_toSeconds(x_76); +x_81 = lean_int_add(x_80, x_79); +lean_dec(x_80); +x_82 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_83 = lean_int_ediv(x_81, x_82); +lean_dec(x_81); +x_84 = l_Std_Time_PlainTime_toNanoseconds(x_76); +lean_dec(x_76); +x_85 = lean_int_mul(x_79, x_65); +lean_dec(x_79); +x_86 = lean_int_add(x_84, x_85); +lean_dec(x_85); +lean_dec(x_84); +x_87 = l_Std_Time_PlainTime_ofNanoseconds(x_86); +lean_dec(x_86); +x_88 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_61); +x_89 = lean_int_add(x_88, x_83); +lean_dec(x_83); +lean_dec(x_88); +x_90 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_89); +lean_dec(x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_87); +x_92 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_91); +x_93 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_93, 0, x_77); +x_94 = lean_mk_thunk(x_93); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subNanoseconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addDays(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_10 = lean_int_add(x_9, x_3); +lean_dec(x_9); +x_11 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_10); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_31); +x_34 = lean_int_add(x_33, x_3); +lean_dec(x_33); +x_35 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_34); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addDays___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addDays(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subDays(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_int_neg(x_3); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_11 = lean_int_add(x_10, x_9); +lean_dec(x_9); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_11); +lean_dec(x_11); +lean_inc(x_8); +lean_inc(x_12); +lean_ctor_set(x_5, 0, x_12); +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_int_neg(x_13); +x_15 = l_Std_Time_PlainTime_toSeconds(x_8); +x_16 = lean_int_add(x_15, x_14); +lean_dec(x_15); +x_17 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_18 = lean_int_ediv(x_16, x_17); +lean_dec(x_16); +x_19 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_20 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_21 = lean_int_mul(x_14, x_20); +lean_dec(x_14); +x_22 = lean_int_add(x_19, x_21); +lean_dec(x_21); +lean_dec(x_19); +x_23 = l_Std_Time_PlainTime_ofNanoseconds(x_22); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_25 = lean_int_add(x_24, x_18); +lean_dec(x_18); +lean_dec(x_24); +x_26 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_25); +lean_dec(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +x_28 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_27); +x_29 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_29, 0, x_5); +x_30 = lean_mk_thunk(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = lean_int_neg(x_3); +x_35 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_32); +x_36 = lean_int_add(x_35, x_34); +lean_dec(x_34); +lean_dec(x_35); +x_37 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_36); +lean_dec(x_36); +lean_inc(x_33); +lean_inc(x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_33); +x_39 = lean_ctor_get(x_1, 0); +x_40 = lean_int_neg(x_39); +x_41 = l_Std_Time_PlainTime_toSeconds(x_33); +x_42 = lean_int_add(x_41, x_40); +lean_dec(x_41); +x_43 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_44 = lean_int_ediv(x_42, x_43); +lean_dec(x_42); +x_45 = l_Std_Time_PlainTime_toNanoseconds(x_33); +lean_dec(x_33); +x_46 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_47 = lean_int_mul(x_40, x_46); +lean_dec(x_40); +x_48 = lean_int_add(x_45, x_47); +lean_dec(x_47); +lean_dec(x_45); +x_49 = l_Std_Time_PlainTime_ofNanoseconds(x_48); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_37); +x_51 = lean_int_add(x_50, x_44); +lean_dec(x_44); +lean_dec(x_50); +x_52 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_51); +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_49); +x_54 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_53); +x_55 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_55, 0, x_38); +x_56 = lean_mk_thunk(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subDays___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subDays(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_addWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addWeeks(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_10 = l_Std_Time_DateTime_addWeeks___closed__1; +x_11 = lean_int_mul(x_3, x_10); +x_12 = lean_int_add(x_9, x_11); +lean_dec(x_11); +lean_dec(x_9); +x_13 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_12); +lean_dec(x_12); +lean_inc(x_8); +lean_inc(x_13); +lean_ctor_set(x_5, 0, x_13); +x_14 = lean_ctor_get(x_1, 0); +x_15 = lean_int_neg(x_14); +x_16 = l_Std_Time_PlainTime_toSeconds(x_8); +x_17 = lean_int_add(x_16, x_15); +lean_dec(x_16); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_19 = lean_int_ediv(x_17, x_18); +lean_dec(x_17); +x_20 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_21 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_22 = lean_int_mul(x_15, x_21); +lean_dec(x_15); +x_23 = lean_int_add(x_20, x_22); +lean_dec(x_22); +lean_dec(x_20); +x_24 = l_Std_Time_PlainTime_ofNanoseconds(x_23); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_26 = lean_int_add(x_25, x_19); +lean_dec(x_19); +lean_dec(x_25); +x_27 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +x_29 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_28); +x_30 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_30, 0, x_5); +x_31 = lean_mk_thunk(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_29); +lean_ctor_set(x_32, 1, x_31); +return x_32; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_33 = lean_ctor_get(x_5, 0); +x_34 = lean_ctor_get(x_5, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_5); +x_35 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_33); +x_36 = l_Std_Time_DateTime_addWeeks___closed__1; +x_37 = lean_int_mul(x_3, x_36); +x_38 = lean_int_add(x_35, x_37); +lean_dec(x_37); +lean_dec(x_35); +x_39 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_38); +lean_dec(x_38); +lean_inc(x_34); +lean_inc(x_39); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_34); +x_41 = lean_ctor_get(x_1, 0); +x_42 = lean_int_neg(x_41); +x_43 = l_Std_Time_PlainTime_toSeconds(x_34); +x_44 = lean_int_add(x_43, x_42); +lean_dec(x_43); +x_45 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_46 = lean_int_ediv(x_44, x_45); +lean_dec(x_44); +x_47 = l_Std_Time_PlainTime_toNanoseconds(x_34); +lean_dec(x_34); +x_48 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_49 = lean_int_mul(x_42, x_48); +lean_dec(x_42); +x_50 = lean_int_add(x_47, x_49); +lean_dec(x_49); +lean_dec(x_47); +x_51 = l_Std_Time_PlainTime_ofNanoseconds(x_50); +lean_dec(x_50); +x_52 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_39); +x_53 = lean_int_add(x_52, x_46); +lean_dec(x_46); +lean_dec(x_52); +x_54 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_53); +lean_dec(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_51); +x_56 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_55); +x_57 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_57, 0, x_40); +x_58 = lean_mk_thunk(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_58); +return x_59; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addWeeks___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addWeeks(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subWeeks(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_int_neg(x_3); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_7); +x_11 = l_Std_Time_DateTime_addWeeks___closed__1; +x_12 = lean_int_mul(x_9, x_11); +lean_dec(x_9); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_13); +lean_dec(x_13); +lean_inc(x_8); +lean_inc(x_14); +lean_ctor_set(x_5, 0, x_14); +x_15 = lean_ctor_get(x_1, 0); +x_16 = lean_int_neg(x_15); +x_17 = l_Std_Time_PlainTime_toSeconds(x_8); +x_18 = lean_int_add(x_17, x_16); +lean_dec(x_17); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_20 = lean_int_ediv(x_18, x_19); +lean_dec(x_18); +x_21 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_22 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_23 = lean_int_mul(x_16, x_22); +lean_dec(x_16); +x_24 = lean_int_add(x_21, x_23); +lean_dec(x_23); +lean_dec(x_21); +x_25 = l_Std_Time_PlainTime_ofNanoseconds(x_24); +lean_dec(x_24); +x_26 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +x_27 = lean_int_add(x_26, x_20); +lean_dec(x_20); +lean_dec(x_26); +x_28 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_27); +lean_dec(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +x_30 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_29); +x_31 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_31, 0, x_5); +x_32 = lean_mk_thunk(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_30); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_34 = lean_ctor_get(x_5, 0); +x_35 = lean_ctor_get(x_5, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_5); +x_36 = lean_int_neg(x_3); +x_37 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_34); +x_38 = l_Std_Time_DateTime_addWeeks___closed__1; +x_39 = lean_int_mul(x_36, x_38); +lean_dec(x_36); +x_40 = lean_int_add(x_37, x_39); +lean_dec(x_39); +lean_dec(x_37); +x_41 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_40); +lean_dec(x_40); +lean_inc(x_35); +lean_inc(x_41); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_35); +x_43 = lean_ctor_get(x_1, 0); +x_44 = lean_int_neg(x_43); +x_45 = l_Std_Time_PlainTime_toSeconds(x_35); +x_46 = lean_int_add(x_45, x_44); +lean_dec(x_45); +x_47 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_48 = lean_int_ediv(x_46, x_47); +lean_dec(x_46); +x_49 = l_Std_Time_PlainTime_toNanoseconds(x_35); +lean_dec(x_35); +x_50 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_51 = lean_int_mul(x_44, x_50); +lean_dec(x_44); +x_52 = lean_int_add(x_49, x_51); +lean_dec(x_51); +lean_dec(x_49); +x_53 = l_Std_Time_PlainTime_ofNanoseconds(x_52); +lean_dec(x_52); +x_54 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_41); +x_55 = lean_int_add(x_54, x_48); +lean_dec(x_48); +lean_dec(x_54); +x_56 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_55); +lean_dec(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_53); +x_58 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_57); +x_59 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_59, 0, x_42); +x_60 = lean_mk_thunk(x_59); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subWeeks___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subWeeks(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +lean_inc(x_5); +x_6 = l_Std_Time_PlainDateTime_addMonthsClip(x_5, x_3); +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_int_neg(x_7); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +lean_dec(x_5); +x_10 = l_Std_Time_PlainTime_toSeconds(x_9); +x_11 = lean_int_add(x_10, x_8); +lean_dec(x_10); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_8, x_15); +lean_dec(x_8); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_23); +x_25 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_25, 0, x_6); +x_26 = lean_mk_thunk(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addMonthsClip(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_int_neg(x_3); +x_10 = l_Std_Time_PlainDate_addMonthsClip(x_7, x_9); +lean_dec(x_9); +lean_inc(x_8); +lean_inc(x_10); +lean_ctor_set(x_5, 0, x_10); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_8); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_10); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_5, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_5); +x_32 = lean_int_neg(x_3); +x_33 = l_Std_Time_PlainDate_addMonthsClip(x_30, x_32); +lean_dec(x_32); +lean_inc(x_31); +lean_inc(x_33); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_31); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_31); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_31); +lean_dec(x_31); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_33); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_34); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subMonthsClip(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +lean_inc(x_5); +x_6 = l_Std_Time_PlainDateTime_addMonthsRollOver(x_5, x_3); +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_int_neg(x_7); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +lean_dec(x_5); +x_10 = l_Std_Time_PlainTime_toSeconds(x_9); +x_11 = lean_int_add(x_10, x_8); +lean_dec(x_10); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_8, x_15); +lean_dec(x_8); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_23); +x_25 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_25, 0, x_6); +x_26 = lean_mk_thunk(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addMonthsRollOver(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_int_neg(x_3); +x_10 = l_Std_Time_PlainDate_addMonthsRollOver(x_7, x_9); +lean_dec(x_9); +lean_inc(x_8); +lean_inc(x_10); +lean_ctor_set(x_5, 0, x_10); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_8); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_10); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_5, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_5); +x_32 = lean_int_neg(x_3); +x_33 = l_Std_Time_PlainDate_addMonthsRollOver(x_30, x_32); +lean_dec(x_32); +lean_inc(x_31); +lean_inc(x_33); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_31); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_31); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_31); +lean_dec(x_31); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_33); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_34); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subMonthsRollOver(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_addYearsRollOver___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_3, x_9); +x_11 = l_Std_Time_PlainDate_addMonthsRollOver(x_7, x_10); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_34 = lean_int_mul(x_3, x_33); +x_35 = l_Std_Time_PlainDate_addMonthsRollOver(x_31, x_34); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addYearsRollOver(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_3, x_9); +x_11 = l_Std_Time_PlainDate_addMonthsClip(x_7, x_10); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_34 = lean_int_mul(x_3, x_33); +x_35 = l_Std_Time_PlainDate_addMonthsClip(x_31, x_34); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_addYearsClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_addYearsClip(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_3, x_9); +x_11 = lean_int_neg(x_10); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_addMonthsRollOver(x_7, x_11); +lean_dec(x_11); +lean_inc(x_8); +lean_inc(x_12); +lean_ctor_set(x_5, 0, x_12); +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_int_neg(x_13); +x_15 = l_Std_Time_PlainTime_toSeconds(x_8); +x_16 = lean_int_add(x_15, x_14); +lean_dec(x_15); +x_17 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_18 = lean_int_ediv(x_16, x_17); +lean_dec(x_16); +x_19 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_20 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_21 = lean_int_mul(x_14, x_20); +lean_dec(x_14); +x_22 = lean_int_add(x_19, x_21); +lean_dec(x_21); +lean_dec(x_19); +x_23 = l_Std_Time_PlainTime_ofNanoseconds(x_22); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_25 = lean_int_add(x_24, x_18); +lean_dec(x_18); +lean_dec(x_24); +x_26 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_25); +lean_dec(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +x_28 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_27); +x_29 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_29, 0, x_5); +x_30 = lean_mk_thunk(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_35 = lean_int_mul(x_3, x_34); +x_36 = lean_int_neg(x_35); +lean_dec(x_35); +x_37 = l_Std_Time_PlainDate_addMonthsRollOver(x_32, x_36); +lean_dec(x_36); +lean_inc(x_33); +lean_inc(x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_33); +x_39 = lean_ctor_get(x_1, 0); +x_40 = lean_int_neg(x_39); +x_41 = l_Std_Time_PlainTime_toSeconds(x_33); +x_42 = lean_int_add(x_41, x_40); +lean_dec(x_41); +x_43 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_44 = lean_int_ediv(x_42, x_43); +lean_dec(x_42); +x_45 = l_Std_Time_PlainTime_toNanoseconds(x_33); +lean_dec(x_33); +x_46 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_47 = lean_int_mul(x_40, x_46); +lean_dec(x_40); +x_48 = lean_int_add(x_45, x_47); +lean_dec(x_47); +lean_dec(x_45); +x_49 = l_Std_Time_PlainTime_ofNanoseconds(x_48); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_37); +x_51 = lean_int_add(x_50, x_44); +lean_dec(x_44); +lean_dec(x_50); +x_52 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_51); +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_49); +x_54 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_53); +x_55 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_55, 0, x_38); +x_56 = lean_mk_thunk(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subYearsRollOver(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_3, x_9); +x_11 = lean_int_neg(x_10); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_addMonthsClip(x_7, x_11); +lean_dec(x_11); +lean_inc(x_8); +lean_inc(x_12); +lean_ctor_set(x_5, 0, x_12); +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_int_neg(x_13); +x_15 = l_Std_Time_PlainTime_toSeconds(x_8); +x_16 = lean_int_add(x_15, x_14); +lean_dec(x_15); +x_17 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_18 = lean_int_ediv(x_16, x_17); +lean_dec(x_16); +x_19 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_20 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_21 = lean_int_mul(x_14, x_20); +lean_dec(x_14); +x_22 = lean_int_add(x_19, x_21); +lean_dec(x_21); +lean_dec(x_19); +x_23 = l_Std_Time_PlainTime_ofNanoseconds(x_22); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_25 = lean_int_add(x_24, x_18); +lean_dec(x_18); +lean_dec(x_24); +x_26 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_25); +lean_dec(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +x_28 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_27); +x_29 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_29, 0, x_5); +x_30 = lean_mk_thunk(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_28); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_32 = lean_ctor_get(x_5, 0); +x_33 = lean_ctor_get(x_5, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_5); +x_34 = l_Std_Time_DateTime_addYearsRollOver___closed__1; +x_35 = lean_int_mul(x_3, x_34); +x_36 = lean_int_neg(x_35); +lean_dec(x_35); +x_37 = l_Std_Time_PlainDate_addMonthsClip(x_32, x_36); +lean_dec(x_36); +lean_inc(x_33); +lean_inc(x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_33); +x_39 = lean_ctor_get(x_1, 0); +x_40 = lean_int_neg(x_39); +x_41 = l_Std_Time_PlainTime_toSeconds(x_33); +x_42 = lean_int_add(x_41, x_40); +lean_dec(x_41); +x_43 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_44 = lean_int_ediv(x_42, x_43); +lean_dec(x_42); +x_45 = l_Std_Time_PlainTime_toNanoseconds(x_33); +lean_dec(x_33); +x_46 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_47 = lean_int_mul(x_40, x_46); +lean_dec(x_40); +x_48 = lean_int_add(x_45, x_47); +lean_dec(x_47); +lean_dec(x_45); +x_49 = l_Std_Time_PlainTime_ofNanoseconds(x_48); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_37); +x_51 = lean_int_add(x_50, x_44); +lean_dec(x_44); +lean_dec(x_50); +x_52 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_51); +lean_dec(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_49); +x_54 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_53); +x_55 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_55, 0, x_38); +x_56 = lean_mk_thunk(x_55); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_56); +return x_57; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_subYearsClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_subYearsClip(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_withDaysClip___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_DateTime_withDaysClip___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_DateTime_withDaysClip___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + x_8 = x_5; +} else { + lean_dec_ref(x_5); + x_8 = lean_box(0); +} +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +x_12 = lean_ctor_get(x_6, 2); +lean_dec(x_12); +x_13 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_14 = lean_int_mod(x_10, x_13); +x_15 = l_Std_Time_instInhabitedDateTime___closed__2; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_int_neg(x_17); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_18, x_19); +if (x_16 == 0) +{ +uint8_t x_40; lean_object* x_41; uint8_t x_42; +x_40 = 0; +x_41 = l_Std_Time_Month_Ordinal_days(x_40, x_11); +x_42 = lean_int_dec_lt(x_41, x_3); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_ctor_set(x_6, 2, x_3); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_6); +lean_ctor_set(x_43, 1, x_7); +x_21 = x_43; +goto block_39; +} +else +{ +lean_object* x_44; +lean_dec(x_3); +lean_ctor_set(x_6, 2, x_41); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_6); +lean_ctor_set(x_44, 1, x_7); +x_21 = x_44; +goto block_39; +} +} +else +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; uint8_t x_48; +x_45 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_46 = lean_int_mod(x_10, x_45); +x_47 = lean_int_dec_eq(x_46, x_15); +lean_dec(x_46); +x_48 = l_instDecidableNot___rarg(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_50 = lean_int_mod(x_10, x_49); +x_51 = lean_int_dec_eq(x_50, x_15); +lean_dec(x_50); +if (x_51 == 0) +{ +uint8_t x_52; lean_object* x_53; uint8_t x_54; +x_52 = 0; +x_53 = l_Std_Time_Month_Ordinal_days(x_52, x_11); +x_54 = lean_int_dec_lt(x_53, x_3); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_53); +lean_ctor_set(x_6, 2, x_3); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_6); +lean_ctor_set(x_55, 1, x_7); +x_21 = x_55; +goto block_39; +} +else +{ +lean_object* x_56; +lean_dec(x_3); +lean_ctor_set(x_6, 2, x_53); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_6); +lean_ctor_set(x_56, 1, x_7); +x_21 = x_56; +goto block_39; +} +} +else +{ +uint8_t x_57; lean_object* x_58; uint8_t x_59; +x_57 = 1; +x_58 = l_Std_Time_Month_Ordinal_days(x_57, x_11); +x_59 = lean_int_dec_lt(x_58, x_3); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_58); +lean_ctor_set(x_6, 2, x_3); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_6); +lean_ctor_set(x_60, 1, x_7); +x_21 = x_60; +goto block_39; +} +else +{ +lean_object* x_61; +lean_dec(x_3); +lean_ctor_set(x_6, 2, x_58); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_6); +lean_ctor_set(x_61, 1, x_7); +x_21 = x_61; +goto block_39; +} +} +} +else +{ +uint8_t x_62; lean_object* x_63; uint8_t x_64; +x_62 = 1; +x_63 = l_Std_Time_Month_Ordinal_days(x_62, x_11); +x_64 = lean_int_dec_lt(x_63, x_3); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_63); +lean_ctor_set(x_6, 2, x_3); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_6); +lean_ctor_set(x_65, 1, x_7); +x_21 = x_65; +goto block_39; +} +else +{ +lean_object* x_66; +lean_dec(x_3); +lean_ctor_set(x_6, 2, x_63); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_6); +lean_ctor_set(x_66, 1, x_7); +x_21 = x_66; +goto block_39; +} +} +} +block_39: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +x_23 = l_Std_Time_PlainTime_toSeconds(x_22); +x_24 = lean_int_add(x_23, x_18); +lean_dec(x_18); +lean_dec(x_23); +x_25 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_26 = lean_int_ediv(x_24, x_25); +lean_dec(x_24); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_22); +lean_dec(x_22); +x_28 = lean_int_add(x_27, x_20); +lean_dec(x_20); +lean_dec(x_27); +x_29 = l_Std_Time_PlainTime_ofNanoseconds(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +x_31 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_32 = lean_int_add(x_31, x_26); +lean_dec(x_26); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +if (lean_is_scalar(x_8)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_8; +} +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_29); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_36 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_36, 0, x_21); +x_37 = lean_mk_thunk(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_67 = lean_ctor_get(x_6, 0); +x_68 = lean_ctor_get(x_6, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_70 = lean_int_mod(x_67, x_69); +x_71 = l_Std_Time_instInhabitedDateTime___closed__2; +x_72 = lean_int_dec_eq(x_70, x_71); +lean_dec(x_70); +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_int_neg(x_73); +x_75 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_74, x_75); +if (x_72 == 0) +{ +uint8_t x_96; lean_object* x_97; uint8_t x_98; +x_96 = 0; +x_97 = l_Std_Time_Month_Ordinal_days(x_96, x_68); +x_98 = lean_int_dec_lt(x_97, x_3); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; +lean_dec(x_97); +x_99 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_99, 0, x_67); +lean_ctor_set(x_99, 1, x_68); +lean_ctor_set(x_99, 2, x_3); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_7); +x_77 = x_100; +goto block_95; +} +else +{ +lean_object* x_101; lean_object* x_102; +lean_dec(x_3); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_67); +lean_ctor_set(x_101, 1, x_68); +lean_ctor_set(x_101, 2, x_97); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_7); +x_77 = x_102; +goto block_95; +} +} +else +{ +lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; +x_103 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_104 = lean_int_mod(x_67, x_103); +x_105 = lean_int_dec_eq(x_104, x_71); +lean_dec(x_104); +x_106 = l_instDecidableNot___rarg(x_105); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_108 = lean_int_mod(x_67, x_107); +x_109 = lean_int_dec_eq(x_108, x_71); +lean_dec(x_108); +if (x_109 == 0) +{ +uint8_t x_110; lean_object* x_111; uint8_t x_112; +x_110 = 0; +x_111 = l_Std_Time_Month_Ordinal_days(x_110, x_68); +x_112 = lean_int_dec_lt(x_111, x_3); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +lean_dec(x_111); +x_113 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_113, 0, x_67); +lean_ctor_set(x_113, 1, x_68); +lean_ctor_set(x_113, 2, x_3); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_7); +x_77 = x_114; +goto block_95; +} +else +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_3); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_67); +lean_ctor_set(x_115, 1, x_68); +lean_ctor_set(x_115, 2, x_111); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_7); +x_77 = x_116; +goto block_95; +} +} +else +{ +uint8_t x_117; lean_object* x_118; uint8_t x_119; +x_117 = 1; +x_118 = l_Std_Time_Month_Ordinal_days(x_117, x_68); +x_119 = lean_int_dec_lt(x_118, x_3); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_118); +x_120 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_120, 0, x_67); +lean_ctor_set(x_120, 1, x_68); +lean_ctor_set(x_120, 2, x_3); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_7); +x_77 = x_121; +goto block_95; +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_3); +x_122 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_122, 0, x_67); +lean_ctor_set(x_122, 1, x_68); +lean_ctor_set(x_122, 2, x_118); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_7); +x_77 = x_123; +goto block_95; +} +} +} +else +{ +uint8_t x_124; lean_object* x_125; uint8_t x_126; +x_124 = 1; +x_125 = l_Std_Time_Month_Ordinal_days(x_124, x_68); +x_126 = lean_int_dec_lt(x_125, x_3); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; +lean_dec(x_125); +x_127 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_127, 0, x_67); +lean_ctor_set(x_127, 1, x_68); +lean_ctor_set(x_127, 2, x_3); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_7); +x_77 = x_128; +goto block_95; +} +else +{ +lean_object* x_129; lean_object* x_130; +lean_dec(x_3); +x_129 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_129, 0, x_67); +lean_ctor_set(x_129, 1, x_68); +lean_ctor_set(x_129, 2, x_125); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_7); +x_77 = x_130; +goto block_95; +} +} +} +block_95: +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +x_79 = l_Std_Time_PlainTime_toSeconds(x_78); +x_80 = lean_int_add(x_79, x_74); +lean_dec(x_74); +lean_dec(x_79); +x_81 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_82 = lean_int_ediv(x_80, x_81); +lean_dec(x_80); +x_83 = l_Std_Time_PlainTime_toNanoseconds(x_78); +lean_dec(x_78); +x_84 = lean_int_add(x_83, x_76); +lean_dec(x_76); +lean_dec(x_83); +x_85 = l_Std_Time_PlainTime_ofNanoseconds(x_84); +lean_dec(x_84); +x_86 = lean_ctor_get(x_77, 0); +lean_inc(x_86); +x_87 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_86); +x_88 = lean_int_add(x_87, x_82); +lean_dec(x_82); +lean_dec(x_87); +x_89 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_88); +lean_dec(x_88); +if (lean_is_scalar(x_8)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_8; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_85); +x_91 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_90); +x_92 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_92, 0, x_77); +x_93 = lean_mk_thunk(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withDaysClip(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = l_Std_Time_PlainDate_rollOver(x_9, x_10, x_3); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 1); +lean_inc(x_34); +lean_dec(x_31); +x_35 = l_Std_Time_PlainDate_rollOver(x_33, x_34, x_3); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withDaysRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withDaysRollOver(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + x_8 = x_5; +} else { + lean_dec_ref(x_5); + x_8 = lean_box(0); +} +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 2); +x_12 = lean_ctor_get(x_6, 1); +lean_dec(x_12); +x_13 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_14 = lean_int_mod(x_10, x_13); +x_15 = l_Std_Time_instInhabitedDateTime___closed__2; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_int_neg(x_17); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_18, x_19); +if (x_16 == 0) +{ +uint8_t x_40; lean_object* x_41; uint8_t x_42; +x_40 = 0; +x_41 = l_Std_Time_Month_Ordinal_days(x_40, x_3); +x_42 = lean_int_dec_lt(x_41, x_11); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_ctor_set(x_6, 1, x_3); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_6); +lean_ctor_set(x_43, 1, x_7); +x_21 = x_43; +goto block_39; +} +else +{ +lean_object* x_44; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_41); +lean_ctor_set(x_6, 1, x_3); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_6); +lean_ctor_set(x_44, 1, x_7); +x_21 = x_44; +goto block_39; +} +} +else +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; uint8_t x_48; +x_45 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_46 = lean_int_mod(x_10, x_45); +x_47 = lean_int_dec_eq(x_46, x_15); +lean_dec(x_46); +x_48 = l_instDecidableNot___rarg(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_50 = lean_int_mod(x_10, x_49); +x_51 = lean_int_dec_eq(x_50, x_15); +lean_dec(x_50); +if (x_51 == 0) +{ +uint8_t x_52; lean_object* x_53; uint8_t x_54; +x_52 = 0; +x_53 = l_Std_Time_Month_Ordinal_days(x_52, x_3); +x_54 = lean_int_dec_lt(x_53, x_11); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_53); +lean_ctor_set(x_6, 1, x_3); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_6); +lean_ctor_set(x_55, 1, x_7); +x_21 = x_55; +goto block_39; +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_53); +lean_ctor_set(x_6, 1, x_3); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_6); +lean_ctor_set(x_56, 1, x_7); +x_21 = x_56; +goto block_39; +} +} +else +{ +uint8_t x_57; lean_object* x_58; uint8_t x_59; +x_57 = 1; +x_58 = l_Std_Time_Month_Ordinal_days(x_57, x_3); +x_59 = lean_int_dec_lt(x_58, x_11); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_58); +lean_ctor_set(x_6, 1, x_3); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_6); +lean_ctor_set(x_60, 1, x_7); +x_21 = x_60; +goto block_39; +} +else +{ +lean_object* x_61; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_58); +lean_ctor_set(x_6, 1, x_3); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_6); +lean_ctor_set(x_61, 1, x_7); +x_21 = x_61; +goto block_39; +} +} +} +else +{ +uint8_t x_62; lean_object* x_63; uint8_t x_64; +x_62 = 1; +x_63 = l_Std_Time_Month_Ordinal_days(x_62, x_3); +x_64 = lean_int_dec_lt(x_63, x_11); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_63); +lean_ctor_set(x_6, 1, x_3); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_6); +lean_ctor_set(x_65, 1, x_7); +x_21 = x_65; +goto block_39; +} +else +{ +lean_object* x_66; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_63); +lean_ctor_set(x_6, 1, x_3); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_6); +lean_ctor_set(x_66, 1, x_7); +x_21 = x_66; +goto block_39; +} +} +} +block_39: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +x_23 = l_Std_Time_PlainTime_toSeconds(x_22); +x_24 = lean_int_add(x_23, x_18); +lean_dec(x_18); +lean_dec(x_23); +x_25 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_26 = lean_int_ediv(x_24, x_25); +lean_dec(x_24); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_22); +lean_dec(x_22); +x_28 = lean_int_add(x_27, x_20); +lean_dec(x_20); +lean_dec(x_27); +x_29 = l_Std_Time_PlainTime_ofNanoseconds(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +x_31 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_32 = lean_int_add(x_31, x_26); +lean_dec(x_26); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +if (lean_is_scalar(x_8)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_8; +} +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_29); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_36 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_36, 0, x_21); +x_37 = lean_mk_thunk(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_67 = lean_ctor_get(x_6, 0); +x_68 = lean_ctor_get(x_6, 2); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_70 = lean_int_mod(x_67, x_69); +x_71 = l_Std_Time_instInhabitedDateTime___closed__2; +x_72 = lean_int_dec_eq(x_70, x_71); +lean_dec(x_70); +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_int_neg(x_73); +x_75 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_74, x_75); +if (x_72 == 0) +{ +uint8_t x_96; lean_object* x_97; uint8_t x_98; +x_96 = 0; +x_97 = l_Std_Time_Month_Ordinal_days(x_96, x_3); +x_98 = lean_int_dec_lt(x_97, x_68); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; +lean_dec(x_97); +x_99 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_99, 0, x_67); +lean_ctor_set(x_99, 1, x_3); +lean_ctor_set(x_99, 2, x_68); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_7); +x_77 = x_100; +goto block_95; +} +else +{ +lean_object* x_101; lean_object* x_102; +lean_dec(x_68); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_67); +lean_ctor_set(x_101, 1, x_3); +lean_ctor_set(x_101, 2, x_97); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_7); +x_77 = x_102; +goto block_95; +} +} +else +{ +lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; +x_103 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_104 = lean_int_mod(x_67, x_103); +x_105 = lean_int_dec_eq(x_104, x_71); +lean_dec(x_104); +x_106 = l_instDecidableNot___rarg(x_105); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_108 = lean_int_mod(x_67, x_107); +x_109 = lean_int_dec_eq(x_108, x_71); +lean_dec(x_108); +if (x_109 == 0) +{ +uint8_t x_110; lean_object* x_111; uint8_t x_112; +x_110 = 0; +x_111 = l_Std_Time_Month_Ordinal_days(x_110, x_3); +x_112 = lean_int_dec_lt(x_111, x_68); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +lean_dec(x_111); +x_113 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_113, 0, x_67); +lean_ctor_set(x_113, 1, x_3); +lean_ctor_set(x_113, 2, x_68); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_7); +x_77 = x_114; +goto block_95; +} +else +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_68); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_67); +lean_ctor_set(x_115, 1, x_3); +lean_ctor_set(x_115, 2, x_111); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_7); +x_77 = x_116; +goto block_95; +} +} +else +{ +uint8_t x_117; lean_object* x_118; uint8_t x_119; +x_117 = 1; +x_118 = l_Std_Time_Month_Ordinal_days(x_117, x_3); +x_119 = lean_int_dec_lt(x_118, x_68); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_118); +x_120 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_120, 0, x_67); +lean_ctor_set(x_120, 1, x_3); +lean_ctor_set(x_120, 2, x_68); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_7); +x_77 = x_121; +goto block_95; +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_68); +x_122 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_122, 0, x_67); +lean_ctor_set(x_122, 1, x_3); +lean_ctor_set(x_122, 2, x_118); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_7); +x_77 = x_123; +goto block_95; +} +} +} +else +{ +uint8_t x_124; lean_object* x_125; uint8_t x_126; +x_124 = 1; +x_125 = l_Std_Time_Month_Ordinal_days(x_124, x_3); +x_126 = lean_int_dec_lt(x_125, x_68); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; +lean_dec(x_125); +x_127 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_127, 0, x_67); +lean_ctor_set(x_127, 1, x_3); +lean_ctor_set(x_127, 2, x_68); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_7); +x_77 = x_128; +goto block_95; +} +else +{ +lean_object* x_129; lean_object* x_130; +lean_dec(x_68); +x_129 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_129, 0, x_67); +lean_ctor_set(x_129, 1, x_3); +lean_ctor_set(x_129, 2, x_125); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_7); +x_77 = x_130; +goto block_95; +} +} +} +block_95: +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +x_79 = l_Std_Time_PlainTime_toSeconds(x_78); +x_80 = lean_int_add(x_79, x_74); +lean_dec(x_74); +lean_dec(x_79); +x_81 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_82 = lean_int_ediv(x_80, x_81); +lean_dec(x_80); +x_83 = l_Std_Time_PlainTime_toNanoseconds(x_78); +lean_dec(x_78); +x_84 = lean_int_add(x_83, x_76); +lean_dec(x_76); +lean_dec(x_83); +x_85 = l_Std_Time_PlainTime_ofNanoseconds(x_84); +lean_dec(x_84); +x_86 = lean_ctor_get(x_77, 0); +lean_inc(x_86); +x_87 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_86); +x_88 = lean_int_add(x_87, x_82); +lean_dec(x_82); +lean_dec(x_87); +x_89 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_88); +lean_dec(x_88); +if (lean_is_scalar(x_8)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_8; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_85); +x_91 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_90); +x_92 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_92, 0, x_77); +x_93 = lean_mk_thunk(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withMonthClip(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_dec(x_7); +x_11 = l_Std_Time_PlainDate_rollOver(x_9, x_3, x_10); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 2); +lean_inc(x_34); +lean_dec(x_31); +x_35 = l_Std_Time_PlainDate_rollOver(x_33, x_3, x_34); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMonthRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withMonthRollOver(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearClip(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +if (lean_is_exclusive(x_5)) { + lean_ctor_release(x_5, 0); + lean_ctor_release(x_5, 1); + x_8 = x_5; +} else { + lean_dec_ref(x_5); + x_8 = lean_box(0); +} +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_6, 1); +x_11 = lean_ctor_get(x_6, 2); +x_12 = lean_ctor_get(x_6, 0); +lean_dec(x_12); +x_13 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_14 = lean_int_mod(x_3, x_13); +x_15 = l_Std_Time_instInhabitedDateTime___closed__2; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_1, 0); +x_18 = lean_int_neg(x_17); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_18, x_19); +if (x_16 == 0) +{ +uint8_t x_40; lean_object* x_41; uint8_t x_42; +x_40 = 0; +x_41 = l_Std_Time_Month_Ordinal_days(x_40, x_10); +x_42 = lean_int_dec_lt(x_41, x_11); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_41); +lean_ctor_set(x_6, 0, x_3); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_6); +lean_ctor_set(x_43, 1, x_7); +x_21 = x_43; +goto block_39; +} +else +{ +lean_object* x_44; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_41); +lean_ctor_set(x_6, 0, x_3); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_6); +lean_ctor_set(x_44, 1, x_7); +x_21 = x_44; +goto block_39; +} +} +else +{ +lean_object* x_45; lean_object* x_46; uint8_t x_47; uint8_t x_48; +x_45 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_46 = lean_int_mod(x_3, x_45); +x_47 = lean_int_dec_eq(x_46, x_15); +lean_dec(x_46); +x_48 = l_instDecidableNot___rarg(x_47); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_49 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_50 = lean_int_mod(x_3, x_49); +x_51 = lean_int_dec_eq(x_50, x_15); +lean_dec(x_50); +if (x_51 == 0) +{ +uint8_t x_52; lean_object* x_53; uint8_t x_54; +x_52 = 0; +x_53 = l_Std_Time_Month_Ordinal_days(x_52, x_10); +x_54 = lean_int_dec_lt(x_53, x_11); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_53); +lean_ctor_set(x_6, 0, x_3); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_6); +lean_ctor_set(x_55, 1, x_7); +x_21 = x_55; +goto block_39; +} +else +{ +lean_object* x_56; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_53); +lean_ctor_set(x_6, 0, x_3); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_6); +lean_ctor_set(x_56, 1, x_7); +x_21 = x_56; +goto block_39; +} +} +else +{ +uint8_t x_57; lean_object* x_58; uint8_t x_59; +x_57 = 1; +x_58 = l_Std_Time_Month_Ordinal_days(x_57, x_10); +x_59 = lean_int_dec_lt(x_58, x_11); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_58); +lean_ctor_set(x_6, 0, x_3); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_6); +lean_ctor_set(x_60, 1, x_7); +x_21 = x_60; +goto block_39; +} +else +{ +lean_object* x_61; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_58); +lean_ctor_set(x_6, 0, x_3); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_6); +lean_ctor_set(x_61, 1, x_7); +x_21 = x_61; +goto block_39; +} +} +} +else +{ +uint8_t x_62; lean_object* x_63; uint8_t x_64; +x_62 = 1; +x_63 = l_Std_Time_Month_Ordinal_days(x_62, x_10); +x_64 = lean_int_dec_lt(x_63, x_11); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_63); +lean_ctor_set(x_6, 0, x_3); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_6); +lean_ctor_set(x_65, 1, x_7); +x_21 = x_65; +goto block_39; +} +else +{ +lean_object* x_66; +lean_dec(x_11); +lean_ctor_set(x_6, 2, x_63); +lean_ctor_set(x_6, 0, x_3); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_6); +lean_ctor_set(x_66, 1, x_7); +x_21 = x_66; +goto block_39; +} +} +} +block_39: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +x_23 = l_Std_Time_PlainTime_toSeconds(x_22); +x_24 = lean_int_add(x_23, x_18); +lean_dec(x_18); +lean_dec(x_23); +x_25 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_26 = lean_int_ediv(x_24, x_25); +lean_dec(x_24); +x_27 = l_Std_Time_PlainTime_toNanoseconds(x_22); +lean_dec(x_22); +x_28 = lean_int_add(x_27, x_20); +lean_dec(x_20); +lean_dec(x_27); +x_29 = l_Std_Time_PlainTime_ofNanoseconds(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_21, 0); +lean_inc(x_30); +x_31 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_32 = lean_int_add(x_31, x_26); +lean_dec(x_26); +lean_dec(x_31); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +if (lean_is_scalar(x_8)) { + x_34 = lean_alloc_ctor(0, 2, 0); +} else { + x_34 = x_8; +} +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_29); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_36 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_36, 0, x_21); +x_37 = lean_mk_thunk(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_35); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_67 = lean_ctor_get(x_6, 1); +x_68 = lean_ctor_get(x_6, 2); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_6); +x_69 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_70 = lean_int_mod(x_3, x_69); +x_71 = l_Std_Time_instInhabitedDateTime___closed__2; +x_72 = lean_int_dec_eq(x_70, x_71); +lean_dec(x_70); +x_73 = lean_ctor_get(x_1, 0); +x_74 = lean_int_neg(x_73); +x_75 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_74, x_75); +if (x_72 == 0) +{ +uint8_t x_96; lean_object* x_97; uint8_t x_98; +x_96 = 0; +x_97 = l_Std_Time_Month_Ordinal_days(x_96, x_67); +x_98 = lean_int_dec_lt(x_97, x_68); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; +lean_dec(x_97); +x_99 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_99, 0, x_3); +lean_ctor_set(x_99, 1, x_67); +lean_ctor_set(x_99, 2, x_68); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_7); +x_77 = x_100; +goto block_95; +} +else +{ +lean_object* x_101; lean_object* x_102; +lean_dec(x_68); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_3); +lean_ctor_set(x_101, 1, x_67); +lean_ctor_set(x_101, 2, x_97); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_7); +x_77 = x_102; +goto block_95; +} +} +else +{ +lean_object* x_103; lean_object* x_104; uint8_t x_105; uint8_t x_106; +x_103 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_104 = lean_int_mod(x_3, x_103); +x_105 = lean_int_dec_eq(x_104, x_71); +lean_dec(x_104); +x_106 = l_instDecidableNot___rarg(x_105); +if (x_106 == 0) +{ +lean_object* x_107; lean_object* x_108; uint8_t x_109; +x_107 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_108 = lean_int_mod(x_3, x_107); +x_109 = lean_int_dec_eq(x_108, x_71); +lean_dec(x_108); +if (x_109 == 0) +{ +uint8_t x_110; lean_object* x_111; uint8_t x_112; +x_110 = 0; +x_111 = l_Std_Time_Month_Ordinal_days(x_110, x_67); +x_112 = lean_int_dec_lt(x_111, x_68); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; +lean_dec(x_111); +x_113 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_113, 0, x_3); +lean_ctor_set(x_113, 1, x_67); +lean_ctor_set(x_113, 2, x_68); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_7); +x_77 = x_114; +goto block_95; +} +else +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_68); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_3); +lean_ctor_set(x_115, 1, x_67); +lean_ctor_set(x_115, 2, x_111); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_7); +x_77 = x_116; +goto block_95; +} +} +else +{ +uint8_t x_117; lean_object* x_118; uint8_t x_119; +x_117 = 1; +x_118 = l_Std_Time_Month_Ordinal_days(x_117, x_67); +x_119 = lean_int_dec_lt(x_118, x_68); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; +lean_dec(x_118); +x_120 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_120, 0, x_3); +lean_ctor_set(x_120, 1, x_67); +lean_ctor_set(x_120, 2, x_68); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_7); +x_77 = x_121; +goto block_95; +} +else +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_68); +x_122 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_122, 0, x_3); +lean_ctor_set(x_122, 1, x_67); +lean_ctor_set(x_122, 2, x_118); +x_123 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_7); +x_77 = x_123; +goto block_95; +} +} +} +else +{ +uint8_t x_124; lean_object* x_125; uint8_t x_126; +x_124 = 1; +x_125 = l_Std_Time_Month_Ordinal_days(x_124, x_67); +x_126 = lean_int_dec_lt(x_125, x_68); +if (x_126 == 0) +{ +lean_object* x_127; lean_object* x_128; +lean_dec(x_125); +x_127 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_127, 0, x_3); +lean_ctor_set(x_127, 1, x_67); +lean_ctor_set(x_127, 2, x_68); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_7); +x_77 = x_128; +goto block_95; +} +else +{ +lean_object* x_129; lean_object* x_130; +lean_dec(x_68); +x_129 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_129, 0, x_3); +lean_ctor_set(x_129, 1, x_67); +lean_ctor_set(x_129, 2, x_125); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_7); +x_77 = x_130; +goto block_95; +} +} +} +block_95: +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +x_79 = l_Std_Time_PlainTime_toSeconds(x_78); +x_80 = lean_int_add(x_79, x_74); +lean_dec(x_74); +lean_dec(x_79); +x_81 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_82 = lean_int_ediv(x_80, x_81); +lean_dec(x_80); +x_83 = l_Std_Time_PlainTime_toNanoseconds(x_78); +lean_dec(x_78); +x_84 = lean_int_add(x_83, x_76); +lean_dec(x_76); +lean_dec(x_83); +x_85 = l_Std_Time_PlainTime_ofNanoseconds(x_84); +lean_dec(x_84); +x_86 = lean_ctor_get(x_77, 0); +lean_inc(x_86); +x_87 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_86); +x_88 = lean_int_add(x_87, x_82); +lean_dec(x_82); +lean_dec(x_87); +x_89 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_88); +lean_dec(x_88); +if (lean_is_scalar(x_8)) { + x_90 = lean_alloc_ctor(0, 2, 0); +} else { + x_90 = x_8; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_85); +x_91 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_90); +x_92 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_92, 0, x_77); +x_93 = lean_mk_thunk(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_91); +lean_ctor_set(x_94, 1, x_93); +return x_94; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearClip___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withYearClip(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearRollOver(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_ctor_get(x_5, 1); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_dec(x_7); +x_11 = l_Std_Time_PlainDate_rollOver(x_3, x_9, x_10); +lean_dec(x_10); +lean_inc(x_8); +lean_inc(x_11); +lean_ctor_set(x_5, 0, x_11); +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_int_neg(x_12); +x_14 = l_Std_Time_PlainTime_toSeconds(x_8); +x_15 = lean_int_add(x_14, x_13); +lean_dec(x_14); +x_16 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_17 = lean_int_ediv(x_15, x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_19 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_20 = lean_int_mul(x_13, x_19); +lean_dec(x_13); +x_21 = lean_int_add(x_18, x_20); +lean_dec(x_20); +lean_dec(x_18); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_11); +x_24 = lean_int_add(x_23, x_17); +lean_dec(x_17); +lean_dec(x_23); +x_25 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_28 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_28, 0, x_5); +x_29 = lean_mk_thunk(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_5, 0); +x_32 = lean_ctor_get(x_5, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_5); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +x_34 = lean_ctor_get(x_31, 2); +lean_inc(x_34); +lean_dec(x_31); +x_35 = l_Std_Time_PlainDate_rollOver(x_3, x_33, x_34); +lean_dec(x_34); +lean_inc(x_32); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +x_37 = lean_ctor_get(x_1, 0); +x_38 = lean_int_neg(x_37); +x_39 = l_Std_Time_PlainTime_toSeconds(x_32); +x_40 = lean_int_add(x_39, x_38); +lean_dec(x_39); +x_41 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_42 = lean_int_ediv(x_40, x_41); +lean_dec(x_40); +x_43 = l_Std_Time_PlainTime_toNanoseconds(x_32); +lean_dec(x_32); +x_44 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_45 = lean_int_mul(x_38, x_44); +lean_dec(x_38); +x_46 = lean_int_add(x_43, x_45); +lean_dec(x_45); +lean_dec(x_43); +x_47 = l_Std_Time_PlainTime_ofNanoseconds(x_46); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_49 = lean_int_add(x_48, x_42); +lean_dec(x_42); +lean_dec(x_48); +x_50 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_49); +lean_dec(x_49); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_47); +x_52 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_51); +x_53 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_53, 0, x_36); +x_54 = lean_mk_thunk(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_54); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withYearRollOver___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withYearRollOver(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withHours(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_7, 0); +lean_dec(x_10); +lean_ctor_set(x_7, 0, x_3); +lean_inc(x_7); +lean_inc(x_9); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_7); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_7, 1); +x_32 = lean_ctor_get(x_7, 2); +x_33 = lean_ctor_get(x_7, 3); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_31); +lean_ctor_set(x_34, 2, x_32); +lean_ctor_set(x_34, 3, x_33); +lean_inc(x_34); +lean_inc(x_30); +lean_ctor_set(x_5, 1, x_34); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_34); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_34); +lean_dec(x_34); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_5); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_54 = lean_ctor_get(x_5, 1); +x_55 = lean_ctor_get(x_5, 0); +lean_inc(x_54); +lean_inc(x_55); +lean_dec(x_5); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 2); +lean_inc(x_57); +x_58 = lean_ctor_get(x_54, 3); +lean_inc(x_58); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + lean_ctor_release(x_54, 2); + lean_ctor_release(x_54, 3); + x_59 = x_54; +} else { + lean_dec_ref(x_54); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 4, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_3); +lean_ctor_set(x_60, 1, x_56); +lean_ctor_set(x_60, 2, x_57); +lean_ctor_set(x_60, 3, x_58); +lean_inc(x_60); +lean_inc(x_55); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_55); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_ctor_get(x_1, 0); +x_63 = lean_int_neg(x_62); +x_64 = l_Std_Time_PlainTime_toSeconds(x_60); +x_65 = lean_int_add(x_64, x_63); +lean_dec(x_64); +x_66 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_67 = lean_int_ediv(x_65, x_66); +lean_dec(x_65); +x_68 = l_Std_Time_PlainTime_toNanoseconds(x_60); +lean_dec(x_60); +x_69 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_70 = lean_int_mul(x_63, x_69); +lean_dec(x_63); +x_71 = lean_int_add(x_68, x_70); +lean_dec(x_70); +lean_dec(x_68); +x_72 = l_Std_Time_PlainTime_ofNanoseconds(x_71); +lean_dec(x_71); +x_73 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_55); +x_74 = lean_int_add(x_73, x_67); +lean_dec(x_67); +lean_dec(x_73); +x_75 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_74); +lean_dec(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_72); +x_77 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_76); +x_78 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_78, 0, x_61); +x_79 = lean_mk_thunk(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withHours___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withHours(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMinutes(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_7, 1); +lean_dec(x_10); +lean_ctor_set(x_7, 1, x_3); +lean_inc(x_7); +lean_inc(x_9); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_7); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 2); +x_33 = lean_ctor_get(x_7, 3); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_3); +lean_ctor_set(x_34, 2, x_32); +lean_ctor_set(x_34, 3, x_33); +lean_inc(x_34); +lean_inc(x_30); +lean_ctor_set(x_5, 1, x_34); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_34); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_34); +lean_dec(x_34); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_5); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_54 = lean_ctor_get(x_5, 1); +x_55 = lean_ctor_get(x_5, 0); +lean_inc(x_54); +lean_inc(x_55); +lean_dec(x_5); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 2); +lean_inc(x_57); +x_58 = lean_ctor_get(x_54, 3); +lean_inc(x_58); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + lean_ctor_release(x_54, 2); + lean_ctor_release(x_54, 3); + x_59 = x_54; +} else { + lean_dec_ref(x_54); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 4, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_56); +lean_ctor_set(x_60, 1, x_3); +lean_ctor_set(x_60, 2, x_57); +lean_ctor_set(x_60, 3, x_58); +lean_inc(x_60); +lean_inc(x_55); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_55); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_ctor_get(x_1, 0); +x_63 = lean_int_neg(x_62); +x_64 = l_Std_Time_PlainTime_toSeconds(x_60); +x_65 = lean_int_add(x_64, x_63); +lean_dec(x_64); +x_66 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_67 = lean_int_ediv(x_65, x_66); +lean_dec(x_65); +x_68 = l_Std_Time_PlainTime_toNanoseconds(x_60); +lean_dec(x_60); +x_69 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_70 = lean_int_mul(x_63, x_69); +lean_dec(x_63); +x_71 = lean_int_add(x_68, x_70); +lean_dec(x_70); +lean_dec(x_68); +x_72 = l_Std_Time_PlainTime_ofNanoseconds(x_71); +lean_dec(x_71); +x_73 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_55); +x_74 = lean_int_add(x_73, x_67); +lean_dec(x_67); +lean_dec(x_73); +x_75 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_74); +lean_dec(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_72); +x_77 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_76); +x_78 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_78, 0, x_61); +x_79 = lean_mk_thunk(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMinutes___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withMinutes(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withSeconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_7, 2); +lean_dec(x_10); +lean_ctor_set(x_7, 2, x_3); +lean_inc(x_7); +lean_inc(x_9); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_7); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 1); +x_33 = lean_ctor_get(x_7, 3); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_3); +lean_ctor_set(x_34, 3, x_33); +lean_inc(x_34); +lean_inc(x_30); +lean_ctor_set(x_5, 1, x_34); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_34); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_34); +lean_dec(x_34); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_5); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_54 = lean_ctor_get(x_5, 1); +x_55 = lean_ctor_get(x_5, 0); +lean_inc(x_54); +lean_inc(x_55); +lean_dec(x_5); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_54, 3); +lean_inc(x_58); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + lean_ctor_release(x_54, 2); + lean_ctor_release(x_54, 3); + x_59 = x_54; +} else { + lean_dec_ref(x_54); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 4, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_56); +lean_ctor_set(x_60, 1, x_57); +lean_ctor_set(x_60, 2, x_3); +lean_ctor_set(x_60, 3, x_58); +lean_inc(x_60); +lean_inc(x_55); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_55); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_ctor_get(x_1, 0); +x_63 = lean_int_neg(x_62); +x_64 = l_Std_Time_PlainTime_toSeconds(x_60); +x_65 = lean_int_add(x_64, x_63); +lean_dec(x_64); +x_66 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_67 = lean_int_ediv(x_65, x_66); +lean_dec(x_65); +x_68 = l_Std_Time_PlainTime_toNanoseconds(x_60); +lean_dec(x_60); +x_69 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_70 = lean_int_mul(x_63, x_69); +lean_dec(x_63); +x_71 = lean_int_add(x_68, x_70); +lean_dec(x_70); +lean_dec(x_68); +x_72 = l_Std_Time_PlainTime_ofNanoseconds(x_71); +lean_dec(x_71); +x_73 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_55); +x_74 = lean_int_add(x_73, x_67); +lean_dec(x_67); +lean_dec(x_73); +x_75 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_74); +lean_dec(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_72); +x_77 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_76); +x_78 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_78, 0, x_61); +x_79 = lean_mk_thunk(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withSeconds(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withNanoseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_7, 3); +lean_dec(x_10); +lean_ctor_set(x_7, 3, x_3); +lean_inc(x_7); +lean_inc(x_9); +x_11 = lean_ctor_get(x_1, 0); +x_12 = lean_int_neg(x_11); +x_13 = l_Std_Time_PlainTime_toSeconds(x_7); +x_14 = lean_int_add(x_13, x_12); +lean_dec(x_13); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_16 = lean_int_ediv(x_14, x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_18 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_12, x_18); +lean_dec(x_12); +x_20 = lean_int_add(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_23 = lean_int_add(x_22, x_16); +lean_dec(x_16); +lean_dec(x_22); +x_24 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_23); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +x_26 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_27, 0, x_5); +x_28 = lean_mk_thunk(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_26); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_30 = lean_ctor_get(x_5, 0); +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 1); +x_33 = lean_ctor_get(x_7, 2); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_7); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_33); +lean_ctor_set(x_34, 3, x_3); +lean_inc(x_34); +lean_inc(x_30); +lean_ctor_set(x_5, 1, x_34); +x_35 = lean_ctor_get(x_1, 0); +x_36 = lean_int_neg(x_35); +x_37 = l_Std_Time_PlainTime_toSeconds(x_34); +x_38 = lean_int_add(x_37, x_36); +lean_dec(x_37); +x_39 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_40 = lean_int_ediv(x_38, x_39); +lean_dec(x_38); +x_41 = l_Std_Time_PlainTime_toNanoseconds(x_34); +lean_dec(x_34); +x_42 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_43 = lean_int_mul(x_36, x_42); +lean_dec(x_36); +x_44 = lean_int_add(x_41, x_43); +lean_dec(x_43); +lean_dec(x_41); +x_45 = l_Std_Time_PlainTime_ofNanoseconds(x_44); +lean_dec(x_44); +x_46 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_30); +x_47 = lean_int_add(x_46, x_40); +lean_dec(x_40); +lean_dec(x_46); +x_48 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_47); +lean_dec(x_47); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_45); +x_50 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_49); +x_51 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_51, 0, x_5); +x_52 = lean_mk_thunk(x_51); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +return x_53; +} +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_54 = lean_ctor_get(x_5, 1); +x_55 = lean_ctor_get(x_5, 0); +lean_inc(x_54); +lean_inc(x_55); +lean_dec(x_5); +x_56 = lean_ctor_get(x_54, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_54, 2); +lean_inc(x_58); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + lean_ctor_release(x_54, 2); + lean_ctor_release(x_54, 3); + x_59 = x_54; +} else { + lean_dec_ref(x_54); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(0, 4, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_56); +lean_ctor_set(x_60, 1, x_57); +lean_ctor_set(x_60, 2, x_58); +lean_ctor_set(x_60, 3, x_3); +lean_inc(x_60); +lean_inc(x_55); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_55); +lean_ctor_set(x_61, 1, x_60); +x_62 = lean_ctor_get(x_1, 0); +x_63 = lean_int_neg(x_62); +x_64 = l_Std_Time_PlainTime_toSeconds(x_60); +x_65 = lean_int_add(x_64, x_63); +lean_dec(x_64); +x_66 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_67 = lean_int_ediv(x_65, x_66); +lean_dec(x_65); +x_68 = l_Std_Time_PlainTime_toNanoseconds(x_60); +lean_dec(x_60); +x_69 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_70 = lean_int_mul(x_63, x_69); +lean_dec(x_63); +x_71 = lean_int_add(x_68, x_70); +lean_dec(x_70); +lean_dec(x_68); +x_72 = l_Std_Time_PlainTime_ofNanoseconds(x_71); +lean_dec(x_71); +x_73 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_55); +x_74 = lean_int_add(x_73, x_67); +lean_dec(x_67); +lean_dec(x_73); +x_75 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_74); +lean_dec(x_74); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_72); +x_77 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_76); +x_78 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_78, 0, x_61); +x_79 = lean_mk_thunk(x_78); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_77); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withNanoseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withNanoseconds(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_DateTime_withMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_DateTime_withMilliseconds___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMilliseconds(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_ctor_get(x_5, 1); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_7, 3); +x_11 = l_Std_Time_DateTime_withMilliseconds___closed__1; +x_12 = lean_int_emod(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_DateTime_withMilliseconds___closed__2; +x_14 = lean_int_mul(x_3, x_13); +x_15 = lean_int_add(x_14, x_12); +lean_dec(x_12); +lean_dec(x_14); +lean_ctor_set(x_7, 3, x_15); +lean_inc(x_7); +lean_inc(x_9); +x_16 = lean_ctor_get(x_1, 0); +x_17 = lean_int_neg(x_16); +x_18 = l_Std_Time_PlainTime_toSeconds(x_7); +x_19 = lean_int_add(x_18, x_17); +lean_dec(x_18); +x_20 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_21 = lean_int_ediv(x_19, x_20); +lean_dec(x_19); +x_22 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_23 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_24 = lean_int_mul(x_17, x_23); +lean_dec(x_17); +x_25 = lean_int_add(x_22, x_24); +lean_dec(x_24); +lean_dec(x_22); +x_26 = l_Std_Time_PlainTime_ofNanoseconds(x_25); +lean_dec(x_25); +x_27 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_9); +x_28 = lean_int_add(x_27, x_21); +lean_dec(x_21); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_26); +x_31 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_30); +x_32 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_32, 0, x_5); +x_33 = lean_mk_thunk(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_35 = lean_ctor_get(x_5, 0); +x_36 = lean_ctor_get(x_7, 0); +x_37 = lean_ctor_get(x_7, 1); +x_38 = lean_ctor_get(x_7, 2); +x_39 = lean_ctor_get(x_7, 3); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_7); +x_40 = l_Std_Time_DateTime_withMilliseconds___closed__1; +x_41 = lean_int_emod(x_39, x_40); +lean_dec(x_39); +x_42 = l_Std_Time_DateTime_withMilliseconds___closed__2; +x_43 = lean_int_mul(x_3, x_42); +x_44 = lean_int_add(x_43, x_41); +lean_dec(x_41); +lean_dec(x_43); +x_45 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_45, 0, x_36); +lean_ctor_set(x_45, 1, x_37); +lean_ctor_set(x_45, 2, x_38); +lean_ctor_set(x_45, 3, x_44); +lean_inc(x_45); +lean_inc(x_35); +lean_ctor_set(x_5, 1, x_45); +x_46 = lean_ctor_get(x_1, 0); +x_47 = lean_int_neg(x_46); +x_48 = l_Std_Time_PlainTime_toSeconds(x_45); +x_49 = lean_int_add(x_48, x_47); +lean_dec(x_48); +x_50 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_51 = lean_int_ediv(x_49, x_50); +lean_dec(x_49); +x_52 = l_Std_Time_PlainTime_toNanoseconds(x_45); +lean_dec(x_45); +x_53 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_54 = lean_int_mul(x_47, x_53); +lean_dec(x_47); +x_55 = lean_int_add(x_52, x_54); +lean_dec(x_54); +lean_dec(x_52); +x_56 = l_Std_Time_PlainTime_ofNanoseconds(x_55); +lean_dec(x_55); +x_57 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_35); +x_58 = lean_int_add(x_57, x_51); +lean_dec(x_51); +lean_dec(x_57); +x_59 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_58); +lean_dec(x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_56); +x_61 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_60); +x_62 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_62, 0, x_5); +x_63 = lean_mk_thunk(x_62); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_65 = lean_ctor_get(x_5, 1); +x_66 = lean_ctor_get(x_5, 0); +lean_inc(x_65); +lean_inc(x_66); +lean_dec(x_5); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +x_69 = lean_ctor_get(x_65, 2); +lean_inc(x_69); +x_70 = lean_ctor_get(x_65, 3); +lean_inc(x_70); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + lean_ctor_release(x_65, 2); + lean_ctor_release(x_65, 3); + x_71 = x_65; +} else { + lean_dec_ref(x_65); + x_71 = lean_box(0); +} +x_72 = l_Std_Time_DateTime_withMilliseconds___closed__1; +x_73 = lean_int_emod(x_70, x_72); +lean_dec(x_70); +x_74 = l_Std_Time_DateTime_withMilliseconds___closed__2; +x_75 = lean_int_mul(x_3, x_74); +x_76 = lean_int_add(x_75, x_73); +lean_dec(x_73); +lean_dec(x_75); +if (lean_is_scalar(x_71)) { + x_77 = lean_alloc_ctor(0, 4, 0); +} else { + x_77 = x_71; +} +lean_ctor_set(x_77, 0, x_67); +lean_ctor_set(x_77, 1, x_68); +lean_ctor_set(x_77, 2, x_69); +lean_ctor_set(x_77, 3, x_76); +lean_inc(x_77); +lean_inc(x_66); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_66); +lean_ctor_set(x_78, 1, x_77); +x_79 = lean_ctor_get(x_1, 0); +x_80 = lean_int_neg(x_79); +x_81 = l_Std_Time_PlainTime_toSeconds(x_77); +x_82 = lean_int_add(x_81, x_80); +lean_dec(x_81); +x_83 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_84 = lean_int_ediv(x_82, x_83); +lean_dec(x_82); +x_85 = l_Std_Time_PlainTime_toNanoseconds(x_77); +lean_dec(x_77); +x_86 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_87 = lean_int_mul(x_80, x_86); +lean_dec(x_80); +x_88 = lean_int_add(x_85, x_87); +lean_dec(x_87); +lean_dec(x_85); +x_89 = l_Std_Time_PlainTime_ofNanoseconds(x_88); +lean_dec(x_88); +x_90 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_66); +x_91 = lean_int_add(x_90, x_84); +lean_dec(x_84); +lean_dec(x_90); +x_92 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_91); +lean_dec(x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_89); +x_94 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_93); +x_95 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_95, 0, x_78); +x_96 = lean_mk_thunk(x_95); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withMilliseconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_withMilliseconds(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_toPlainDateTime___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainDateTime___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_toPlainDateTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_toPlainDateTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_year___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_year___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_year___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_year(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_month___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_month___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_month___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_month(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_day___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_day___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_day___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_day(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_hour___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_hour___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_hour___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_hour(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_minute___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_minute___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_minute___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_minute(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_second___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_second___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_second___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_second(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 3); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_DateTime_withMilliseconds___closed__1; +x_7 = lean_int_emod(x_5, x_6); +lean_dec(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_millisecond___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_millisecond___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_millisecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_millisecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 3); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_nanosecond___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_nanosecond___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_nanosecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_nanosecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_DateTime_weekday___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_weekday(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_weekday___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday___rarg___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_DateTime_weekday___rarg(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekday___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_weekday(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_DateTime_era___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_Year_Offset_era(x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_era___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era___rarg___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_DateTime_era___rarg(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_era___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_era(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withWeekday(lean_object* x_1, lean_object* x_2, uint8_t x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_thunk_get_own(x_4); +lean_inc(x_5); +x_6 = l_Std_Time_PlainDateTime_withWeekday(x_5, x_3); +x_7 = lean_ctor_get(x_1, 0); +x_8 = lean_int_neg(x_7); +x_9 = lean_ctor_get(x_5, 1); +lean_inc(x_9); +lean_dec(x_5); +x_10 = l_Std_Time_PlainTime_toSeconds(x_9); +x_11 = lean_int_add(x_10, x_8); +lean_dec(x_10); +x_12 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_15 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_8, x_15); +lean_dec(x_8); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_23); +x_25 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_25, 0, x_6); +x_26 = lean_mk_thunk(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_withWeekday___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l_Std_Time_DateTime_withWeekday(x_1, x_2, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Std_Time_DateTime_inLeapYear___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_7 = lean_int_mod(x_5, x_6); +x_8 = l_Std_Time_instInhabitedDateTime___closed__2; +x_9 = lean_int_dec_eq(x_7, x_8); +lean_dec(x_7); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_5); +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; +x_11 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_12 = lean_int_mod(x_5, x_11); +x_13 = lean_int_dec_eq(x_12, x_8); +lean_dec(x_12); +x_14 = l_instDecidableNot___rarg(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_16 = lean_int_mod(x_5, x_15); +lean_dec(x_5); +x_17 = lean_int_dec_eq(x_16, x_8); +lean_dec(x_16); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_5); +x_18 = 1; +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_inLeapYear___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear___rarg___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_DateTime_inLeapYear___rarg(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_inLeapYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_inLeapYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = l_Std_Time_DateTime_withDaysClip___closed__1; +x_7 = lean_int_mod(x_5, x_6); +x_8 = l_Std_Time_instInhabitedDateTime___closed__2; +x_9 = lean_int_dec_eq(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_4, 2); +lean_inc(x_11); +lean_dec(x_4); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +if (x_9 == 0) +{ +uint8_t x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = 0; +x_14 = l_Std_Time_ValidDate_dayOfYear(x_13, x_12); +lean_dec(x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; +x_15 = l_Std_Time_DateTime_withDaysClip___closed__2; +x_16 = lean_int_mod(x_5, x_15); +x_17 = lean_int_dec_eq(x_16, x_8); +lean_dec(x_16); +x_18 = l_instDecidableNot___rarg(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Std_Time_DateTime_withDaysClip___closed__3; +x_20 = lean_int_mod(x_5, x_19); +lean_dec(x_5); +x_21 = lean_int_dec_eq(x_20, x_8); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; lean_object* x_23; +x_22 = 0; +x_23 = l_Std_Time_ValidDate_dayOfYear(x_22, x_12); +lean_dec(x_12); +return x_23; +} +else +{ +uint8_t x_24; lean_object* x_25; +x_24 = 1; +x_25 = l_Std_Time_ValidDate_dayOfYear(x_24, x_12); +lean_dec(x_12); +return x_25; +} +} +else +{ +uint8_t x_26; lean_object* x_27; +lean_dec(x_5); +x_26 = 1; +x_27 = l_Std_Time_ValidDate_dayOfYear(x_26, x_12); +lean_dec(x_12); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_dayOfYear___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_dayOfYear___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_dayOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_dayOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_weekOfYear(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_weekOfYear___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_weekOfYear___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_weekOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = l_Std_Time_PlainDateTime_weekOfMonth(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_weekOfMonth___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_weekOfMonth___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_weekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_weekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_alignedWeekOfMonth___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_alignedWeekOfMonth___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_alignedWeekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_alignedWeekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_quarter(x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_quarter___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_quarter___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_quarter___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_quarter(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_time___rarg___boxed), 1, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___rarg___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_time___rarg(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_time___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_DateTime_time(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofDaysSinceUNIXEpoch(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_4 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_1); +lean_inc(x_2); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_int_neg(x_6); +x_8 = l_Std_Time_PlainTime_toSeconds(x_2); +x_9 = lean_int_add(x_8, x_7); +lean_dec(x_8); +x_10 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_2); +lean_dec(x_2); +x_13 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_14 = lean_int_mul(x_7, x_13); +lean_dec(x_7); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_PlainTime_ofNanoseconds(x_15); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +x_18 = lean_int_add(x_17, x_11); +lean_dec(x_11); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_16); +x_21 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_20); +x_22 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_22, 0, x_5); +x_23 = lean_mk_thunk(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_21); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_ofDaysSinceUNIXEpoch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_ofDaysSinceUNIXEpoch(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addDays___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subDays___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addWeeks___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subWeeks___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addHours___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subHours___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addMinutes___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subMinutes___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addSeconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subSeconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__5(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addMilliseconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__5(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subMilliseconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddOffset__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_addNanoseconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubOffset__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_DateTime_subNanoseconds___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_4, 0); +x_6 = lean_int_neg(x_5); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_int_neg(x_7); +x_9 = lean_ctor_get(x_3, 0); +x_10 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_11 = lean_int_mul(x_9, x_10); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_mul(x_6, x_10); +lean_dec(x_6); +x_15 = lean_int_add(x_14, x_8); +lean_dec(x_8); +lean_dec(x_14); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_Duration_ofNanoseconds(x_16); +lean_dec(x_16); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Std_Time_DateTime_instHSubDuration___rarg___boxed), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_instHSubDuration___rarg(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_DateTime_instHSubDuration(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddDuration(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_4 = lean_ctor_get(x_3, 0); +x_5 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_thunk_get_own(x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_ctor_get(x_13, 3); +lean_inc(x_14); +x_15 = lean_int_add(x_14, x_8); +lean_dec(x_8); +lean_dec(x_14); +x_16 = lean_int_ediv(x_15, x_5); +x_17 = lean_int_emod(x_15, x_5); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_toNanoseconds(x_13); +lean_dec(x_13); +x_19 = lean_int_mul(x_16, x_5); +lean_dec(x_16); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +x_21 = l_Std_Time_PlainTime_ofNanoseconds(x_20); +lean_dec(x_20); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_23 = lean_ctor_get(x_21, 3); +lean_dec(x_23); +lean_ctor_set(x_21, 3, x_17); +lean_inc(x_21); +lean_inc(x_12); +lean_ctor_set(x_10, 1, x_21); +x_24 = lean_ctor_get(x_1, 0); +x_25 = lean_int_neg(x_24); +x_26 = l_Std_Time_PlainTime_toSeconds(x_21); +x_27 = lean_int_add(x_26, x_25); +lean_dec(x_26); +x_28 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_29 = lean_int_ediv(x_27, x_28); +lean_dec(x_27); +x_30 = l_Std_Time_PlainTime_toNanoseconds(x_21); +lean_dec(x_21); +x_31 = lean_int_mul(x_25, x_5); +lean_dec(x_25); +x_32 = lean_int_add(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +x_33 = l_Std_Time_PlainTime_ofNanoseconds(x_32); +lean_dec(x_32); +x_34 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_35 = lean_int_add(x_34, x_29); +lean_dec(x_29); +lean_dec(x_34); +x_36 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_35); +lean_dec(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_33); +x_38 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_37); +x_39 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_39, 0, x_10); +x_40 = lean_mk_thunk(x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +return x_41; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_42 = lean_ctor_get(x_21, 0); +x_43 = lean_ctor_get(x_21, 1); +x_44 = lean_ctor_get(x_21, 2); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_21); +x_45 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_44); +lean_ctor_set(x_45, 3, x_17); +lean_inc(x_45); +lean_inc(x_12); +lean_ctor_set(x_10, 1, x_45); +x_46 = lean_ctor_get(x_1, 0); +x_47 = lean_int_neg(x_46); +x_48 = l_Std_Time_PlainTime_toSeconds(x_45); +x_49 = lean_int_add(x_48, x_47); +lean_dec(x_48); +x_50 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_51 = lean_int_ediv(x_49, x_50); +lean_dec(x_49); +x_52 = l_Std_Time_PlainTime_toNanoseconds(x_45); +lean_dec(x_45); +x_53 = lean_int_mul(x_47, x_5); +lean_dec(x_47); +x_54 = lean_int_add(x_52, x_53); +lean_dec(x_53); +lean_dec(x_52); +x_55 = l_Std_Time_PlainTime_ofNanoseconds(x_54); +lean_dec(x_54); +x_56 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_12); +x_57 = lean_int_add(x_56, x_51); +lean_dec(x_51); +lean_dec(x_56); +x_58 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_57); +lean_dec(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_55); +x_60 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_59); +x_61 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_61, 0, x_10); +x_62 = lean_mk_thunk(x_61); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_62); +return x_63; +} +} +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +x_64 = lean_ctor_get(x_10, 0); +x_65 = lean_ctor_get(x_10, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_10); +x_66 = lean_ctor_get(x_65, 3); +lean_inc(x_66); +x_67 = lean_int_add(x_66, x_8); +lean_dec(x_8); +lean_dec(x_66); +x_68 = lean_int_ediv(x_67, x_5); +x_69 = lean_int_emod(x_67, x_5); +lean_dec(x_67); +x_70 = l_Std_Time_PlainTime_toNanoseconds(x_65); +lean_dec(x_65); +x_71 = lean_int_mul(x_68, x_5); +lean_dec(x_68); +x_72 = lean_int_add(x_70, x_71); +lean_dec(x_71); +lean_dec(x_70); +x_73 = l_Std_Time_PlainTime_ofNanoseconds(x_72); +lean_dec(x_72); +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_73, 2); +lean_inc(x_76); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + lean_ctor_release(x_73, 2); + lean_ctor_release(x_73, 3); + x_77 = x_73; +} else { + lean_dec_ref(x_73); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(0, 4, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_74); +lean_ctor_set(x_78, 1, x_75); +lean_ctor_set(x_78, 2, x_76); +lean_ctor_set(x_78, 3, x_69); +lean_inc(x_78); +lean_inc(x_64); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_64); +lean_ctor_set(x_79, 1, x_78); +x_80 = lean_ctor_get(x_1, 0); +x_81 = lean_int_neg(x_80); +x_82 = l_Std_Time_PlainTime_toSeconds(x_78); +x_83 = lean_int_add(x_82, x_81); +lean_dec(x_82); +x_84 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_85 = lean_int_ediv(x_83, x_84); +lean_dec(x_83); +x_86 = l_Std_Time_PlainTime_toNanoseconds(x_78); +lean_dec(x_78); +x_87 = lean_int_mul(x_81, x_5); +lean_dec(x_81); +x_88 = lean_int_add(x_86, x_87); +lean_dec(x_87); +lean_dec(x_86); +x_89 = l_Std_Time_PlainTime_ofNanoseconds(x_88); +lean_dec(x_88); +x_90 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_64); +x_91 = lean_int_add(x_90, x_85); +lean_dec(x_85); +lean_dec(x_90); +x_92 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_91); +lean_dec(x_91); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_89); +x_94 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_93); +x_95 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_95, 0, x_79); +x_96 = lean_mk_thunk(x_95); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_94); +lean_ctor_set(x_97, 1, x_96); +return x_97; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHAddDuration___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_instHAddDuration(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_4 = lean_ctor_get(x_3, 0); +x_5 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2; +x_6 = lean_int_mul(x_4, x_5); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_int_add(x_6, x_7); +lean_dec(x_6); +x_9 = lean_ctor_get(x_2, 1); +x_10 = lean_thunk_get_own(x_9); +x_11 = lean_int_neg(x_8); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_13 = lean_ctor_get(x_10, 0); +x_14 = lean_ctor_get(x_10, 1); +x_15 = lean_ctor_get(x_14, 3); +lean_inc(x_15); +x_16 = lean_int_add(x_15, x_11); +lean_dec(x_11); +lean_dec(x_15); +x_17 = lean_int_ediv(x_16, x_5); +x_18 = lean_int_emod(x_16, x_5); +lean_dec(x_16); +x_19 = l_Std_Time_PlainTime_toNanoseconds(x_14); +lean_dec(x_14); +x_20 = lean_int_mul(x_17, x_5); +lean_dec(x_17); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +x_22 = l_Std_Time_PlainTime_ofNanoseconds(x_21); +lean_dec(x_21); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_24 = lean_ctor_get(x_22, 3); +lean_dec(x_24); +lean_ctor_set(x_22, 3, x_18); +lean_inc(x_22); +lean_inc(x_13); +lean_ctor_set(x_10, 1, x_22); +x_25 = lean_ctor_get(x_1, 0); +x_26 = lean_int_neg(x_25); +x_27 = l_Std_Time_PlainTime_toSeconds(x_22); +x_28 = lean_int_add(x_27, x_26); +lean_dec(x_27); +x_29 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_30 = lean_int_ediv(x_28, x_29); +lean_dec(x_28); +x_31 = l_Std_Time_PlainTime_toNanoseconds(x_22); +lean_dec(x_22); +x_32 = lean_int_mul(x_26, x_5); +lean_dec(x_26); +x_33 = lean_int_add(x_31, x_32); +lean_dec(x_32); +lean_dec(x_31); +x_34 = l_Std_Time_PlainTime_ofNanoseconds(x_33); +lean_dec(x_33); +x_35 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_36 = lean_int_add(x_35, x_30); +lean_dec(x_30); +lean_dec(x_35); +x_37 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_36); +lean_dec(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_34); +x_39 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_38); +x_40 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_40, 0, x_10); +x_41 = lean_mk_thunk(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_39); +lean_ctor_set(x_42, 1, x_41); +return x_42; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_43 = lean_ctor_get(x_22, 0); +x_44 = lean_ctor_get(x_22, 1); +x_45 = lean_ctor_get(x_22, 2); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_22); +x_46 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_46, 0, x_43); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set(x_46, 2, x_45); +lean_ctor_set(x_46, 3, x_18); +lean_inc(x_46); +lean_inc(x_13); +lean_ctor_set(x_10, 1, x_46); +x_47 = lean_ctor_get(x_1, 0); +x_48 = lean_int_neg(x_47); +x_49 = l_Std_Time_PlainTime_toSeconds(x_46); +x_50 = lean_int_add(x_49, x_48); +lean_dec(x_49); +x_51 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_52 = lean_int_ediv(x_50, x_51); +lean_dec(x_50); +x_53 = l_Std_Time_PlainTime_toNanoseconds(x_46); +lean_dec(x_46); +x_54 = lean_int_mul(x_48, x_5); +lean_dec(x_48); +x_55 = lean_int_add(x_53, x_54); +lean_dec(x_54); +lean_dec(x_53); +x_56 = l_Std_Time_PlainTime_ofNanoseconds(x_55); +lean_dec(x_55); +x_57 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_58 = lean_int_add(x_57, x_52); +lean_dec(x_52); +lean_dec(x_57); +x_59 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_58); +lean_dec(x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_56); +x_61 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_60); +x_62 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_62, 0, x_10); +x_63 = lean_mk_thunk(x_62); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_61); +lean_ctor_set(x_64, 1, x_63); +return x_64; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_65 = lean_ctor_get(x_10, 0); +x_66 = lean_ctor_get(x_10, 1); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_10); +x_67 = lean_ctor_get(x_66, 3); +lean_inc(x_67); +x_68 = lean_int_add(x_67, x_11); +lean_dec(x_11); +lean_dec(x_67); +x_69 = lean_int_ediv(x_68, x_5); +x_70 = lean_int_emod(x_68, x_5); +lean_dec(x_68); +x_71 = l_Std_Time_PlainTime_toNanoseconds(x_66); +lean_dec(x_66); +x_72 = lean_int_mul(x_69, x_5); +lean_dec(x_69); +x_73 = lean_int_add(x_71, x_72); +lean_dec(x_72); +lean_dec(x_71); +x_74 = l_Std_Time_PlainTime_ofNanoseconds(x_73); +lean_dec(x_73); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +x_77 = lean_ctor_get(x_74, 2); +lean_inc(x_77); +if (lean_is_exclusive(x_74)) { + lean_ctor_release(x_74, 0); + lean_ctor_release(x_74, 1); + lean_ctor_release(x_74, 2); + lean_ctor_release(x_74, 3); + x_78 = x_74; +} else { + lean_dec_ref(x_74); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(0, 4, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_75); +lean_ctor_set(x_79, 1, x_76); +lean_ctor_set(x_79, 2, x_77); +lean_ctor_set(x_79, 3, x_70); +lean_inc(x_79); +lean_inc(x_65); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_65); +lean_ctor_set(x_80, 1, x_79); +x_81 = lean_ctor_get(x_1, 0); +x_82 = lean_int_neg(x_81); +x_83 = l_Std_Time_PlainTime_toSeconds(x_79); +x_84 = lean_int_add(x_83, x_82); +lean_dec(x_83); +x_85 = l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1; +x_86 = lean_int_ediv(x_84, x_85); +lean_dec(x_84); +x_87 = l_Std_Time_PlainTime_toNanoseconds(x_79); +lean_dec(x_79); +x_88 = lean_int_mul(x_82, x_5); +lean_dec(x_82); +x_89 = lean_int_add(x_87, x_88); +lean_dec(x_88); +lean_dec(x_87); +x_90 = l_Std_Time_PlainTime_ofNanoseconds(x_89); +lean_dec(x_89); +x_91 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_65); +x_92 = lean_int_add(x_91, x_86); +lean_dec(x_86); +lean_dec(x_91); +x_93 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_92); +lean_dec(x_92); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_90); +x_95 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_94); +x_96 = lean_alloc_closure((void*)(l_Std_Time_DateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_96, 0, x_80); +x_97 = lean_mk_thunk(x_96); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_97); +return x_98; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_DateTime_instHSubDuration__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_DateTime_instHSubDuration__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_TimeZone(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Internal(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_DateTime(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_TimeZone(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Internal(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_instInhabitedDateTime___lambda__1___closed__1 = _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___lambda__1___closed__1); +l_Std_Time_instInhabitedDateTime___lambda__1___closed__2 = _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___lambda__1___closed__2); +l_Std_Time_instInhabitedDateTime___lambda__1___closed__3 = _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___lambda__1___closed__3); +l_Std_Time_instInhabitedDateTime___lambda__1___closed__4 = _init_l_Std_Time_instInhabitedDateTime___lambda__1___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___lambda__1___closed__4); +l_Std_Time_instInhabitedDateTime___closed__1 = _init_l_Std_Time_instInhabitedDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___closed__1); +l_Std_Time_instInhabitedDateTime___closed__2 = _init_l_Std_Time_instInhabitedDateTime___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___closed__2); +l_Std_Time_instInhabitedDateTime___closed__3 = _init_l_Std_Time_instInhabitedDateTime___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedDateTime___closed__3); +l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1 = _init_l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__1); +l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2 = _init_l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_DateTime_ofTimestamp___lambda__1___closed__2); +l_Std_Time_DateTime_addHours___closed__1 = _init_l_Std_Time_DateTime_addHours___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_addHours___closed__1); +l_Std_Time_DateTime_addMinutes___closed__1 = _init_l_Std_Time_DateTime_addMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_addMinutes___closed__1); +l_Std_Time_DateTime_addMilliseconds___closed__1 = _init_l_Std_Time_DateTime_addMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_addMilliseconds___closed__1); +l_Std_Time_DateTime_addWeeks___closed__1 = _init_l_Std_Time_DateTime_addWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_addWeeks___closed__1); +l_Std_Time_DateTime_addYearsRollOver___closed__1 = _init_l_Std_Time_DateTime_addYearsRollOver___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_addYearsRollOver___closed__1); +l_Std_Time_DateTime_withDaysClip___closed__1 = _init_l_Std_Time_DateTime_withDaysClip___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_withDaysClip___closed__1); +l_Std_Time_DateTime_withDaysClip___closed__2 = _init_l_Std_Time_DateTime_withDaysClip___closed__2(); +lean_mark_persistent(l_Std_Time_DateTime_withDaysClip___closed__2); +l_Std_Time_DateTime_withDaysClip___closed__3 = _init_l_Std_Time_DateTime_withDaysClip___closed__3(); +lean_mark_persistent(l_Std_Time_DateTime_withDaysClip___closed__3); +l_Std_Time_DateTime_withMilliseconds___closed__1 = _init_l_Std_Time_DateTime_withMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_DateTime_withMilliseconds___closed__1); +l_Std_Time_DateTime_withMilliseconds___closed__2 = _init_l_Std_Time_DateTime_withMilliseconds___closed__2(); +lean_mark_persistent(l_Std_Time_DateTime_withMilliseconds___closed__2); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/Offset.c b/stage0/stdlib/Std/Time/Zoned/Offset.c new file mode 100644 index 000000000000..546b7208bf53 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/Offset.c @@ -0,0 +1,846 @@ +// Lean compiler output +// Module: Std.Time.Zoned.Offset +// Imports: Std.Time.Time +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__2; +lean_object* lean_int_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHoursAndMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object*, uint8_t); +uint8_t lean_int_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +lean_object* l_Int_repr(lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_instBEqOffset(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7; +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__6; +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__4; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10; +lean_object* lean_nat_abs(lean_object*); +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +static lean_object* l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3; +lean_object* lean_int_mul(lean_object*, lean_object*); +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_zero; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2; +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHours___boxed(lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHoursAndMinutes___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instInhabitedOffset; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_toIsoString___boxed(lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHours(lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instBEqOffset___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprOffset; +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_instReprOffset___closed__1; +lean_object* lean_int_neg(lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5; +static lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4; +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("second", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3; +x_2 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_4 = lean_int_dec_lt(x_1, x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_5 = l_Int_repr(x_1); +x_6 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_6, 0, x_5); +x_7 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7; +x_8 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +x_9 = 0; +x_10 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set_uint8(x_10, sizeof(void*)*1, x_9); +x_11 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6; +x_12 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12; +x_14 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14; +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11; +x_18 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +x_19 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set_uint8(x_19, sizeof(void*)*1, x_9); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_20 = l_Int_repr(x_1); +x_21 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Repr_addAppParen(x_21, x_22); +x_24 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7; +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = 0; +x_27 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_26); +x_28 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6; +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12; +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14; +x_33 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11; +x_35 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +x_36 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_26); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedOffset() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_instBEqOffset(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = lean_int_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instBEqOffset___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_TimeZone_instBEqOffset(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(":", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("-", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("0", 1, 1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("+", 1, 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object* x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_134; uint8_t x_135; +x_134 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_135 = lean_int_dec_le(x_134, x_1); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; +x_136 = lean_int_neg(x_1); +lean_dec(x_1); +x_137 = l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +x_3 = x_137; +x_4 = x_136; +goto block_133; +} +else +{ +lean_object* x_138; +x_138 = l_Std_Time_TimeZone_Offset_toIsoString___closed__7; +x_3 = x_138; +x_4 = x_1; +goto block_133; +} +block_133: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_53; lean_object* x_100; uint8_t x_101; uint8_t x_102; +x_5 = l_Std_Time_TimeZone_Offset_toIsoString___closed__1; +x_6 = lean_int_ediv(x_4, x_5); +x_7 = lean_int_mod(x_4, x_5); +lean_dec(x_4); +x_8 = l_Std_Time_TimeZone_Offset_toIsoString___closed__2; +x_9 = lean_int_ediv(x_7, x_8); +lean_dec(x_7); +x_100 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7; +x_101 = lean_int_dec_lt(x_6, x_100); +x_102 = lean_int_dec_lt(x_9, x_100); +if (x_101 == 0) +{ +lean_object* x_103; uint8_t x_104; +x_103 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_104 = lean_int_dec_lt(x_6, x_103); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +x_105 = lean_nat_abs(x_6); +lean_dec(x_6); +x_106 = l___private_Init_Data_Repr_0__Nat_reprFast(x_105); +if (x_102 == 0) +{ +x_10 = x_106; +goto block_52; +} +else +{ +x_53 = x_106; +goto block_99; +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_107 = lean_nat_abs(x_6); +lean_dec(x_6); +x_108 = lean_unsigned_to_nat(1u); +x_109 = lean_nat_sub(x_107, x_108); +lean_dec(x_107); +x_110 = lean_nat_add(x_109, x_108); +lean_dec(x_109); +x_111 = l___private_Init_Data_Repr_0__Nat_reprFast(x_110); +x_112 = l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +x_113 = lean_string_append(x_112, x_111); +lean_dec(x_111); +if (x_102 == 0) +{ +x_10 = x_113; +goto block_52; +} +else +{ +x_53 = x_113; +goto block_99; +} +} +} +else +{ +lean_object* x_114; uint8_t x_115; +x_114 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_115 = lean_int_dec_lt(x_6, x_114); +if (x_115 == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_116 = lean_nat_abs(x_6); +lean_dec(x_6); +x_117 = l___private_Init_Data_Repr_0__Nat_reprFast(x_116); +x_118 = l_Std_Time_TimeZone_Offset_toIsoString___closed__6; +x_119 = lean_string_append(x_118, x_117); +lean_dec(x_117); +x_120 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_121 = lean_string_append(x_119, x_120); +if (x_102 == 0) +{ +x_10 = x_121; +goto block_52; +} +else +{ +x_53 = x_121; +goto block_99; +} +} +else +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_122 = lean_nat_abs(x_6); +lean_dec(x_6); +x_123 = lean_unsigned_to_nat(1u); +x_124 = lean_nat_sub(x_122, x_123); +lean_dec(x_122); +x_125 = lean_nat_add(x_124, x_123); +lean_dec(x_124); +x_126 = l___private_Init_Data_Repr_0__Nat_reprFast(x_125); +x_127 = l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +x_128 = lean_string_append(x_127, x_126); +lean_dec(x_126); +x_129 = l_Std_Time_TimeZone_Offset_toIsoString___closed__6; +x_130 = lean_string_append(x_129, x_128); +lean_dec(x_128); +x_131 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_132 = lean_string_append(x_130, x_131); +if (x_102 == 0) +{ +x_10 = x_132; +goto block_52; +} +else +{ +x_53 = x_132; +goto block_99; +} +} +} +block_52: +{ +lean_object* x_11; uint8_t x_12; +x_11 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_12 = lean_int_dec_lt(x_9, x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_nat_abs(x_9); +lean_dec(x_9); +x_14 = l___private_Init_Data_Repr_0__Nat_reprFast(x_13); +if (x_2 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_15 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_16 = lean_string_append(x_15, x_3); +lean_dec(x_3); +x_17 = lean_string_append(x_16, x_15); +x_18 = lean_string_append(x_17, x_10); +lean_dec(x_10); +x_19 = lean_string_append(x_18, x_15); +x_20 = lean_string_append(x_19, x_14); +lean_dec(x_14); +x_21 = lean_string_append(x_20, x_15); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_22 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_23 = lean_string_append(x_22, x_3); +lean_dec(x_3); +x_24 = lean_string_append(x_23, x_22); +x_25 = lean_string_append(x_24, x_10); +lean_dec(x_10); +x_26 = l_Std_Time_TimeZone_Offset_toIsoString___closed__4; +x_27 = lean_string_append(x_25, x_26); +x_28 = lean_string_append(x_27, x_14); +lean_dec(x_14); +x_29 = lean_string_append(x_28, x_22); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_30 = lean_nat_abs(x_9); +lean_dec(x_9); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_sub(x_30, x_31); +lean_dec(x_30); +x_33 = lean_nat_add(x_32, x_31); +lean_dec(x_32); +x_34 = l___private_Init_Data_Repr_0__Nat_reprFast(x_33); +x_35 = l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +x_36 = lean_string_append(x_35, x_34); +lean_dec(x_34); +if (x_2 == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_37 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_38 = lean_string_append(x_37, x_3); +lean_dec(x_3); +x_39 = lean_string_append(x_38, x_37); +x_40 = lean_string_append(x_39, x_10); +lean_dec(x_10); +x_41 = lean_string_append(x_40, x_37); +x_42 = lean_string_append(x_41, x_36); +lean_dec(x_36); +x_43 = lean_string_append(x_42, x_37); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_44 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_45 = lean_string_append(x_44, x_3); +lean_dec(x_3); +x_46 = lean_string_append(x_45, x_44); +x_47 = lean_string_append(x_46, x_10); +lean_dec(x_10); +x_48 = l_Std_Time_TimeZone_Offset_toIsoString___closed__4; +x_49 = lean_string_append(x_47, x_48); +x_50 = lean_string_append(x_49, x_36); +lean_dec(x_36); +x_51 = lean_string_append(x_50, x_44); +return x_51; +} +} +} +block_99: +{ +lean_object* x_54; uint8_t x_55; +x_54 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +x_55 = lean_int_dec_lt(x_9, x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_56 = lean_nat_abs(x_9); +lean_dec(x_9); +x_57 = l___private_Init_Data_Repr_0__Nat_reprFast(x_56); +x_58 = l_Std_Time_TimeZone_Offset_toIsoString___closed__6; +x_59 = lean_string_append(x_58, x_57); +lean_dec(x_57); +x_60 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_61 = lean_string_append(x_59, x_60); +if (x_2 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_62 = lean_string_append(x_60, x_3); +lean_dec(x_3); +x_63 = lean_string_append(x_62, x_60); +x_64 = lean_string_append(x_63, x_53); +lean_dec(x_53); +x_65 = lean_string_append(x_64, x_60); +x_66 = lean_string_append(x_65, x_61); +lean_dec(x_61); +x_67 = lean_string_append(x_66, x_60); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_68 = lean_string_append(x_60, x_3); +lean_dec(x_3); +x_69 = lean_string_append(x_68, x_60); +x_70 = lean_string_append(x_69, x_53); +lean_dec(x_53); +x_71 = l_Std_Time_TimeZone_Offset_toIsoString___closed__4; +x_72 = lean_string_append(x_70, x_71); +x_73 = lean_string_append(x_72, x_61); +lean_dec(x_61); +x_74 = lean_string_append(x_73, x_60); +return x_74; +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_75 = lean_nat_abs(x_9); +lean_dec(x_9); +x_76 = lean_unsigned_to_nat(1u); +x_77 = lean_nat_sub(x_75, x_76); +lean_dec(x_75); +x_78 = lean_nat_add(x_77, x_76); +lean_dec(x_77); +x_79 = l___private_Init_Data_Repr_0__Nat_reprFast(x_78); +x_80 = l_Std_Time_TimeZone_Offset_toIsoString___closed__5; +x_81 = lean_string_append(x_80, x_79); +lean_dec(x_79); +x_82 = l_Std_Time_TimeZone_Offset_toIsoString___closed__6; +x_83 = lean_string_append(x_82, x_81); +lean_dec(x_81); +x_84 = l_Std_Time_TimeZone_Offset_toIsoString___closed__3; +x_85 = lean_string_append(x_83, x_84); +if (x_2 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_86 = lean_string_append(x_84, x_3); +lean_dec(x_3); +x_87 = lean_string_append(x_86, x_84); +x_88 = lean_string_append(x_87, x_53); +lean_dec(x_53); +x_89 = lean_string_append(x_88, x_84); +x_90 = lean_string_append(x_89, x_85); +lean_dec(x_85); +x_91 = lean_string_append(x_90, x_84); +return x_91; +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_92 = lean_string_append(x_84, x_3); +lean_dec(x_3); +x_93 = lean_string_append(x_92, x_84); +x_94 = lean_string_append(x_93, x_53); +lean_dec(x_53); +x_95 = l_Std_Time_TimeZone_Offset_toIsoString___closed__4; +x_96 = lean_string_append(x_94, x_95); +x_97 = lean_string_append(x_96, x_85); +lean_dec(x_85); +x_98 = lean_string_append(x_97, x_84); +return x_98; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_toIsoString___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Std_Time_TimeZone_Offset_toIsoString(x_1, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Offset_zero() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHours(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Std_Time_TimeZone_Offset_toIsoString___closed__1; +x_3 = lean_int_mul(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHours___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_Offset_ofHours(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHoursAndMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_3 = l_Std_Time_TimeZone_Offset_toIsoString___closed__1; +x_4 = lean_int_mul(x_1, x_3); +x_5 = l_Std_Time_TimeZone_Offset_toIsoString___closed__2; +x_6 = lean_int_mul(x_2, x_5); +x_7 = lean_int_add(x_4, x_6); +lean_dec(x_6); +lean_dec(x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Offset_ofHoursAndMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Offset_ofHoursAndMinutes(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +lean_object* initialize_Std_Time_Time(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_Offset(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Time(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__1); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__2); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__3); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__4); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__5); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__6); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__7); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__8); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__9); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__10); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__11); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__12); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__13); +l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14 = _init_l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19____closed__14); +l_Std_Time_TimeZone_instReprOffset___closed__1 = _init_l_Std_Time_TimeZone_instReprOffset___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprOffset___closed__1); +l_Std_Time_TimeZone_instReprOffset = _init_l_Std_Time_TimeZone_instReprOffset(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprOffset); +l_Std_Time_TimeZone_instInhabitedOffset = _init_l_Std_Time_TimeZone_instInhabitedOffset(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedOffset); +l_Std_Time_TimeZone_Offset_toIsoString___closed__1 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__1); +l_Std_Time_TimeZone_Offset_toIsoString___closed__2 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__2); +l_Std_Time_TimeZone_Offset_toIsoString___closed__3 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__3); +l_Std_Time_TimeZone_Offset_toIsoString___closed__4 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__4(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__4); +l_Std_Time_TimeZone_Offset_toIsoString___closed__5 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__5(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__5); +l_Std_Time_TimeZone_Offset_toIsoString___closed__6 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__6(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__6); +l_Std_Time_TimeZone_Offset_toIsoString___closed__7 = _init_l_Std_Time_TimeZone_Offset_toIsoString___closed__7(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_toIsoString___closed__7); +l_Std_Time_TimeZone_Offset_zero = _init_l_Std_Time_TimeZone_Offset_zero(); +lean_mark_persistent(l_Std_Time_TimeZone_Offset_zero); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/TimeZone.c b/stage0/stdlib/Std/Time/Zoned/TimeZone.c new file mode 100644 index 000000000000..f01576e534e3 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/TimeZone.c @@ -0,0 +1,927 @@ +// Lean compiler output +// Module: Std.Time.Zoned.TimeZone +// Imports: Std.Time.Zoned.Offset +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofSeconds___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22; +static lean_object* l_Std_Time_TimeZone_UTC___closed__2; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_GMT; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24; +lean_object* l_String_quote(lean_object*); +static lean_object* l_Std_Time_TimeZone_UTC___closed__1; +static lean_object* l_Std_Time_instInhabitedTimeZone___closed__1; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32; +static lean_object* l_Std_Time_TimeZone_GMT___closed__3; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54_(lean_object*, lean_object*); +uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19; +LEAN_EXPORT uint8_t l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26; +lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instReprTimeZone; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28; +static lean_object* l_Std_Time_TimeZone_GMT___closed__2; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2; +static lean_object* l_Std_Time_instReprTimeZone___closed__1; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21; +static lean_object* l_Std_Time_instInhabitedTimeZone___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_instBEqTimeZone; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofHours(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16; +lean_object* lean_string_length(lean_object*); +static lean_object* l_Std_Time_instInhabitedTimeZone___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofHours___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29; +extern lean_object* l_Std_Time_TimeZone_Offset_zero; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_toSeconds___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTC; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofSeconds(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Std_Time_instBEqTimeZone___closed__1; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8; +lean_object* l_Std_Time_TimeZone_Offset_ofHours(lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30; +static lean_object* l_Std_Time_TimeZone_GMT___closed__1; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedTimeZone; +static lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3; +static lean_object* _init_l_Std_Time_instInhabitedTimeZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedTimeZone___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedTimeZone___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Std_Time_instInhabitedTimeZone___closed__1; +x_2 = l_Std_Time_instInhabitedTimeZone___closed__2; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*3, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedTimeZone() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedTimeZone___closed__3; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("offset", 6, 6); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3; +x_2 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("name", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("abbreviation", 12, 12); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(16u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isDST", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18; +x_2 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18; +x_2 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(x_3, x_4); +x_6 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +x_21 = l_String_quote(x_20); +x_22 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12; +x_24 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*1, x_8); +x_26 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_26, 0, x_19); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_12); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_14); +x_29 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14; +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_18); +x_32 = lean_ctor_get(x_1, 2); +x_33 = l_String_quote(x_32); +x_34 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15; +x_36 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set_uint8(x_37, sizeof(void*)*1, x_8); +x_38 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_12); +x_40 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_14); +x_41 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17; +x_42 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_18); +x_44 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_45 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_43); +lean_ctor_set(x_46, 1, x_45); +x_47 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22; +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24; +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +x_51 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21; +x_52 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +x_53 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set_uint8(x_53, sizeof(void*)*1, x_8); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_54 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32; +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_43); +lean_ctor_set(x_55, 1, x_54); +x_56 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22; +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +x_58 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24; +x_59 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21; +x_61 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_59); +x_62 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*1, x_8); +return x_62; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instReprTimeZone___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instReprTimeZone() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instReprTimeZone___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; uint8_t x_11; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_ctor_get(x_2, 1); +x_9 = lean_ctor_get(x_2, 2); +x_10 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_11 = lean_int_dec_eq(x_3, x_7); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = lean_string_dec_eq(x_4, x_8); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = 0; +return x_14; +} +else +{ +uint8_t x_15; +x_15 = lean_string_dec_eq(x_5, x_9); +if (x_15 == 0) +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +else +{ +if (x_6 == 0) +{ +if (x_10 == 0) +{ +uint8_t x_17; +x_17 = 1; +return x_17; +} +else +{ +uint8_t x_18; +x_18 = 0; +return x_18; +} +} +else +{ +return x_10; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instBEqTimeZone___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_beqTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_152____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instBEqTimeZone() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instBEqTimeZone___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_UTC___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UTC", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_UTC___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Std_Time_TimeZone_Offset_zero; +x_2 = l_Std_Time_TimeZone_UTC___closed__1; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*3, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_UTC() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_UTC___closed__2; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_GMT___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Greenwich Mean Time", 19, 19); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_GMT___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("GMT", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_GMT___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_1 = l_Std_Time_TimeZone_Offset_zero; +x_2 = l_Std_Time_TimeZone_GMT___closed__1; +x_3 = l_Std_Time_TimeZone_GMT___closed__2; +x_4 = 0; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_TimeZone_GMT() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_GMT___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofHours(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Std_Time_TimeZone_Offset_ofHours(x_3); +x_6 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofHours___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_4); +lean_dec(x_4); +x_6 = l_Std_Time_TimeZone_ofHours(x_1, x_2, x_3, x_5); +lean_dec(x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofSeconds(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_1); +lean_ctor_set(x_5, 2, x_2); +lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ofSeconds___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_4); +lean_dec(x_4); +x_6 = l_Std_Time_TimeZone_ofSeconds(x_1, x_2, x_3, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_toSeconds(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_toSeconds___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_toSeconds(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_Zoned_Offset(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_TimeZone(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Zoned_Offset(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_instInhabitedTimeZone___closed__1 = _init_l_Std_Time_instInhabitedTimeZone___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedTimeZone___closed__1); +l_Std_Time_instInhabitedTimeZone___closed__2 = _init_l_Std_Time_instInhabitedTimeZone___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedTimeZone___closed__2); +l_Std_Time_instInhabitedTimeZone___closed__3 = _init_l_Std_Time_instInhabitedTimeZone___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedTimeZone___closed__3); +l_Std_Time_instInhabitedTimeZone = _init_l_Std_Time_instInhabitedTimeZone(); +lean_mark_persistent(l_Std_Time_instInhabitedTimeZone); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__1); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__2); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__3); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__4); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__5); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__6); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__7); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__8); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__9); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__10); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__11); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__12); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__13); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__14); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__15); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__16); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__17); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__18); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__19); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__20); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__21); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__22); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__23); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__24); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__25); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__26); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__27); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__28); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__29); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__30); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__31); +l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32 = _init_l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32(); +lean_mark_persistent(l___private_Std_Time_Zoned_TimeZone_0__Std_Time_reprTimeZone____x40_Std_Time_Zoned_TimeZone___hyg_54____closed__32); +l_Std_Time_instReprTimeZone___closed__1 = _init_l_Std_Time_instReprTimeZone___closed__1(); +lean_mark_persistent(l_Std_Time_instReprTimeZone___closed__1); +l_Std_Time_instReprTimeZone = _init_l_Std_Time_instReprTimeZone(); +lean_mark_persistent(l_Std_Time_instReprTimeZone); +l_Std_Time_instBEqTimeZone___closed__1 = _init_l_Std_Time_instBEqTimeZone___closed__1(); +lean_mark_persistent(l_Std_Time_instBEqTimeZone___closed__1); +l_Std_Time_instBEqTimeZone = _init_l_Std_Time_instBEqTimeZone(); +lean_mark_persistent(l_Std_Time_instBEqTimeZone); +l_Std_Time_TimeZone_UTC___closed__1 = _init_l_Std_Time_TimeZone_UTC___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_UTC___closed__1); +l_Std_Time_TimeZone_UTC___closed__2 = _init_l_Std_Time_TimeZone_UTC___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_UTC___closed__2); +l_Std_Time_TimeZone_UTC = _init_l_Std_Time_TimeZone_UTC(); +lean_mark_persistent(l_Std_Time_TimeZone_UTC); +l_Std_Time_TimeZone_GMT___closed__1 = _init_l_Std_Time_TimeZone_GMT___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_GMT___closed__1); +l_Std_Time_TimeZone_GMT___closed__2 = _init_l_Std_Time_TimeZone_GMT___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_GMT___closed__2); +l_Std_Time_TimeZone_GMT___closed__3 = _init_l_Std_Time_TimeZone_GMT___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_GMT___closed__3); +l_Std_Time_TimeZone_GMT = _init_l_Std_Time_TimeZone_GMT(); +lean_mark_persistent(l_Std_Time_TimeZone_GMT); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/ZoneRules.c b/stage0/stdlib/Std/Time/Zoned/ZoneRules.c new file mode 100644 index 000000000000..10559cc242ac --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/ZoneRules.c @@ -0,0 +1,2841 @@ +// Lean compiler output +// Module: Std.Time.Zoned.ZoneRules +// Imports: Std.Time.DateTime Std.Time.Zoned.TimeZone +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(lean_object*, lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_ofTimeZone(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instInhabitedLocalTimeType; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1___boxed(lean_object*); +lean_object* l_String_quote(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_findLocalTimeTypeForTimestamp(lean_object*, lean_object*); +lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17; +static lean_object* l_Std_Time_TimeZone_ZoneRules_UTC___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_timezoneAt(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_ZoneRules_UTC___closed__3; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8; +static lean_object* l_Std_Time_TimeZone_instReprLocalTimeType___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110_(uint8_t, lean_object*); +lean_object* l_Std_Time_TimeZone_Offset_toIsoString(lean_object*, uint8_t); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_UTC; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprLocalTimeType; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11; +static lean_object* l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_timezoneAt___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion(lean_object*); +static lean_object* l_Std_Time_TimeZone_instReprZoneRules___closed__1; +lean_object* l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionForTimestamp___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_instInhabitedStdWall; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instInhabitedZoneRules; +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__1(lean_object*); +lean_object* l_Array_back_x3f___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7; +lean_object* lean_array_to_list(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprUTLocal; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4; +lean_object* l_Int_repr(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7; +static lean_object* l_Std_Time_TimeZone_Transition_apply___closed__1; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530_(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9; +static lean_object* l_Std_Time_TimeZone_instReprStdWall___closed__1; +static lean_object* l_Std_Time_TimeZone_Transition_timezoneAt___closed__2; +static lean_object* l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10; +static lean_object* l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_findLocalTimeTypeForTimestamp___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_toCtorIdx(uint8_t); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19; +static lean_object* l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_apply___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instInhabitedTransition; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_timezoneAt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprStdWall; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5; +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +lean_object* lean_string_length(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13; +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17; +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3; +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition___boxed(lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15; +lean_object* lean_array_mk(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_ZoneRules_UTC___closed__2; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprZoneRules; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_instReprTransition; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_ofTimeZone___boxed(lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__1___boxed(lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22; +static lean_object* l_Std_Time_TimeZone_instReprUTLocal___closed__1; +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_instInhabitedUTLocal; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12; +static lean_object* l_Std_Time_TimeZone_Transition_timezoneAt___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18; +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____boxed(lean_object*, lean_object*); +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_instInhabitedTransition___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21; +static lean_object* l_Std_Time_TimeZone_instReprTransition___closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19; +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_apply(lean_object*, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10; +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_timezoneAt___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12_(uint8_t, lean_object*); +static lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18; +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3; +static lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1; +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_TimeZone_UTLocal_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_TimeZone_UTLocal_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.UTLocal.ut", 28, 28); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.UTLocal.local", 31, 31); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprUTLocal___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprUTLocal() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprUTLocal___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_TimeZone_instInhabitedUTLocal() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_toCtorIdx(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +else +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Std_Time_TimeZone_StdWall_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_StdWall_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_StdWall_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Std_Time_TimeZone_StdWall_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.StdWall.wall", 30, 30); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("Std.Time.TimeZone.StdWall.standard", 34, 34); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110_(uint8_t x_1, lean_object* x_2) { +_start: +{ +if (x_1 == 0) +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +else +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprStdWall___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprStdWall() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprStdWall___closed__1; +return x_1; +} +} +static uint8_t _init_l_Std_Time_TimeZone_instInhabitedStdWall() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("gmtOffset", 9, 9); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" := ", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(13u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(",", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("isDst", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(9u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("abbreviation", 12, 12); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(16u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("wall", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(8u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("utLocal", 7, 7); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("identifier", 10, 10); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(14u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("{ ", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked(" }", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("false", 5, 5); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("true", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_Offset_0__Std_Time_TimeZone_reprOffset____x40_Std_Time_Zoned_Offset___hyg_19_(x_3, x_4); +x_6 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_21 = lean_ctor_get(x_1, 1); +x_22 = l_String_quote(x_21); +x_23 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15; +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set_uint8(x_26, sizeof(void*)*1, x_8); +x_27 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 1); +x_28 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110_(x_27, x_4); +x_29 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18; +x_30 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set_uint8(x_31, sizeof(void*)*1, x_8); +x_32 = lean_ctor_get_uint8(x_1, sizeof(void*)*3 + 2); +x_33 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12_(x_32, x_4); +x_34 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21; +x_35 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +x_36 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_8); +x_37 = lean_ctor_get(x_1, 2); +x_38 = l_String_quote(x_37); +x_39 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24; +x_41 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set_uint8(x_42, sizeof(void*)*1, x_8); +if (x_20 == 0) +{ +lean_object* x_80; +x_80 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32; +x_43 = x_80; +goto block_79; +} +else +{ +lean_object* x_81; +x_81 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34; +x_43 = x_81; +goto block_79; +} +block_79: +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_44 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12; +x_45 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_43); +x_46 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set_uint8(x_46, sizeof(void*)*1, x_8); +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_19); +lean_ctor_set(x_47, 1, x_46); +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_12); +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_14); +x_50 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14; +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_18); +x_53 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_26); +x_54 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_12); +x_55 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_14); +x_56 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17; +x_57 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_18); +x_59 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_31); +x_60 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_12); +x_61 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_14); +x_62 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20; +x_63 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +x_64 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_18); +x_65 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_36); +x_66 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_12); +x_67 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_14); +x_68 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23; +x_69 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_18); +x_71 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_42); +x_72 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +x_73 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +x_75 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +x_77 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_75); +x_78 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*1, x_8); +return x_78; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprLocalTimeType___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprLocalTimeType() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprLocalTimeType___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +x_2 = 0; +x_3 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2; +x_4 = 0; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 2, x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_2); +x_6 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_5); +lean_ctor_set(x_6, 2, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("time", 4, 4); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("localTimeType", 13, 13); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(17u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +x_3 = lean_ctor_get(x_1, 0); +x_4 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +x_5 = lean_int_dec_lt(x_3, x_4); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255_(x_6, x_7); +x_9 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7; +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = 0; +x_12 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set_uint8(x_12, sizeof(void*)*1, x_11); +if (x_5 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_13 = l_Int_repr(x_3); +x_14 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18; +x_16 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*1, x_11); +x_18 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +x_21 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_box(1); +x_23 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6; +x_25 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_12); +x_29 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +x_30 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_31 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +x_32 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +x_34 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_11); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_36 = l_Int_repr(x_3); +x_37 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Repr_addAppParen(x_37, x_7); +x_39 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18; +x_40 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*1, x_11); +x_42 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4; +x_43 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +x_45 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_box(1); +x_47 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +x_48 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6; +x_49 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_51 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_12); +x_53 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +x_54 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_52); +x_55 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +x_56 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +x_57 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +x_58 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +x_59 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set_uint8(x_59, sizeof(void*)*1, x_11); +return x_59; +} +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprTransition___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprTransition() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprTransition___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedTransition___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +x_2 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedTransition() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instInhabitedTransition___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_List_foldl___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_1); +return x_2; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_3); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_3, 1); +lean_inc(x_1); +lean_ctor_set_tag(x_3, 5); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 0, x_2); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_5, x_7); +lean_dec(x_5); +x_9 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +x_2 = x_9; +x_3 = x_6; +goto _start; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +lean_inc(x_1); +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_2); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_unsigned_to_nat(0u); +x_15 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_11, x_14); +lean_dec(x_11); +x_16 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_16, 0, x_13); +lean_ctor_set(x_16, 1, x_15); +x_2 = x_16; +x_3 = x_12; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +lean_dec(x_2); +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_2); +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +lean_dec(x_1); +x_6 = lean_unsigned_to_nat(0u); +x_7 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_5, x_6); +lean_dec(x_5); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +lean_dec(x_1); +x_9 = lean_unsigned_to_nat(0u); +x_10 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440_(x_8, x_9); +lean_dec(x_8); +x_11 = l_List_foldl___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__3(x_2, x_10, x_4); +return x_11; +} +} +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("initialLocalTimeType", 20, 20); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(24u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("transitions", 11, 11); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(15u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +x_2 = lean_box(1); +x_3 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[", 2, 2); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10; +x_2 = lean_string_length(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11; +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("]", 1, 1); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("#[]", 3, 3); +return x_1; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8; +x_2 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); +x_5 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255_(x_3, x_4); +lean_dec(x_3); +x_6 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5; +x_7 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +x_8 = 0; +x_9 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); +x_10 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4; +x_11 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9; +x_13 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +x_14 = lean_box(1); +x_15 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +x_16 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7; +x_17 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +x_18 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5; +x_19 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_ctor_get(x_1, 1); +lean_inc(x_20); +lean_dec(x_1); +x_21 = lean_array_get_size(x_20); +x_22 = lean_nat_dec_eq(x_21, x_4); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_23 = lean_array_to_list(x_20); +x_24 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9; +x_25 = l_Std_Format_joinSep___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__2(x_23, x_24); +x_26 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13; +x_27 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15; +x_29 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12; +x_31 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = 1; +x_33 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); +x_34 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8; +x_35 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_33); +x_36 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set_uint8(x_36, sizeof(void*)*1, x_8); +x_37 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_37, 0, x_19); +lean_ctor_set(x_37, 1, x_36); +x_38 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +x_39 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +x_41 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +x_43 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set_uint8(x_44, sizeof(void*)*1, x_8); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +lean_dec(x_20); +x_45 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19; +x_46 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_45); +x_47 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28; +x_48 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30; +x_50 = lean_alloc_ctor(5, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +x_51 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27; +x_52 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +x_53 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set_uint8(x_53, sizeof(void*)*1, x_8); +return x_53; +} +} +} +LEAN_EXPORT lean_object* l_repr___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_repr___at___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____spec__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530_(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprZoneRules___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instReprZoneRules() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instReprZoneRules___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3; +x_2 = l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_instInhabitedZoneRules() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_ctor_get(x_1, 1); +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +lean_inc(x_6); +lean_inc(x_3); +x_7 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_7, 0, x_3); +lean_ctor_set(x_7, 1, x_6); +lean_ctor_set(x_7, 2, x_5); +lean_ctor_set_uint8(x_7, sizeof(void*)*3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Transition_apply___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_apply(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_2, 1); +x_4 = lean_ctor_get(x_3, 0); +x_5 = lean_int_add(x_4, x_4); +x_6 = lean_ctor_get(x_1, 0); +x_7 = l_Std_Time_TimeZone_Transition_apply___closed__1; +x_8 = lean_int_mul(x_6, x_7); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_int_add(x_8, x_9); +lean_dec(x_8); +x_11 = lean_int_mul(x_5, x_7); +lean_dec(x_5); +x_12 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_add(x_10, x_13); +lean_dec(x_13); +lean_dec(x_10); +x_15 = l_Std_Time_Duration_ofNanoseconds(x_14); +lean_dec(x_14); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_apply___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_apply(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_dec_lt(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +lean_dec(x_2); +x_4 = lean_alloc_closure((void*)(l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1___boxed), 2, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Array_findIdx_x3f_loop___rarg(x_4, x_1, x_5); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; +x_7 = lean_box(0); +return x_7; +} +else +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 0) +{ +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_findTransitionIndexForTimestamp(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = l_Array_back_x3f___rarg(x_1); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_ctor_get(x_3, 0); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_sub(x_5, x_6); +lean_dec(x_5); +x_8 = l_Array_get_x3f___rarg(x_1, x_7); +lean_dec(x_7); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_findTransitionForTimestamp___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Transition_timezoneAt___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("cannot find local timezone.", 27, 27); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_Transition_timezoneAt___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_TimeZone_Transition_timezoneAt___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_timezoneAt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_Transition_timezoneAt___closed__2; +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 0); +x_7 = l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition(x_6); +lean_dec(x_6); +lean_ctor_set(x_3, 0, x_7); +return x_3; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +lean_dec(x_3); +x_9 = l_Std_Time_TimeZone_Transition_createTimeZoneFromTransition(x_8); +lean_dec(x_8); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_Transition_timezoneAt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_Transition_timezoneAt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +uint8_t x_4; lean_object* x_5; +x_4 = 1; +lean_inc(x_1); +x_5 = l_Std_Time_TimeZone_Offset_toIsoString(x_1, x_4); +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = 0; +x_7 = 1; +x_8 = 0; +lean_inc(x_5); +x_9 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_5); +lean_ctor_set(x_9, 2, x_5); +lean_ctor_set_uint8(x_9, sizeof(void*)*3, x_6); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 1, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 2, x_8); +x_10 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +else +{ +lean_object* x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_12 = lean_ctor_get(x_2, 0); +x_13 = 0; +x_14 = 1; +x_15 = 0; +lean_inc(x_12); +x_16 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_5); +lean_ctor_set(x_16, 2, x_12); +lean_ctor_set_uint8(x_16, sizeof(void*)*3, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*3 + 1, x_14); +lean_ctor_set_uint8(x_16, sizeof(void*)*3 + 2, x_15); +x_17 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +else +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_19 = lean_ctor_get(x_3, 0); +x_20 = 1; +lean_inc(x_1); +x_21 = l_Std_Time_TimeZone_Offset_toIsoString(x_1, x_20); +x_22 = 0; +x_23 = 1; +x_24 = 0; +lean_inc(x_19); +x_25 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_19); +lean_ctor_set(x_25, 2, x_21); +lean_ctor_set_uint8(x_25, sizeof(void*)*3, x_22); +lean_ctor_set_uint8(x_25, sizeof(void*)*3 + 1, x_23); +lean_ctor_set_uint8(x_25, sizeof(void*)*3 + 2, x_24); +x_26 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_28 = lean_ctor_get(x_3, 0); +x_29 = lean_ctor_get(x_2, 0); +x_30 = 0; +x_31 = 1; +x_32 = 0; +lean_inc(x_29); +lean_inc(x_28); +x_33 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_33, 0, x_1); +lean_ctor_set(x_33, 1, x_28); +lean_ctor_set(x_33, 2, x_29); +lean_ctor_set_uint8(x_33, sizeof(void*)*3, x_30); +lean_ctor_set_uint8(x_33, sizeof(void*)*3 + 1, x_31); +lean_ctor_set_uint8(x_33, sizeof(void*)*3 + 2, x_32); +x_34 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("UTC", 3, 3); +return x_1; +} +} +static lean_object* _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_TimeZone_ZoneRules_UTC___closed__1; +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1; +x_2 = l_Std_Time_TimeZone_ZoneRules_UTC___closed__2; +x_3 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone(x_1, x_2, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_TimeZone_ZoneRules_UTC() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_TimeZone_ZoneRules_UTC___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_findLocalTimeTypeForTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 1); +x_4 = l_Std_Time_TimeZone_Transition_findTransitionForTimestamp(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_findLocalTimeTypeForTimestamp___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_ZoneRules_findLocalTimeTypeForTimestamp(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_timezoneAt(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 1); +x_4 = l_Std_Time_TimeZone_Transition_timezoneAt(x_3, x_2); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; +lean_dec(x_4); +x_5 = lean_ctor_get(x_1, 0); +x_6 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_5); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +lean_dec(x_4); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_timezoneAt___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_TimeZone_ZoneRules_timezoneAt(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_ofTimeZone(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; uint8_t x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_1, 2); +x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_6 = 0; +x_7 = 1; +lean_inc(x_3); +lean_inc(x_4); +lean_inc(x_2); +x_8 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_4); +lean_ctor_set(x_8, 2, x_3); +lean_ctor_set_uint8(x_8, sizeof(void*)*3, x_5); +lean_ctor_set_uint8(x_8, sizeof(void*)*3 + 1, x_6); +lean_ctor_set_uint8(x_8, sizeof(void*)*3 + 2, x_7); +x_9 = l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Std_Time_TimeZone_ZoneRules_ofTimeZone___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_TimeZone_ZoneRules_ofTimeZone(x_1); +lean_dec(x_1); +return x_2; +} +} +lean_object* initialize_Std_Time_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_TimeZone(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_TimeZone(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1 = _init_l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_UTLocal_noConfusion___rarg___closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__2); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__3); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__4); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__5); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__6); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__7); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__8); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__9); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__10); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__11); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__12); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__13); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprUTLocal____x40_Std_Time_Zoned_ZoneRules___hyg_12____closed__14); +l_Std_Time_TimeZone_instReprUTLocal___closed__1 = _init_l_Std_Time_TimeZone_instReprUTLocal___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprUTLocal___closed__1); +l_Std_Time_TimeZone_instReprUTLocal = _init_l_Std_Time_TimeZone_instReprUTLocal(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprUTLocal); +l_Std_Time_TimeZone_instInhabitedUTLocal = _init_l_Std_Time_TimeZone_instInhabitedUTLocal(); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__2); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__3); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__4); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__5); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__6); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__7); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__8); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__9); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__10); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__11); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprStdWall____x40_Std_Time_Zoned_ZoneRules___hyg_110____closed__12); +l_Std_Time_TimeZone_instReprStdWall___closed__1 = _init_l_Std_Time_TimeZone_instReprStdWall___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprStdWall___closed__1); +l_Std_Time_TimeZone_instReprStdWall = _init_l_Std_Time_TimeZone_instReprStdWall(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprStdWall); +l_Std_Time_TimeZone_instInhabitedStdWall = _init_l_Std_Time_TimeZone_instInhabitedStdWall(); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__2); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__3); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__4); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__5); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__6); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__7); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__8); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__9); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__10); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__11); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__12); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__13); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__14); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__15); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__16); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__17); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__18); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__19); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__20); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__21); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__22); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__23); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__24); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__25); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__26); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__27); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__28); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__29); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__30); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__31); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__32); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__33); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprLocalTimeType____x40_Std_Time_Zoned_ZoneRules___hyg_255____closed__34); +l_Std_Time_TimeZone_instReprLocalTimeType___closed__1 = _init_l_Std_Time_TimeZone_instReprLocalTimeType___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprLocalTimeType___closed__1); +l_Std_Time_TimeZone_instReprLocalTimeType = _init_l_Std_Time_TimeZone_instReprLocalTimeType(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprLocalTimeType); +l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1 = _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__1); +l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2 = _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__2); +l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3 = _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedLocalTimeType___closed__3); +l_Std_Time_TimeZone_instInhabitedLocalTimeType = _init_l_Std_Time_TimeZone_instInhabitedLocalTimeType(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedLocalTimeType); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__2); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__3); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__4); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__5); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__6); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprTransition____x40_Std_Time_Zoned_ZoneRules___hyg_440____closed__7); +l_Std_Time_TimeZone_instReprTransition___closed__1 = _init_l_Std_Time_TimeZone_instReprTransition___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprTransition___closed__1); +l_Std_Time_TimeZone_instReprTransition = _init_l_Std_Time_TimeZone_instReprTransition(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprTransition); +l_Std_Time_TimeZone_instInhabitedTransition___closed__1 = _init_l_Std_Time_TimeZone_instInhabitedTransition___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedTransition___closed__1); +l_Std_Time_TimeZone_instInhabitedTransition = _init_l_Std_Time_TimeZone_instInhabitedTransition(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedTransition); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__1); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__2); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__3); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__4); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__5); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__6); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__7); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__8); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__9); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__10); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__11); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__12); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__13); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__14); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__15); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__16); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__17); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__18); +l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19 = _init_l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19(); +lean_mark_persistent(l___private_Std_Time_Zoned_ZoneRules_0__Std_Time_TimeZone_reprZoneRules____x40_Std_Time_Zoned_ZoneRules___hyg_530____closed__19); +l_Std_Time_TimeZone_instReprZoneRules___closed__1 = _init_l_Std_Time_TimeZone_instReprZoneRules___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprZoneRules___closed__1); +l_Std_Time_TimeZone_instReprZoneRules = _init_l_Std_Time_TimeZone_instReprZoneRules(); +lean_mark_persistent(l_Std_Time_TimeZone_instReprZoneRules); +l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1 = _init_l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedZoneRules___closed__1); +l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2 = _init_l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedZoneRules___closed__2); +l_Std_Time_TimeZone_instInhabitedZoneRules = _init_l_Std_Time_TimeZone_instInhabitedZoneRules(); +lean_mark_persistent(l_Std_Time_TimeZone_instInhabitedZoneRules); +l_Std_Time_TimeZone_Transition_apply___closed__1 = _init_l_Std_Time_TimeZone_Transition_apply___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_Transition_apply___closed__1); +l_Std_Time_TimeZone_Transition_timezoneAt___closed__1 = _init_l_Std_Time_TimeZone_Transition_timezoneAt___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_Transition_timezoneAt___closed__1); +l_Std_Time_TimeZone_Transition_timezoneAt___closed__2 = _init_l_Std_Time_TimeZone_Transition_timezoneAt___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_Transition_timezoneAt___closed__2); +l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1 = _init_l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_ZoneRules_fixedOffsetZone___closed__1); +l_Std_Time_TimeZone_ZoneRules_UTC___closed__1 = _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__1(); +lean_mark_persistent(l_Std_Time_TimeZone_ZoneRules_UTC___closed__1); +l_Std_Time_TimeZone_ZoneRules_UTC___closed__2 = _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__2(); +lean_mark_persistent(l_Std_Time_TimeZone_ZoneRules_UTC___closed__2); +l_Std_Time_TimeZone_ZoneRules_UTC___closed__3 = _init_l_Std_Time_TimeZone_ZoneRules_UTC___closed__3(); +lean_mark_persistent(l_Std_Time_TimeZone_ZoneRules_UTC___closed__3); +l_Std_Time_TimeZone_ZoneRules_UTC = _init_l_Std_Time_TimeZone_ZoneRules_UTC(); +lean_mark_persistent(l_Std_Time_TimeZone_ZoneRules_UTC); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Std/Time/Zoned/ZonedDateTime.c b/stage0/stdlib/Std/Time/Zoned/ZonedDateTime.c new file mode 100644 index 000000000000..b058484d1c43 --- /dev/null +++ b/stage0/stdlib/Std/Time/Zoned/ZonedDateTime.c @@ -0,0 +1,14945 @@ +// Lean compiler output +// Module: Std.Time.Zoned.ZonedDateTime +// Imports: Std.Time.Zoned.DateTime Std.Time.Zoned.ZoneRules +#include +#if defined(__clang__) +#pragma clang diagnostic ignored "-Wunused-parameter" +#pragma clang diagnostic ignored "-Wunused-label" +#elif defined(__GNUC__) && !defined(__CLANG__) +#pragma GCC diagnostic ignored "-Wunused-parameter" +#pragma GCC diagnostic ignored "-Wunused-label" +#pragma GCC diagnostic ignored "-Wunused-but-set-variable" +#endif +#ifdef __cplusplus +extern "C" { +#endif +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addWeeks(lean_object*, lean_object*); +lean_object* lean_int_mod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfMonth(lean_object*); +lean_object* l_Std_Time_Duration_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nanosecond(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_TimeZone_LocalTimeType_getTimeZone(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toNanoseconds(lean_object*); +lean_object* l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekday___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toTimestamp___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +static lean_object* l_Std_Time_ZonedDateTime_addMinutes___closed__1; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsRollOver___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_Month_Ordinal_days(uint8_t, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofDaysSinceUNIXEpoch___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subDays(lean_object*, lean_object*); +lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_inLeapYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDateTime(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15; +static lean_object* l_Std_Time_ZonedDateTime_addWeeks___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMinutes(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subSeconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMinutes___boxed(lean_object*, lean_object*); +lean_object* lean_int_emod(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDateTime(lean_object*); +lean_object* l_Std_Time_PlainDate_alignedWeekOfMonth(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__5; +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_inLeapYear(lean_object*); +lean_object* l_Std_Time_PlainDateTime_addMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__6; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofDaysSinceUNIXEpoch(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__4; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_alignedWeekOfMonth(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18; +uint8_t lean_int_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMilliseconds___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toTimestamp(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_millisecond___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_month(lean_object*); +lean_object* l_Std_Time_PlainDate_quarter(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMonthRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1; +static lean_object* l_Std_Time_ZonedDateTime_withMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_time___boxed(lean_object*); +uint8_t l_instDecidableNot___rarg(uint8_t); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDaysSinceUNIXEpoch(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withWeekday(lean_object*, uint8_t); +lean_object* l_Std_Time_PlainDateTime_weekOfMonth(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_year___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMinutes___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMilliseconds(lean_object*, lean_object*); +lean_object* lean_nat_to_int(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__3; +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_time(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addSeconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__6; +lean_object* l_Std_Time_ValidDate_dayOfYear(uint8_t, lean_object*); +lean_object* l_Std_Time_PlainTime_ofNanoseconds(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__4; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_year(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subHours(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_offset(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subWeeks(lean_object*, lean_object*); +lean_object* l_Array_back_x3f___rarg(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeAssumingUTC(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subSeconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22; +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMilliseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_second(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13; +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__1; +lean_object* l_Std_Time_PlainDateTime_toTimestampAssumingUTC(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_hour(lean_object*); +lean_object* l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__8; +lean_object* lean_thunk_get_own(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_era___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDateTime___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedZonedDateTime; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_hour___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subHours___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsClip___boxed(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainTime_toMilliseconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_weekday(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withYearClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsRollOver___boxed(lean_object*, lean_object*); +lean_object* lean_mk_thunk(lean_object*); +lean_object* lean_int_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMonthClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestampWithZone___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addDays(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_rollOver(lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDateTime_withWeekday(lean_object*, uint8_t); +static lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1; +lean_object* l_Std_Time_PlainDate_weekOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_millisecond(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysClip(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfMonth___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsRollOver(lean_object*, lean_object*); +lean_object* l_Std_Time_TimeZone_Transition_timezoneAt(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_convertZoneRules(lean_object*, lean_object*); +lean_object* lean_nat_abs(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nanosecond___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestampWithZone(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp(lean_object*, lean_object*); +lean_object* lean_int_mul(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_second___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__6; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_day___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_dayOfYear___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddDuration___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1; +lean_object* l_Std_Time_PlainTime_addMilliseconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration__1___boxed(lean_object*, lean_object*); +uint8_t l_Std_Time_Year_Offset_era(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_addMilliseconds___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfYear___boxed(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Std_Time_PlainDate_weekday(lean_object*); +lean_object* l_Std_Time_PlainDateTime_addMonthsClip(lean_object*, lean_object*); +uint8_t lean_int_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_dayOfYear(lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsClip___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMinutes(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMilliseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__9; +static lean_object* l_Std_Time_ZonedDateTime_millisecond___closed__1; +lean_object* lean_array_mk(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withWeekday___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withSeconds(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsClip(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subDays___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38; +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___boxed(lean_object*); +lean_object* l_Std_Time_PlainTime_toSeconds(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__5; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_int_add(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_day(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_month___boxed(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_offset___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_quarter(lean_object*); +uint8_t lean_int_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__3; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subWeeks___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_era(lean_object*); +lean_object* l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysRollOver___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subNanoseconds(lean_object*, lean_object*); +lean_object* lean_int_ediv(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +lean_object* lean_int_neg(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsClip___boxed(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17; +static lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_alignedWeekOfMonth___boxed(lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withYearRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_addHours___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddOffset__5; +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21; +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfYear(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddDuration(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addSeconds___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_minute(lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsRollOver(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_minute___boxed(lean_object*); +lean_object* l_Std_Time_PlainDate_addMonthsRollOver(lean_object*, lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___closed__2; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_quarter___boxed(lean_object*); +static lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withNanoseconds(lean_object*, lean_object*); +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1; +static lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubOffset__4; +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDaysSinceUNIXEpoch___boxed(lean_object*); +lean_object* l_Array_get_x3f___rarg(lean_object*, lean_object*); +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(11u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(30u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11; +x_3 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19; +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(23u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = lean_int_sub(x_1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(59u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_3 = lean_int_sub(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33; +x_3 = lean_int_emod(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_3 = lean_int_add(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37; +x_3 = lean_box(x_1); +x_4 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_4, 0, x_3); +lean_ctor_set(x_4, 1, x_2); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37; +x_3 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38; +x_4 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40; +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_instInhabitedZonedDateTime___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedZonedDateTime___closed__1; +x_2 = lean_mk_thunk(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_unchecked("", 0, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = 0; +x_3 = l_Std_Time_instInhabitedZonedDateTime___closed__4; +x_4 = 0; +x_5 = 0; +x_6 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set_uint8(x_6, sizeof(void*)*3, x_2); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 1, x_4); +lean_ctor_set_uint8(x_6, sizeof(void*)*3 + 2, x_5); +return x_6; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Time_instInhabitedZonedDateTime___closed__5; +x_2 = l_Std_Time_instInhabitedZonedDateTime___closed__6; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = l_Std_Time_instInhabitedZonedDateTime___closed__4; +x_3 = 0; +x_4 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set_uint8(x_4, sizeof(void*)*3, x_3); +return x_4; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Std_Time_instInhabitedZonedDateTime___closed__2; +x_2 = l_Std_Time_instInhabitedZonedDateTime___closed__3; +x_3 = l_Std_Time_instInhabitedZonedDateTime___closed__7; +x_4 = l_Std_Time_instInhabitedZonedDateTime___closed__8; +x_5 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_4); +return x_5; +} +} +static lean_object* _init_l_Std_Time_instInhabitedZonedDateTime() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_instInhabitedZonedDateTime___closed__9; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_instInhabitedZonedDateTime___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_instInhabitedZonedDateTime___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_4 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_5); +lean_dec(x_7); +x_9 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_mul(x_5, x_12); +x_14 = lean_int_add(x_11, x_13); +lean_dec(x_13); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +lean_dec(x_4); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_inc(x_1); +x_9 = l_Std_Time_TimeZone_Transition_timezoneAt(x_8, x_1); +lean_dec(x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_9); +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_10); +lean_dec(x_10); +x_3 = x_11; +goto block_7; +} +else +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_3 = x_12; +goto block_7; +} +block_7: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +lean_inc(x_3); +lean_inc(x_1); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_1); +lean_closure_set(x_4, 1, x_3); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_1); +lean_ctor_set(x_6, 2, x_2); +lean_ctor_set(x_6, 3, x_3); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_int_dec_le(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +x_7 = l_Std_Time_PlainTime_toSeconds(x_6); +x_8 = lean_int_add(x_7, x_2); +lean_dec(x_7); +x_9 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_10 = lean_int_ediv(x_8, x_9); +lean_dec(x_8); +x_11 = l_Std_Time_PlainTime_toNanoseconds(x_6); +lean_dec(x_6); +x_12 = lean_int_mul(x_2, x_3); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_ofNanoseconds(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_5, 0); +lean_inc(x_15); +lean_dec(x_5); +x_16 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_15); +x_17 = lean_int_add(x_16, x_10); +lean_dec(x_10); +lean_dec(x_16); +x_18 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_17); +lean_dec(x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_14); +return x_19; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_2 = lean_int_neg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_inc(x_4); +x_5 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_5, 0, x_4); +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +x_7 = lean_unsigned_to_nat(0u); +x_8 = l_Array_findIdx_x3f_loop___rarg(x_5, x_6, x_7); +x_9 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_10 = lean_int_mul(x_4, x_9); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +lean_dec(x_3); +x_12 = lean_int_add(x_10, x_11); +lean_dec(x_11); +lean_dec(x_10); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_26; +lean_dec(x_4); +x_26 = l_Array_back_x3f___rarg(x_6); +lean_dec(x_6); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; +x_27 = lean_ctor_get(x_2, 0); +lean_inc(x_27); +x_13 = x_27; +goto block_25; +} +else +{ +lean_object* x_28; lean_object* x_29; +x_28 = lean_ctor_get(x_26, 0); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_13 = x_29; +goto block_25; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_30 = lean_ctor_get(x_8, 0); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_unsigned_to_nat(1u); +x_32 = lean_nat_sub(x_30, x_31); +x_33 = l_Array_get_x3f___rarg(x_6, x_32); +lean_dec(x_32); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; +lean_dec(x_30); +lean_dec(x_6); +lean_dec(x_4); +x_34 = lean_ctor_get(x_2, 0); +lean_inc(x_34); +x_13 = x_34; +goto block_25; +} +else +{ +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Array_get_x3f___rarg(x_6, x_30); +lean_dec(x_30); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +x_37 = l_Array_back_x3f___rarg(x_6); +lean_dec(x_6); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; +lean_dec(x_35); +lean_dec(x_4); +x_38 = lean_ctor_get(x_2, 0); +lean_inc(x_38); +x_13 = x_38; +goto block_25; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_39 = lean_ctor_get(x_37, 0); +lean_inc(x_39); +lean_dec(x_37); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_35, 1); +lean_inc(x_41); +lean_dec(x_35); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_nat_abs(x_42); +lean_dec(x_42); +x_44 = lean_nat_to_int(x_43); +x_45 = lean_int_sub(x_40, x_44); +lean_dec(x_44); +lean_dec(x_40); +x_46 = lean_int_dec_lt(x_4, x_45); +lean_dec(x_45); +lean_dec(x_4); +if (x_46 == 0) +{ +lean_object* x_47; +lean_dec(x_41); +x_47 = lean_ctor_get(x_39, 1); +lean_inc(x_47); +lean_dec(x_39); +x_13 = x_47; +goto block_25; +} +else +{ +lean_dec(x_39); +x_13 = x_41; +goto block_25; +} +} +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +lean_dec(x_6); +x_48 = lean_ctor_get(x_36, 0); +lean_inc(x_48); +lean_dec(x_36); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_35, 1); +lean_inc(x_50); +lean_dec(x_35); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_nat_abs(x_51); +lean_dec(x_51); +x_53 = lean_nat_to_int(x_52); +x_54 = lean_int_sub(x_49, x_53); +lean_dec(x_53); +lean_dec(x_49); +x_55 = lean_int_dec_lt(x_4, x_54); +lean_dec(x_54); +lean_dec(x_4); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_50); +x_56 = lean_ctor_get(x_48, 1); +lean_inc(x_56); +lean_dec(x_48); +x_13 = x_56; +goto block_25; +} +else +{ +lean_dec(x_48); +x_13 = x_50; +goto block_25; +} +} +} +} +block_25: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_13); +lean_dec(x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_int_neg(x_15); +x_17 = lean_int_mul(x_16, x_9); +lean_dec(x_16); +x_18 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_19 = lean_int_add(x_17, x_18); +lean_dec(x_17); +x_20 = lean_int_add(x_12, x_19); +lean_dec(x_19); +lean_dec(x_12); +x_21 = l_Std_Time_Duration_ofNanoseconds(x_20); +lean_dec(x_20); +lean_inc(x_21); +x_22 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_22, 0, x_21); +lean_closure_set(x_22, 1, x_15); +lean_closure_set(x_22, 2, x_9); +x_23 = lean_mk_thunk(x_22); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +lean_ctor_set(x_24, 2, x_2); +lean_ctor_set(x_24, 3, x_14); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_array_mk(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestampWithZone(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_17; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_7 = 0; +x_8 = 1; +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_3); +x_9 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_5); +lean_ctor_set(x_9, 2, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*3, x_6); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 1, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 2, x_8); +x_10 = l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1; +lean_inc(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +lean_inc(x_1); +x_17 = l_Std_Time_TimeZone_Transition_timezoneAt(x_10, x_1); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; +lean_dec(x_17); +x_18 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_9); +lean_dec(x_9); +x_12 = x_18; +goto block_16; +} +else +{ +lean_object* x_19; +lean_dec(x_9); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_12 = x_19; +goto block_16; +} +block_16: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_12); +lean_inc(x_1); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_12); +x_14 = lean_mk_thunk(x_13); +x_15 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_1); +lean_ctor_set(x_15, 2, x_11); +lean_ctor_set(x_15, 3, x_12); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofTimestampWithZone___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_ofTimestampWithZone(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1; +x_2 = l_Array_back_x3f___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; uint8_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_3 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get_uint8(x_2, sizeof(void*)*3); +x_7 = 0; +x_8 = 1; +lean_inc(x_4); +lean_inc(x_5); +lean_inc(x_3); +x_9 = lean_alloc_ctor(0, 3, 3); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_5); +lean_ctor_set(x_9, 2, x_4); +lean_ctor_set_uint8(x_9, sizeof(void*)*3, x_6); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 1, x_7); +lean_ctor_set_uint8(x_9, sizeof(void*)*3 + 2, x_8); +x_10 = l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1; +lean_inc(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_findIdx_x3f_loop___rarg(x_14, x_10, x_15); +x_17 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_18 = lean_int_mul(x_13, x_17); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_34; +lean_dec(x_13); +x_34 = l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1; +if (lean_obj_tag(x_34) == 0) +{ +x_21 = x_9; +goto block_33; +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_9); +x_35 = lean_ctor_get(x_34, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_21 = x_36; +goto block_33; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_16, 0); +lean_inc(x_37); +lean_dec(x_16); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_sub(x_37, x_38); +x_40 = l_Array_get_x3f___rarg(x_10, x_39); +lean_dec(x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_dec(x_37); +lean_dec(x_13); +x_21 = x_9; +goto block_33; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_Array_get_x3f___rarg(x_10, x_37); +lean_dec(x_37); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +x_43 = l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1; +if (lean_obj_tag(x_43) == 0) +{ +lean_dec(x_41); +lean_dec(x_13); +x_21 = x_9; +goto block_33; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +lean_dec(x_9); +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_41, 1); +lean_inc(x_46); +lean_dec(x_41); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_nat_abs(x_47); +lean_dec(x_47); +x_49 = lean_nat_to_int(x_48); +x_50 = lean_int_sub(x_45, x_49); +lean_dec(x_49); +lean_dec(x_45); +x_51 = lean_int_dec_lt(x_13, x_50); +lean_dec(x_50); +lean_dec(x_13); +if (x_51 == 0) +{ +lean_object* x_52; +lean_dec(x_46); +x_52 = lean_ctor_get(x_44, 1); +lean_inc(x_52); +lean_dec(x_44); +x_21 = x_52; +goto block_33; +} +else +{ +lean_dec(x_44); +x_21 = x_46; +goto block_33; +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +lean_dec(x_9); +x_53 = lean_ctor_get(x_42, 0); +lean_inc(x_53); +lean_dec(x_42); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_41, 1); +lean_inc(x_55); +lean_dec(x_41); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_nat_abs(x_56); +lean_dec(x_56); +x_58 = lean_nat_to_int(x_57); +x_59 = lean_int_sub(x_54, x_58); +lean_dec(x_58); +lean_dec(x_54); +x_60 = lean_int_dec_lt(x_13, x_59); +lean_dec(x_59); +lean_dec(x_13); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_55); +x_61 = lean_ctor_get(x_53, 1); +lean_inc(x_61); +lean_dec(x_53); +x_21 = x_61; +goto block_33; +} +else +{ +lean_dec(x_53); +x_21 = x_55; +goto block_33; +} +} +} +} +block_33: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_int_neg(x_23); +x_25 = lean_int_mul(x_24, x_17); +lean_dec(x_24); +x_26 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_25); +x_28 = lean_int_add(x_20, x_27); +lean_dec(x_27); +lean_dec(x_20); +x_29 = l_Std_Time_Duration_ofNanoseconds(x_28); +lean_dec(x_28); +lean_inc(x_29); +x_30 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_17); +x_31 = lean_mk_thunk(x_30); +x_32 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_11); +lean_ctor_set(x_32, 3, x_22); +return x_32; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toTimestamp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toTimestamp___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_toTimestamp(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_convertZoneRules(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_10; lean_object* x_11; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_4 = x_1; +} else { + lean_dec_ref(x_1); + x_4 = lean_box(0); +} +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_inc(x_3); +x_11 = l_Std_Time_TimeZone_Transition_timezoneAt(x_10, x_3); +lean_dec(x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_11); +x_12 = lean_ctor_get(x_2, 0); +lean_inc(x_12); +x_13 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_12); +lean_dec(x_12); +x_5 = x_13; +goto block_9; +} +else +{ +lean_object* x_14; +x_14 = lean_ctor_get(x_11, 0); +lean_inc(x_14); +lean_dec(x_11); +x_5 = x_14; +goto block_9; +} +block_9: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_inc(x_5); +lean_inc(x_3); +x_6 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_6, 0, x_3); +lean_closure_set(x_6, 1, x_5); +x_7 = lean_mk_thunk(x_6); +if (lean_is_scalar(x_4)) { + x_8 = lean_alloc_ctor(0, 4, 0); +} else { + x_8 = x_4; +} +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_3); +lean_ctor_set(x_8, 2, x_2); +lean_ctor_set(x_8, 3, x_5); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofPlainDateTimeAssumingUTC(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_9; lean_object* x_10; +x_3 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_1); +x_9 = lean_ctor_get(x_2, 1); +lean_inc(x_9); +lean_inc(x_3); +x_10 = l_Std_Time_TimeZone_Transition_timezoneAt(x_9, x_3); +lean_dec(x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_10); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_11); +lean_dec(x_11); +x_4 = x_12; +goto block_8; +} +else +{ +lean_object* x_13; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_4 = x_13; +goto block_8; +} +block_8: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_inc(x_4); +lean_inc(x_3); +x_5 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_5, 0, x_3); +lean_closure_set(x_5, 1, x_4); +x_6 = lean_mk_thunk(x_5); +x_7 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +lean_ctor_set(x_7, 2, x_2); +lean_ctor_set(x_7, 3, x_4); +return x_7; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDateTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toPlainDateTime___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_toPlainDateTime(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDateTime(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 1); +lean_inc(x_2); +x_3 = lean_ctor_get(x_1, 3); +lean_inc(x_3); +lean_dec(x_1); +lean_inc(x_2); +x_4 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_4, 0, x_2); +lean_closure_set(x_4, 1, x_3); +x_5 = lean_mk_thunk(x_4); +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_time(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_time___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_time(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_year(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_year___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_year(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_month(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_month___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_month(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_day(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_day___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_day(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_hour(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_hour___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_hour(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_minute(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 1); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_minute___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_minute(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_second(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +lean_dec(x_4); +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_second___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_second(x_1); +lean_dec(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_millisecond___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_millisecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 3); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_ZonedDateTime_millisecond___closed__1; +x_7 = lean_int_ediv(x_5, x_6); +lean_dec(x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_millisecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_millisecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nanosecond(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 3); +lean_inc(x_5); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_nanosecond___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_nanosecond(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_offset(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_ctor_get(x_1, 3); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_offset___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_offset(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_weekday(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_weekday(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekday___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_ZonedDateTime_weekday(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(4u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(100u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(400u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_dayOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_7 = lean_int_mod(x_5, x_6); +x_8 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_9 = lean_int_dec_eq(x_7, x_8); +lean_dec(x_7); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +x_11 = lean_ctor_get(x_4, 2); +lean_inc(x_11); +lean_dec(x_4); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +if (x_9 == 0) +{ +uint8_t x_13; lean_object* x_14; +lean_dec(x_5); +x_13 = 0; +x_14 = l_Std_Time_ValidDate_dayOfYear(x_13, x_12); +lean_dec(x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; +x_15 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_16 = lean_int_mod(x_5, x_15); +x_17 = lean_int_dec_eq(x_16, x_8); +lean_dec(x_16); +x_18 = l_instDecidableNot___rarg(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_20 = lean_int_mod(x_5, x_19); +lean_dec(x_5); +x_21 = lean_int_dec_eq(x_20, x_8); +lean_dec(x_20); +if (x_21 == 0) +{ +uint8_t x_22; lean_object* x_23; +x_22 = 0; +x_23 = l_Std_Time_ValidDate_dayOfYear(x_22, x_12); +lean_dec(x_12); +return x_23; +} +else +{ +uint8_t x_24; lean_object* x_25; +x_24 = 1; +x_25 = l_Std_Time_ValidDate_dayOfYear(x_24, x_12); +lean_dec(x_12); +return x_25; +} +} +else +{ +uint8_t x_26; lean_object* x_27; +lean_dec(x_5); +x_26 = 1; +x_27 = l_Std_Time_ValidDate_dayOfYear(x_26, x_12); +lean_dec(x_12); +return x_27; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_dayOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_dayOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_weekOfYear(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfYear___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_weekOfYear(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = l_Std_Time_PlainDateTime_weekOfMonth(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_weekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_weekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_alignedWeekOfMonth(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_alignedWeekOfMonth(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_alignedWeekOfMonth___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_alignedWeekOfMonth(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_quarter(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_quarter(x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_quarter___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_quarter(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_19; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_10 = lean_int_add(x_9, x_2); +lean_dec(x_9); +x_11 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_18 = lean_ctor_get(x_4, 1); +lean_inc(x_18); +lean_inc(x_12); +x_19 = l_Std_Time_TimeZone_Transition_timezoneAt(x_18, x_12); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_19); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); +x_21 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_20); +lean_dec(x_20); +x_13 = x_21; +goto block_17; +} +else +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_13 = x_22; +goto block_17; +} +block_17: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_inc(x_13); +lean_inc(x_12); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = lean_mk_thunk(x_14); +if (lean_is_scalar(x_5)) { + x_16 = lean_alloc_ctor(0, 4, 0); +} else { + x_16 = x_5; +} +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_4); +lean_ctor_set(x_16, 3, x_13); +return x_16; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_35; lean_object* x_36; +x_23 = lean_ctor_get(x_6, 0); +x_24 = lean_ctor_get(x_6, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_6); +x_25 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_23); +x_26 = lean_int_add(x_25, x_2); +lean_dec(x_25); +x_27 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +x_29 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_28); +x_35 = lean_ctor_get(x_4, 1); +lean_inc(x_35); +lean_inc(x_29); +x_36 = l_Std_Time_TimeZone_Transition_timezoneAt(x_35, x_29); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_36); +x_37 = lean_ctor_get(x_4, 0); +lean_inc(x_37); +x_38 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_37); +lean_dec(x_37); +x_30 = x_38; +goto block_34; +} +else +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_30 = x_39; +goto block_34; +} +block_34: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_inc(x_30); +lean_inc(x_29); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_31, 0, x_29); +lean_closure_set(x_31, 1, x_30); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_30); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subDays(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_11 = lean_int_add(x_10, x_9); +lean_dec(x_9); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_11); +lean_dec(x_11); +lean_ctor_set(x_6, 0, x_12); +x_13 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +lean_inc(x_13); +x_20 = l_Std_Time_TimeZone_Transition_timezoneAt(x_19, x_13); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_20); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_14 = x_22; +goto block_18; +} +else +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_dec(x_20); +x_14 = x_23; +goto block_18; +} +block_18: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_14); +lean_inc(x_13); +x_15 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = lean_mk_thunk(x_15); +if (lean_is_scalar(x_5)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_5; +} +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +lean_ctor_set(x_17, 2, x_4); +lean_ctor_set(x_17, 3, x_14); +return x_17; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_37; lean_object* x_38; +x_24 = lean_ctor_get(x_6, 0); +x_25 = lean_ctor_get(x_6, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_6); +x_26 = lean_int_neg(x_2); +x_27 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_24); +x_28 = lean_int_add(x_27, x_26); +lean_dec(x_26); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +x_31 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_30); +x_37 = lean_ctor_get(x_4, 1); +lean_inc(x_37); +lean_inc(x_31); +x_38 = l_Std_Time_TimeZone_Transition_timezoneAt(x_37, x_31); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_38); +x_39 = lean_ctor_get(x_4, 0); +lean_inc(x_39); +x_40 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_39); +lean_dec(x_39); +x_32 = x_40; +goto block_36; +} +else +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_32 = x_41; +goto block_36; +} +block_36: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_inc(x_32); +lean_inc(x_31); +x_33 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = lean_mk_thunk(x_33); +if (lean_is_scalar(x_5)) { + x_35 = lean_alloc_ctor(0, 4, 0); +} else { + x_35 = x_5; +} +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +lean_ctor_set(x_35, 2, x_4); +lean_ctor_set(x_35, 3, x_32); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subDays___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subDays(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_addWeeks___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(7u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_20; lean_object* x_21; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_10 = l_Std_Time_ZonedDateTime_addWeeks___closed__1; +x_11 = lean_int_mul(x_2, x_10); +x_12 = lean_int_add(x_9, x_11); +lean_dec(x_11); +lean_dec(x_9); +x_13 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_12); +lean_dec(x_12); +lean_ctor_set(x_6, 0, x_13); +x_14 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_20 = lean_ctor_get(x_4, 1); +lean_inc(x_20); +lean_inc(x_14); +x_21 = l_Std_Time_TimeZone_Transition_timezoneAt(x_20, x_14); +lean_dec(x_20); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_21); +x_22 = lean_ctor_get(x_4, 0); +lean_inc(x_22); +x_23 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_22); +lean_dec(x_22); +x_15 = x_23; +goto block_19; +} +else +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_21, 0); +lean_inc(x_24); +lean_dec(x_21); +x_15 = x_24; +goto block_19; +} +block_19: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_inc(x_15); +lean_inc(x_14); +x_16 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = lean_mk_thunk(x_16); +if (lean_is_scalar(x_5)) { + x_18 = lean_alloc_ctor(0, 4, 0); +} else { + x_18 = x_5; +} +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_14); +lean_ctor_set(x_18, 2, x_4); +lean_ctor_set(x_18, 3, x_15); +return x_18; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_39; lean_object* x_40; +x_25 = lean_ctor_get(x_6, 0); +x_26 = lean_ctor_get(x_6, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_6); +x_27 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_25); +x_28 = l_Std_Time_ZonedDateTime_addWeeks___closed__1; +x_29 = lean_int_mul(x_2, x_28); +x_30 = lean_int_add(x_27, x_29); +lean_dec(x_29); +lean_dec(x_27); +x_31 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_30); +lean_dec(x_30); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_26); +x_33 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_32); +x_39 = lean_ctor_get(x_4, 1); +lean_inc(x_39); +lean_inc(x_33); +x_40 = l_Std_Time_TimeZone_Transition_timezoneAt(x_39, x_33); +lean_dec(x_39); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; lean_object* x_42; +lean_dec(x_40); +x_41 = lean_ctor_get(x_4, 0); +lean_inc(x_41); +x_42 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_41); +lean_dec(x_41); +x_34 = x_42; +goto block_38; +} +else +{ +lean_object* x_43; +x_43 = lean_ctor_get(x_40, 0); +lean_inc(x_43); +lean_dec(x_40); +x_34 = x_43; +goto block_38; +} +block_38: +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_inc(x_34); +lean_inc(x_33); +x_35 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); +x_36 = lean_mk_thunk(x_35); +if (lean_is_scalar(x_5)) { + x_37 = lean_alloc_ctor(0, 4, 0); +} else { + x_37 = x_5; +} +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_33); +lean_ctor_set(x_37, 2, x_4); +lean_ctor_set(x_37, 3, x_34); +return x_37; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subWeeks(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_21; lean_object* x_22; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_8); +x_11 = l_Std_Time_ZonedDateTime_addWeeks___closed__1; +x_12 = lean_int_mul(x_9, x_11); +lean_dec(x_9); +x_13 = lean_int_add(x_10, x_12); +lean_dec(x_12); +lean_dec(x_10); +x_14 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_13); +lean_dec(x_13); +lean_ctor_set(x_6, 0, x_14); +x_15 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_21 = lean_ctor_get(x_4, 1); +lean_inc(x_21); +lean_inc(x_15); +x_22 = l_Std_Time_TimeZone_Transition_timezoneAt(x_21, x_15); +lean_dec(x_21); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_22); +x_23 = lean_ctor_get(x_4, 0); +lean_inc(x_23); +x_24 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_23); +lean_dec(x_23); +x_16 = x_24; +goto block_20; +} +else +{ +lean_object* x_25; +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +x_16 = x_25; +goto block_20; +} +block_20: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_inc(x_16); +lean_inc(x_15); +x_17 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_17, 0, x_15); +lean_closure_set(x_17, 1, x_16); +x_18 = lean_mk_thunk(x_17); +if (lean_is_scalar(x_5)) { + x_19 = lean_alloc_ctor(0, 4, 0); +} else { + x_19 = x_5; +} +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_15); +lean_ctor_set(x_19, 2, x_4); +lean_ctor_set(x_19, 3, x_16); +return x_19; +} +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_41; lean_object* x_42; +x_26 = lean_ctor_get(x_6, 0); +x_27 = lean_ctor_get(x_6, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_6); +x_28 = lean_int_neg(x_2); +x_29 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_26); +x_30 = l_Std_Time_ZonedDateTime_addWeeks___closed__1; +x_31 = lean_int_mul(x_28, x_30); +lean_dec(x_28); +x_32 = lean_int_add(x_29, x_31); +lean_dec(x_31); +lean_dec(x_29); +x_33 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_32); +lean_dec(x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_27); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_34); +x_41 = lean_ctor_get(x_4, 1); +lean_inc(x_41); +lean_inc(x_35); +x_42 = l_Std_Time_TimeZone_Transition_timezoneAt(x_41, x_35); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_42); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_44 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_43); +lean_dec(x_43); +x_36 = x_44; +goto block_40; +} +else +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_42, 0); +lean_inc(x_45); +lean_dec(x_42); +x_36 = x_45; +goto block_40; +} +block_40: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_inc(x_36); +lean_inc(x_35); +x_37 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_37, 0, x_35); +lean_closure_set(x_37, 1, x_36); +x_38 = lean_mk_thunk(x_37); +if (lean_is_scalar(x_5)) { + x_39 = lean_alloc_ctor(0, 4, 0); +} else { + x_39 = x_5; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_35); +lean_ctor_set(x_39, 2, x_4); +lean_ctor_set(x_39, 3, x_36); +return x_39; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subWeeks___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subWeeks(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = l_Std_Time_PlainDateTime_addMonthsClip(x_6, x_2); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_7); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_8); +x_15 = l_Std_Time_TimeZone_Transition_timezoneAt(x_14, x_8); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_15); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +x_17 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_16); +lean_dec(x_16); +x_9 = x_17; +goto block_13; +} +else +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +x_9 = x_18; +goto block_13; +} +block_13: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_inc(x_9); +lean_inc(x_8); +x_10 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_10, 0, x_8); +lean_closure_set(x_10, 1, x_9); +x_11 = lean_mk_thunk(x_10); +if (lean_is_scalar(x_5)) { + x_12 = lean_alloc_ctor(0, 4, 0); +} else { + x_12 = x_5; +} +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +lean_ctor_set(x_12, 2, x_4); +lean_ctor_set(x_12, 3, x_9); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_17; lean_object* x_18; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_addMonthsClip(x_8, x_9); +lean_dec(x_9); +lean_ctor_set(x_6, 0, x_10); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_11); +x_18 = l_Std_Time_TimeZone_Transition_timezoneAt(x_17, x_11); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_18); +x_19 = lean_ctor_get(x_4, 0); +lean_inc(x_19); +x_20 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_19); +lean_dec(x_19); +x_12 = x_20; +goto block_16; +} +else +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_18, 0); +lean_inc(x_21); +lean_dec(x_18); +x_12 = x_21; +goto block_16; +} +block_16: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_12); +lean_inc(x_11); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = lean_mk_thunk(x_13); +if (lean_is_scalar(x_5)) { + x_15 = lean_alloc_ctor(0, 4, 0); +} else { + x_15 = x_5; +} +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +lean_ctor_set(x_15, 2, x_4); +lean_ctor_set(x_15, 3, x_12); +return x_15; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_33; lean_object* x_34; +x_22 = lean_ctor_get(x_6, 0); +x_23 = lean_ctor_get(x_6, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_6); +x_24 = lean_int_neg(x_2); +x_25 = l_Std_Time_PlainDate_addMonthsClip(x_22, x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_33 = lean_ctor_get(x_4, 1); +lean_inc(x_33); +lean_inc(x_27); +x_34 = l_Std_Time_TimeZone_Transition_timezoneAt(x_33, x_27); +lean_dec(x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_34); +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_36 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_35); +lean_dec(x_35); +x_28 = x_36; +goto block_32; +} +else +{ +lean_object* x_37; +x_37 = lean_ctor_get(x_34, 0); +lean_inc(x_37); +lean_dec(x_34); +x_28 = x_37; +goto block_32; +} +block_32: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_inc(x_28); +lean_inc(x_27); +x_29 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_29, 0, x_27); +lean_closure_set(x_29, 1, x_28); +x_30 = lean_mk_thunk(x_29); +if (lean_is_scalar(x_5)) { + x_31 = lean_alloc_ctor(0, 4, 0); +} else { + x_31 = x_5; +} +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +lean_ctor_set(x_31, 2, x_4); +lean_ctor_set(x_31, 3, x_28); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subMonthsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_14; lean_object* x_15; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = l_Std_Time_PlainDateTime_addMonthsRollOver(x_6, x_2); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_7); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_8); +x_15 = l_Std_Time_TimeZone_Transition_timezoneAt(x_14, x_8); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; +lean_dec(x_15); +x_16 = lean_ctor_get(x_4, 0); +lean_inc(x_16); +x_17 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_16); +lean_dec(x_16); +x_9 = x_17; +goto block_13; +} +else +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(x_15); +x_9 = x_18; +goto block_13; +} +block_13: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_inc(x_9); +lean_inc(x_8); +x_10 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_10, 0, x_8); +lean_closure_set(x_10, 1, x_9); +x_11 = lean_mk_thunk(x_10); +if (lean_is_scalar(x_5)) { + x_12 = lean_alloc_ctor(0, 4, 0); +} else { + x_12 = x_5; +} +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +lean_ctor_set(x_12, 2, x_4); +lean_ctor_set(x_12, 3, x_9); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_17; lean_object* x_18; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_int_neg(x_2); +x_10 = l_Std_Time_PlainDate_addMonthsRollOver(x_8, x_9); +lean_dec(x_9); +lean_ctor_set(x_6, 0, x_10); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_inc(x_11); +x_18 = l_Std_Time_TimeZone_Transition_timezoneAt(x_17, x_11); +lean_dec(x_17); +if (lean_obj_tag(x_18) == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_18); +x_19 = lean_ctor_get(x_4, 0); +lean_inc(x_19); +x_20 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_19); +lean_dec(x_19); +x_12 = x_20; +goto block_16; +} +else +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_18, 0); +lean_inc(x_21); +lean_dec(x_18); +x_12 = x_21; +goto block_16; +} +block_16: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_12); +lean_inc(x_11); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = lean_mk_thunk(x_13); +if (lean_is_scalar(x_5)) { + x_15 = lean_alloc_ctor(0, 4, 0); +} else { + x_15 = x_5; +} +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +lean_ctor_set(x_15, 2, x_4); +lean_ctor_set(x_15, 3, x_12); +return x_15; +} +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_33; lean_object* x_34; +x_22 = lean_ctor_get(x_6, 0); +x_23 = lean_ctor_get(x_6, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_6); +x_24 = lean_int_neg(x_2); +x_25 = l_Std_Time_PlainDate_addMonthsRollOver(x_22, x_24); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +x_27 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_26); +x_33 = lean_ctor_get(x_4, 1); +lean_inc(x_33); +lean_inc(x_27); +x_34 = l_Std_Time_TimeZone_Transition_timezoneAt(x_33, x_27); +lean_dec(x_33); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_34); +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_36 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_35); +lean_dec(x_35); +x_28 = x_36; +goto block_32; +} +else +{ +lean_object* x_37; +x_37 = lean_ctor_get(x_34, 0); +lean_inc(x_37); +lean_dec(x_34); +x_28 = x_37; +goto block_32; +} +block_32: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_inc(x_28); +lean_inc(x_27); +x_29 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_29, 0, x_27); +lean_closure_set(x_29, 1, x_28); +x_30 = lean_mk_thunk(x_29); +if (lean_is_scalar(x_5)) { + x_31 = lean_alloc_ctor(0, 4, 0); +} else { + x_31 = x_5; +} +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +lean_ctor_set(x_31, 2, x_4); +lean_ctor_set(x_31, 3, x_28); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMonthsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subMonthsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(12u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_19; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = l_Std_Time_PlainDate_addMonthsRollOver(x_8, x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_18 = lean_ctor_get(x_4, 1); +lean_inc(x_18); +lean_inc(x_12); +x_19 = l_Std_Time_TimeZone_Transition_timezoneAt(x_18, x_12); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_19); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); +x_21 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_20); +lean_dec(x_20); +x_13 = x_21; +goto block_17; +} +else +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_13 = x_22; +goto block_17; +} +block_17: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_inc(x_13); +lean_inc(x_12); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = lean_mk_thunk(x_14); +if (lean_is_scalar(x_5)) { + x_16 = lean_alloc_ctor(0, 4, 0); +} else { + x_16 = x_5; +} +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_4); +lean_ctor_set(x_16, 3, x_13); +return x_16; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_35; lean_object* x_36; +x_23 = lean_ctor_get(x_6, 0); +x_24 = lean_ctor_get(x_6, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_6); +x_25 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_26 = lean_int_mul(x_2, x_25); +x_27 = l_Std_Time_PlainDate_addMonthsRollOver(x_23, x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +x_29 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_28); +x_35 = lean_ctor_get(x_4, 1); +lean_inc(x_35); +lean_inc(x_29); +x_36 = l_Std_Time_TimeZone_Transition_timezoneAt(x_35, x_29); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_36); +x_37 = lean_ctor_get(x_4, 0); +lean_inc(x_37); +x_38 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_37); +lean_dec(x_37); +x_30 = x_38; +goto block_34; +} +else +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_30 = x_39; +goto block_34; +} +block_34: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_inc(x_30); +lean_inc(x_29); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_31, 0, x_29); +lean_closure_set(x_31, 1, x_30); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_30); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_18; lean_object* x_19; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = l_Std_Time_PlainDate_addMonthsClip(x_8, x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_18 = lean_ctor_get(x_4, 1); +lean_inc(x_18); +lean_inc(x_12); +x_19 = l_Std_Time_TimeZone_Transition_timezoneAt(x_18, x_12); +lean_dec(x_18); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_19); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); +x_21 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_20); +lean_dec(x_20); +x_13 = x_21; +goto block_17; +} +else +{ +lean_object* x_22; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_13 = x_22; +goto block_17; +} +block_17: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_inc(x_13); +lean_inc(x_12); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = lean_mk_thunk(x_14); +if (lean_is_scalar(x_5)) { + x_16 = lean_alloc_ctor(0, 4, 0); +} else { + x_16 = x_5; +} +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_4); +lean_ctor_set(x_16, 3, x_13); +return x_16; +} +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_35; lean_object* x_36; +x_23 = lean_ctor_get(x_6, 0); +x_24 = lean_ctor_get(x_6, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_6); +x_25 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_26 = lean_int_mul(x_2, x_25); +x_27 = l_Std_Time_PlainDate_addMonthsClip(x_23, x_26); +lean_dec(x_26); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +x_29 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_28); +x_35 = lean_ctor_get(x_4, 1); +lean_inc(x_35); +lean_inc(x_29); +x_36 = l_Std_Time_TimeZone_Transition_timezoneAt(x_35, x_29); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; +lean_dec(x_36); +x_37 = lean_ctor_get(x_4, 0); +lean_inc(x_37); +x_38 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_37); +lean_dec(x_37); +x_30 = x_38; +goto block_34; +} +else +{ +lean_object* x_39; +x_39 = lean_ctor_get(x_36, 0); +lean_inc(x_39); +lean_dec(x_36); +x_30 = x_39; +goto block_34; +} +block_34: +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_inc(x_30); +lean_inc(x_29); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_31, 0, x_29); +lean_closure_set(x_31, 1, x_30); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_30); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_neg(x_10); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_addMonthsClip(x_8, x_11); +lean_dec(x_11); +lean_ctor_set(x_6, 0, x_12); +x_13 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +lean_inc(x_13); +x_20 = l_Std_Time_TimeZone_Transition_timezoneAt(x_19, x_13); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_20); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_14 = x_22; +goto block_18; +} +else +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_dec(x_20); +x_14 = x_23; +goto block_18; +} +block_18: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_14); +lean_inc(x_13); +x_15 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = lean_mk_thunk(x_15); +if (lean_is_scalar(x_5)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_5; +} +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +lean_ctor_set(x_17, 2, x_4); +lean_ctor_set(x_17, 3, x_14); +return x_17; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_37; lean_object* x_38; +x_24 = lean_ctor_get(x_6, 0); +x_25 = lean_ctor_get(x_6, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_6); +x_26 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_27 = lean_int_mul(x_2, x_26); +x_28 = lean_int_neg(x_27); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_addMonthsClip(x_24, x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +x_31 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_30); +x_37 = lean_ctor_get(x_4, 1); +lean_inc(x_37); +lean_inc(x_31); +x_38 = l_Std_Time_TimeZone_Transition_timezoneAt(x_37, x_31); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_38); +x_39 = lean_ctor_get(x_4, 0); +lean_inc(x_39); +x_40 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_39); +lean_dec(x_39); +x_32 = x_40; +goto block_36; +} +else +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_32 = x_41; +goto block_36; +} +block_36: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_inc(x_32); +lean_inc(x_31); +x_33 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = lean_mk_thunk(x_33); +if (lean_is_scalar(x_5)) { + x_35 = lean_alloc_ctor(0, 4, 0); +} else { + x_35 = x_5; +} +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +lean_ctor_set(x_35, 2, x_4); +lean_ctor_set(x_35, 3, x_32); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsClip___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subYearsClip(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_19; lean_object* x_20; +x_8 = lean_ctor_get(x_6, 0); +x_9 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_neg(x_10); +lean_dec(x_10); +x_12 = l_Std_Time_PlainDate_addMonthsRollOver(x_8, x_11); +lean_dec(x_11); +lean_ctor_set(x_6, 0, x_12); +x_13 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +lean_inc(x_13); +x_20 = l_Std_Time_TimeZone_Transition_timezoneAt(x_19, x_13); +lean_dec(x_19); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_20); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_14 = x_22; +goto block_18; +} +else +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_dec(x_20); +x_14 = x_23; +goto block_18; +} +block_18: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_14); +lean_inc(x_13); +x_15 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_15, 0, x_13); +lean_closure_set(x_15, 1, x_14); +x_16 = lean_mk_thunk(x_15); +if (lean_is_scalar(x_5)) { + x_17 = lean_alloc_ctor(0, 4, 0); +} else { + x_17 = x_5; +} +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +lean_ctor_set(x_17, 2, x_4); +lean_ctor_set(x_17, 3, x_14); +return x_17; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_37; lean_object* x_38; +x_24 = lean_ctor_get(x_6, 0); +x_25 = lean_ctor_get(x_6, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_6); +x_26 = l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1; +x_27 = lean_int_mul(x_2, x_26); +x_28 = lean_int_neg(x_27); +lean_dec(x_27); +x_29 = l_Std_Time_PlainDate_addMonthsRollOver(x_24, x_28); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_25); +x_31 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_30); +x_37 = lean_ctor_get(x_4, 1); +lean_inc(x_37); +lean_inc(x_31); +x_38 = l_Std_Time_TimeZone_Transition_timezoneAt(x_37, x_31); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_38); +x_39 = lean_ctor_get(x_4, 0); +lean_inc(x_39); +x_40 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_39); +lean_dec(x_39); +x_32 = x_40; +goto block_36; +} +else +{ +lean_object* x_41; +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +lean_dec(x_38); +x_32 = x_41; +goto block_36; +} +block_36: +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_inc(x_32); +lean_inc(x_31); +x_33 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_33, 0, x_31); +lean_closure_set(x_33, 1, x_32); +x_34 = lean_mk_thunk(x_33); +if (lean_is_scalar(x_5)) { + x_35 = lean_alloc_ctor(0, 4, 0); +} else { + x_35 = x_5; +} +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_31); +lean_ctor_set(x_35, 2, x_4); +lean_ctor_set(x_35, 3, x_32); +return x_35; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subYearsRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subYearsRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_7 = lean_ctor_get(x_2, 0); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = l_Std_Time_PlainTime_toSeconds(x_8); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = lean_int_ediv(x_10, x_3); +lean_dec(x_10); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_13 = lean_int_mul(x_7, x_4); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_11); +lean_dec(x_11); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_addHours___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(3600u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_30; lean_object* x_31; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = l_Std_Time_ZonedDateTime_addHours___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_8); +x_12 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_15 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_10, x_15); +lean_dec(x_10); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +lean_dec(x_6); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_23); +x_30 = lean_ctor_get(x_4, 1); +lean_inc(x_30); +lean_inc(x_24); +x_31 = l_Std_Time_TimeZone_Transition_timezoneAt(x_30, x_24); +lean_dec(x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_31); +x_32 = lean_ctor_get(x_4, 0); +lean_inc(x_32); +x_33 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_32); +lean_dec(x_32); +x_25 = x_33; +goto block_29; +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_25 = x_34; +goto block_29; +} +block_29: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_inc(x_25); +lean_inc(x_24); +x_26 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_26, 0, x_24); +lean_closure_set(x_26, 1, x_25); +lean_closure_set(x_26, 2, x_12); +lean_closure_set(x_26, 3, x_15); +x_27 = lean_mk_thunk(x_26); +if (lean_is_scalar(x_5)) { + x_28 = lean_alloc_ctor(0, 4, 0); +} else { + x_28 = x_5; +} +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +lean_ctor_set(x_28, 2, x_4); +lean_ctor_set(x_28, 3, x_25); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l_Std_Time_ZonedDateTime_addHours___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addHours(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_31; lean_object* x_32; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_int_neg(x_2); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = l_Std_Time_PlainTime_toSeconds(x_8); +x_10 = l_Std_Time_ZonedDateTime_addHours___closed__1; +x_11 = lean_int_mul(x_7, x_10); +lean_dec(x_7); +x_12 = lean_int_add(x_9, x_11); +lean_dec(x_9); +x_13 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_14 = lean_int_ediv(x_12, x_13); +lean_dec(x_12); +x_15 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_16 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_17 = lean_int_mul(x_11, x_16); +lean_dec(x_11); +x_18 = lean_int_add(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +x_19 = l_Std_Time_PlainTime_ofNanoseconds(x_18); +lean_dec(x_18); +x_20 = lean_ctor_get(x_6, 0); +lean_inc(x_20); +lean_dec(x_6); +x_21 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_20); +x_22 = lean_int_add(x_21, x_14); +lean_dec(x_14); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +x_25 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_24); +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); +lean_inc(x_25); +x_32 = l_Std_Time_TimeZone_Transition_timezoneAt(x_31, x_25); +lean_dec(x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_32); +x_33 = lean_ctor_get(x_4, 0); +lean_inc(x_33); +x_34 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_33); +lean_dec(x_33); +x_26 = x_34; +goto block_30; +} +else +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_26 = x_35; +goto block_30; +} +block_30: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_inc(x_26); +lean_inc(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_27, 0, x_25); +lean_closure_set(x_27, 1, x_26); +lean_closure_set(x_27, 2, x_13); +lean_closure_set(x_27, 3, x_16); +x_28 = lean_mk_thunk(x_27); +if (lean_is_scalar(x_5)) { + x_29 = lean_alloc_ctor(0, 4, 0); +} else { + x_29 = x_5; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +lean_ctor_set(x_29, 2, x_4); +lean_ctor_set(x_29, 3, x_26); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subHours___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subHours(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_addMinutes___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(60u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_30; lean_object* x_31; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = l_Std_Time_ZonedDateTime_addMinutes___closed__1; +x_10 = lean_int_mul(x_2, x_9); +x_11 = lean_int_add(x_8, x_10); +lean_dec(x_8); +x_12 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_13 = lean_int_ediv(x_11, x_12); +lean_dec(x_11); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_15 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_16 = lean_int_mul(x_10, x_15); +lean_dec(x_10); +x_17 = lean_int_add(x_14, x_16); +lean_dec(x_16); +lean_dec(x_14); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = lean_ctor_get(x_6, 0); +lean_inc(x_19); +lean_dec(x_6); +x_20 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_19); +x_21 = lean_int_add(x_20, x_13); +lean_dec(x_13); +lean_dec(x_20); +x_22 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_21); +lean_dec(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_23); +x_30 = lean_ctor_get(x_4, 1); +lean_inc(x_30); +lean_inc(x_24); +x_31 = l_Std_Time_TimeZone_Transition_timezoneAt(x_30, x_24); +lean_dec(x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_31); +x_32 = lean_ctor_get(x_4, 0); +lean_inc(x_32); +x_33 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_32); +lean_dec(x_32); +x_25 = x_33; +goto block_29; +} +else +{ +lean_object* x_34; +x_34 = lean_ctor_get(x_31, 0); +lean_inc(x_34); +lean_dec(x_31); +x_25 = x_34; +goto block_29; +} +block_29: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_inc(x_25); +lean_inc(x_24); +x_26 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_26, 0, x_24); +lean_closure_set(x_26, 1, x_25); +lean_closure_set(x_26, 2, x_12); +lean_closure_set(x_26, 3, x_15); +x_27 = lean_mk_thunk(x_26); +if (lean_is_scalar(x_5)) { + x_28 = lean_alloc_ctor(0, 4, 0); +} else { + x_28 = x_5; +} +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_24); +lean_ctor_set(x_28, 2, x_4); +lean_ctor_set(x_28, 3, x_25); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addMinutes(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_31; lean_object* x_32; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_int_neg(x_2); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = l_Std_Time_PlainTime_toSeconds(x_8); +x_10 = l_Std_Time_ZonedDateTime_addMinutes___closed__1; +x_11 = lean_int_mul(x_7, x_10); +lean_dec(x_7); +x_12 = lean_int_add(x_9, x_11); +lean_dec(x_9); +x_13 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_14 = lean_int_ediv(x_12, x_13); +lean_dec(x_12); +x_15 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_16 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_17 = lean_int_mul(x_11, x_16); +lean_dec(x_11); +x_18 = lean_int_add(x_15, x_17); +lean_dec(x_17); +lean_dec(x_15); +x_19 = l_Std_Time_PlainTime_ofNanoseconds(x_18); +lean_dec(x_18); +x_20 = lean_ctor_get(x_6, 0); +lean_inc(x_20); +lean_dec(x_6); +x_21 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_20); +x_22 = lean_int_add(x_21, x_14); +lean_dec(x_14); +lean_dec(x_21); +x_23 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_22); +lean_dec(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +x_25 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_24); +x_31 = lean_ctor_get(x_4, 1); +lean_inc(x_31); +lean_inc(x_25); +x_32 = l_Std_Time_TimeZone_Transition_timezoneAt(x_31, x_25); +lean_dec(x_31); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_32); +x_33 = lean_ctor_get(x_4, 0); +lean_inc(x_33); +x_34 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_33); +lean_dec(x_33); +x_26 = x_34; +goto block_30; +} +else +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_32, 0); +lean_inc(x_35); +lean_dec(x_32); +x_26 = x_35; +goto block_30; +} +block_30: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_inc(x_26); +lean_inc(x_25); +x_27 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_27, 0, x_25); +lean_closure_set(x_27, 1, x_26); +lean_closure_set(x_27, 2, x_13); +lean_closure_set(x_27, 3, x_16); +x_28 = lean_mk_thunk(x_27); +if (lean_is_scalar(x_5)) { + x_29 = lean_alloc_ctor(0, 4, 0); +} else { + x_29 = x_5; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +lean_ctor_set(x_29, 2, x_4); +lean_ctor_set(x_29, 3, x_26); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMinutes___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subMinutes(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_addMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(86400000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_24; lean_object* x_25; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toMilliseconds(x_7); +x_9 = lean_int_add(x_8, x_2); +lean_dec(x_8); +x_10 = l_Std_Time_ZonedDateTime_addMilliseconds___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_addMilliseconds(x_7, x_2); +lean_dec(x_7); +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +lean_dec(x_6); +x_14 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_13); +x_15 = lean_int_add(x_14, x_11); +lean_dec(x_11); +lean_dec(x_14); +x_16 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_15); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_12); +x_18 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_17); +x_24 = lean_ctor_get(x_4, 1); +lean_inc(x_24); +lean_inc(x_18); +x_25 = l_Std_Time_TimeZone_Transition_timezoneAt(x_24, x_18); +lean_dec(x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_25); +x_26 = lean_ctor_get(x_4, 0); +lean_inc(x_26); +x_27 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_26); +lean_dec(x_26); +x_19 = x_27; +goto block_23; +} +else +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_25, 0); +lean_inc(x_28); +lean_dec(x_25); +x_19 = x_28; +goto block_23; +} +block_23: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_inc(x_19); +lean_inc(x_18); +x_20 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_20, 0, x_18); +lean_closure_set(x_20, 1, x_19); +x_21 = lean_mk_thunk(x_20); +if (lean_is_scalar(x_5)) { + x_22 = lean_alloc_ctor(0, 4, 0); +} else { + x_22 = x_5; +} +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_18); +lean_ctor_set(x_22, 2, x_4); +lean_ctor_set(x_22, 3, x_19); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_25; lean_object* x_26; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_int_neg(x_2); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = l_Std_Time_PlainTime_toMilliseconds(x_8); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = l_Std_Time_ZonedDateTime_addMilliseconds___closed__1; +x_12 = lean_int_ediv(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_addMilliseconds(x_8, x_7); +lean_dec(x_7); +lean_dec(x_8); +x_14 = lean_ctor_get(x_6, 0); +lean_inc(x_14); +lean_dec(x_6); +x_15 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_14); +x_16 = lean_int_add(x_15, x_12); +lean_dec(x_12); +lean_dec(x_15); +x_17 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_16); +lean_dec(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_13); +x_19 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_18); +x_25 = lean_ctor_get(x_4, 1); +lean_inc(x_25); +lean_inc(x_19); +x_26 = l_Std_Time_TimeZone_Transition_timezoneAt(x_25, x_19); +lean_dec(x_25); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; +lean_dec(x_26); +x_27 = lean_ctor_get(x_4, 0); +lean_inc(x_27); +x_28 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_27); +lean_dec(x_27); +x_20 = x_28; +goto block_24; +} +else +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_26, 0); +lean_inc(x_29); +lean_dec(x_26); +x_20 = x_29; +goto block_24; +} +block_24: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_inc(x_20); +lean_inc(x_19); +x_21 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___boxed), 3, 2); +lean_closure_set(x_21, 0, x_19); +lean_closure_set(x_21, 1, x_20); +x_22 = lean_mk_thunk(x_21); +if (lean_is_scalar(x_5)) { + x_23 = lean_alloc_ctor(0, 4, 0); +} else { + x_23 = x_5; +} +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_19); +lean_ctor_set(x_23, 2, x_4); +lean_ctor_set(x_23, 3, x_20); +return x_23; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_28; lean_object* x_29; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = lean_int_add(x_8, x_2); +lean_dec(x_8); +x_10 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_13 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_14 = lean_int_mul(x_2, x_13); +x_15 = lean_int_add(x_12, x_14); +lean_dec(x_14); +lean_dec(x_12); +x_16 = l_Std_Time_PlainTime_ofNanoseconds(x_15); +lean_dec(x_15); +x_17 = lean_ctor_get(x_6, 0); +lean_inc(x_17); +lean_dec(x_6); +x_18 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_17); +x_19 = lean_int_add(x_18, x_11); +lean_dec(x_11); +lean_dec(x_18); +x_20 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_19); +lean_dec(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_16); +x_22 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_21); +x_28 = lean_ctor_get(x_4, 1); +lean_inc(x_28); +lean_inc(x_22); +x_29 = l_Std_Time_TimeZone_Transition_timezoneAt(x_28, x_22); +lean_dec(x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_29); +x_30 = lean_ctor_get(x_4, 0); +lean_inc(x_30); +x_31 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_30); +lean_dec(x_30); +x_23 = x_31; +goto block_27; +} +else +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_29, 0); +lean_inc(x_32); +lean_dec(x_29); +x_23 = x_32; +goto block_27; +} +block_27: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_inc(x_23); +lean_inc(x_22); +x_24 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_23); +lean_closure_set(x_24, 2, x_10); +lean_closure_set(x_24, 3, x_13); +x_25 = lean_mk_thunk(x_24); +if (lean_is_scalar(x_5)) { + x_26 = lean_alloc_ctor(0, 4, 0); +} else { + x_26 = x_5; +} +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +lean_ctor_set(x_26, 2, x_4); +lean_ctor_set(x_26, 3, x_23); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addSeconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_29; lean_object* x_30; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_int_neg(x_2); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +x_9 = l_Std_Time_PlainTime_toSeconds(x_8); +x_10 = lean_int_add(x_9, x_7); +lean_dec(x_9); +x_11 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_12 = lean_int_ediv(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_14 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_15 = lean_int_mul(x_7, x_14); +lean_dec(x_7); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_6, 0); +lean_inc(x_18); +lean_dec(x_6); +x_19 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_18); +x_20 = lean_int_add(x_19, x_12); +lean_dec(x_12); +lean_dec(x_19); +x_21 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_20); +lean_dec(x_20); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_17); +x_23 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_22); +x_29 = lean_ctor_get(x_4, 1); +lean_inc(x_29); +lean_inc(x_23); +x_30 = l_Std_Time_TimeZone_Transition_timezoneAt(x_29, x_23); +lean_dec(x_29); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; +lean_dec(x_30); +x_31 = lean_ctor_get(x_4, 0); +lean_inc(x_31); +x_32 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_31); +lean_dec(x_31); +x_24 = x_32; +goto block_28; +} +else +{ +lean_object* x_33; +x_33 = lean_ctor_get(x_30, 0); +lean_inc(x_33); +lean_dec(x_30); +x_24 = x_33; +goto block_28; +} +block_28: +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_inc(x_24); +lean_inc(x_23); +x_25 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___lambda__1___boxed), 5, 4); +lean_closure_set(x_25, 0, x_23); +lean_closure_set(x_25, 1, x_24); +lean_closure_set(x_25, 2, x_11); +lean_closure_set(x_25, 3, x_14); +x_26 = lean_mk_thunk(x_25); +if (lean_is_scalar(x_5)) { + x_27 = lean_alloc_ctor(0, 4, 0); +} else { + x_27 = x_5; +} +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +lean_ctor_set(x_27, 2, x_4); +lean_ctor_set(x_27, 3, x_24); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subSeconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subSeconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_5 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_1); +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +x_8 = l_Std_Time_PlainTime_toSeconds(x_7); +x_9 = lean_int_add(x_8, x_6); +lean_dec(x_8); +x_10 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1; +x_11 = lean_int_ediv(x_9, x_10); +lean_dec(x_9); +x_12 = l_Std_Time_PlainTime_toNanoseconds(x_7); +lean_dec(x_7); +x_13 = lean_int_mul(x_6, x_3); +x_14 = lean_int_add(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +x_15 = l_Std_Time_PlainTime_ofNanoseconds(x_14); +lean_dec(x_14); +x_16 = lean_ctor_get(x_5, 0); +lean_inc(x_16); +lean_dec(x_5); +x_17 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_16); +x_18 = lean_int_add(x_17, x_11); +lean_dec(x_11); +lean_dec(x_17); +x_19 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_18); +lean_dec(x_18); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_15); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_8 = lean_ctor_get(x_6, 1); +x_9 = lean_ctor_get(x_8, 3); +lean_inc(x_9); +x_10 = lean_int_add(x_9, x_2); +lean_dec(x_9); +x_11 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_12 = lean_int_ediv(x_10, x_11); +x_13 = lean_int_emod(x_10, x_11); +lean_dec(x_10); +x_14 = l_Std_Time_PlainTime_toNanoseconds(x_8); +lean_dec(x_8); +x_15 = lean_int_mul(x_12, x_11); +lean_dec(x_12); +x_16 = lean_int_add(x_14, x_15); +lean_dec(x_15); +lean_dec(x_14); +x_17 = l_Std_Time_PlainTime_ofNanoseconds(x_16); +lean_dec(x_16); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_26; lean_object* x_27; +x_19 = lean_ctor_get(x_17, 3); +lean_dec(x_19); +lean_ctor_set(x_17, 3, x_13); +lean_ctor_set(x_6, 1, x_17); +x_20 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_26 = lean_ctor_get(x_4, 1); +lean_inc(x_26); +lean_inc(x_20); +x_27 = l_Std_Time_TimeZone_Transition_timezoneAt(x_26, x_20); +lean_dec(x_26); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; +lean_dec(x_27); +x_28 = lean_ctor_get(x_4, 0); +lean_inc(x_28); +x_29 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_28); +lean_dec(x_28); +x_21 = x_29; +goto block_25; +} +else +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_27, 0); +lean_inc(x_30); +lean_dec(x_27); +x_21 = x_30; +goto block_25; +} +block_25: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_inc(x_21); +lean_inc(x_20); +x_22 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +lean_closure_set(x_22, 2, x_11); +x_23 = lean_mk_thunk(x_22); +if (lean_is_scalar(x_5)) { + x_24 = lean_alloc_ctor(0, 4, 0); +} else { + x_24 = x_5; +} +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_20); +lean_ctor_set(x_24, 2, x_4); +lean_ctor_set(x_24, 3, x_21); +return x_24; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_41; lean_object* x_42; +x_31 = lean_ctor_get(x_17, 0); +x_32 = lean_ctor_get(x_17, 1); +x_33 = lean_ctor_get(x_17, 2); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_17); +x_34 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_33); +lean_ctor_set(x_34, 3, x_13); +lean_ctor_set(x_6, 1, x_34); +x_35 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_41 = lean_ctor_get(x_4, 1); +lean_inc(x_41); +lean_inc(x_35); +x_42 = l_Std_Time_TimeZone_Transition_timezoneAt(x_41, x_35); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; lean_object* x_44; +lean_dec(x_42); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_44 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_43); +lean_dec(x_43); +x_36 = x_44; +goto block_40; +} +else +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_42, 0); +lean_inc(x_45); +lean_dec(x_42); +x_36 = x_45; +goto block_40; +} +block_40: +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_inc(x_36); +lean_inc(x_35); +x_37 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_37, 0, x_35); +lean_closure_set(x_37, 1, x_36); +lean_closure_set(x_37, 2, x_11); +x_38 = lean_mk_thunk(x_37); +if (lean_is_scalar(x_5)) { + x_39 = lean_alloc_ctor(0, 4, 0); +} else { + x_39 = x_5; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_35); +lean_ctor_set(x_39, 2, x_4); +lean_ctor_set(x_39, 3, x_36); +return x_39; +} +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_69; lean_object* x_70; +x_46 = lean_ctor_get(x_6, 0); +x_47 = lean_ctor_get(x_6, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_6); +x_48 = lean_ctor_get(x_47, 3); +lean_inc(x_48); +x_49 = lean_int_add(x_48, x_2); +lean_dec(x_48); +x_50 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_51 = lean_int_ediv(x_49, x_50); +x_52 = lean_int_emod(x_49, x_50); +lean_dec(x_49); +x_53 = l_Std_Time_PlainTime_toNanoseconds(x_47); +lean_dec(x_47); +x_54 = lean_int_mul(x_51, x_50); +lean_dec(x_51); +x_55 = lean_int_add(x_53, x_54); +lean_dec(x_54); +lean_dec(x_53); +x_56 = l_Std_Time_PlainTime_ofNanoseconds(x_55); +lean_dec(x_55); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +x_59 = lean_ctor_get(x_56, 2); +lean_inc(x_59); +if (lean_is_exclusive(x_56)) { + lean_ctor_release(x_56, 0); + lean_ctor_release(x_56, 1); + lean_ctor_release(x_56, 2); + lean_ctor_release(x_56, 3); + x_60 = x_56; +} else { + lean_dec_ref(x_56); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(0, 4, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_57); +lean_ctor_set(x_61, 1, x_58); +lean_ctor_set(x_61, 2, x_59); +lean_ctor_set(x_61, 3, x_52); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_46); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_62); +x_69 = lean_ctor_get(x_4, 1); +lean_inc(x_69); +lean_inc(x_63); +x_70 = l_Std_Time_TimeZone_Transition_timezoneAt(x_69, x_63); +lean_dec(x_69); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; +lean_dec(x_70); +x_71 = lean_ctor_get(x_4, 0); +lean_inc(x_71); +x_72 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_71); +lean_dec(x_71); +x_64 = x_72; +goto block_68; +} +else +{ +lean_object* x_73; +x_73 = lean_ctor_get(x_70, 0); +lean_inc(x_73); +lean_dec(x_70); +x_64 = x_73; +goto block_68; +} +block_68: +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_inc(x_64); +lean_inc(x_63); +x_65 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_65, 0, x_63); +lean_closure_set(x_65, 1, x_64); +lean_closure_set(x_65, 2, x_50); +x_66 = lean_mk_thunk(x_65); +if (lean_is_scalar(x_5)) { + x_67 = lean_alloc_ctor(0, 4, 0); +} else { + x_67 = x_5; +} +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_63); +lean_ctor_set(x_67, 2, x_4); +lean_ctor_set(x_67, 3, x_64); +return x_67; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_addNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_addNanoseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = l_Std_Time_PlainDateTime_ofTimestampAssumingUTC(x_3); +lean_dec(x_3); +x_7 = lean_int_neg(x_2); +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_9 = lean_ctor_get(x_6, 1); +x_10 = lean_ctor_get(x_9, 3); +lean_inc(x_10); +x_11 = lean_int_add(x_10, x_7); +lean_dec(x_7); +lean_dec(x_10); +x_12 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_ediv(x_11, x_12); +x_14 = lean_int_emod(x_11, x_12); +lean_dec(x_11); +x_15 = l_Std_Time_PlainTime_toNanoseconds(x_9); +lean_dec(x_9); +x_16 = lean_int_mul(x_13, x_12); +lean_dec(x_13); +x_17 = lean_int_add(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +x_18 = l_Std_Time_PlainTime_ofNanoseconds(x_17); +lean_dec(x_17); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_27; lean_object* x_28; +x_20 = lean_ctor_get(x_18, 3); +lean_dec(x_20); +lean_ctor_set(x_18, 3, x_14); +lean_ctor_set(x_6, 1, x_18); +x_21 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_27 = lean_ctor_get(x_4, 1); +lean_inc(x_27); +lean_inc(x_21); +x_28 = l_Std_Time_TimeZone_Transition_timezoneAt(x_27, x_21); +lean_dec(x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; +lean_dec(x_28); +x_29 = lean_ctor_get(x_4, 0); +lean_inc(x_29); +x_30 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_29); +lean_dec(x_29); +x_22 = x_30; +goto block_26; +} +else +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_28, 0); +lean_inc(x_31); +lean_dec(x_28); +x_22 = x_31; +goto block_26; +} +block_26: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_inc(x_22); +lean_inc(x_21); +x_23 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_12); +x_24 = lean_mk_thunk(x_23); +if (lean_is_scalar(x_5)) { + x_25 = lean_alloc_ctor(0, 4, 0); +} else { + x_25 = x_5; +} +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_21); +lean_ctor_set(x_25, 2, x_4); +lean_ctor_set(x_25, 3, x_22); +return x_25; +} +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_42; lean_object* x_43; +x_32 = lean_ctor_get(x_18, 0); +x_33 = lean_ctor_get(x_18, 1); +x_34 = lean_ctor_get(x_18, 2); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_18); +x_35 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_35, 0, x_32); +lean_ctor_set(x_35, 1, x_33); +lean_ctor_set(x_35, 2, x_34); +lean_ctor_set(x_35, 3, x_14); +lean_ctor_set(x_6, 1, x_35); +x_36 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_42 = lean_ctor_get(x_4, 1); +lean_inc(x_42); +lean_inc(x_36); +x_43 = l_Std_Time_TimeZone_Transition_timezoneAt(x_42, x_36); +lean_dec(x_42); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; lean_object* x_45; +lean_dec(x_43); +x_44 = lean_ctor_get(x_4, 0); +lean_inc(x_44); +x_45 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_44); +lean_dec(x_44); +x_37 = x_45; +goto block_41; +} +else +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_43, 0); +lean_inc(x_46); +lean_dec(x_43); +x_37 = x_46; +goto block_41; +} +block_41: +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_inc(x_37); +lean_inc(x_36); +x_38 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_38, 0, x_36); +lean_closure_set(x_38, 1, x_37); +lean_closure_set(x_38, 2, x_12); +x_39 = lean_mk_thunk(x_38); +if (lean_is_scalar(x_5)) { + x_40 = lean_alloc_ctor(0, 4, 0); +} else { + x_40 = x_5; +} +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_36); +lean_ctor_set(x_40, 2, x_4); +lean_ctor_set(x_40, 3, x_37); +return x_40; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_70; lean_object* x_71; +x_47 = lean_ctor_get(x_6, 0); +x_48 = lean_ctor_get(x_6, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_6); +x_49 = lean_ctor_get(x_48, 3); +lean_inc(x_49); +x_50 = lean_int_add(x_49, x_7); +lean_dec(x_7); +lean_dec(x_49); +x_51 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_52 = lean_int_ediv(x_50, x_51); +x_53 = lean_int_emod(x_50, x_51); +lean_dec(x_50); +x_54 = l_Std_Time_PlainTime_toNanoseconds(x_48); +lean_dec(x_48); +x_55 = lean_int_mul(x_52, x_51); +lean_dec(x_52); +x_56 = lean_int_add(x_54, x_55); +lean_dec(x_55); +lean_dec(x_54); +x_57 = l_Std_Time_PlainTime_ofNanoseconds(x_56); +lean_dec(x_56); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +x_60 = lean_ctor_get(x_57, 2); +lean_inc(x_60); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + lean_ctor_release(x_57, 2); + lean_ctor_release(x_57, 3); + x_61 = x_57; +} else { + lean_dec_ref(x_57); + x_61 = lean_box(0); +} +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(0, 4, 0); +} else { + x_62 = x_61; +} +lean_ctor_set(x_62, 0, x_58); +lean_ctor_set(x_62, 1, x_59); +lean_ctor_set(x_62, 2, x_60); +lean_ctor_set(x_62, 3, x_53); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_47); +lean_ctor_set(x_63, 1, x_62); +x_64 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_63); +x_70 = lean_ctor_get(x_4, 1); +lean_inc(x_70); +lean_inc(x_64); +x_71 = l_Std_Time_TimeZone_Transition_timezoneAt(x_70, x_64); +lean_dec(x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; +lean_dec(x_71); +x_72 = lean_ctor_get(x_4, 0); +lean_inc(x_72); +x_73 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_72); +lean_dec(x_72); +x_65 = x_73; +goto block_69; +} +else +{ +lean_object* x_74; +x_74 = lean_ctor_get(x_71, 0); +lean_inc(x_74); +lean_dec(x_71); +x_65 = x_74; +goto block_69; +} +block_69: +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_inc(x_65); +lean_inc(x_64); +x_66 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___lambda__1___boxed), 4, 3); +lean_closure_set(x_66, 0, x_64); +lean_closure_set(x_66, 1, x_65); +lean_closure_set(x_66, 2, x_51); +x_67 = lean_mk_thunk(x_66); +if (lean_is_scalar(x_5)) { + x_68 = lean_alloc_ctor(0, 4, 0); +} else { + x_68 = x_5; +} +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_64); +lean_ctor_set(x_68, 2, x_4); +lean_ctor_set(x_68, 3, x_65); +return x_68; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_subNanoseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_subNanoseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_era(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_Year_Offset_era(x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_era___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_ZonedDateTime_era(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withWeekday(lean_object* x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = l_Std_Time_PlainDateTime_withWeekday(x_6, x_2); +x_8 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +lean_inc(x_9); +x_10 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_10, 0, x_9); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_findIdx_x3f_loop___rarg(x_10, x_11, x_12); +x_14 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_15 = lean_int_mul(x_9, x_14); +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_int_add(x_15, x_16); +lean_dec(x_16); +lean_dec(x_15); +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_31; +lean_dec(x_9); +x_31 = l_Array_back_x3f___rarg(x_11); +lean_dec(x_11); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_4, 0); +lean_inc(x_32); +x_18 = x_32; +goto block_30; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +lean_inc(x_33); +lean_dec(x_31); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_18 = x_34; +goto block_30; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_13, 0); +lean_inc(x_35); +lean_dec(x_13); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_sub(x_35, x_36); +x_38 = l_Array_get_x3f___rarg(x_11, x_37); +lean_dec(x_37); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; +lean_dec(x_35); +lean_dec(x_11); +lean_dec(x_9); +x_39 = lean_ctor_get(x_4, 0); +lean_inc(x_39); +x_18 = x_39; +goto block_30; +} +else +{ +lean_object* x_40; lean_object* x_41; +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = l_Array_get_x3f___rarg(x_11, x_35); +lean_dec(x_35); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +x_42 = l_Array_back_x3f___rarg(x_11); +lean_dec(x_11); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +lean_dec(x_40); +lean_dec(x_9); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_18 = x_43; +goto block_30; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_40, 1); +lean_inc(x_46); +lean_dec(x_40); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_nat_abs(x_47); +lean_dec(x_47); +x_49 = lean_nat_to_int(x_48); +x_50 = lean_int_sub(x_45, x_49); +lean_dec(x_49); +lean_dec(x_45); +x_51 = lean_int_dec_lt(x_9, x_50); +lean_dec(x_50); +lean_dec(x_9); +if (x_51 == 0) +{ +lean_object* x_52; +lean_dec(x_46); +x_52 = lean_ctor_get(x_44, 1); +lean_inc(x_52); +lean_dec(x_44); +x_18 = x_52; +goto block_30; +} +else +{ +lean_dec(x_44); +x_18 = x_46; +goto block_30; +} +} +} +else +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +lean_dec(x_11); +x_53 = lean_ctor_get(x_41, 0); +lean_inc(x_53); +lean_dec(x_41); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_40, 1); +lean_inc(x_55); +lean_dec(x_40); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_nat_abs(x_56); +lean_dec(x_56); +x_58 = lean_nat_to_int(x_57); +x_59 = lean_int_sub(x_54, x_58); +lean_dec(x_58); +lean_dec(x_54); +x_60 = lean_int_dec_lt(x_9, x_59); +lean_dec(x_59); +lean_dec(x_9); +if (x_60 == 0) +{ +lean_object* x_61; +lean_dec(x_55); +x_61 = lean_ctor_get(x_53, 1); +lean_inc(x_61); +lean_dec(x_53); +x_18 = x_61; +goto block_30; +} +else +{ +lean_dec(x_53); +x_18 = x_55; +goto block_30; +} +} +} +} +block_30: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_19 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_18); +lean_dec(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_int_neg(x_20); +x_22 = lean_int_mul(x_21, x_14); +lean_dec(x_21); +x_23 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_24 = lean_int_add(x_22, x_23); +lean_dec(x_22); +x_25 = lean_int_add(x_17, x_24); +lean_dec(x_24); +lean_dec(x_17); +x_26 = l_Std_Time_Duration_ofNanoseconds(x_25); +lean_dec(x_25); +lean_inc(x_26); +x_27 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_27, 0, x_26); +lean_closure_set(x_27, 1, x_20); +lean_closure_set(x_27, 2, x_14); +x_28 = lean_mk_thunk(x_27); +if (lean_is_scalar(x_5)) { + x_29 = lean_alloc_ctor(0, 4, 0); +} else { + x_29 = x_5; +} +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_26); +lean_ctor_set(x_29, 2, x_4); +lean_ctor_set(x_29, 3, x_19); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withWeekday___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Std_Time_ZonedDateTime_withWeekday(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 0); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +x_12 = lean_ctor_get(x_8, 2); +lean_dec(x_12); +x_13 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_14 = lean_int_mod(x_10, x_13); +x_15 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +if (x_16 == 0) +{ +uint8_t x_73; lean_object* x_74; uint8_t x_75; +x_73 = 0; +x_74 = l_Std_Time_Month_Ordinal_days(x_73, x_11); +x_75 = lean_int_dec_lt(x_74, x_2); +if (x_75 == 0) +{ +lean_dec(x_74); +lean_ctor_set(x_8, 2, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_8, 2, x_74); +x_18 = x_6; +goto block_72; +} +} +else +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; +x_76 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_77 = lean_int_mod(x_10, x_76); +x_78 = lean_int_dec_eq(x_77, x_15); +lean_dec(x_77); +x_79 = l_instDecidableNot___rarg(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_81 = lean_int_mod(x_10, x_80); +x_82 = lean_int_dec_eq(x_81, x_15); +lean_dec(x_81); +if (x_82 == 0) +{ +uint8_t x_83; lean_object* x_84; uint8_t x_85; +x_83 = 0; +x_84 = l_Std_Time_Month_Ordinal_days(x_83, x_11); +x_85 = lean_int_dec_lt(x_84, x_2); +if (x_85 == 0) +{ +lean_dec(x_84); +lean_ctor_set(x_8, 2, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_8, 2, x_84); +x_18 = x_6; +goto block_72; +} +} +else +{ +uint8_t x_86; lean_object* x_87; uint8_t x_88; +x_86 = 1; +x_87 = l_Std_Time_Month_Ordinal_days(x_86, x_11); +x_88 = lean_int_dec_lt(x_87, x_2); +if (x_88 == 0) +{ +lean_dec(x_87); +lean_ctor_set(x_8, 2, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_8, 2, x_87); +x_18 = x_6; +goto block_72; +} +} +} +else +{ +uint8_t x_89; lean_object* x_90; uint8_t x_91; +x_89 = 1; +x_90 = l_Std_Time_Month_Ordinal_days(x_89, x_11); +x_91 = lean_int_dec_lt(x_90, x_2); +if (x_91 == 0) +{ +lean_dec(x_90); +lean_ctor_set(x_8, 2, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_2); +lean_ctor_set(x_8, 2, x_90); +x_18 = x_6; +goto block_72; +} +} +} +block_72: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_inc(x_20); +x_21 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_21, 0, x_20); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Array_findIdx_x3f_loop___rarg(x_21, x_17, x_22); +x_24 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_25 = lean_int_mul(x_20, x_24); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_41; +lean_dec(x_20); +x_41 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_28 = x_42; +goto block_40; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_28 = x_44; +goto block_40; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_23, 0); +lean_inc(x_45); +lean_dec(x_23); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_sub(x_45, x_46); +x_48 = l_Array_get_x3f___rarg(x_17, x_47); +lean_dec(x_47); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; +lean_dec(x_45); +lean_dec(x_20); +lean_dec(x_17); +x_49 = lean_ctor_get(x_4, 0); +lean_inc(x_49); +x_28 = x_49; +goto block_40; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 0); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Array_get_x3f___rarg(x_17, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; +x_52 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; +lean_dec(x_50); +lean_dec(x_20); +x_53 = lean_ctor_get(x_4, 0); +lean_inc(x_53); +x_28 = x_53; +goto block_40; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); +lean_dec(x_50); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_nat_abs(x_57); +lean_dec(x_57); +x_59 = lean_nat_to_int(x_58); +x_60 = lean_int_sub(x_55, x_59); +lean_dec(x_59); +lean_dec(x_55); +x_61 = lean_int_dec_lt(x_20, x_60); +lean_dec(x_60); +lean_dec(x_20); +if (x_61 == 0) +{ +lean_object* x_62; +lean_dec(x_56); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_28 = x_62; +goto block_40; +} +else +{ +lean_dec(x_54); +x_28 = x_56; +goto block_40; +} +} +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +lean_dec(x_17); +x_63 = lean_ctor_get(x_51, 0); +lean_inc(x_63); +lean_dec(x_51); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_dec(x_50); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_nat_abs(x_66); +lean_dec(x_66); +x_68 = lean_nat_to_int(x_67); +x_69 = lean_int_sub(x_64, x_68); +lean_dec(x_68); +lean_dec(x_64); +x_70 = lean_int_dec_lt(x_20, x_69); +lean_dec(x_69); +lean_dec(x_20); +if (x_70 == 0) +{ +lean_object* x_71; +lean_dec(x_65); +x_71 = lean_ctor_get(x_63, 1); +lean_inc(x_71); +lean_dec(x_63); +x_28 = x_71; +goto block_40; +} +else +{ +lean_dec(x_63); +x_28 = x_65; +goto block_40; +} +} +} +} +block_40: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_int_neg(x_30); +x_32 = lean_int_mul(x_31, x_24); +lean_dec(x_31); +x_33 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_34 = lean_int_add(x_32, x_33); +lean_dec(x_32); +x_35 = lean_int_add(x_27, x_34); +lean_dec(x_34); +lean_dec(x_27); +x_36 = l_Std_Time_Duration_ofNanoseconds(x_35); +lean_dec(x_35); +lean_inc(x_36); +x_37 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_37, 0, x_36); +lean_closure_set(x_37, 1, x_30); +lean_closure_set(x_37, 2, x_24); +x_38 = lean_mk_thunk(x_37); +if (lean_is_scalar(x_5)) { + x_39 = lean_alloc_ctor(0, 4, 0); +} else { + x_39 = x_5; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +lean_ctor_set(x_39, 2, x_4); +lean_ctor_set(x_39, 3, x_29); +return x_39; +} +} +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +x_92 = lean_ctor_get(x_8, 0); +x_93 = lean_ctor_get(x_8, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_8); +x_94 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_95 = lean_int_mod(x_92, x_94); +x_96 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_97 = lean_int_dec_eq(x_95, x_96); +lean_dec(x_95); +x_98 = lean_ctor_get(x_4, 1); +lean_inc(x_98); +if (x_97 == 0) +{ +uint8_t x_154; lean_object* x_155; uint8_t x_156; +x_154 = 0; +x_155 = l_Std_Time_Month_Ordinal_days(x_154, x_93); +x_156 = lean_int_dec_lt(x_155, x_2); +if (x_156 == 0) +{ +lean_object* x_157; +lean_dec(x_155); +x_157 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_157, 0, x_92); +lean_ctor_set(x_157, 1, x_93); +lean_ctor_set(x_157, 2, x_2); +lean_ctor_set(x_6, 0, x_157); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_158; +lean_dec(x_2); +x_158 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_158, 0, x_92); +lean_ctor_set(x_158, 1, x_93); +lean_ctor_set(x_158, 2, x_155); +lean_ctor_set(x_6, 0, x_158); +x_99 = x_6; +goto block_153; +} +} +else +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; uint8_t x_162; +x_159 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_160 = lean_int_mod(x_92, x_159); +x_161 = lean_int_dec_eq(x_160, x_96); +lean_dec(x_160); +x_162 = l_instDecidableNot___rarg(x_161); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; uint8_t x_165; +x_163 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_164 = lean_int_mod(x_92, x_163); +x_165 = lean_int_dec_eq(x_164, x_96); +lean_dec(x_164); +if (x_165 == 0) +{ +uint8_t x_166; lean_object* x_167; uint8_t x_168; +x_166 = 0; +x_167 = l_Std_Time_Month_Ordinal_days(x_166, x_93); +x_168 = lean_int_dec_lt(x_167, x_2); +if (x_168 == 0) +{ +lean_object* x_169; +lean_dec(x_167); +x_169 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_169, 0, x_92); +lean_ctor_set(x_169, 1, x_93); +lean_ctor_set(x_169, 2, x_2); +lean_ctor_set(x_6, 0, x_169); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_170; +lean_dec(x_2); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_92); +lean_ctor_set(x_170, 1, x_93); +lean_ctor_set(x_170, 2, x_167); +lean_ctor_set(x_6, 0, x_170); +x_99 = x_6; +goto block_153; +} +} +else +{ +uint8_t x_171; lean_object* x_172; uint8_t x_173; +x_171 = 1; +x_172 = l_Std_Time_Month_Ordinal_days(x_171, x_93); +x_173 = lean_int_dec_lt(x_172, x_2); +if (x_173 == 0) +{ +lean_object* x_174; +lean_dec(x_172); +x_174 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_174, 0, x_92); +lean_ctor_set(x_174, 1, x_93); +lean_ctor_set(x_174, 2, x_2); +lean_ctor_set(x_6, 0, x_174); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_175; +lean_dec(x_2); +x_175 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_175, 0, x_92); +lean_ctor_set(x_175, 1, x_93); +lean_ctor_set(x_175, 2, x_172); +lean_ctor_set(x_6, 0, x_175); +x_99 = x_6; +goto block_153; +} +} +} +else +{ +uint8_t x_176; lean_object* x_177; uint8_t x_178; +x_176 = 1; +x_177 = l_Std_Time_Month_Ordinal_days(x_176, x_93); +x_178 = lean_int_dec_lt(x_177, x_2); +if (x_178 == 0) +{ +lean_object* x_179; +lean_dec(x_177); +x_179 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_179, 0, x_92); +lean_ctor_set(x_179, 1, x_93); +lean_ctor_set(x_179, 2, x_2); +lean_ctor_set(x_6, 0, x_179); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_180; +lean_dec(x_2); +x_180 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_180, 0, x_92); +lean_ctor_set(x_180, 1, x_93); +lean_ctor_set(x_180, 2, x_177); +lean_ctor_set(x_6, 0, x_180); +x_99 = x_6; +goto block_153; +} +} +} +block_153: +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_100 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +lean_inc(x_101); +x_102 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_102, 0, x_101); +x_103 = lean_unsigned_to_nat(0u); +x_104 = l_Array_findIdx_x3f_loop___rarg(x_102, x_98, x_103); +x_105 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_106 = lean_int_mul(x_101, x_105); +x_107 = lean_ctor_get(x_100, 1); +lean_inc(x_107); +lean_dec(x_100); +x_108 = lean_int_add(x_106, x_107); +lean_dec(x_107); +lean_dec(x_106); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_122; +lean_dec(x_101); +x_122 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_122) == 0) +{ +lean_object* x_123; +x_123 = lean_ctor_get(x_4, 0); +lean_inc(x_123); +x_109 = x_123; +goto block_121; +} +else +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_122, 0); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_109 = x_125; +goto block_121; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_104, 0); +lean_inc(x_126); +lean_dec(x_104); +x_127 = lean_unsigned_to_nat(1u); +x_128 = lean_nat_sub(x_126, x_127); +x_129 = l_Array_get_x3f___rarg(x_98, x_128); +lean_dec(x_128); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; +lean_dec(x_126); +lean_dec(x_101); +lean_dec(x_98); +x_130 = lean_ctor_get(x_4, 0); +lean_inc(x_130); +x_109 = x_130; +goto block_121; +} +else +{ +lean_object* x_131; lean_object* x_132; +x_131 = lean_ctor_get(x_129, 0); +lean_inc(x_131); +lean_dec(x_129); +x_132 = l_Array_get_x3f___rarg(x_98, x_126); +lean_dec(x_126); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; +x_133 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; +lean_dec(x_131); +lean_dec(x_101); +x_134 = lean_ctor_get(x_4, 0); +lean_inc(x_134); +x_109 = x_134; +goto block_121; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_135 = lean_ctor_get(x_133, 0); +lean_inc(x_135); +lean_dec(x_133); +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +lean_dec(x_131); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_nat_abs(x_138); +lean_dec(x_138); +x_140 = lean_nat_to_int(x_139); +x_141 = lean_int_sub(x_136, x_140); +lean_dec(x_140); +lean_dec(x_136); +x_142 = lean_int_dec_lt(x_101, x_141); +lean_dec(x_141); +lean_dec(x_101); +if (x_142 == 0) +{ +lean_object* x_143; +lean_dec(x_137); +x_143 = lean_ctor_get(x_135, 1); +lean_inc(x_143); +lean_dec(x_135); +x_109 = x_143; +goto block_121; +} +else +{ +lean_dec(x_135); +x_109 = x_137; +goto block_121; +} +} +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +lean_dec(x_98); +x_144 = lean_ctor_get(x_132, 0); +lean_inc(x_144); +lean_dec(x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_131, 1); +lean_inc(x_146); +lean_dec(x_131); +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_nat_abs(x_147); +lean_dec(x_147); +x_149 = lean_nat_to_int(x_148); +x_150 = lean_int_sub(x_145, x_149); +lean_dec(x_149); +lean_dec(x_145); +x_151 = lean_int_dec_lt(x_101, x_150); +lean_dec(x_150); +lean_dec(x_101); +if (x_151 == 0) +{ +lean_object* x_152; +lean_dec(x_146); +x_152 = lean_ctor_get(x_144, 1); +lean_inc(x_152); +lean_dec(x_144); +x_109 = x_152; +goto block_121; +} +else +{ +lean_dec(x_144); +x_109 = x_146; +goto block_121; +} +} +} +} +block_121: +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_110 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_109); +lean_dec(x_109); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_int_neg(x_111); +x_113 = lean_int_mul(x_112, x_105); +lean_dec(x_112); +x_114 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_115 = lean_int_add(x_113, x_114); +lean_dec(x_113); +x_116 = lean_int_add(x_108, x_115); +lean_dec(x_115); +lean_dec(x_108); +x_117 = l_Std_Time_Duration_ofNanoseconds(x_116); +lean_dec(x_116); +lean_inc(x_117); +x_118 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_118, 0, x_117); +lean_closure_set(x_118, 1, x_111); +lean_closure_set(x_118, 2, x_105); +x_119 = lean_mk_thunk(x_118); +if (lean_is_scalar(x_5)) { + x_120 = lean_alloc_ctor(0, 4, 0); +} else { + x_120 = x_5; +} +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +lean_ctor_set(x_120, 2, x_4); +lean_ctor_set(x_120, 3, x_110); +return x_120; +} +} +} +} +else +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; +x_181 = lean_ctor_get(x_6, 0); +x_182 = lean_ctor_get(x_6, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_6); +x_183 = lean_ctor_get(x_181, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_181, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + lean_ctor_release(x_181, 2); + x_185 = x_181; +} else { + lean_dec_ref(x_181); + x_185 = lean_box(0); +} +x_186 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_187 = lean_int_mod(x_183, x_186); +x_188 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_189 = lean_int_dec_eq(x_187, x_188); +lean_dec(x_187); +x_190 = lean_ctor_get(x_4, 1); +lean_inc(x_190); +if (x_189 == 0) +{ +uint8_t x_246; lean_object* x_247; uint8_t x_248; +x_246 = 0; +x_247 = l_Std_Time_Month_Ordinal_days(x_246, x_184); +x_248 = lean_int_dec_lt(x_247, x_2); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; +lean_dec(x_247); +if (lean_is_scalar(x_185)) { + x_249 = lean_alloc_ctor(0, 3, 0); +} else { + x_249 = x_185; +} +lean_ctor_set(x_249, 0, x_183); +lean_ctor_set(x_249, 1, x_184); +lean_ctor_set(x_249, 2, x_2); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_182); +x_191 = x_250; +goto block_245; +} +else +{ +lean_object* x_251; lean_object* x_252; +lean_dec(x_2); +if (lean_is_scalar(x_185)) { + x_251 = lean_alloc_ctor(0, 3, 0); +} else { + x_251 = x_185; +} +lean_ctor_set(x_251, 0, x_183); +lean_ctor_set(x_251, 1, x_184); +lean_ctor_set(x_251, 2, x_247); +x_252 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_182); +x_191 = x_252; +goto block_245; +} +} +else +{ +lean_object* x_253; lean_object* x_254; uint8_t x_255; uint8_t x_256; +x_253 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_254 = lean_int_mod(x_183, x_253); +x_255 = lean_int_dec_eq(x_254, x_188); +lean_dec(x_254); +x_256 = l_instDecidableNot___rarg(x_255); +if (x_256 == 0) +{ +lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_257 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_258 = lean_int_mod(x_183, x_257); +x_259 = lean_int_dec_eq(x_258, x_188); +lean_dec(x_258); +if (x_259 == 0) +{ +uint8_t x_260; lean_object* x_261; uint8_t x_262; +x_260 = 0; +x_261 = l_Std_Time_Month_Ordinal_days(x_260, x_184); +x_262 = lean_int_dec_lt(x_261, x_2); +if (x_262 == 0) +{ +lean_object* x_263; lean_object* x_264; +lean_dec(x_261); +if (lean_is_scalar(x_185)) { + x_263 = lean_alloc_ctor(0, 3, 0); +} else { + x_263 = x_185; +} +lean_ctor_set(x_263, 0, x_183); +lean_ctor_set(x_263, 1, x_184); +lean_ctor_set(x_263, 2, x_2); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_182); +x_191 = x_264; +goto block_245; +} +else +{ +lean_object* x_265; lean_object* x_266; +lean_dec(x_2); +if (lean_is_scalar(x_185)) { + x_265 = lean_alloc_ctor(0, 3, 0); +} else { + x_265 = x_185; +} +lean_ctor_set(x_265, 0, x_183); +lean_ctor_set(x_265, 1, x_184); +lean_ctor_set(x_265, 2, x_261); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_182); +x_191 = x_266; +goto block_245; +} +} +else +{ +uint8_t x_267; lean_object* x_268; uint8_t x_269; +x_267 = 1; +x_268 = l_Std_Time_Month_Ordinal_days(x_267, x_184); +x_269 = lean_int_dec_lt(x_268, x_2); +if (x_269 == 0) +{ +lean_object* x_270; lean_object* x_271; +lean_dec(x_268); +if (lean_is_scalar(x_185)) { + x_270 = lean_alloc_ctor(0, 3, 0); +} else { + x_270 = x_185; +} +lean_ctor_set(x_270, 0, x_183); +lean_ctor_set(x_270, 1, x_184); +lean_ctor_set(x_270, 2, x_2); +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_182); +x_191 = x_271; +goto block_245; +} +else +{ +lean_object* x_272; lean_object* x_273; +lean_dec(x_2); +if (lean_is_scalar(x_185)) { + x_272 = lean_alloc_ctor(0, 3, 0); +} else { + x_272 = x_185; +} +lean_ctor_set(x_272, 0, x_183); +lean_ctor_set(x_272, 1, x_184); +lean_ctor_set(x_272, 2, x_268); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_182); +x_191 = x_273; +goto block_245; +} +} +} +else +{ +uint8_t x_274; lean_object* x_275; uint8_t x_276; +x_274 = 1; +x_275 = l_Std_Time_Month_Ordinal_days(x_274, x_184); +x_276 = lean_int_dec_lt(x_275, x_2); +if (x_276 == 0) +{ +lean_object* x_277; lean_object* x_278; +lean_dec(x_275); +if (lean_is_scalar(x_185)) { + x_277 = lean_alloc_ctor(0, 3, 0); +} else { + x_277 = x_185; +} +lean_ctor_set(x_277, 0, x_183); +lean_ctor_set(x_277, 1, x_184); +lean_ctor_set(x_277, 2, x_2); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_182); +x_191 = x_278; +goto block_245; +} +else +{ +lean_object* x_279; lean_object* x_280; +lean_dec(x_2); +if (lean_is_scalar(x_185)) { + x_279 = lean_alloc_ctor(0, 3, 0); +} else { + x_279 = x_185; +} +lean_ctor_set(x_279, 0, x_183); +lean_ctor_set(x_279, 1, x_184); +lean_ctor_set(x_279, 2, x_275); +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_182); +x_191 = x_280; +goto block_245; +} +} +} +block_245: +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_192 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_191); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +lean_inc(x_193); +x_194 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_194, 0, x_193); +x_195 = lean_unsigned_to_nat(0u); +x_196 = l_Array_findIdx_x3f_loop___rarg(x_194, x_190, x_195); +x_197 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_198 = lean_int_mul(x_193, x_197); +x_199 = lean_ctor_get(x_192, 1); +lean_inc(x_199); +lean_dec(x_192); +x_200 = lean_int_add(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_214; +lean_dec(x_193); +x_214 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; +x_215 = lean_ctor_get(x_4, 0); +lean_inc(x_215); +x_201 = x_215; +goto block_213; +} +else +{ +lean_object* x_216; lean_object* x_217; +x_216 = lean_ctor_get(x_214, 0); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +lean_dec(x_216); +x_201 = x_217; +goto block_213; +} +} +else +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_218 = lean_ctor_get(x_196, 0); +lean_inc(x_218); +lean_dec(x_196); +x_219 = lean_unsigned_to_nat(1u); +x_220 = lean_nat_sub(x_218, x_219); +x_221 = l_Array_get_x3f___rarg(x_190, x_220); +lean_dec(x_220); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; +lean_dec(x_218); +lean_dec(x_193); +lean_dec(x_190); +x_222 = lean_ctor_get(x_4, 0); +lean_inc(x_222); +x_201 = x_222; +goto block_213; +} +else +{ +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_221, 0); +lean_inc(x_223); +lean_dec(x_221); +x_224 = l_Array_get_x3f___rarg(x_190, x_218); +lean_dec(x_218); +if (lean_obj_tag(x_224) == 0) +{ +lean_object* x_225; +x_225 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_225) == 0) +{ +lean_object* x_226; +lean_dec(x_223); +lean_dec(x_193); +x_226 = lean_ctor_get(x_4, 0); +lean_inc(x_226); +x_201 = x_226; +goto block_213; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; +x_227 = lean_ctor_get(x_225, 0); +lean_inc(x_227); +lean_dec(x_225); +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_223, 1); +lean_inc(x_229); +lean_dec(x_223); +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_nat_abs(x_230); +lean_dec(x_230); +x_232 = lean_nat_to_int(x_231); +x_233 = lean_int_sub(x_228, x_232); +lean_dec(x_232); +lean_dec(x_228); +x_234 = lean_int_dec_lt(x_193, x_233); +lean_dec(x_233); +lean_dec(x_193); +if (x_234 == 0) +{ +lean_object* x_235; +lean_dec(x_229); +x_235 = lean_ctor_get(x_227, 1); +lean_inc(x_235); +lean_dec(x_227); +x_201 = x_235; +goto block_213; +} +else +{ +lean_dec(x_227); +x_201 = x_229; +goto block_213; +} +} +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; +lean_dec(x_190); +x_236 = lean_ctor_get(x_224, 0); +lean_inc(x_236); +lean_dec(x_224); +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_223, 1); +lean_inc(x_238); +lean_dec(x_223); +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_nat_abs(x_239); +lean_dec(x_239); +x_241 = lean_nat_to_int(x_240); +x_242 = lean_int_sub(x_237, x_241); +lean_dec(x_241); +lean_dec(x_237); +x_243 = lean_int_dec_lt(x_193, x_242); +lean_dec(x_242); +lean_dec(x_193); +if (x_243 == 0) +{ +lean_object* x_244; +lean_dec(x_238); +x_244 = lean_ctor_get(x_236, 1); +lean_inc(x_244); +lean_dec(x_236); +x_201 = x_244; +goto block_213; +} +else +{ +lean_dec(x_236); +x_201 = x_238; +goto block_213; +} +} +} +} +block_213: +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_202 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_201); +lean_dec(x_201); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_int_neg(x_203); +x_205 = lean_int_mul(x_204, x_197); +lean_dec(x_204); +x_206 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_207 = lean_int_add(x_205, x_206); +lean_dec(x_205); +x_208 = lean_int_add(x_200, x_207); +lean_dec(x_207); +lean_dec(x_200); +x_209 = l_Std_Time_Duration_ofNanoseconds(x_208); +lean_dec(x_208); +lean_inc(x_209); +x_210 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_210, 0, x_209); +lean_closure_set(x_210, 1, x_203); +lean_closure_set(x_210, 2, x_197); +x_211 = lean_mk_thunk(x_210); +if (lean_is_scalar(x_5)) { + x_212 = lean_alloc_ctor(0, 4, 0); +} else { + x_212 = x_5; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_209); +lean_ctor_set(x_212, 2, x_4); +lean_ctor_set(x_212, 3, x_202); +return x_212; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Std_Time_PlainDate_rollOver(x_9, x_10, x_2); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_findIdx_x3f_loop___rarg(x_14, x_15, x_16); +x_18 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_13, x_18); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_35; +lean_dec(x_13); +x_35 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_4, 0); +lean_inc(x_36); +x_22 = x_36; +goto block_34; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_22 = x_38; +goto block_34; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_17, 0); +lean_inc(x_39); +lean_dec(x_17); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_sub(x_39, x_40); +x_42 = l_Array_get_x3f___rarg(x_15, x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +lean_dec(x_39); +lean_dec(x_15); +lean_dec(x_13); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_22 = x_43; +goto block_34; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Array_get_x3f___rarg(x_15, x_39); +lean_dec(x_39); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +x_46 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; +lean_dec(x_44); +lean_dec(x_13); +x_47 = lean_ctor_get(x_4, 0); +lean_inc(x_47); +x_22 = x_47; +goto block_34; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_dec(x_44); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_nat_abs(x_51); +lean_dec(x_51); +x_53 = lean_nat_to_int(x_52); +x_54 = lean_int_sub(x_49, x_53); +lean_dec(x_53); +lean_dec(x_49); +x_55 = lean_int_dec_lt(x_13, x_54); +lean_dec(x_54); +lean_dec(x_13); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_50); +x_56 = lean_ctor_get(x_48, 1); +lean_inc(x_56); +lean_dec(x_48); +x_22 = x_56; +goto block_34; +} +else +{ +lean_dec(x_48); +x_22 = x_50; +goto block_34; +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_15); +x_57 = lean_ctor_get(x_45, 0); +lean_inc(x_57); +lean_dec(x_45); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_44, 1); +lean_inc(x_59); +lean_dec(x_44); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_nat_abs(x_60); +lean_dec(x_60); +x_62 = lean_nat_to_int(x_61); +x_63 = lean_int_sub(x_58, x_62); +lean_dec(x_62); +lean_dec(x_58); +x_64 = lean_int_dec_lt(x_13, x_63); +lean_dec(x_63); +lean_dec(x_13); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_59); +x_65 = lean_ctor_get(x_57, 1); +lean_inc(x_65); +lean_dec(x_57); +x_22 = x_65; +goto block_34; +} +else +{ +lean_dec(x_57); +x_22 = x_59; +goto block_34; +} +} +} +} +block_34: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_23 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_22); +lean_dec(x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_int_neg(x_24); +x_26 = lean_int_mul(x_25, x_18); +lean_dec(x_25); +x_27 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_26); +x_29 = lean_int_add(x_21, x_28); +lean_dec(x_28); +lean_dec(x_21); +x_30 = l_Std_Time_Duration_ofNanoseconds(x_29); +lean_dec(x_29); +lean_inc(x_30); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_31, 0, x_30); +lean_closure_set(x_31, 1, x_24); +lean_closure_set(x_31, 2, x_18); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_23); +return x_33; +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_66 = lean_ctor_get(x_6, 0); +x_67 = lean_ctor_get(x_6, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_6); +x_68 = lean_ctor_get(x_66, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = l_Std_Time_PlainDate_rollOver(x_68, x_69, x_2); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_67); +x_72 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_71); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +lean_inc(x_73); +x_74 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_74, 0, x_73); +x_75 = lean_ctor_get(x_4, 1); +lean_inc(x_75); +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Array_findIdx_x3f_loop___rarg(x_74, x_75, x_76); +x_78 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_79 = lean_int_mul(x_73, x_78); +x_80 = lean_ctor_get(x_72, 1); +lean_inc(x_80); +lean_dec(x_72); +x_81 = lean_int_add(x_79, x_80); +lean_dec(x_80); +lean_dec(x_79); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_95; +lean_dec(x_73); +x_95 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; +x_96 = lean_ctor_get(x_4, 0); +lean_inc(x_96); +x_82 = x_96; +goto block_94; +} +else +{ +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_95, 0); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_82 = x_98; +goto block_94; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_77, 0); +lean_inc(x_99); +lean_dec(x_77); +x_100 = lean_unsigned_to_nat(1u); +x_101 = lean_nat_sub(x_99, x_100); +x_102 = l_Array_get_x3f___rarg(x_75, x_101); +lean_dec(x_101); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +lean_dec(x_99); +lean_dec(x_75); +lean_dec(x_73); +x_103 = lean_ctor_get(x_4, 0); +lean_inc(x_103); +x_82 = x_103; +goto block_94; +} +else +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_102, 0); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Array_get_x3f___rarg(x_75, x_99); +lean_dec(x_99); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; +x_106 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; +lean_dec(x_104); +lean_dec(x_73); +x_107 = lean_ctor_get(x_4, 0); +lean_inc(x_107); +x_82 = x_107; +goto block_94; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_108 = lean_ctor_get(x_106, 0); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_nat_abs(x_111); +lean_dec(x_111); +x_113 = lean_nat_to_int(x_112); +x_114 = lean_int_sub(x_109, x_113); +lean_dec(x_113); +lean_dec(x_109); +x_115 = lean_int_dec_lt(x_73, x_114); +lean_dec(x_114); +lean_dec(x_73); +if (x_115 == 0) +{ +lean_object* x_116; +lean_dec(x_110); +x_116 = lean_ctor_get(x_108, 1); +lean_inc(x_116); +lean_dec(x_108); +x_82 = x_116; +goto block_94; +} +else +{ +lean_dec(x_108); +x_82 = x_110; +goto block_94; +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +lean_dec(x_75); +x_117 = lean_ctor_get(x_105, 0); +lean_inc(x_117); +lean_dec(x_105); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_104, 1); +lean_inc(x_119); +lean_dec(x_104); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_nat_abs(x_120); +lean_dec(x_120); +x_122 = lean_nat_to_int(x_121); +x_123 = lean_int_sub(x_118, x_122); +lean_dec(x_122); +lean_dec(x_118); +x_124 = lean_int_dec_lt(x_73, x_123); +lean_dec(x_123); +lean_dec(x_73); +if (x_124 == 0) +{ +lean_object* x_125; +lean_dec(x_119); +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +lean_dec(x_117); +x_82 = x_125; +goto block_94; +} +else +{ +lean_dec(x_117); +x_82 = x_119; +goto block_94; +} +} +} +} +block_94: +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_83 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_82); +lean_dec(x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_int_neg(x_84); +x_86 = lean_int_mul(x_85, x_78); +lean_dec(x_85); +x_87 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_88 = lean_int_add(x_86, x_87); +lean_dec(x_86); +x_89 = lean_int_add(x_81, x_88); +lean_dec(x_88); +lean_dec(x_81); +x_90 = l_Std_Time_Duration_ofNanoseconds(x_89); +lean_dec(x_89); +lean_inc(x_90); +x_91 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_91, 0, x_90); +lean_closure_set(x_91, 1, x_84); +lean_closure_set(x_91, 2, x_78); +x_92 = lean_mk_thunk(x_91); +if (lean_is_scalar(x_5)) { + x_93 = lean_alloc_ctor(0, 4, 0); +} else { + x_93 = x_5; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +lean_ctor_set(x_93, 2, x_4); +lean_ctor_set(x_93, 3, x_83); +return x_93; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withDaysRollOver___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_withDaysRollOver(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMonthClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 0); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 2); +x_12 = lean_ctor_get(x_8, 1); +lean_dec(x_12); +x_13 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_14 = lean_int_mod(x_10, x_13); +x_15 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +if (x_16 == 0) +{ +uint8_t x_73; lean_object* x_74; uint8_t x_75; +x_73 = 0; +x_74 = l_Std_Time_Month_Ordinal_days(x_73, x_2); +x_75 = lean_int_dec_lt(x_74, x_11); +if (x_75 == 0) +{ +lean_dec(x_74); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_74); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +} +else +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; +x_76 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_77 = lean_int_mod(x_10, x_76); +x_78 = lean_int_dec_eq(x_77, x_15); +lean_dec(x_77); +x_79 = l_instDecidableNot___rarg(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_81 = lean_int_mod(x_10, x_80); +x_82 = lean_int_dec_eq(x_81, x_15); +lean_dec(x_81); +if (x_82 == 0) +{ +uint8_t x_83; lean_object* x_84; uint8_t x_85; +x_83 = 0; +x_84 = l_Std_Time_Month_Ordinal_days(x_83, x_2); +x_85 = lean_int_dec_lt(x_84, x_11); +if (x_85 == 0) +{ +lean_dec(x_84); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_84); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +} +else +{ +uint8_t x_86; lean_object* x_87; uint8_t x_88; +x_86 = 1; +x_87 = l_Std_Time_Month_Ordinal_days(x_86, x_2); +x_88 = lean_int_dec_lt(x_87, x_11); +if (x_88 == 0) +{ +lean_dec(x_87); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_87); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +} +} +else +{ +uint8_t x_89; lean_object* x_90; uint8_t x_91; +x_89 = 1; +x_90 = l_Std_Time_Month_Ordinal_days(x_89, x_2); +x_91 = lean_int_dec_lt(x_90, x_11); +if (x_91 == 0) +{ +lean_dec(x_90); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_90); +lean_ctor_set(x_8, 1, x_2); +x_18 = x_6; +goto block_72; +} +} +} +block_72: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_inc(x_20); +x_21 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_21, 0, x_20); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Array_findIdx_x3f_loop___rarg(x_21, x_17, x_22); +x_24 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_25 = lean_int_mul(x_20, x_24); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_41; +lean_dec(x_20); +x_41 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_28 = x_42; +goto block_40; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_28 = x_44; +goto block_40; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_23, 0); +lean_inc(x_45); +lean_dec(x_23); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_sub(x_45, x_46); +x_48 = l_Array_get_x3f___rarg(x_17, x_47); +lean_dec(x_47); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; +lean_dec(x_45); +lean_dec(x_20); +lean_dec(x_17); +x_49 = lean_ctor_get(x_4, 0); +lean_inc(x_49); +x_28 = x_49; +goto block_40; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 0); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Array_get_x3f___rarg(x_17, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; +x_52 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; +lean_dec(x_50); +lean_dec(x_20); +x_53 = lean_ctor_get(x_4, 0); +lean_inc(x_53); +x_28 = x_53; +goto block_40; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); +lean_dec(x_50); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_nat_abs(x_57); +lean_dec(x_57); +x_59 = lean_nat_to_int(x_58); +x_60 = lean_int_sub(x_55, x_59); +lean_dec(x_59); +lean_dec(x_55); +x_61 = lean_int_dec_lt(x_20, x_60); +lean_dec(x_60); +lean_dec(x_20); +if (x_61 == 0) +{ +lean_object* x_62; +lean_dec(x_56); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_28 = x_62; +goto block_40; +} +else +{ +lean_dec(x_54); +x_28 = x_56; +goto block_40; +} +} +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +lean_dec(x_17); +x_63 = lean_ctor_get(x_51, 0); +lean_inc(x_63); +lean_dec(x_51); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_dec(x_50); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_nat_abs(x_66); +lean_dec(x_66); +x_68 = lean_nat_to_int(x_67); +x_69 = lean_int_sub(x_64, x_68); +lean_dec(x_68); +lean_dec(x_64); +x_70 = lean_int_dec_lt(x_20, x_69); +lean_dec(x_69); +lean_dec(x_20); +if (x_70 == 0) +{ +lean_object* x_71; +lean_dec(x_65); +x_71 = lean_ctor_get(x_63, 1); +lean_inc(x_71); +lean_dec(x_63); +x_28 = x_71; +goto block_40; +} +else +{ +lean_dec(x_63); +x_28 = x_65; +goto block_40; +} +} +} +} +block_40: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_int_neg(x_30); +x_32 = lean_int_mul(x_31, x_24); +lean_dec(x_31); +x_33 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_34 = lean_int_add(x_32, x_33); +lean_dec(x_32); +x_35 = lean_int_add(x_27, x_34); +lean_dec(x_34); +lean_dec(x_27); +x_36 = l_Std_Time_Duration_ofNanoseconds(x_35); +lean_dec(x_35); +lean_inc(x_36); +x_37 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_37, 0, x_36); +lean_closure_set(x_37, 1, x_30); +lean_closure_set(x_37, 2, x_24); +x_38 = lean_mk_thunk(x_37); +if (lean_is_scalar(x_5)) { + x_39 = lean_alloc_ctor(0, 4, 0); +} else { + x_39 = x_5; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +lean_ctor_set(x_39, 2, x_4); +lean_ctor_set(x_39, 3, x_29); +return x_39; +} +} +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +x_92 = lean_ctor_get(x_8, 0); +x_93 = lean_ctor_get(x_8, 2); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_8); +x_94 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_95 = lean_int_mod(x_92, x_94); +x_96 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_97 = lean_int_dec_eq(x_95, x_96); +lean_dec(x_95); +x_98 = lean_ctor_get(x_4, 1); +lean_inc(x_98); +if (x_97 == 0) +{ +uint8_t x_154; lean_object* x_155; uint8_t x_156; +x_154 = 0; +x_155 = l_Std_Time_Month_Ordinal_days(x_154, x_2); +x_156 = lean_int_dec_lt(x_155, x_93); +if (x_156 == 0) +{ +lean_object* x_157; +lean_dec(x_155); +x_157 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_157, 0, x_92); +lean_ctor_set(x_157, 1, x_2); +lean_ctor_set(x_157, 2, x_93); +lean_ctor_set(x_6, 0, x_157); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_158; +lean_dec(x_93); +x_158 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_158, 0, x_92); +lean_ctor_set(x_158, 1, x_2); +lean_ctor_set(x_158, 2, x_155); +lean_ctor_set(x_6, 0, x_158); +x_99 = x_6; +goto block_153; +} +} +else +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; uint8_t x_162; +x_159 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_160 = lean_int_mod(x_92, x_159); +x_161 = lean_int_dec_eq(x_160, x_96); +lean_dec(x_160); +x_162 = l_instDecidableNot___rarg(x_161); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; uint8_t x_165; +x_163 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_164 = lean_int_mod(x_92, x_163); +x_165 = lean_int_dec_eq(x_164, x_96); +lean_dec(x_164); +if (x_165 == 0) +{ +uint8_t x_166; lean_object* x_167; uint8_t x_168; +x_166 = 0; +x_167 = l_Std_Time_Month_Ordinal_days(x_166, x_2); +x_168 = lean_int_dec_lt(x_167, x_93); +if (x_168 == 0) +{ +lean_object* x_169; +lean_dec(x_167); +x_169 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_169, 0, x_92); +lean_ctor_set(x_169, 1, x_2); +lean_ctor_set(x_169, 2, x_93); +lean_ctor_set(x_6, 0, x_169); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_170; +lean_dec(x_93); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_92); +lean_ctor_set(x_170, 1, x_2); +lean_ctor_set(x_170, 2, x_167); +lean_ctor_set(x_6, 0, x_170); +x_99 = x_6; +goto block_153; +} +} +else +{ +uint8_t x_171; lean_object* x_172; uint8_t x_173; +x_171 = 1; +x_172 = l_Std_Time_Month_Ordinal_days(x_171, x_2); +x_173 = lean_int_dec_lt(x_172, x_93); +if (x_173 == 0) +{ +lean_object* x_174; +lean_dec(x_172); +x_174 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_174, 0, x_92); +lean_ctor_set(x_174, 1, x_2); +lean_ctor_set(x_174, 2, x_93); +lean_ctor_set(x_6, 0, x_174); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_175; +lean_dec(x_93); +x_175 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_175, 0, x_92); +lean_ctor_set(x_175, 1, x_2); +lean_ctor_set(x_175, 2, x_172); +lean_ctor_set(x_6, 0, x_175); +x_99 = x_6; +goto block_153; +} +} +} +else +{ +uint8_t x_176; lean_object* x_177; uint8_t x_178; +x_176 = 1; +x_177 = l_Std_Time_Month_Ordinal_days(x_176, x_2); +x_178 = lean_int_dec_lt(x_177, x_93); +if (x_178 == 0) +{ +lean_object* x_179; +lean_dec(x_177); +x_179 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_179, 0, x_92); +lean_ctor_set(x_179, 1, x_2); +lean_ctor_set(x_179, 2, x_93); +lean_ctor_set(x_6, 0, x_179); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_180; +lean_dec(x_93); +x_180 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_180, 0, x_92); +lean_ctor_set(x_180, 1, x_2); +lean_ctor_set(x_180, 2, x_177); +lean_ctor_set(x_6, 0, x_180); +x_99 = x_6; +goto block_153; +} +} +} +block_153: +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_100 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +lean_inc(x_101); +x_102 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_102, 0, x_101); +x_103 = lean_unsigned_to_nat(0u); +x_104 = l_Array_findIdx_x3f_loop___rarg(x_102, x_98, x_103); +x_105 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_106 = lean_int_mul(x_101, x_105); +x_107 = lean_ctor_get(x_100, 1); +lean_inc(x_107); +lean_dec(x_100); +x_108 = lean_int_add(x_106, x_107); +lean_dec(x_107); +lean_dec(x_106); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_122; +lean_dec(x_101); +x_122 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_122) == 0) +{ +lean_object* x_123; +x_123 = lean_ctor_get(x_4, 0); +lean_inc(x_123); +x_109 = x_123; +goto block_121; +} +else +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_122, 0); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_109 = x_125; +goto block_121; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_104, 0); +lean_inc(x_126); +lean_dec(x_104); +x_127 = lean_unsigned_to_nat(1u); +x_128 = lean_nat_sub(x_126, x_127); +x_129 = l_Array_get_x3f___rarg(x_98, x_128); +lean_dec(x_128); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; +lean_dec(x_126); +lean_dec(x_101); +lean_dec(x_98); +x_130 = lean_ctor_get(x_4, 0); +lean_inc(x_130); +x_109 = x_130; +goto block_121; +} +else +{ +lean_object* x_131; lean_object* x_132; +x_131 = lean_ctor_get(x_129, 0); +lean_inc(x_131); +lean_dec(x_129); +x_132 = l_Array_get_x3f___rarg(x_98, x_126); +lean_dec(x_126); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; +x_133 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; +lean_dec(x_131); +lean_dec(x_101); +x_134 = lean_ctor_get(x_4, 0); +lean_inc(x_134); +x_109 = x_134; +goto block_121; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_135 = lean_ctor_get(x_133, 0); +lean_inc(x_135); +lean_dec(x_133); +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +lean_dec(x_131); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_nat_abs(x_138); +lean_dec(x_138); +x_140 = lean_nat_to_int(x_139); +x_141 = lean_int_sub(x_136, x_140); +lean_dec(x_140); +lean_dec(x_136); +x_142 = lean_int_dec_lt(x_101, x_141); +lean_dec(x_141); +lean_dec(x_101); +if (x_142 == 0) +{ +lean_object* x_143; +lean_dec(x_137); +x_143 = lean_ctor_get(x_135, 1); +lean_inc(x_143); +lean_dec(x_135); +x_109 = x_143; +goto block_121; +} +else +{ +lean_dec(x_135); +x_109 = x_137; +goto block_121; +} +} +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +lean_dec(x_98); +x_144 = lean_ctor_get(x_132, 0); +lean_inc(x_144); +lean_dec(x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_131, 1); +lean_inc(x_146); +lean_dec(x_131); +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_nat_abs(x_147); +lean_dec(x_147); +x_149 = lean_nat_to_int(x_148); +x_150 = lean_int_sub(x_145, x_149); +lean_dec(x_149); +lean_dec(x_145); +x_151 = lean_int_dec_lt(x_101, x_150); +lean_dec(x_150); +lean_dec(x_101); +if (x_151 == 0) +{ +lean_object* x_152; +lean_dec(x_146); +x_152 = lean_ctor_get(x_144, 1); +lean_inc(x_152); +lean_dec(x_144); +x_109 = x_152; +goto block_121; +} +else +{ +lean_dec(x_144); +x_109 = x_146; +goto block_121; +} +} +} +} +block_121: +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_110 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_109); +lean_dec(x_109); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_int_neg(x_111); +x_113 = lean_int_mul(x_112, x_105); +lean_dec(x_112); +x_114 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_115 = lean_int_add(x_113, x_114); +lean_dec(x_113); +x_116 = lean_int_add(x_108, x_115); +lean_dec(x_115); +lean_dec(x_108); +x_117 = l_Std_Time_Duration_ofNanoseconds(x_116); +lean_dec(x_116); +lean_inc(x_117); +x_118 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_118, 0, x_117); +lean_closure_set(x_118, 1, x_111); +lean_closure_set(x_118, 2, x_105); +x_119 = lean_mk_thunk(x_118); +if (lean_is_scalar(x_5)) { + x_120 = lean_alloc_ctor(0, 4, 0); +} else { + x_120 = x_5; +} +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +lean_ctor_set(x_120, 2, x_4); +lean_ctor_set(x_120, 3, x_110); +return x_120; +} +} +} +} +else +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; +x_181 = lean_ctor_get(x_6, 0); +x_182 = lean_ctor_get(x_6, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_6); +x_183 = lean_ctor_get(x_181, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_181, 2); +lean_inc(x_184); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + lean_ctor_release(x_181, 2); + x_185 = x_181; +} else { + lean_dec_ref(x_181); + x_185 = lean_box(0); +} +x_186 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_187 = lean_int_mod(x_183, x_186); +x_188 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_189 = lean_int_dec_eq(x_187, x_188); +lean_dec(x_187); +x_190 = lean_ctor_get(x_4, 1); +lean_inc(x_190); +if (x_189 == 0) +{ +uint8_t x_246; lean_object* x_247; uint8_t x_248; +x_246 = 0; +x_247 = l_Std_Time_Month_Ordinal_days(x_246, x_2); +x_248 = lean_int_dec_lt(x_247, x_184); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; +lean_dec(x_247); +if (lean_is_scalar(x_185)) { + x_249 = lean_alloc_ctor(0, 3, 0); +} else { + x_249 = x_185; +} +lean_ctor_set(x_249, 0, x_183); +lean_ctor_set(x_249, 1, x_2); +lean_ctor_set(x_249, 2, x_184); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_182); +x_191 = x_250; +goto block_245; +} +else +{ +lean_object* x_251; lean_object* x_252; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_251 = lean_alloc_ctor(0, 3, 0); +} else { + x_251 = x_185; +} +lean_ctor_set(x_251, 0, x_183); +lean_ctor_set(x_251, 1, x_2); +lean_ctor_set(x_251, 2, x_247); +x_252 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_182); +x_191 = x_252; +goto block_245; +} +} +else +{ +lean_object* x_253; lean_object* x_254; uint8_t x_255; uint8_t x_256; +x_253 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_254 = lean_int_mod(x_183, x_253); +x_255 = lean_int_dec_eq(x_254, x_188); +lean_dec(x_254); +x_256 = l_instDecidableNot___rarg(x_255); +if (x_256 == 0) +{ +lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_257 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_258 = lean_int_mod(x_183, x_257); +x_259 = lean_int_dec_eq(x_258, x_188); +lean_dec(x_258); +if (x_259 == 0) +{ +uint8_t x_260; lean_object* x_261; uint8_t x_262; +x_260 = 0; +x_261 = l_Std_Time_Month_Ordinal_days(x_260, x_2); +x_262 = lean_int_dec_lt(x_261, x_184); +if (x_262 == 0) +{ +lean_object* x_263; lean_object* x_264; +lean_dec(x_261); +if (lean_is_scalar(x_185)) { + x_263 = lean_alloc_ctor(0, 3, 0); +} else { + x_263 = x_185; +} +lean_ctor_set(x_263, 0, x_183); +lean_ctor_set(x_263, 1, x_2); +lean_ctor_set(x_263, 2, x_184); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_182); +x_191 = x_264; +goto block_245; +} +else +{ +lean_object* x_265; lean_object* x_266; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_265 = lean_alloc_ctor(0, 3, 0); +} else { + x_265 = x_185; +} +lean_ctor_set(x_265, 0, x_183); +lean_ctor_set(x_265, 1, x_2); +lean_ctor_set(x_265, 2, x_261); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_182); +x_191 = x_266; +goto block_245; +} +} +else +{ +uint8_t x_267; lean_object* x_268; uint8_t x_269; +x_267 = 1; +x_268 = l_Std_Time_Month_Ordinal_days(x_267, x_2); +x_269 = lean_int_dec_lt(x_268, x_184); +if (x_269 == 0) +{ +lean_object* x_270; lean_object* x_271; +lean_dec(x_268); +if (lean_is_scalar(x_185)) { + x_270 = lean_alloc_ctor(0, 3, 0); +} else { + x_270 = x_185; +} +lean_ctor_set(x_270, 0, x_183); +lean_ctor_set(x_270, 1, x_2); +lean_ctor_set(x_270, 2, x_184); +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_182); +x_191 = x_271; +goto block_245; +} +else +{ +lean_object* x_272; lean_object* x_273; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_272 = lean_alloc_ctor(0, 3, 0); +} else { + x_272 = x_185; +} +lean_ctor_set(x_272, 0, x_183); +lean_ctor_set(x_272, 1, x_2); +lean_ctor_set(x_272, 2, x_268); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_182); +x_191 = x_273; +goto block_245; +} +} +} +else +{ +uint8_t x_274; lean_object* x_275; uint8_t x_276; +x_274 = 1; +x_275 = l_Std_Time_Month_Ordinal_days(x_274, x_2); +x_276 = lean_int_dec_lt(x_275, x_184); +if (x_276 == 0) +{ +lean_object* x_277; lean_object* x_278; +lean_dec(x_275); +if (lean_is_scalar(x_185)) { + x_277 = lean_alloc_ctor(0, 3, 0); +} else { + x_277 = x_185; +} +lean_ctor_set(x_277, 0, x_183); +lean_ctor_set(x_277, 1, x_2); +lean_ctor_set(x_277, 2, x_184); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_182); +x_191 = x_278; +goto block_245; +} +else +{ +lean_object* x_279; lean_object* x_280; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_279 = lean_alloc_ctor(0, 3, 0); +} else { + x_279 = x_185; +} +lean_ctor_set(x_279, 0, x_183); +lean_ctor_set(x_279, 1, x_2); +lean_ctor_set(x_279, 2, x_275); +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_182); +x_191 = x_280; +goto block_245; +} +} +} +block_245: +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_192 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_191); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +lean_inc(x_193); +x_194 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_194, 0, x_193); +x_195 = lean_unsigned_to_nat(0u); +x_196 = l_Array_findIdx_x3f_loop___rarg(x_194, x_190, x_195); +x_197 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_198 = lean_int_mul(x_193, x_197); +x_199 = lean_ctor_get(x_192, 1); +lean_inc(x_199); +lean_dec(x_192); +x_200 = lean_int_add(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_214; +lean_dec(x_193); +x_214 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; +x_215 = lean_ctor_get(x_4, 0); +lean_inc(x_215); +x_201 = x_215; +goto block_213; +} +else +{ +lean_object* x_216; lean_object* x_217; +x_216 = lean_ctor_get(x_214, 0); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +lean_dec(x_216); +x_201 = x_217; +goto block_213; +} +} +else +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_218 = lean_ctor_get(x_196, 0); +lean_inc(x_218); +lean_dec(x_196); +x_219 = lean_unsigned_to_nat(1u); +x_220 = lean_nat_sub(x_218, x_219); +x_221 = l_Array_get_x3f___rarg(x_190, x_220); +lean_dec(x_220); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; +lean_dec(x_218); +lean_dec(x_193); +lean_dec(x_190); +x_222 = lean_ctor_get(x_4, 0); +lean_inc(x_222); +x_201 = x_222; +goto block_213; +} +else +{ +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_221, 0); +lean_inc(x_223); +lean_dec(x_221); +x_224 = l_Array_get_x3f___rarg(x_190, x_218); +lean_dec(x_218); +if (lean_obj_tag(x_224) == 0) +{ +lean_object* x_225; +x_225 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_225) == 0) +{ +lean_object* x_226; +lean_dec(x_223); +lean_dec(x_193); +x_226 = lean_ctor_get(x_4, 0); +lean_inc(x_226); +x_201 = x_226; +goto block_213; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; +x_227 = lean_ctor_get(x_225, 0); +lean_inc(x_227); +lean_dec(x_225); +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_223, 1); +lean_inc(x_229); +lean_dec(x_223); +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_nat_abs(x_230); +lean_dec(x_230); +x_232 = lean_nat_to_int(x_231); +x_233 = lean_int_sub(x_228, x_232); +lean_dec(x_232); +lean_dec(x_228); +x_234 = lean_int_dec_lt(x_193, x_233); +lean_dec(x_233); +lean_dec(x_193); +if (x_234 == 0) +{ +lean_object* x_235; +lean_dec(x_229); +x_235 = lean_ctor_get(x_227, 1); +lean_inc(x_235); +lean_dec(x_227); +x_201 = x_235; +goto block_213; +} +else +{ +lean_dec(x_227); +x_201 = x_229; +goto block_213; +} +} +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; +lean_dec(x_190); +x_236 = lean_ctor_get(x_224, 0); +lean_inc(x_236); +lean_dec(x_224); +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_223, 1); +lean_inc(x_238); +lean_dec(x_223); +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_nat_abs(x_239); +lean_dec(x_239); +x_241 = lean_nat_to_int(x_240); +x_242 = lean_int_sub(x_237, x_241); +lean_dec(x_241); +lean_dec(x_237); +x_243 = lean_int_dec_lt(x_193, x_242); +lean_dec(x_242); +lean_dec(x_193); +if (x_243 == 0) +{ +lean_object* x_244; +lean_dec(x_238); +x_244 = lean_ctor_get(x_236, 1); +lean_inc(x_244); +lean_dec(x_236); +x_201 = x_244; +goto block_213; +} +else +{ +lean_dec(x_236); +x_201 = x_238; +goto block_213; +} +} +} +} +block_213: +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_202 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_201); +lean_dec(x_201); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_int_neg(x_203); +x_205 = lean_int_mul(x_204, x_197); +lean_dec(x_204); +x_206 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_207 = lean_int_add(x_205, x_206); +lean_dec(x_205); +x_208 = lean_int_add(x_200, x_207); +lean_dec(x_207); +lean_dec(x_200); +x_209 = l_Std_Time_Duration_ofNanoseconds(x_208); +lean_dec(x_208); +lean_inc(x_209); +x_210 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_210, 0, x_209); +lean_closure_set(x_210, 1, x_203); +lean_closure_set(x_210, 2, x_197); +x_211 = lean_mk_thunk(x_210); +if (lean_is_scalar(x_5)) { + x_212 = lean_alloc_ctor(0, 4, 0); +} else { + x_212 = x_5; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_209); +lean_ctor_set(x_212, 2, x_4); +lean_ctor_set(x_212, 3, x_202); +return x_212; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMonthRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 2); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Std_Time_PlainDate_rollOver(x_9, x_2, x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_findIdx_x3f_loop___rarg(x_14, x_15, x_16); +x_18 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_13, x_18); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_35; +lean_dec(x_13); +x_35 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_4, 0); +lean_inc(x_36); +x_22 = x_36; +goto block_34; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_22 = x_38; +goto block_34; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_17, 0); +lean_inc(x_39); +lean_dec(x_17); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_sub(x_39, x_40); +x_42 = l_Array_get_x3f___rarg(x_15, x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +lean_dec(x_39); +lean_dec(x_15); +lean_dec(x_13); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_22 = x_43; +goto block_34; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Array_get_x3f___rarg(x_15, x_39); +lean_dec(x_39); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +x_46 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; +lean_dec(x_44); +lean_dec(x_13); +x_47 = lean_ctor_get(x_4, 0); +lean_inc(x_47); +x_22 = x_47; +goto block_34; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_dec(x_44); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_nat_abs(x_51); +lean_dec(x_51); +x_53 = lean_nat_to_int(x_52); +x_54 = lean_int_sub(x_49, x_53); +lean_dec(x_53); +lean_dec(x_49); +x_55 = lean_int_dec_lt(x_13, x_54); +lean_dec(x_54); +lean_dec(x_13); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_50); +x_56 = lean_ctor_get(x_48, 1); +lean_inc(x_56); +lean_dec(x_48); +x_22 = x_56; +goto block_34; +} +else +{ +lean_dec(x_48); +x_22 = x_50; +goto block_34; +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_15); +x_57 = lean_ctor_get(x_45, 0); +lean_inc(x_57); +lean_dec(x_45); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_44, 1); +lean_inc(x_59); +lean_dec(x_44); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_nat_abs(x_60); +lean_dec(x_60); +x_62 = lean_nat_to_int(x_61); +x_63 = lean_int_sub(x_58, x_62); +lean_dec(x_62); +lean_dec(x_58); +x_64 = lean_int_dec_lt(x_13, x_63); +lean_dec(x_63); +lean_dec(x_13); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_59); +x_65 = lean_ctor_get(x_57, 1); +lean_inc(x_65); +lean_dec(x_57); +x_22 = x_65; +goto block_34; +} +else +{ +lean_dec(x_57); +x_22 = x_59; +goto block_34; +} +} +} +} +block_34: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_23 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_22); +lean_dec(x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_int_neg(x_24); +x_26 = lean_int_mul(x_25, x_18); +lean_dec(x_25); +x_27 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_26); +x_29 = lean_int_add(x_21, x_28); +lean_dec(x_28); +lean_dec(x_21); +x_30 = l_Std_Time_Duration_ofNanoseconds(x_29); +lean_dec(x_29); +lean_inc(x_30); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_31, 0, x_30); +lean_closure_set(x_31, 1, x_24); +lean_closure_set(x_31, 2, x_18); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_23); +return x_33; +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_66 = lean_ctor_get(x_6, 0); +x_67 = lean_ctor_get(x_6, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_6); +x_68 = lean_ctor_get(x_66, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_66, 2); +lean_inc(x_69); +lean_dec(x_66); +x_70 = l_Std_Time_PlainDate_rollOver(x_68, x_2, x_69); +lean_dec(x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_67); +x_72 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_71); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +lean_inc(x_73); +x_74 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_74, 0, x_73); +x_75 = lean_ctor_get(x_4, 1); +lean_inc(x_75); +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Array_findIdx_x3f_loop___rarg(x_74, x_75, x_76); +x_78 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_79 = lean_int_mul(x_73, x_78); +x_80 = lean_ctor_get(x_72, 1); +lean_inc(x_80); +lean_dec(x_72); +x_81 = lean_int_add(x_79, x_80); +lean_dec(x_80); +lean_dec(x_79); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_95; +lean_dec(x_73); +x_95 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; +x_96 = lean_ctor_get(x_4, 0); +lean_inc(x_96); +x_82 = x_96; +goto block_94; +} +else +{ +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_95, 0); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_82 = x_98; +goto block_94; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_77, 0); +lean_inc(x_99); +lean_dec(x_77); +x_100 = lean_unsigned_to_nat(1u); +x_101 = lean_nat_sub(x_99, x_100); +x_102 = l_Array_get_x3f___rarg(x_75, x_101); +lean_dec(x_101); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +lean_dec(x_99); +lean_dec(x_75); +lean_dec(x_73); +x_103 = lean_ctor_get(x_4, 0); +lean_inc(x_103); +x_82 = x_103; +goto block_94; +} +else +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_102, 0); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Array_get_x3f___rarg(x_75, x_99); +lean_dec(x_99); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; +x_106 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; +lean_dec(x_104); +lean_dec(x_73); +x_107 = lean_ctor_get(x_4, 0); +lean_inc(x_107); +x_82 = x_107; +goto block_94; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_108 = lean_ctor_get(x_106, 0); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_nat_abs(x_111); +lean_dec(x_111); +x_113 = lean_nat_to_int(x_112); +x_114 = lean_int_sub(x_109, x_113); +lean_dec(x_113); +lean_dec(x_109); +x_115 = lean_int_dec_lt(x_73, x_114); +lean_dec(x_114); +lean_dec(x_73); +if (x_115 == 0) +{ +lean_object* x_116; +lean_dec(x_110); +x_116 = lean_ctor_get(x_108, 1); +lean_inc(x_116); +lean_dec(x_108); +x_82 = x_116; +goto block_94; +} +else +{ +lean_dec(x_108); +x_82 = x_110; +goto block_94; +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +lean_dec(x_75); +x_117 = lean_ctor_get(x_105, 0); +lean_inc(x_117); +lean_dec(x_105); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_104, 1); +lean_inc(x_119); +lean_dec(x_104); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_nat_abs(x_120); +lean_dec(x_120); +x_122 = lean_nat_to_int(x_121); +x_123 = lean_int_sub(x_118, x_122); +lean_dec(x_122); +lean_dec(x_118); +x_124 = lean_int_dec_lt(x_73, x_123); +lean_dec(x_123); +lean_dec(x_73); +if (x_124 == 0) +{ +lean_object* x_125; +lean_dec(x_119); +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +lean_dec(x_117); +x_82 = x_125; +goto block_94; +} +else +{ +lean_dec(x_117); +x_82 = x_119; +goto block_94; +} +} +} +} +block_94: +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_83 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_82); +lean_dec(x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_int_neg(x_84); +x_86 = lean_int_mul(x_85, x_78); +lean_dec(x_85); +x_87 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_88 = lean_int_add(x_86, x_87); +lean_dec(x_86); +x_89 = lean_int_add(x_81, x_88); +lean_dec(x_88); +lean_dec(x_81); +x_90 = l_Std_Time_Duration_ofNanoseconds(x_89); +lean_dec(x_89); +lean_inc(x_90); +x_91 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_91, 0, x_90); +lean_closure_set(x_91, 1, x_84); +lean_closure_set(x_91, 2, x_78); +x_92 = lean_mk_thunk(x_91); +if (lean_is_scalar(x_5)) { + x_93 = lean_alloc_ctor(0, 4, 0); +} else { + x_93 = x_5; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +lean_ctor_set(x_93, 2, x_4); +lean_ctor_set(x_93, 3, x_83); +return x_93; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withYearClip(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 0); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_8, 1); +x_11 = lean_ctor_get(x_8, 2); +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_14 = lean_int_mod(x_2, x_13); +x_15 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_16 = lean_int_dec_eq(x_14, x_15); +lean_dec(x_14); +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +if (x_16 == 0) +{ +uint8_t x_73; lean_object* x_74; uint8_t x_75; +x_73 = 0; +x_74 = l_Std_Time_Month_Ordinal_days(x_73, x_10); +x_75 = lean_int_dec_lt(x_74, x_11); +if (x_75 == 0) +{ +lean_dec(x_74); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_74); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +} +else +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; uint8_t x_79; +x_76 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_77 = lean_int_mod(x_2, x_76); +x_78 = lean_int_dec_eq(x_77, x_15); +lean_dec(x_77); +x_79 = l_instDecidableNot___rarg(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_80 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_81 = lean_int_mod(x_2, x_80); +x_82 = lean_int_dec_eq(x_81, x_15); +lean_dec(x_81); +if (x_82 == 0) +{ +uint8_t x_83; lean_object* x_84; uint8_t x_85; +x_83 = 0; +x_84 = l_Std_Time_Month_Ordinal_days(x_83, x_10); +x_85 = lean_int_dec_lt(x_84, x_11); +if (x_85 == 0) +{ +lean_dec(x_84); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_84); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +} +else +{ +uint8_t x_86; lean_object* x_87; uint8_t x_88; +x_86 = 1; +x_87 = l_Std_Time_Month_Ordinal_days(x_86, x_10); +x_88 = lean_int_dec_lt(x_87, x_11); +if (x_88 == 0) +{ +lean_dec(x_87); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_87); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +} +} +else +{ +uint8_t x_89; lean_object* x_90; uint8_t x_91; +x_89 = 1; +x_90 = l_Std_Time_Month_Ordinal_days(x_89, x_10); +x_91 = lean_int_dec_lt(x_90, x_11); +if (x_91 == 0) +{ +lean_dec(x_90); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +else +{ +lean_dec(x_11); +lean_ctor_set(x_8, 2, x_90); +lean_ctor_set(x_8, 0, x_2); +x_18 = x_6; +goto block_72; +} +} +} +block_72: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_19 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_18); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_inc(x_20); +x_21 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_21, 0, x_20); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Array_findIdx_x3f_loop___rarg(x_21, x_17, x_22); +x_24 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_25 = lean_int_mul(x_20, x_24); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_dec(x_19); +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_41; +lean_dec(x_20); +x_41 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_28 = x_42; +goto block_40; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_28 = x_44; +goto block_40; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_23, 0); +lean_inc(x_45); +lean_dec(x_23); +x_46 = lean_unsigned_to_nat(1u); +x_47 = lean_nat_sub(x_45, x_46); +x_48 = l_Array_get_x3f___rarg(x_17, x_47); +lean_dec(x_47); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; +lean_dec(x_45); +lean_dec(x_20); +lean_dec(x_17); +x_49 = lean_ctor_get(x_4, 0); +lean_inc(x_49); +x_28 = x_49; +goto block_40; +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 0); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Array_get_x3f___rarg(x_17, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; +x_52 = l_Array_back_x3f___rarg(x_17); +lean_dec(x_17); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; +lean_dec(x_50); +lean_dec(x_20); +x_53 = lean_ctor_get(x_4, 0); +lean_inc(x_53); +x_28 = x_53; +goto block_40; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_50, 1); +lean_inc(x_56); +lean_dec(x_50); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_nat_abs(x_57); +lean_dec(x_57); +x_59 = lean_nat_to_int(x_58); +x_60 = lean_int_sub(x_55, x_59); +lean_dec(x_59); +lean_dec(x_55); +x_61 = lean_int_dec_lt(x_20, x_60); +lean_dec(x_60); +lean_dec(x_20); +if (x_61 == 0) +{ +lean_object* x_62; +lean_dec(x_56); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_dec(x_54); +x_28 = x_62; +goto block_40; +} +else +{ +lean_dec(x_54); +x_28 = x_56; +goto block_40; +} +} +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +lean_dec(x_17); +x_63 = lean_ctor_get(x_51, 0); +lean_inc(x_63); +lean_dec(x_51); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_dec(x_50); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_nat_abs(x_66); +lean_dec(x_66); +x_68 = lean_nat_to_int(x_67); +x_69 = lean_int_sub(x_64, x_68); +lean_dec(x_68); +lean_dec(x_64); +x_70 = lean_int_dec_lt(x_20, x_69); +lean_dec(x_69); +lean_dec(x_20); +if (x_70 == 0) +{ +lean_object* x_71; +lean_dec(x_65); +x_71 = lean_ctor_get(x_63, 1); +lean_inc(x_71); +lean_dec(x_63); +x_28 = x_71; +goto block_40; +} +else +{ +lean_dec(x_63); +x_28 = x_65; +goto block_40; +} +} +} +} +block_40: +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_29 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_28); +lean_dec(x_28); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +x_31 = lean_int_neg(x_30); +x_32 = lean_int_mul(x_31, x_24); +lean_dec(x_31); +x_33 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_34 = lean_int_add(x_32, x_33); +lean_dec(x_32); +x_35 = lean_int_add(x_27, x_34); +lean_dec(x_34); +lean_dec(x_27); +x_36 = l_Std_Time_Duration_ofNanoseconds(x_35); +lean_dec(x_35); +lean_inc(x_36); +x_37 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_37, 0, x_36); +lean_closure_set(x_37, 1, x_30); +lean_closure_set(x_37, 2, x_24); +x_38 = lean_mk_thunk(x_37); +if (lean_is_scalar(x_5)) { + x_39 = lean_alloc_ctor(0, 4, 0); +} else { + x_39 = x_5; +} +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +lean_ctor_set(x_39, 2, x_4); +lean_ctor_set(x_39, 3, x_29); +return x_39; +} +} +} +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; lean_object* x_98; lean_object* x_99; +x_92 = lean_ctor_get(x_8, 1); +x_93 = lean_ctor_get(x_8, 2); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_8); +x_94 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_95 = lean_int_mod(x_2, x_94); +x_96 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_97 = lean_int_dec_eq(x_95, x_96); +lean_dec(x_95); +x_98 = lean_ctor_get(x_4, 1); +lean_inc(x_98); +if (x_97 == 0) +{ +uint8_t x_154; lean_object* x_155; uint8_t x_156; +x_154 = 0; +x_155 = l_Std_Time_Month_Ordinal_days(x_154, x_92); +x_156 = lean_int_dec_lt(x_155, x_93); +if (x_156 == 0) +{ +lean_object* x_157; +lean_dec(x_155); +x_157 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_157, 0, x_2); +lean_ctor_set(x_157, 1, x_92); +lean_ctor_set(x_157, 2, x_93); +lean_ctor_set(x_6, 0, x_157); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_158; +lean_dec(x_93); +x_158 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_158, 0, x_2); +lean_ctor_set(x_158, 1, x_92); +lean_ctor_set(x_158, 2, x_155); +lean_ctor_set(x_6, 0, x_158); +x_99 = x_6; +goto block_153; +} +} +else +{ +lean_object* x_159; lean_object* x_160; uint8_t x_161; uint8_t x_162; +x_159 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_160 = lean_int_mod(x_2, x_159); +x_161 = lean_int_dec_eq(x_160, x_96); +lean_dec(x_160); +x_162 = l_instDecidableNot___rarg(x_161); +if (x_162 == 0) +{ +lean_object* x_163; lean_object* x_164; uint8_t x_165; +x_163 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_164 = lean_int_mod(x_2, x_163); +x_165 = lean_int_dec_eq(x_164, x_96); +lean_dec(x_164); +if (x_165 == 0) +{ +uint8_t x_166; lean_object* x_167; uint8_t x_168; +x_166 = 0; +x_167 = l_Std_Time_Month_Ordinal_days(x_166, x_92); +x_168 = lean_int_dec_lt(x_167, x_93); +if (x_168 == 0) +{ +lean_object* x_169; +lean_dec(x_167); +x_169 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_169, 0, x_2); +lean_ctor_set(x_169, 1, x_92); +lean_ctor_set(x_169, 2, x_93); +lean_ctor_set(x_6, 0, x_169); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_170; +lean_dec(x_93); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_2); +lean_ctor_set(x_170, 1, x_92); +lean_ctor_set(x_170, 2, x_167); +lean_ctor_set(x_6, 0, x_170); +x_99 = x_6; +goto block_153; +} +} +else +{ +uint8_t x_171; lean_object* x_172; uint8_t x_173; +x_171 = 1; +x_172 = l_Std_Time_Month_Ordinal_days(x_171, x_92); +x_173 = lean_int_dec_lt(x_172, x_93); +if (x_173 == 0) +{ +lean_object* x_174; +lean_dec(x_172); +x_174 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_174, 0, x_2); +lean_ctor_set(x_174, 1, x_92); +lean_ctor_set(x_174, 2, x_93); +lean_ctor_set(x_6, 0, x_174); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_175; +lean_dec(x_93); +x_175 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_175, 0, x_2); +lean_ctor_set(x_175, 1, x_92); +lean_ctor_set(x_175, 2, x_172); +lean_ctor_set(x_6, 0, x_175); +x_99 = x_6; +goto block_153; +} +} +} +else +{ +uint8_t x_176; lean_object* x_177; uint8_t x_178; +x_176 = 1; +x_177 = l_Std_Time_Month_Ordinal_days(x_176, x_92); +x_178 = lean_int_dec_lt(x_177, x_93); +if (x_178 == 0) +{ +lean_object* x_179; +lean_dec(x_177); +x_179 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_179, 0, x_2); +lean_ctor_set(x_179, 1, x_92); +lean_ctor_set(x_179, 2, x_93); +lean_ctor_set(x_6, 0, x_179); +x_99 = x_6; +goto block_153; +} +else +{ +lean_object* x_180; +lean_dec(x_93); +x_180 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_180, 0, x_2); +lean_ctor_set(x_180, 1, x_92); +lean_ctor_set(x_180, 2, x_177); +lean_ctor_set(x_6, 0, x_180); +x_99 = x_6; +goto block_153; +} +} +} +block_153: +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_100 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_99); +x_101 = lean_ctor_get(x_100, 0); +lean_inc(x_101); +lean_inc(x_101); +x_102 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_102, 0, x_101); +x_103 = lean_unsigned_to_nat(0u); +x_104 = l_Array_findIdx_x3f_loop___rarg(x_102, x_98, x_103); +x_105 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_106 = lean_int_mul(x_101, x_105); +x_107 = lean_ctor_get(x_100, 1); +lean_inc(x_107); +lean_dec(x_100); +x_108 = lean_int_add(x_106, x_107); +lean_dec(x_107); +lean_dec(x_106); +if (lean_obj_tag(x_104) == 0) +{ +lean_object* x_122; +lean_dec(x_101); +x_122 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_122) == 0) +{ +lean_object* x_123; +x_123 = lean_ctor_get(x_4, 0); +lean_inc(x_123); +x_109 = x_123; +goto block_121; +} +else +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_122, 0); +lean_inc(x_124); +lean_dec(x_122); +x_125 = lean_ctor_get(x_124, 1); +lean_inc(x_125); +lean_dec(x_124); +x_109 = x_125; +goto block_121; +} +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_104, 0); +lean_inc(x_126); +lean_dec(x_104); +x_127 = lean_unsigned_to_nat(1u); +x_128 = lean_nat_sub(x_126, x_127); +x_129 = l_Array_get_x3f___rarg(x_98, x_128); +lean_dec(x_128); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; +lean_dec(x_126); +lean_dec(x_101); +lean_dec(x_98); +x_130 = lean_ctor_get(x_4, 0); +lean_inc(x_130); +x_109 = x_130; +goto block_121; +} +else +{ +lean_object* x_131; lean_object* x_132; +x_131 = lean_ctor_get(x_129, 0); +lean_inc(x_131); +lean_dec(x_129); +x_132 = l_Array_get_x3f___rarg(x_98, x_126); +lean_dec(x_126); +if (lean_obj_tag(x_132) == 0) +{ +lean_object* x_133; +x_133 = l_Array_back_x3f___rarg(x_98); +lean_dec(x_98); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; +lean_dec(x_131); +lean_dec(x_101); +x_134 = lean_ctor_get(x_4, 0); +lean_inc(x_134); +x_109 = x_134; +goto block_121; +} +else +{ +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; uint8_t x_142; +x_135 = lean_ctor_get(x_133, 0); +lean_inc(x_135); +lean_dec(x_133); +x_136 = lean_ctor_get(x_135, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_131, 1); +lean_inc(x_137); +lean_dec(x_131); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_nat_abs(x_138); +lean_dec(x_138); +x_140 = lean_nat_to_int(x_139); +x_141 = lean_int_sub(x_136, x_140); +lean_dec(x_140); +lean_dec(x_136); +x_142 = lean_int_dec_lt(x_101, x_141); +lean_dec(x_141); +lean_dec(x_101); +if (x_142 == 0) +{ +lean_object* x_143; +lean_dec(x_137); +x_143 = lean_ctor_get(x_135, 1); +lean_inc(x_143); +lean_dec(x_135); +x_109 = x_143; +goto block_121; +} +else +{ +lean_dec(x_135); +x_109 = x_137; +goto block_121; +} +} +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; +lean_dec(x_98); +x_144 = lean_ctor_get(x_132, 0); +lean_inc(x_144); +lean_dec(x_132); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_131, 1); +lean_inc(x_146); +lean_dec(x_131); +x_147 = lean_ctor_get(x_146, 0); +lean_inc(x_147); +x_148 = lean_nat_abs(x_147); +lean_dec(x_147); +x_149 = lean_nat_to_int(x_148); +x_150 = lean_int_sub(x_145, x_149); +lean_dec(x_149); +lean_dec(x_145); +x_151 = lean_int_dec_lt(x_101, x_150); +lean_dec(x_150); +lean_dec(x_101); +if (x_151 == 0) +{ +lean_object* x_152; +lean_dec(x_146); +x_152 = lean_ctor_get(x_144, 1); +lean_inc(x_152); +lean_dec(x_144); +x_109 = x_152; +goto block_121; +} +else +{ +lean_dec(x_144); +x_109 = x_146; +goto block_121; +} +} +} +} +block_121: +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_110 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_109); +lean_dec(x_109); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_int_neg(x_111); +x_113 = lean_int_mul(x_112, x_105); +lean_dec(x_112); +x_114 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_115 = lean_int_add(x_113, x_114); +lean_dec(x_113); +x_116 = lean_int_add(x_108, x_115); +lean_dec(x_115); +lean_dec(x_108); +x_117 = l_Std_Time_Duration_ofNanoseconds(x_116); +lean_dec(x_116); +lean_inc(x_117); +x_118 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_118, 0, x_117); +lean_closure_set(x_118, 1, x_111); +lean_closure_set(x_118, 2, x_105); +x_119 = lean_mk_thunk(x_118); +if (lean_is_scalar(x_5)) { + x_120 = lean_alloc_ctor(0, 4, 0); +} else { + x_120 = x_5; +} +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +lean_ctor_set(x_120, 2, x_4); +lean_ctor_set(x_120, 3, x_110); +return x_120; +} +} +} +} +else +{ +lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; lean_object* x_190; lean_object* x_191; +x_181 = lean_ctor_get(x_6, 0); +x_182 = lean_ctor_get(x_6, 1); +lean_inc(x_182); +lean_inc(x_181); +lean_dec(x_6); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +x_184 = lean_ctor_get(x_181, 2); +lean_inc(x_184); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + lean_ctor_release(x_181, 2); + x_185 = x_181; +} else { + lean_dec_ref(x_181); + x_185 = lean_box(0); +} +x_186 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_187 = lean_int_mod(x_2, x_186); +x_188 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_189 = lean_int_dec_eq(x_187, x_188); +lean_dec(x_187); +x_190 = lean_ctor_get(x_4, 1); +lean_inc(x_190); +if (x_189 == 0) +{ +uint8_t x_246; lean_object* x_247; uint8_t x_248; +x_246 = 0; +x_247 = l_Std_Time_Month_Ordinal_days(x_246, x_183); +x_248 = lean_int_dec_lt(x_247, x_184); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; +lean_dec(x_247); +if (lean_is_scalar(x_185)) { + x_249 = lean_alloc_ctor(0, 3, 0); +} else { + x_249 = x_185; +} +lean_ctor_set(x_249, 0, x_2); +lean_ctor_set(x_249, 1, x_183); +lean_ctor_set(x_249, 2, x_184); +x_250 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_182); +x_191 = x_250; +goto block_245; +} +else +{ +lean_object* x_251; lean_object* x_252; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_251 = lean_alloc_ctor(0, 3, 0); +} else { + x_251 = x_185; +} +lean_ctor_set(x_251, 0, x_2); +lean_ctor_set(x_251, 1, x_183); +lean_ctor_set(x_251, 2, x_247); +x_252 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_182); +x_191 = x_252; +goto block_245; +} +} +else +{ +lean_object* x_253; lean_object* x_254; uint8_t x_255; uint8_t x_256; +x_253 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_254 = lean_int_mod(x_2, x_253); +x_255 = lean_int_dec_eq(x_254, x_188); +lean_dec(x_254); +x_256 = l_instDecidableNot___rarg(x_255); +if (x_256 == 0) +{ +lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_257 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_258 = lean_int_mod(x_2, x_257); +x_259 = lean_int_dec_eq(x_258, x_188); +lean_dec(x_258); +if (x_259 == 0) +{ +uint8_t x_260; lean_object* x_261; uint8_t x_262; +x_260 = 0; +x_261 = l_Std_Time_Month_Ordinal_days(x_260, x_183); +x_262 = lean_int_dec_lt(x_261, x_184); +if (x_262 == 0) +{ +lean_object* x_263; lean_object* x_264; +lean_dec(x_261); +if (lean_is_scalar(x_185)) { + x_263 = lean_alloc_ctor(0, 3, 0); +} else { + x_263 = x_185; +} +lean_ctor_set(x_263, 0, x_2); +lean_ctor_set(x_263, 1, x_183); +lean_ctor_set(x_263, 2, x_184); +x_264 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_264, 0, x_263); +lean_ctor_set(x_264, 1, x_182); +x_191 = x_264; +goto block_245; +} +else +{ +lean_object* x_265; lean_object* x_266; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_265 = lean_alloc_ctor(0, 3, 0); +} else { + x_265 = x_185; +} +lean_ctor_set(x_265, 0, x_2); +lean_ctor_set(x_265, 1, x_183); +lean_ctor_set(x_265, 2, x_261); +x_266 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_182); +x_191 = x_266; +goto block_245; +} +} +else +{ +uint8_t x_267; lean_object* x_268; uint8_t x_269; +x_267 = 1; +x_268 = l_Std_Time_Month_Ordinal_days(x_267, x_183); +x_269 = lean_int_dec_lt(x_268, x_184); +if (x_269 == 0) +{ +lean_object* x_270; lean_object* x_271; +lean_dec(x_268); +if (lean_is_scalar(x_185)) { + x_270 = lean_alloc_ctor(0, 3, 0); +} else { + x_270 = x_185; +} +lean_ctor_set(x_270, 0, x_2); +lean_ctor_set(x_270, 1, x_183); +lean_ctor_set(x_270, 2, x_184); +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_182); +x_191 = x_271; +goto block_245; +} +else +{ +lean_object* x_272; lean_object* x_273; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_272 = lean_alloc_ctor(0, 3, 0); +} else { + x_272 = x_185; +} +lean_ctor_set(x_272, 0, x_2); +lean_ctor_set(x_272, 1, x_183); +lean_ctor_set(x_272, 2, x_268); +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_272); +lean_ctor_set(x_273, 1, x_182); +x_191 = x_273; +goto block_245; +} +} +} +else +{ +uint8_t x_274; lean_object* x_275; uint8_t x_276; +x_274 = 1; +x_275 = l_Std_Time_Month_Ordinal_days(x_274, x_183); +x_276 = lean_int_dec_lt(x_275, x_184); +if (x_276 == 0) +{ +lean_object* x_277; lean_object* x_278; +lean_dec(x_275); +if (lean_is_scalar(x_185)) { + x_277 = lean_alloc_ctor(0, 3, 0); +} else { + x_277 = x_185; +} +lean_ctor_set(x_277, 0, x_2); +lean_ctor_set(x_277, 1, x_183); +lean_ctor_set(x_277, 2, x_184); +x_278 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_182); +x_191 = x_278; +goto block_245; +} +else +{ +lean_object* x_279; lean_object* x_280; +lean_dec(x_184); +if (lean_is_scalar(x_185)) { + x_279 = lean_alloc_ctor(0, 3, 0); +} else { + x_279 = x_185; +} +lean_ctor_set(x_279, 0, x_2); +lean_ctor_set(x_279, 1, x_183); +lean_ctor_set(x_279, 2, x_275); +x_280 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_280, 0, x_279); +lean_ctor_set(x_280, 1, x_182); +x_191 = x_280; +goto block_245; +} +} +} +block_245: +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_192 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_191); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +lean_inc(x_193); +x_194 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_194, 0, x_193); +x_195 = lean_unsigned_to_nat(0u); +x_196 = l_Array_findIdx_x3f_loop___rarg(x_194, x_190, x_195); +x_197 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_198 = lean_int_mul(x_193, x_197); +x_199 = lean_ctor_get(x_192, 1); +lean_inc(x_199); +lean_dec(x_192); +x_200 = lean_int_add(x_198, x_199); +lean_dec(x_199); +lean_dec(x_198); +if (lean_obj_tag(x_196) == 0) +{ +lean_object* x_214; +lean_dec(x_193); +x_214 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_214) == 0) +{ +lean_object* x_215; +x_215 = lean_ctor_get(x_4, 0); +lean_inc(x_215); +x_201 = x_215; +goto block_213; +} +else +{ +lean_object* x_216; lean_object* x_217; +x_216 = lean_ctor_get(x_214, 0); +lean_inc(x_216); +lean_dec(x_214); +x_217 = lean_ctor_get(x_216, 1); +lean_inc(x_217); +lean_dec(x_216); +x_201 = x_217; +goto block_213; +} +} +else +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +x_218 = lean_ctor_get(x_196, 0); +lean_inc(x_218); +lean_dec(x_196); +x_219 = lean_unsigned_to_nat(1u); +x_220 = lean_nat_sub(x_218, x_219); +x_221 = l_Array_get_x3f___rarg(x_190, x_220); +lean_dec(x_220); +if (lean_obj_tag(x_221) == 0) +{ +lean_object* x_222; +lean_dec(x_218); +lean_dec(x_193); +lean_dec(x_190); +x_222 = lean_ctor_get(x_4, 0); +lean_inc(x_222); +x_201 = x_222; +goto block_213; +} +else +{ +lean_object* x_223; lean_object* x_224; +x_223 = lean_ctor_get(x_221, 0); +lean_inc(x_223); +lean_dec(x_221); +x_224 = l_Array_get_x3f___rarg(x_190, x_218); +lean_dec(x_218); +if (lean_obj_tag(x_224) == 0) +{ +lean_object* x_225; +x_225 = l_Array_back_x3f___rarg(x_190); +lean_dec(x_190); +if (lean_obj_tag(x_225) == 0) +{ +lean_object* x_226; +lean_dec(x_223); +lean_dec(x_193); +x_226 = lean_ctor_get(x_4, 0); +lean_inc(x_226); +x_201 = x_226; +goto block_213; +} +else +{ +lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; uint8_t x_234; +x_227 = lean_ctor_get(x_225, 0); +lean_inc(x_227); +lean_dec(x_225); +x_228 = lean_ctor_get(x_227, 0); +lean_inc(x_228); +x_229 = lean_ctor_get(x_223, 1); +lean_inc(x_229); +lean_dec(x_223); +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_nat_abs(x_230); +lean_dec(x_230); +x_232 = lean_nat_to_int(x_231); +x_233 = lean_int_sub(x_228, x_232); +lean_dec(x_232); +lean_dec(x_228); +x_234 = lean_int_dec_lt(x_193, x_233); +lean_dec(x_233); +lean_dec(x_193); +if (x_234 == 0) +{ +lean_object* x_235; +lean_dec(x_229); +x_235 = lean_ctor_get(x_227, 1); +lean_inc(x_235); +lean_dec(x_227); +x_201 = x_235; +goto block_213; +} +else +{ +lean_dec(x_227); +x_201 = x_229; +goto block_213; +} +} +} +else +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; +lean_dec(x_190); +x_236 = lean_ctor_get(x_224, 0); +lean_inc(x_236); +lean_dec(x_224); +x_237 = lean_ctor_get(x_236, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_223, 1); +lean_inc(x_238); +lean_dec(x_223); +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_nat_abs(x_239); +lean_dec(x_239); +x_241 = lean_nat_to_int(x_240); +x_242 = lean_int_sub(x_237, x_241); +lean_dec(x_241); +lean_dec(x_237); +x_243 = lean_int_dec_lt(x_193, x_242); +lean_dec(x_242); +lean_dec(x_193); +if (x_243 == 0) +{ +lean_object* x_244; +lean_dec(x_238); +x_244 = lean_ctor_get(x_236, 1); +lean_inc(x_244); +lean_dec(x_236); +x_201 = x_244; +goto block_213; +} +else +{ +lean_dec(x_236); +x_201 = x_238; +goto block_213; +} +} +} +} +block_213: +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +x_202 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_201); +lean_dec(x_201); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_int_neg(x_203); +x_205 = lean_int_mul(x_204, x_197); +lean_dec(x_204); +x_206 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_207 = lean_int_add(x_205, x_206); +lean_dec(x_205); +x_208 = lean_int_add(x_200, x_207); +lean_dec(x_207); +lean_dec(x_200); +x_209 = l_Std_Time_Duration_ofNanoseconds(x_208); +lean_dec(x_208); +lean_inc(x_209); +x_210 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_210, 0, x_209); +lean_closure_set(x_210, 1, x_203); +lean_closure_set(x_210, 2, x_197); +x_211 = lean_mk_thunk(x_210); +if (lean_is_scalar(x_5)) { + x_212 = lean_alloc_ctor(0, 4, 0); +} else { + x_212 = x_5; +} +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_209); +lean_ctor_set(x_212, 2, x_4); +lean_ctor_set(x_212, 3, x_202); +return x_212; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withYearRollOver(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 2); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Std_Time_PlainDate_rollOver(x_2, x_9, x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_11); +x_12 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_inc(x_13); +x_14 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_ctor_get(x_4, 1); +lean_inc(x_15); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Array_findIdx_x3f_loop___rarg(x_14, x_15, x_16); +x_18 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_19 = lean_int_mul(x_13, x_18); +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_int_add(x_19, x_20); +lean_dec(x_20); +lean_dec(x_19); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_35; +lean_dec(x_13); +x_35 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_35) == 0) +{ +lean_object* x_36; +x_36 = lean_ctor_get(x_4, 0); +lean_inc(x_36); +x_22 = x_36; +goto block_34; +} +else +{ +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_22 = x_38; +goto block_34; +} +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_17, 0); +lean_inc(x_39); +lean_dec(x_17); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_sub(x_39, x_40); +x_42 = l_Array_get_x3f___rarg(x_15, x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 0) +{ +lean_object* x_43; +lean_dec(x_39); +lean_dec(x_15); +lean_dec(x_13); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_22 = x_43; +goto block_34; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Array_get_x3f___rarg(x_15, x_39); +lean_dec(x_39); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +x_46 = l_Array_back_x3f___rarg(x_15); +lean_dec(x_15); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; +lean_dec(x_44); +lean_dec(x_13); +x_47 = lean_ctor_get(x_4, 0); +lean_inc(x_47); +x_22 = x_47; +goto block_34; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_44, 1); +lean_inc(x_50); +lean_dec(x_44); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +x_52 = lean_nat_abs(x_51); +lean_dec(x_51); +x_53 = lean_nat_to_int(x_52); +x_54 = lean_int_sub(x_49, x_53); +lean_dec(x_53); +lean_dec(x_49); +x_55 = lean_int_dec_lt(x_13, x_54); +lean_dec(x_54); +lean_dec(x_13); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_50); +x_56 = lean_ctor_get(x_48, 1); +lean_inc(x_56); +lean_dec(x_48); +x_22 = x_56; +goto block_34; +} +else +{ +lean_dec(x_48); +x_22 = x_50; +goto block_34; +} +} +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_15); +x_57 = lean_ctor_get(x_45, 0); +lean_inc(x_57); +lean_dec(x_45); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_44, 1); +lean_inc(x_59); +lean_dec(x_44); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_nat_abs(x_60); +lean_dec(x_60); +x_62 = lean_nat_to_int(x_61); +x_63 = lean_int_sub(x_58, x_62); +lean_dec(x_62); +lean_dec(x_58); +x_64 = lean_int_dec_lt(x_13, x_63); +lean_dec(x_63); +lean_dec(x_13); +if (x_64 == 0) +{ +lean_object* x_65; +lean_dec(x_59); +x_65 = lean_ctor_get(x_57, 1); +lean_inc(x_65); +lean_dec(x_57); +x_22 = x_65; +goto block_34; +} +else +{ +lean_dec(x_57); +x_22 = x_59; +goto block_34; +} +} +} +} +block_34: +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_23 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_22); +lean_dec(x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_int_neg(x_24); +x_26 = lean_int_mul(x_25, x_18); +lean_dec(x_25); +x_27 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_28 = lean_int_add(x_26, x_27); +lean_dec(x_26); +x_29 = lean_int_add(x_21, x_28); +lean_dec(x_28); +lean_dec(x_21); +x_30 = l_Std_Time_Duration_ofNanoseconds(x_29); +lean_dec(x_29); +lean_inc(x_30); +x_31 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_31, 0, x_30); +lean_closure_set(x_31, 1, x_24); +lean_closure_set(x_31, 2, x_18); +x_32 = lean_mk_thunk(x_31); +if (lean_is_scalar(x_5)) { + x_33 = lean_alloc_ctor(0, 4, 0); +} else { + x_33 = x_5; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +lean_ctor_set(x_33, 2, x_4); +lean_ctor_set(x_33, 3, x_23); +return x_33; +} +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_66 = lean_ctor_get(x_6, 0); +x_67 = lean_ctor_get(x_6, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_6); +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +x_69 = lean_ctor_get(x_66, 2); +lean_inc(x_69); +lean_dec(x_66); +x_70 = l_Std_Time_PlainDate_rollOver(x_2, x_68, x_69); +lean_dec(x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_67); +x_72 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_71); +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +lean_inc(x_73); +x_74 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_74, 0, x_73); +x_75 = lean_ctor_get(x_4, 1); +lean_inc(x_75); +x_76 = lean_unsigned_to_nat(0u); +x_77 = l_Array_findIdx_x3f_loop___rarg(x_74, x_75, x_76); +x_78 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_79 = lean_int_mul(x_73, x_78); +x_80 = lean_ctor_get(x_72, 1); +lean_inc(x_80); +lean_dec(x_72); +x_81 = lean_int_add(x_79, x_80); +lean_dec(x_80); +lean_dec(x_79); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_95; +lean_dec(x_73); +x_95 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; +x_96 = lean_ctor_get(x_4, 0); +lean_inc(x_96); +x_82 = x_96; +goto block_94; +} +else +{ +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_95, 0); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_ctor_get(x_97, 1); +lean_inc(x_98); +lean_dec(x_97); +x_82 = x_98; +goto block_94; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_99 = lean_ctor_get(x_77, 0); +lean_inc(x_99); +lean_dec(x_77); +x_100 = lean_unsigned_to_nat(1u); +x_101 = lean_nat_sub(x_99, x_100); +x_102 = l_Array_get_x3f___rarg(x_75, x_101); +lean_dec(x_101); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +lean_dec(x_99); +lean_dec(x_75); +lean_dec(x_73); +x_103 = lean_ctor_get(x_4, 0); +lean_inc(x_103); +x_82 = x_103; +goto block_94; +} +else +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_102, 0); +lean_inc(x_104); +lean_dec(x_102); +x_105 = l_Array_get_x3f___rarg(x_75, x_99); +lean_dec(x_99); +if (lean_obj_tag(x_105) == 0) +{ +lean_object* x_106; +x_106 = l_Array_back_x3f___rarg(x_75); +lean_dec(x_75); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; +lean_dec(x_104); +lean_dec(x_73); +x_107 = lean_ctor_get(x_4, 0); +lean_inc(x_107); +x_82 = x_107; +goto block_94; +} +else +{ +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_108 = lean_ctor_get(x_106, 0); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +lean_dec(x_104); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_nat_abs(x_111); +lean_dec(x_111); +x_113 = lean_nat_to_int(x_112); +x_114 = lean_int_sub(x_109, x_113); +lean_dec(x_113); +lean_dec(x_109); +x_115 = lean_int_dec_lt(x_73, x_114); +lean_dec(x_114); +lean_dec(x_73); +if (x_115 == 0) +{ +lean_object* x_116; +lean_dec(x_110); +x_116 = lean_ctor_get(x_108, 1); +lean_inc(x_116); +lean_dec(x_108); +x_82 = x_116; +goto block_94; +} +else +{ +lean_dec(x_108); +x_82 = x_110; +goto block_94; +} +} +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +lean_dec(x_75); +x_117 = lean_ctor_get(x_105, 0); +lean_inc(x_117); +lean_dec(x_105); +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_104, 1); +lean_inc(x_119); +lean_dec(x_104); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_nat_abs(x_120); +lean_dec(x_120); +x_122 = lean_nat_to_int(x_121); +x_123 = lean_int_sub(x_118, x_122); +lean_dec(x_122); +lean_dec(x_118); +x_124 = lean_int_dec_lt(x_73, x_123); +lean_dec(x_123); +lean_dec(x_73); +if (x_124 == 0) +{ +lean_object* x_125; +lean_dec(x_119); +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +lean_dec(x_117); +x_82 = x_125; +goto block_94; +} +else +{ +lean_dec(x_117); +x_82 = x_119; +goto block_94; +} +} +} +} +block_94: +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_83 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_82); +lean_dec(x_82); +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_int_neg(x_84); +x_86 = lean_int_mul(x_85, x_78); +lean_dec(x_85); +x_87 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_88 = lean_int_add(x_86, x_87); +lean_dec(x_86); +x_89 = lean_int_add(x_81, x_88); +lean_dec(x_88); +lean_dec(x_81); +x_90 = l_Std_Time_Duration_ofNanoseconds(x_89); +lean_dec(x_89); +lean_inc(x_90); +x_91 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_91, 0, x_90); +lean_closure_set(x_91, 1, x_84); +lean_closure_set(x_91, 2, x_78); +x_92 = lean_mk_thunk(x_91); +if (lean_is_scalar(x_5)) { + x_93 = lean_alloc_ctor(0, 4, 0); +} else { + x_93 = x_5; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +lean_ctor_set(x_93, 2, x_4); +lean_ctor_set(x_93, 3, x_83); +return x_93; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withHours(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_2); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_findIdx_x3f_loop___rarg(x_13, x_14, x_15); +x_17 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_18 = lean_int_mul(x_12, x_17); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_34; +lean_dec(x_12); +x_34 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_21 = x_35; +goto block_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_21 = x_37; +goto block_33; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_16, 0); +lean_inc(x_38); +lean_dec(x_16); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_sub(x_38, x_39); +x_41 = l_Array_get_x3f___rarg(x_14, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +lean_dec(x_38); +lean_dec(x_14); +lean_dec(x_12); +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_21 = x_42; +goto block_33; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Array_get_x3f___rarg(x_14, x_38); +lean_dec(x_38); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +lean_dec(x_43); +lean_dec(x_12); +x_46 = lean_ctor_get(x_4, 0); +lean_inc(x_46); +x_21 = x_46; +goto block_33; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_nat_abs(x_50); +lean_dec(x_50); +x_52 = lean_nat_to_int(x_51); +x_53 = lean_int_sub(x_48, x_52); +lean_dec(x_52); +lean_dec(x_48); +x_54 = lean_int_dec_lt(x_12, x_53); +lean_dec(x_53); +lean_dec(x_12); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_49); +x_55 = lean_ctor_get(x_47, 1); +lean_inc(x_55); +lean_dec(x_47); +x_21 = x_55; +goto block_33; +} +else +{ +lean_dec(x_47); +x_21 = x_49; +goto block_33; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_dec(x_14); +x_56 = lean_ctor_get(x_44, 0); +lean_inc(x_56); +lean_dec(x_44); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_43, 1); +lean_inc(x_58); +lean_dec(x_43); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_nat_abs(x_59); +lean_dec(x_59); +x_61 = lean_nat_to_int(x_60); +x_62 = lean_int_sub(x_57, x_61); +lean_dec(x_61); +lean_dec(x_57); +x_63 = lean_int_dec_lt(x_12, x_62); +lean_dec(x_62); +lean_dec(x_12); +if (x_63 == 0) +{ +lean_object* x_64; +lean_dec(x_58); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_21 = x_64; +goto block_33; +} +else +{ +lean_dec(x_56); +x_21 = x_58; +goto block_33; +} +} +} +} +block_33: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_int_neg(x_23); +x_25 = lean_int_mul(x_24, x_17); +lean_dec(x_24); +x_26 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_25); +x_28 = lean_int_add(x_20, x_27); +lean_dec(x_27); +lean_dec(x_20); +x_29 = l_Std_Time_Duration_ofNanoseconds(x_28); +lean_dec(x_28); +lean_inc(x_29); +x_30 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_17); +x_31 = lean_mk_thunk(x_30); +if (lean_is_scalar(x_5)) { + x_32 = lean_alloc_ctor(0, 4, 0); +} else { + x_32 = x_5; +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_4); +lean_ctor_set(x_32, 3, x_22); +return x_32; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_65 = lean_ctor_get(x_8, 1); +x_66 = lean_ctor_get(x_8, 2); +x_67 = lean_ctor_get(x_8, 3); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_8); +x_68 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_68, 0, x_2); +lean_ctor_set(x_68, 1, x_65); +lean_ctor_set(x_68, 2, x_66); +lean_ctor_set(x_68, 3, x_67); +lean_ctor_set(x_6, 1, x_68); +x_69 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +lean_inc(x_70); +x_71 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_71, 0, x_70); +x_72 = lean_ctor_get(x_4, 1); +lean_inc(x_72); +x_73 = lean_unsigned_to_nat(0u); +x_74 = l_Array_findIdx_x3f_loop___rarg(x_71, x_72, x_73); +x_75 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_70, x_75); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_dec(x_69); +x_78 = lean_int_add(x_76, x_77); +lean_dec(x_77); +lean_dec(x_76); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_92; +lean_dec(x_70); +x_92 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_4, 0); +lean_inc(x_93); +x_79 = x_93; +goto block_91; +} +else +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_79 = x_95; +goto block_91; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_74, 0); +lean_inc(x_96); +lean_dec(x_74); +x_97 = lean_unsigned_to_nat(1u); +x_98 = lean_nat_sub(x_96, x_97); +x_99 = l_Array_get_x3f___rarg(x_72, x_98); +lean_dec(x_98); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; +lean_dec(x_96); +lean_dec(x_72); +lean_dec(x_70); +x_100 = lean_ctor_get(x_4, 0); +lean_inc(x_100); +x_79 = x_100; +goto block_91; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +lean_dec(x_99); +x_102 = l_Array_get_x3f___rarg(x_72, x_96); +lean_dec(x_96); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +x_103 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; +lean_dec(x_101); +lean_dec(x_70); +x_104 = lean_ctor_get(x_4, 0); +lean_inc(x_104); +x_79 = x_104; +goto block_91; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_101, 1); +lean_inc(x_107); +lean_dec(x_101); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_nat_abs(x_108); +lean_dec(x_108); +x_110 = lean_nat_to_int(x_109); +x_111 = lean_int_sub(x_106, x_110); +lean_dec(x_110); +lean_dec(x_106); +x_112 = lean_int_dec_lt(x_70, x_111); +lean_dec(x_111); +lean_dec(x_70); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_107); +x_113 = lean_ctor_get(x_105, 1); +lean_inc(x_113); +lean_dec(x_105); +x_79 = x_113; +goto block_91; +} +else +{ +lean_dec(x_105); +x_79 = x_107; +goto block_91; +} +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +lean_dec(x_72); +x_114 = lean_ctor_get(x_102, 0); +lean_inc(x_114); +lean_dec(x_102); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_101, 1); +lean_inc(x_116); +lean_dec(x_101); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_nat_abs(x_117); +lean_dec(x_117); +x_119 = lean_nat_to_int(x_118); +x_120 = lean_int_sub(x_115, x_119); +lean_dec(x_119); +lean_dec(x_115); +x_121 = lean_int_dec_lt(x_70, x_120); +lean_dec(x_120); +lean_dec(x_70); +if (x_121 == 0) +{ +lean_object* x_122; +lean_dec(x_116); +x_122 = lean_ctor_get(x_114, 1); +lean_inc(x_122); +lean_dec(x_114); +x_79 = x_122; +goto block_91; +} +else +{ +lean_dec(x_114); +x_79 = x_116; +goto block_91; +} +} +} +} +block_91: +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_79); +lean_dec(x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_int_neg(x_81); +x_83 = lean_int_mul(x_82, x_75); +lean_dec(x_82); +x_84 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_85 = lean_int_add(x_83, x_84); +lean_dec(x_83); +x_86 = lean_int_add(x_78, x_85); +lean_dec(x_85); +lean_dec(x_78); +x_87 = l_Std_Time_Duration_ofNanoseconds(x_86); +lean_dec(x_86); +lean_inc(x_87); +x_88 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_88, 0, x_87); +lean_closure_set(x_88, 1, x_81); +lean_closure_set(x_88, 2, x_75); +x_89 = lean_mk_thunk(x_88); +if (lean_is_scalar(x_5)) { + x_90 = lean_alloc_ctor(0, 4, 0); +} else { + x_90 = x_5; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_87); +lean_ctor_set(x_90, 2, x_4); +lean_ctor_set(x_90, 3, x_80); +return x_90; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 2); +lean_inc(x_126); +x_127 = lean_ctor_get(x_123, 3); +lean_inc(x_127); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + lean_ctor_release(x_123, 2); + lean_ctor_release(x_123, 3); + x_128 = x_123; +} else { + lean_dec_ref(x_123); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 4, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_2); +lean_ctor_set(x_129, 1, x_125); +lean_ctor_set(x_129, 2, x_126); +lean_ctor_set(x_129, 3, x_127); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_129); +x_131 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_130); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_inc(x_132); +x_133 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_133, 0, x_132); +x_134 = lean_ctor_get(x_4, 1); +lean_inc(x_134); +x_135 = lean_unsigned_to_nat(0u); +x_136 = l_Array_findIdx_x3f_loop___rarg(x_133, x_134, x_135); +x_137 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_138 = lean_int_mul(x_132, x_137); +x_139 = lean_ctor_get(x_131, 1); +lean_inc(x_139); +lean_dec(x_131); +x_140 = lean_int_add(x_138, x_139); +lean_dec(x_139); +lean_dec(x_138); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_154; +lean_dec(x_132); +x_154 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; +x_155 = lean_ctor_get(x_4, 0); +lean_inc(x_155); +x_141 = x_155; +goto block_153; +} +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_154, 0); +lean_inc(x_156); +lean_dec(x_154); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_141 = x_157; +goto block_153; +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = lean_ctor_get(x_136, 0); +lean_inc(x_158); +lean_dec(x_136); +x_159 = lean_unsigned_to_nat(1u); +x_160 = lean_nat_sub(x_158, x_159); +x_161 = l_Array_get_x3f___rarg(x_134, x_160); +lean_dec(x_160); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_dec(x_158); +lean_dec(x_134); +lean_dec(x_132); +x_162 = lean_ctor_get(x_4, 0); +lean_inc(x_162); +x_141 = x_162; +goto block_153; +} +else +{ +lean_object* x_163; lean_object* x_164; +x_163 = lean_ctor_get(x_161, 0); +lean_inc(x_163); +lean_dec(x_161); +x_164 = l_Array_get_x3f___rarg(x_134, x_158); +lean_dec(x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; +lean_dec(x_163); +lean_dec(x_132); +x_166 = lean_ctor_get(x_4, 0); +lean_inc(x_166); +x_141 = x_166; +goto block_153; +} +else +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_163, 1); +lean_inc(x_169); +lean_dec(x_163); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_nat_abs(x_170); +lean_dec(x_170); +x_172 = lean_nat_to_int(x_171); +x_173 = lean_int_sub(x_168, x_172); +lean_dec(x_172); +lean_dec(x_168); +x_174 = lean_int_dec_lt(x_132, x_173); +lean_dec(x_173); +lean_dec(x_132); +if (x_174 == 0) +{ +lean_object* x_175; +lean_dec(x_169); +x_175 = lean_ctor_get(x_167, 1); +lean_inc(x_175); +lean_dec(x_167); +x_141 = x_175; +goto block_153; +} +else +{ +lean_dec(x_167); +x_141 = x_169; +goto block_153; +} +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; +lean_dec(x_134); +x_176 = lean_ctor_get(x_164, 0); +lean_inc(x_176); +lean_dec(x_164); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_163, 1); +lean_inc(x_178); +lean_dec(x_163); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_nat_abs(x_179); +lean_dec(x_179); +x_181 = lean_nat_to_int(x_180); +x_182 = lean_int_sub(x_177, x_181); +lean_dec(x_181); +lean_dec(x_177); +x_183 = lean_int_dec_lt(x_132, x_182); +lean_dec(x_182); +lean_dec(x_132); +if (x_183 == 0) +{ +lean_object* x_184; +lean_dec(x_178); +x_184 = lean_ctor_get(x_176, 1); +lean_inc(x_184); +lean_dec(x_176); +x_141 = x_184; +goto block_153; +} +else +{ +lean_dec(x_176); +x_141 = x_178; +goto block_153; +} +} +} +} +block_153: +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_142 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_141); +lean_dec(x_141); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_int_neg(x_143); +x_145 = lean_int_mul(x_144, x_137); +lean_dec(x_144); +x_146 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_147 = lean_int_add(x_145, x_146); +lean_dec(x_145); +x_148 = lean_int_add(x_140, x_147); +lean_dec(x_147); +lean_dec(x_140); +x_149 = l_Std_Time_Duration_ofNanoseconds(x_148); +lean_dec(x_148); +lean_inc(x_149); +x_150 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_150, 0, x_149); +lean_closure_set(x_150, 1, x_143); +lean_closure_set(x_150, 2, x_137); +x_151 = lean_mk_thunk(x_150); +if (lean_is_scalar(x_5)) { + x_152 = lean_alloc_ctor(0, 4, 0); +} else { + x_152 = x_5; +} +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_149); +lean_ctor_set(x_152, 2, x_4); +lean_ctor_set(x_152, 3, x_142); +return x_152; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMinutes(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_8, 1); +lean_dec(x_10); +lean_ctor_set(x_8, 1, x_2); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_findIdx_x3f_loop___rarg(x_13, x_14, x_15); +x_17 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_18 = lean_int_mul(x_12, x_17); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_34; +lean_dec(x_12); +x_34 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_21 = x_35; +goto block_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_21 = x_37; +goto block_33; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_16, 0); +lean_inc(x_38); +lean_dec(x_16); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_sub(x_38, x_39); +x_41 = l_Array_get_x3f___rarg(x_14, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +lean_dec(x_38); +lean_dec(x_14); +lean_dec(x_12); +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_21 = x_42; +goto block_33; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Array_get_x3f___rarg(x_14, x_38); +lean_dec(x_38); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +lean_dec(x_43); +lean_dec(x_12); +x_46 = lean_ctor_get(x_4, 0); +lean_inc(x_46); +x_21 = x_46; +goto block_33; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_nat_abs(x_50); +lean_dec(x_50); +x_52 = lean_nat_to_int(x_51); +x_53 = lean_int_sub(x_48, x_52); +lean_dec(x_52); +lean_dec(x_48); +x_54 = lean_int_dec_lt(x_12, x_53); +lean_dec(x_53); +lean_dec(x_12); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_49); +x_55 = lean_ctor_get(x_47, 1); +lean_inc(x_55); +lean_dec(x_47); +x_21 = x_55; +goto block_33; +} +else +{ +lean_dec(x_47); +x_21 = x_49; +goto block_33; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_dec(x_14); +x_56 = lean_ctor_get(x_44, 0); +lean_inc(x_56); +lean_dec(x_44); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_43, 1); +lean_inc(x_58); +lean_dec(x_43); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_nat_abs(x_59); +lean_dec(x_59); +x_61 = lean_nat_to_int(x_60); +x_62 = lean_int_sub(x_57, x_61); +lean_dec(x_61); +lean_dec(x_57); +x_63 = lean_int_dec_lt(x_12, x_62); +lean_dec(x_62); +lean_dec(x_12); +if (x_63 == 0) +{ +lean_object* x_64; +lean_dec(x_58); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_21 = x_64; +goto block_33; +} +else +{ +lean_dec(x_56); +x_21 = x_58; +goto block_33; +} +} +} +} +block_33: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_int_neg(x_23); +x_25 = lean_int_mul(x_24, x_17); +lean_dec(x_24); +x_26 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_25); +x_28 = lean_int_add(x_20, x_27); +lean_dec(x_27); +lean_dec(x_20); +x_29 = l_Std_Time_Duration_ofNanoseconds(x_28); +lean_dec(x_28); +lean_inc(x_29); +x_30 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_17); +x_31 = lean_mk_thunk(x_30); +if (lean_is_scalar(x_5)) { + x_32 = lean_alloc_ctor(0, 4, 0); +} else { + x_32 = x_5; +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_4); +lean_ctor_set(x_32, 3, x_22); +return x_32; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_65 = lean_ctor_get(x_8, 0); +x_66 = lean_ctor_get(x_8, 2); +x_67 = lean_ctor_get(x_8, 3); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_8); +x_68 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_2); +lean_ctor_set(x_68, 2, x_66); +lean_ctor_set(x_68, 3, x_67); +lean_ctor_set(x_6, 1, x_68); +x_69 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +lean_inc(x_70); +x_71 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_71, 0, x_70); +x_72 = lean_ctor_get(x_4, 1); +lean_inc(x_72); +x_73 = lean_unsigned_to_nat(0u); +x_74 = l_Array_findIdx_x3f_loop___rarg(x_71, x_72, x_73); +x_75 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_70, x_75); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_dec(x_69); +x_78 = lean_int_add(x_76, x_77); +lean_dec(x_77); +lean_dec(x_76); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_92; +lean_dec(x_70); +x_92 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_4, 0); +lean_inc(x_93); +x_79 = x_93; +goto block_91; +} +else +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_79 = x_95; +goto block_91; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_74, 0); +lean_inc(x_96); +lean_dec(x_74); +x_97 = lean_unsigned_to_nat(1u); +x_98 = lean_nat_sub(x_96, x_97); +x_99 = l_Array_get_x3f___rarg(x_72, x_98); +lean_dec(x_98); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; +lean_dec(x_96); +lean_dec(x_72); +lean_dec(x_70); +x_100 = lean_ctor_get(x_4, 0); +lean_inc(x_100); +x_79 = x_100; +goto block_91; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +lean_dec(x_99); +x_102 = l_Array_get_x3f___rarg(x_72, x_96); +lean_dec(x_96); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +x_103 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; +lean_dec(x_101); +lean_dec(x_70); +x_104 = lean_ctor_get(x_4, 0); +lean_inc(x_104); +x_79 = x_104; +goto block_91; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_101, 1); +lean_inc(x_107); +lean_dec(x_101); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_nat_abs(x_108); +lean_dec(x_108); +x_110 = lean_nat_to_int(x_109); +x_111 = lean_int_sub(x_106, x_110); +lean_dec(x_110); +lean_dec(x_106); +x_112 = lean_int_dec_lt(x_70, x_111); +lean_dec(x_111); +lean_dec(x_70); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_107); +x_113 = lean_ctor_get(x_105, 1); +lean_inc(x_113); +lean_dec(x_105); +x_79 = x_113; +goto block_91; +} +else +{ +lean_dec(x_105); +x_79 = x_107; +goto block_91; +} +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +lean_dec(x_72); +x_114 = lean_ctor_get(x_102, 0); +lean_inc(x_114); +lean_dec(x_102); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_101, 1); +lean_inc(x_116); +lean_dec(x_101); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_nat_abs(x_117); +lean_dec(x_117); +x_119 = lean_nat_to_int(x_118); +x_120 = lean_int_sub(x_115, x_119); +lean_dec(x_119); +lean_dec(x_115); +x_121 = lean_int_dec_lt(x_70, x_120); +lean_dec(x_120); +lean_dec(x_70); +if (x_121 == 0) +{ +lean_object* x_122; +lean_dec(x_116); +x_122 = lean_ctor_get(x_114, 1); +lean_inc(x_122); +lean_dec(x_114); +x_79 = x_122; +goto block_91; +} +else +{ +lean_dec(x_114); +x_79 = x_116; +goto block_91; +} +} +} +} +block_91: +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_79); +lean_dec(x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_int_neg(x_81); +x_83 = lean_int_mul(x_82, x_75); +lean_dec(x_82); +x_84 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_85 = lean_int_add(x_83, x_84); +lean_dec(x_83); +x_86 = lean_int_add(x_78, x_85); +lean_dec(x_85); +lean_dec(x_78); +x_87 = l_Std_Time_Duration_ofNanoseconds(x_86); +lean_dec(x_86); +lean_inc(x_87); +x_88 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_88, 0, x_87); +lean_closure_set(x_88, 1, x_81); +lean_closure_set(x_88, 2, x_75); +x_89 = lean_mk_thunk(x_88); +if (lean_is_scalar(x_5)) { + x_90 = lean_alloc_ctor(0, 4, 0); +} else { + x_90 = x_5; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_87); +lean_ctor_set(x_90, 2, x_4); +lean_ctor_set(x_90, 3, x_80); +return x_90; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 2); +lean_inc(x_126); +x_127 = lean_ctor_get(x_123, 3); +lean_inc(x_127); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + lean_ctor_release(x_123, 2); + lean_ctor_release(x_123, 3); + x_128 = x_123; +} else { + lean_dec_ref(x_123); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 4, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_2); +lean_ctor_set(x_129, 2, x_126); +lean_ctor_set(x_129, 3, x_127); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_129); +x_131 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_130); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_inc(x_132); +x_133 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_133, 0, x_132); +x_134 = lean_ctor_get(x_4, 1); +lean_inc(x_134); +x_135 = lean_unsigned_to_nat(0u); +x_136 = l_Array_findIdx_x3f_loop___rarg(x_133, x_134, x_135); +x_137 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_138 = lean_int_mul(x_132, x_137); +x_139 = lean_ctor_get(x_131, 1); +lean_inc(x_139); +lean_dec(x_131); +x_140 = lean_int_add(x_138, x_139); +lean_dec(x_139); +lean_dec(x_138); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_154; +lean_dec(x_132); +x_154 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; +x_155 = lean_ctor_get(x_4, 0); +lean_inc(x_155); +x_141 = x_155; +goto block_153; +} +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_154, 0); +lean_inc(x_156); +lean_dec(x_154); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_141 = x_157; +goto block_153; +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = lean_ctor_get(x_136, 0); +lean_inc(x_158); +lean_dec(x_136); +x_159 = lean_unsigned_to_nat(1u); +x_160 = lean_nat_sub(x_158, x_159); +x_161 = l_Array_get_x3f___rarg(x_134, x_160); +lean_dec(x_160); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_dec(x_158); +lean_dec(x_134); +lean_dec(x_132); +x_162 = lean_ctor_get(x_4, 0); +lean_inc(x_162); +x_141 = x_162; +goto block_153; +} +else +{ +lean_object* x_163; lean_object* x_164; +x_163 = lean_ctor_get(x_161, 0); +lean_inc(x_163); +lean_dec(x_161); +x_164 = l_Array_get_x3f___rarg(x_134, x_158); +lean_dec(x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; +lean_dec(x_163); +lean_dec(x_132); +x_166 = lean_ctor_get(x_4, 0); +lean_inc(x_166); +x_141 = x_166; +goto block_153; +} +else +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_163, 1); +lean_inc(x_169); +lean_dec(x_163); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_nat_abs(x_170); +lean_dec(x_170); +x_172 = lean_nat_to_int(x_171); +x_173 = lean_int_sub(x_168, x_172); +lean_dec(x_172); +lean_dec(x_168); +x_174 = lean_int_dec_lt(x_132, x_173); +lean_dec(x_173); +lean_dec(x_132); +if (x_174 == 0) +{ +lean_object* x_175; +lean_dec(x_169); +x_175 = lean_ctor_get(x_167, 1); +lean_inc(x_175); +lean_dec(x_167); +x_141 = x_175; +goto block_153; +} +else +{ +lean_dec(x_167); +x_141 = x_169; +goto block_153; +} +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; +lean_dec(x_134); +x_176 = lean_ctor_get(x_164, 0); +lean_inc(x_176); +lean_dec(x_164); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_163, 1); +lean_inc(x_178); +lean_dec(x_163); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_nat_abs(x_179); +lean_dec(x_179); +x_181 = lean_nat_to_int(x_180); +x_182 = lean_int_sub(x_177, x_181); +lean_dec(x_181); +lean_dec(x_177); +x_183 = lean_int_dec_lt(x_132, x_182); +lean_dec(x_182); +lean_dec(x_132); +if (x_183 == 0) +{ +lean_object* x_184; +lean_dec(x_178); +x_184 = lean_ctor_get(x_176, 1); +lean_inc(x_184); +lean_dec(x_176); +x_141 = x_184; +goto block_153; +} +else +{ +lean_dec(x_176); +x_141 = x_178; +goto block_153; +} +} +} +} +block_153: +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_142 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_141); +lean_dec(x_141); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_int_neg(x_143); +x_145 = lean_int_mul(x_144, x_137); +lean_dec(x_144); +x_146 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_147 = lean_int_add(x_145, x_146); +lean_dec(x_145); +x_148 = lean_int_add(x_140, x_147); +lean_dec(x_147); +lean_dec(x_140); +x_149 = l_Std_Time_Duration_ofNanoseconds(x_148); +lean_dec(x_148); +lean_inc(x_149); +x_150 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_150, 0, x_149); +lean_closure_set(x_150, 1, x_143); +lean_closure_set(x_150, 2, x_137); +x_151 = lean_mk_thunk(x_150); +if (lean_is_scalar(x_5)) { + x_152 = lean_alloc_ctor(0, 4, 0); +} else { + x_152 = x_5; +} +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_149); +lean_ctor_set(x_152, 2, x_4); +lean_ctor_set(x_152, 3, x_142); +return x_152; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withSeconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_8, 2); +lean_dec(x_10); +lean_ctor_set(x_8, 2, x_2); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_findIdx_x3f_loop___rarg(x_13, x_14, x_15); +x_17 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_18 = lean_int_mul(x_12, x_17); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_34; +lean_dec(x_12); +x_34 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_21 = x_35; +goto block_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_21 = x_37; +goto block_33; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_16, 0); +lean_inc(x_38); +lean_dec(x_16); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_sub(x_38, x_39); +x_41 = l_Array_get_x3f___rarg(x_14, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +lean_dec(x_38); +lean_dec(x_14); +lean_dec(x_12); +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_21 = x_42; +goto block_33; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Array_get_x3f___rarg(x_14, x_38); +lean_dec(x_38); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +lean_dec(x_43); +lean_dec(x_12); +x_46 = lean_ctor_get(x_4, 0); +lean_inc(x_46); +x_21 = x_46; +goto block_33; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_nat_abs(x_50); +lean_dec(x_50); +x_52 = lean_nat_to_int(x_51); +x_53 = lean_int_sub(x_48, x_52); +lean_dec(x_52); +lean_dec(x_48); +x_54 = lean_int_dec_lt(x_12, x_53); +lean_dec(x_53); +lean_dec(x_12); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_49); +x_55 = lean_ctor_get(x_47, 1); +lean_inc(x_55); +lean_dec(x_47); +x_21 = x_55; +goto block_33; +} +else +{ +lean_dec(x_47); +x_21 = x_49; +goto block_33; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_dec(x_14); +x_56 = lean_ctor_get(x_44, 0); +lean_inc(x_56); +lean_dec(x_44); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_43, 1); +lean_inc(x_58); +lean_dec(x_43); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_nat_abs(x_59); +lean_dec(x_59); +x_61 = lean_nat_to_int(x_60); +x_62 = lean_int_sub(x_57, x_61); +lean_dec(x_61); +lean_dec(x_57); +x_63 = lean_int_dec_lt(x_12, x_62); +lean_dec(x_62); +lean_dec(x_12); +if (x_63 == 0) +{ +lean_object* x_64; +lean_dec(x_58); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_21 = x_64; +goto block_33; +} +else +{ +lean_dec(x_56); +x_21 = x_58; +goto block_33; +} +} +} +} +block_33: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_int_neg(x_23); +x_25 = lean_int_mul(x_24, x_17); +lean_dec(x_24); +x_26 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_25); +x_28 = lean_int_add(x_20, x_27); +lean_dec(x_27); +lean_dec(x_20); +x_29 = l_Std_Time_Duration_ofNanoseconds(x_28); +lean_dec(x_28); +lean_inc(x_29); +x_30 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_17); +x_31 = lean_mk_thunk(x_30); +if (lean_is_scalar(x_5)) { + x_32 = lean_alloc_ctor(0, 4, 0); +} else { + x_32 = x_5; +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_4); +lean_ctor_set(x_32, 3, x_22); +return x_32; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_65 = lean_ctor_get(x_8, 0); +x_66 = lean_ctor_get(x_8, 1); +x_67 = lean_ctor_get(x_8, 3); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_8); +x_68 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +lean_ctor_set(x_68, 2, x_2); +lean_ctor_set(x_68, 3, x_67); +lean_ctor_set(x_6, 1, x_68); +x_69 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +lean_inc(x_70); +x_71 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_71, 0, x_70); +x_72 = lean_ctor_get(x_4, 1); +lean_inc(x_72); +x_73 = lean_unsigned_to_nat(0u); +x_74 = l_Array_findIdx_x3f_loop___rarg(x_71, x_72, x_73); +x_75 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_70, x_75); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_dec(x_69); +x_78 = lean_int_add(x_76, x_77); +lean_dec(x_77); +lean_dec(x_76); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_92; +lean_dec(x_70); +x_92 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_4, 0); +lean_inc(x_93); +x_79 = x_93; +goto block_91; +} +else +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_79 = x_95; +goto block_91; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_74, 0); +lean_inc(x_96); +lean_dec(x_74); +x_97 = lean_unsigned_to_nat(1u); +x_98 = lean_nat_sub(x_96, x_97); +x_99 = l_Array_get_x3f___rarg(x_72, x_98); +lean_dec(x_98); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; +lean_dec(x_96); +lean_dec(x_72); +lean_dec(x_70); +x_100 = lean_ctor_get(x_4, 0); +lean_inc(x_100); +x_79 = x_100; +goto block_91; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +lean_dec(x_99); +x_102 = l_Array_get_x3f___rarg(x_72, x_96); +lean_dec(x_96); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +x_103 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; +lean_dec(x_101); +lean_dec(x_70); +x_104 = lean_ctor_get(x_4, 0); +lean_inc(x_104); +x_79 = x_104; +goto block_91; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_101, 1); +lean_inc(x_107); +lean_dec(x_101); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_nat_abs(x_108); +lean_dec(x_108); +x_110 = lean_nat_to_int(x_109); +x_111 = lean_int_sub(x_106, x_110); +lean_dec(x_110); +lean_dec(x_106); +x_112 = lean_int_dec_lt(x_70, x_111); +lean_dec(x_111); +lean_dec(x_70); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_107); +x_113 = lean_ctor_get(x_105, 1); +lean_inc(x_113); +lean_dec(x_105); +x_79 = x_113; +goto block_91; +} +else +{ +lean_dec(x_105); +x_79 = x_107; +goto block_91; +} +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +lean_dec(x_72); +x_114 = lean_ctor_get(x_102, 0); +lean_inc(x_114); +lean_dec(x_102); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_101, 1); +lean_inc(x_116); +lean_dec(x_101); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_nat_abs(x_117); +lean_dec(x_117); +x_119 = lean_nat_to_int(x_118); +x_120 = lean_int_sub(x_115, x_119); +lean_dec(x_119); +lean_dec(x_115); +x_121 = lean_int_dec_lt(x_70, x_120); +lean_dec(x_120); +lean_dec(x_70); +if (x_121 == 0) +{ +lean_object* x_122; +lean_dec(x_116); +x_122 = lean_ctor_get(x_114, 1); +lean_inc(x_122); +lean_dec(x_114); +x_79 = x_122; +goto block_91; +} +else +{ +lean_dec(x_114); +x_79 = x_116; +goto block_91; +} +} +} +} +block_91: +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_79); +lean_dec(x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_int_neg(x_81); +x_83 = lean_int_mul(x_82, x_75); +lean_dec(x_82); +x_84 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_85 = lean_int_add(x_83, x_84); +lean_dec(x_83); +x_86 = lean_int_add(x_78, x_85); +lean_dec(x_85); +lean_dec(x_78); +x_87 = l_Std_Time_Duration_ofNanoseconds(x_86); +lean_dec(x_86); +lean_inc(x_87); +x_88 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_88, 0, x_87); +lean_closure_set(x_88, 1, x_81); +lean_closure_set(x_88, 2, x_75); +x_89 = lean_mk_thunk(x_88); +if (lean_is_scalar(x_5)) { + x_90 = lean_alloc_ctor(0, 4, 0); +} else { + x_90 = x_5; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_87); +lean_ctor_set(x_90, 2, x_4); +lean_ctor_set(x_90, 3, x_80); +return x_90; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +x_127 = lean_ctor_get(x_123, 3); +lean_inc(x_127); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + lean_ctor_release(x_123, 2); + lean_ctor_release(x_123, 3); + x_128 = x_123; +} else { + lean_dec_ref(x_123); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 4, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_126); +lean_ctor_set(x_129, 2, x_2); +lean_ctor_set(x_129, 3, x_127); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_129); +x_131 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_130); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_inc(x_132); +x_133 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_133, 0, x_132); +x_134 = lean_ctor_get(x_4, 1); +lean_inc(x_134); +x_135 = lean_unsigned_to_nat(0u); +x_136 = l_Array_findIdx_x3f_loop___rarg(x_133, x_134, x_135); +x_137 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_138 = lean_int_mul(x_132, x_137); +x_139 = lean_ctor_get(x_131, 1); +lean_inc(x_139); +lean_dec(x_131); +x_140 = lean_int_add(x_138, x_139); +lean_dec(x_139); +lean_dec(x_138); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_154; +lean_dec(x_132); +x_154 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; +x_155 = lean_ctor_get(x_4, 0); +lean_inc(x_155); +x_141 = x_155; +goto block_153; +} +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_154, 0); +lean_inc(x_156); +lean_dec(x_154); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_141 = x_157; +goto block_153; +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = lean_ctor_get(x_136, 0); +lean_inc(x_158); +lean_dec(x_136); +x_159 = lean_unsigned_to_nat(1u); +x_160 = lean_nat_sub(x_158, x_159); +x_161 = l_Array_get_x3f___rarg(x_134, x_160); +lean_dec(x_160); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_dec(x_158); +lean_dec(x_134); +lean_dec(x_132); +x_162 = lean_ctor_get(x_4, 0); +lean_inc(x_162); +x_141 = x_162; +goto block_153; +} +else +{ +lean_object* x_163; lean_object* x_164; +x_163 = lean_ctor_get(x_161, 0); +lean_inc(x_163); +lean_dec(x_161); +x_164 = l_Array_get_x3f___rarg(x_134, x_158); +lean_dec(x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; +lean_dec(x_163); +lean_dec(x_132); +x_166 = lean_ctor_get(x_4, 0); +lean_inc(x_166); +x_141 = x_166; +goto block_153; +} +else +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_163, 1); +lean_inc(x_169); +lean_dec(x_163); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_nat_abs(x_170); +lean_dec(x_170); +x_172 = lean_nat_to_int(x_171); +x_173 = lean_int_sub(x_168, x_172); +lean_dec(x_172); +lean_dec(x_168); +x_174 = lean_int_dec_lt(x_132, x_173); +lean_dec(x_173); +lean_dec(x_132); +if (x_174 == 0) +{ +lean_object* x_175; +lean_dec(x_169); +x_175 = lean_ctor_get(x_167, 1); +lean_inc(x_175); +lean_dec(x_167); +x_141 = x_175; +goto block_153; +} +else +{ +lean_dec(x_167); +x_141 = x_169; +goto block_153; +} +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; +lean_dec(x_134); +x_176 = lean_ctor_get(x_164, 0); +lean_inc(x_176); +lean_dec(x_164); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_163, 1); +lean_inc(x_178); +lean_dec(x_163); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_nat_abs(x_179); +lean_dec(x_179); +x_181 = lean_nat_to_int(x_180); +x_182 = lean_int_sub(x_177, x_181); +lean_dec(x_181); +lean_dec(x_177); +x_183 = lean_int_dec_lt(x_132, x_182); +lean_dec(x_182); +lean_dec(x_132); +if (x_183 == 0) +{ +lean_object* x_184; +lean_dec(x_178); +x_184 = lean_ctor_get(x_176, 1); +lean_inc(x_184); +lean_dec(x_176); +x_141 = x_184; +goto block_153; +} +else +{ +lean_dec(x_176); +x_141 = x_178; +goto block_153; +} +} +} +} +block_153: +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_142 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_141); +lean_dec(x_141); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_int_neg(x_143); +x_145 = lean_int_mul(x_144, x_137); +lean_dec(x_144); +x_146 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_147 = lean_int_add(x_145, x_146); +lean_dec(x_145); +x_148 = lean_int_add(x_140, x_147); +lean_dec(x_147); +lean_dec(x_140); +x_149 = l_Std_Time_Duration_ofNanoseconds(x_148); +lean_dec(x_148); +lean_inc(x_149); +x_150 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_150, 0, x_149); +lean_closure_set(x_150, 1, x_143); +lean_closure_set(x_150, 2, x_137); +x_151 = lean_mk_thunk(x_150); +if (lean_is_scalar(x_5)) { + x_152 = lean_alloc_ctor(0, 4, 0); +} else { + x_152 = x_5; +} +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_149); +lean_ctor_set(x_152, 2, x_4); +lean_ctor_set(x_152, 3, x_142); +return x_152; +} +} +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_withMilliseconds___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMilliseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_10 = lean_ctor_get(x_8, 3); +x_11 = l_Std_Time_ZonedDateTime_withMilliseconds___closed__1; +x_12 = lean_int_emod(x_10, x_11); +lean_dec(x_10); +x_13 = l_Std_Time_ZonedDateTime_millisecond___closed__1; +x_14 = lean_int_mul(x_2, x_13); +x_15 = lean_int_add(x_14, x_12); +lean_dec(x_12); +lean_dec(x_14); +lean_ctor_set(x_8, 3, x_15); +x_16 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_inc(x_17); +x_18 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_18, 0, x_17); +x_19 = lean_ctor_get(x_4, 1); +lean_inc(x_19); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Array_findIdx_x3f_loop___rarg(x_18, x_19, x_20); +x_22 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_23 = lean_int_mul(x_17, x_22); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_dec(x_16); +x_25 = lean_int_add(x_23, x_24); +lean_dec(x_24); +lean_dec(x_23); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_39; +lean_dec(x_17); +x_39 = l_Array_back_x3f___rarg(x_19); +lean_dec(x_19); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_4, 0); +lean_inc(x_40); +x_26 = x_40; +goto block_38; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_39, 0); +lean_inc(x_41); +lean_dec(x_39); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_26 = x_42; +goto block_38; +} +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_21, 0); +lean_inc(x_43); +lean_dec(x_21); +x_44 = lean_unsigned_to_nat(1u); +x_45 = lean_nat_sub(x_43, x_44); +x_46 = l_Array_get_x3f___rarg(x_19, x_45); +lean_dec(x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; +lean_dec(x_43); +lean_dec(x_19); +lean_dec(x_17); +x_47 = lean_ctor_get(x_4, 0); +lean_inc(x_47); +x_26 = x_47; +goto block_38; +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Array_get_x3f___rarg(x_19, x_43); +lean_dec(x_43); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; +x_50 = l_Array_back_x3f___rarg(x_19); +lean_dec(x_19); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; +lean_dec(x_48); +lean_dec(x_17); +x_51 = lean_ctor_get(x_4, 0); +lean_inc(x_51); +x_26 = x_51; +goto block_38; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_52 = lean_ctor_get(x_50, 0); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_48, 1); +lean_inc(x_54); +lean_dec(x_48); +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_nat_abs(x_55); +lean_dec(x_55); +x_57 = lean_nat_to_int(x_56); +x_58 = lean_int_sub(x_53, x_57); +lean_dec(x_57); +lean_dec(x_53); +x_59 = lean_int_dec_lt(x_17, x_58); +lean_dec(x_58); +lean_dec(x_17); +if (x_59 == 0) +{ +lean_object* x_60; +lean_dec(x_54); +x_60 = lean_ctor_get(x_52, 1); +lean_inc(x_60); +lean_dec(x_52); +x_26 = x_60; +goto block_38; +} +else +{ +lean_dec(x_52); +x_26 = x_54; +goto block_38; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +lean_dec(x_19); +x_61 = lean_ctor_get(x_49, 0); +lean_inc(x_61); +lean_dec(x_49); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_48, 1); +lean_inc(x_63); +lean_dec(x_48); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_nat_abs(x_64); +lean_dec(x_64); +x_66 = lean_nat_to_int(x_65); +x_67 = lean_int_sub(x_62, x_66); +lean_dec(x_66); +lean_dec(x_62); +x_68 = lean_int_dec_lt(x_17, x_67); +lean_dec(x_67); +lean_dec(x_17); +if (x_68 == 0) +{ +lean_object* x_69; +lean_dec(x_63); +x_69 = lean_ctor_get(x_61, 1); +lean_inc(x_69); +lean_dec(x_61); +x_26 = x_69; +goto block_38; +} +else +{ +lean_dec(x_61); +x_26 = x_63; +goto block_38; +} +} +} +} +block_38: +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_27 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_26); +lean_dec(x_26); +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_int_neg(x_28); +x_30 = lean_int_mul(x_29, x_22); +lean_dec(x_29); +x_31 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_32 = lean_int_add(x_30, x_31); +lean_dec(x_30); +x_33 = lean_int_add(x_25, x_32); +lean_dec(x_32); +lean_dec(x_25); +x_34 = l_Std_Time_Duration_ofNanoseconds(x_33); +lean_dec(x_33); +lean_inc(x_34); +x_35 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_35, 0, x_34); +lean_closure_set(x_35, 1, x_28); +lean_closure_set(x_35, 2, x_22); +x_36 = lean_mk_thunk(x_35); +if (lean_is_scalar(x_5)) { + x_37 = lean_alloc_ctor(0, 4, 0); +} else { + x_37 = x_5; +} +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_34); +lean_ctor_set(x_37, 2, x_4); +lean_ctor_set(x_37, 3, x_27); +return x_37; +} +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_70 = lean_ctor_get(x_8, 0); +x_71 = lean_ctor_get(x_8, 1); +x_72 = lean_ctor_get(x_8, 2); +x_73 = lean_ctor_get(x_8, 3); +lean_inc(x_73); +lean_inc(x_72); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_8); +x_74 = l_Std_Time_ZonedDateTime_withMilliseconds___closed__1; +x_75 = lean_int_emod(x_73, x_74); +lean_dec(x_73); +x_76 = l_Std_Time_ZonedDateTime_millisecond___closed__1; +x_77 = lean_int_mul(x_2, x_76); +x_78 = lean_int_add(x_77, x_75); +lean_dec(x_75); +lean_dec(x_77); +x_79 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_79, 0, x_70); +lean_ctor_set(x_79, 1, x_71); +lean_ctor_set(x_79, 2, x_72); +lean_ctor_set(x_79, 3, x_78); +lean_ctor_set(x_6, 1, x_79); +x_80 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +lean_inc(x_81); +x_82 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_82, 0, x_81); +x_83 = lean_ctor_get(x_4, 1); +lean_inc(x_83); +x_84 = lean_unsigned_to_nat(0u); +x_85 = l_Array_findIdx_x3f_loop___rarg(x_82, x_83, x_84); +x_86 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_87 = lean_int_mul(x_81, x_86); +x_88 = lean_ctor_get(x_80, 1); +lean_inc(x_88); +lean_dec(x_80); +x_89 = lean_int_add(x_87, x_88); +lean_dec(x_88); +lean_dec(x_87); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_103; +lean_dec(x_81); +x_103 = l_Array_back_x3f___rarg(x_83); +lean_dec(x_83); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; +x_104 = lean_ctor_get(x_4, 0); +lean_inc(x_104); +x_90 = x_104; +goto block_102; +} +else +{ +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_90 = x_106; +goto block_102; +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_107 = lean_ctor_get(x_85, 0); +lean_inc(x_107); +lean_dec(x_85); +x_108 = lean_unsigned_to_nat(1u); +x_109 = lean_nat_sub(x_107, x_108); +x_110 = l_Array_get_x3f___rarg(x_83, x_109); +lean_dec(x_109); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; +lean_dec(x_107); +lean_dec(x_83); +lean_dec(x_81); +x_111 = lean_ctor_get(x_4, 0); +lean_inc(x_111); +x_90 = x_111; +goto block_102; +} +else +{ +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_110, 0); +lean_inc(x_112); +lean_dec(x_110); +x_113 = l_Array_get_x3f___rarg(x_83, x_107); +lean_dec(x_107); +if (lean_obj_tag(x_113) == 0) +{ +lean_object* x_114; +x_114 = l_Array_back_x3f___rarg(x_83); +lean_dec(x_83); +if (lean_obj_tag(x_114) == 0) +{ +lean_object* x_115; +lean_dec(x_112); +lean_dec(x_81); +x_115 = lean_ctor_get(x_4, 0); +lean_inc(x_115); +x_90 = x_115; +goto block_102; +} +else +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_116 = lean_ctor_get(x_114, 0); +lean_inc(x_116); +lean_dec(x_114); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_112, 1); +lean_inc(x_118); +lean_dec(x_112); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_nat_abs(x_119); +lean_dec(x_119); +x_121 = lean_nat_to_int(x_120); +x_122 = lean_int_sub(x_117, x_121); +lean_dec(x_121); +lean_dec(x_117); +x_123 = lean_int_dec_lt(x_81, x_122); +lean_dec(x_122); +lean_dec(x_81); +if (x_123 == 0) +{ +lean_object* x_124; +lean_dec(x_118); +x_124 = lean_ctor_get(x_116, 1); +lean_inc(x_124); +lean_dec(x_116); +x_90 = x_124; +goto block_102; +} +else +{ +lean_dec(x_116); +x_90 = x_118; +goto block_102; +} +} +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +lean_dec(x_83); +x_125 = lean_ctor_get(x_113, 0); +lean_inc(x_125); +lean_dec(x_113); +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_112, 1); +lean_inc(x_127); +lean_dec(x_112); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_nat_abs(x_128); +lean_dec(x_128); +x_130 = lean_nat_to_int(x_129); +x_131 = lean_int_sub(x_126, x_130); +lean_dec(x_130); +lean_dec(x_126); +x_132 = lean_int_dec_lt(x_81, x_131); +lean_dec(x_131); +lean_dec(x_81); +if (x_132 == 0) +{ +lean_object* x_133; +lean_dec(x_127); +x_133 = lean_ctor_get(x_125, 1); +lean_inc(x_133); +lean_dec(x_125); +x_90 = x_133; +goto block_102; +} +else +{ +lean_dec(x_125); +x_90 = x_127; +goto block_102; +} +} +} +} +block_102: +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_91 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_90); +lean_dec(x_90); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_int_neg(x_92); +x_94 = lean_int_mul(x_93, x_86); +lean_dec(x_93); +x_95 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_96 = lean_int_add(x_94, x_95); +lean_dec(x_94); +x_97 = lean_int_add(x_89, x_96); +lean_dec(x_96); +lean_dec(x_89); +x_98 = l_Std_Time_Duration_ofNanoseconds(x_97); +lean_dec(x_97); +lean_inc(x_98); +x_99 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_99, 0, x_98); +lean_closure_set(x_99, 1, x_92); +lean_closure_set(x_99, 2, x_86); +x_100 = lean_mk_thunk(x_99); +if (lean_is_scalar(x_5)) { + x_101 = lean_alloc_ctor(0, 4, 0); +} else { + x_101 = x_5; +} +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +lean_ctor_set(x_101, 2, x_4); +lean_ctor_set(x_101, 3, x_91); +return x_101; +} +} +} +else +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_134 = lean_ctor_get(x_6, 1); +x_135 = lean_ctor_get(x_6, 0); +lean_inc(x_134); +lean_inc(x_135); +lean_dec(x_6); +x_136 = lean_ctor_get(x_134, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_134, 1); +lean_inc(x_137); +x_138 = lean_ctor_get(x_134, 2); +lean_inc(x_138); +x_139 = lean_ctor_get(x_134, 3); +lean_inc(x_139); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + lean_ctor_release(x_134, 2); + lean_ctor_release(x_134, 3); + x_140 = x_134; +} else { + lean_dec_ref(x_134); + x_140 = lean_box(0); +} +x_141 = l_Std_Time_ZonedDateTime_withMilliseconds___closed__1; +x_142 = lean_int_emod(x_139, x_141); +lean_dec(x_139); +x_143 = l_Std_Time_ZonedDateTime_millisecond___closed__1; +x_144 = lean_int_mul(x_2, x_143); +x_145 = lean_int_add(x_144, x_142); +lean_dec(x_142); +lean_dec(x_144); +if (lean_is_scalar(x_140)) { + x_146 = lean_alloc_ctor(0, 4, 0); +} else { + x_146 = x_140; +} +lean_ctor_set(x_146, 0, x_136); +lean_ctor_set(x_146, 1, x_137); +lean_ctor_set(x_146, 2, x_138); +lean_ctor_set(x_146, 3, x_145); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_135); +lean_ctor_set(x_147, 1, x_146); +x_148 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_147); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +lean_inc(x_149); +x_150 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_150, 0, x_149); +x_151 = lean_ctor_get(x_4, 1); +lean_inc(x_151); +x_152 = lean_unsigned_to_nat(0u); +x_153 = l_Array_findIdx_x3f_loop___rarg(x_150, x_151, x_152); +x_154 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_155 = lean_int_mul(x_149, x_154); +x_156 = lean_ctor_get(x_148, 1); +lean_inc(x_156); +lean_dec(x_148); +x_157 = lean_int_add(x_155, x_156); +lean_dec(x_156); +lean_dec(x_155); +if (lean_obj_tag(x_153) == 0) +{ +lean_object* x_171; +lean_dec(x_149); +x_171 = l_Array_back_x3f___rarg(x_151); +lean_dec(x_151); +if (lean_obj_tag(x_171) == 0) +{ +lean_object* x_172; +x_172 = lean_ctor_get(x_4, 0); +lean_inc(x_172); +x_158 = x_172; +goto block_170; +} +else +{ +lean_object* x_173; lean_object* x_174; +x_173 = lean_ctor_get(x_171, 0); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_173, 1); +lean_inc(x_174); +lean_dec(x_173); +x_158 = x_174; +goto block_170; +} +} +else +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_175 = lean_ctor_get(x_153, 0); +lean_inc(x_175); +lean_dec(x_153); +x_176 = lean_unsigned_to_nat(1u); +x_177 = lean_nat_sub(x_175, x_176); +x_178 = l_Array_get_x3f___rarg(x_151, x_177); +lean_dec(x_177); +if (lean_obj_tag(x_178) == 0) +{ +lean_object* x_179; +lean_dec(x_175); +lean_dec(x_151); +lean_dec(x_149); +x_179 = lean_ctor_get(x_4, 0); +lean_inc(x_179); +x_158 = x_179; +goto block_170; +} +else +{ +lean_object* x_180; lean_object* x_181; +x_180 = lean_ctor_get(x_178, 0); +lean_inc(x_180); +lean_dec(x_178); +x_181 = l_Array_get_x3f___rarg(x_151, x_175); +lean_dec(x_175); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; +x_182 = l_Array_back_x3f___rarg(x_151); +lean_dec(x_151); +if (lean_obj_tag(x_182) == 0) +{ +lean_object* x_183; +lean_dec(x_180); +lean_dec(x_149); +x_183 = lean_ctor_get(x_4, 0); +lean_inc(x_183); +x_158 = x_183; +goto block_170; +} +else +{ +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; uint8_t x_191; +x_184 = lean_ctor_get(x_182, 0); +lean_inc(x_184); +lean_dec(x_182); +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +x_186 = lean_ctor_get(x_180, 1); +lean_inc(x_186); +lean_dec(x_180); +x_187 = lean_ctor_get(x_186, 0); +lean_inc(x_187); +x_188 = lean_nat_abs(x_187); +lean_dec(x_187); +x_189 = lean_nat_to_int(x_188); +x_190 = lean_int_sub(x_185, x_189); +lean_dec(x_189); +lean_dec(x_185); +x_191 = lean_int_dec_lt(x_149, x_190); +lean_dec(x_190); +lean_dec(x_149); +if (x_191 == 0) +{ +lean_object* x_192; +lean_dec(x_186); +x_192 = lean_ctor_get(x_184, 1); +lean_inc(x_192); +lean_dec(x_184); +x_158 = x_192; +goto block_170; +} +else +{ +lean_dec(x_184); +x_158 = x_186; +goto block_170; +} +} +} +else +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; +lean_dec(x_151); +x_193 = lean_ctor_get(x_181, 0); +lean_inc(x_193); +lean_dec(x_181); +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_180, 1); +lean_inc(x_195); +lean_dec(x_180); +x_196 = lean_ctor_get(x_195, 0); +lean_inc(x_196); +x_197 = lean_nat_abs(x_196); +lean_dec(x_196); +x_198 = lean_nat_to_int(x_197); +x_199 = lean_int_sub(x_194, x_198); +lean_dec(x_198); +lean_dec(x_194); +x_200 = lean_int_dec_lt(x_149, x_199); +lean_dec(x_199); +lean_dec(x_149); +if (x_200 == 0) +{ +lean_object* x_201; +lean_dec(x_195); +x_201 = lean_ctor_get(x_193, 1); +lean_inc(x_201); +lean_dec(x_193); +x_158 = x_201; +goto block_170; +} +else +{ +lean_dec(x_193); +x_158 = x_195; +goto block_170; +} +} +} +} +block_170: +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_159 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_158); +lean_dec(x_158); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_int_neg(x_160); +x_162 = lean_int_mul(x_161, x_154); +lean_dec(x_161); +x_163 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_164 = lean_int_add(x_162, x_163); +lean_dec(x_162); +x_165 = lean_int_add(x_157, x_164); +lean_dec(x_164); +lean_dec(x_157); +x_166 = l_Std_Time_Duration_ofNanoseconds(x_165); +lean_dec(x_165); +lean_inc(x_166); +x_167 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_167, 0, x_166); +lean_closure_set(x_167, 1, x_160); +lean_closure_set(x_167, 2, x_154); +x_168 = lean_mk_thunk(x_167); +if (lean_is_scalar(x_5)) { + x_169 = lean_alloc_ctor(0, 4, 0); +} else { + x_169 = x_5; +} +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_166); +lean_ctor_set(x_169, 2, x_4); +lean_ctor_set(x_169, 3, x_159); +return x_169; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withMilliseconds___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_withMilliseconds(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_withNanoseconds(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); +if (lean_is_exclusive(x_1)) { + lean_ctor_release(x_1, 0); + lean_ctor_release(x_1, 1); + lean_ctor_release(x_1, 2); + lean_ctor_release(x_1, 3); + x_5 = x_1; +} else { + lean_dec_ref(x_1); + x_5 = lean_box(0); +} +x_6 = lean_thunk_get_own(x_3); +lean_dec(x_3); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_6, 1); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_10 = lean_ctor_get(x_8, 3); +lean_dec(x_10); +lean_ctor_set(x_8, 3, x_2); +x_11 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +x_15 = lean_unsigned_to_nat(0u); +x_16 = l_Array_findIdx_x3f_loop___rarg(x_13, x_14, x_15); +x_17 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_18 = lean_int_mul(x_12, x_17); +x_19 = lean_ctor_get(x_11, 1); +lean_inc(x_19); +lean_dec(x_11); +x_20 = lean_int_add(x_18, x_19); +lean_dec(x_19); +lean_dec(x_18); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_34; +lean_dec(x_12); +x_34 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_34) == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_4, 0); +lean_inc(x_35); +x_21 = x_35; +goto block_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_21 = x_37; +goto block_33; +} +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_38 = lean_ctor_get(x_16, 0); +lean_inc(x_38); +lean_dec(x_16); +x_39 = lean_unsigned_to_nat(1u); +x_40 = lean_nat_sub(x_38, x_39); +x_41 = l_Array_get_x3f___rarg(x_14, x_40); +lean_dec(x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; +lean_dec(x_38); +lean_dec(x_14); +lean_dec(x_12); +x_42 = lean_ctor_get(x_4, 0); +lean_inc(x_42); +x_21 = x_42; +goto block_33; +} +else +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Array_get_x3f___rarg(x_14, x_38); +lean_dec(x_38); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; +x_45 = l_Array_back_x3f___rarg(x_14); +lean_dec(x_14); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; +lean_dec(x_43); +lean_dec(x_12); +x_46 = lean_ctor_get(x_4, 0); +lean_inc(x_46); +x_21 = x_46; +goto block_33; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_43, 1); +lean_inc(x_49); +lean_dec(x_43); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_nat_abs(x_50); +lean_dec(x_50); +x_52 = lean_nat_to_int(x_51); +x_53 = lean_int_sub(x_48, x_52); +lean_dec(x_52); +lean_dec(x_48); +x_54 = lean_int_dec_lt(x_12, x_53); +lean_dec(x_53); +lean_dec(x_12); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_49); +x_55 = lean_ctor_get(x_47, 1); +lean_inc(x_55); +lean_dec(x_47); +x_21 = x_55; +goto block_33; +} +else +{ +lean_dec(x_47); +x_21 = x_49; +goto block_33; +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_dec(x_14); +x_56 = lean_ctor_get(x_44, 0); +lean_inc(x_56); +lean_dec(x_44); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_43, 1); +lean_inc(x_58); +lean_dec(x_43); +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_nat_abs(x_59); +lean_dec(x_59); +x_61 = lean_nat_to_int(x_60); +x_62 = lean_int_sub(x_57, x_61); +lean_dec(x_61); +lean_dec(x_57); +x_63 = lean_int_dec_lt(x_12, x_62); +lean_dec(x_62); +lean_dec(x_12); +if (x_63 == 0) +{ +lean_object* x_64; +lean_dec(x_58); +x_64 = lean_ctor_get(x_56, 1); +lean_inc(x_64); +lean_dec(x_56); +x_21 = x_64; +goto block_33; +} +else +{ +lean_dec(x_56); +x_21 = x_58; +goto block_33; +} +} +} +} +block_33: +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_int_neg(x_23); +x_25 = lean_int_mul(x_24, x_17); +lean_dec(x_24); +x_26 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_27 = lean_int_add(x_25, x_26); +lean_dec(x_25); +x_28 = lean_int_add(x_20, x_27); +lean_dec(x_27); +lean_dec(x_20); +x_29 = l_Std_Time_Duration_ofNanoseconds(x_28); +lean_dec(x_28); +lean_inc(x_29); +x_30 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_17); +x_31 = lean_mk_thunk(x_30); +if (lean_is_scalar(x_5)) { + x_32 = lean_alloc_ctor(0, 4, 0); +} else { + x_32 = x_5; +} +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_29); +lean_ctor_set(x_32, 2, x_4); +lean_ctor_set(x_32, 3, x_22); +return x_32; +} +} +else +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_65 = lean_ctor_get(x_8, 0); +x_66 = lean_ctor_get(x_8, 1); +x_67 = lean_ctor_get(x_8, 2); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_dec(x_8); +x_68 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +lean_ctor_set(x_68, 2, x_67); +lean_ctor_set(x_68, 3, x_2); +lean_ctor_set(x_6, 1, x_68); +x_69 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_6); +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +lean_inc(x_70); +x_71 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_71, 0, x_70); +x_72 = lean_ctor_get(x_4, 1); +lean_inc(x_72); +x_73 = lean_unsigned_to_nat(0u); +x_74 = l_Array_findIdx_x3f_loop___rarg(x_71, x_72, x_73); +x_75 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_76 = lean_int_mul(x_70, x_75); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +lean_dec(x_69); +x_78 = lean_int_add(x_76, x_77); +lean_dec(x_77); +lean_dec(x_76); +if (lean_obj_tag(x_74) == 0) +{ +lean_object* x_92; +lean_dec(x_70); +x_92 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_92) == 0) +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_4, 0); +lean_inc(x_93); +x_79 = x_93; +goto block_91; +} +else +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_92, 0); +lean_inc(x_94); +lean_dec(x_92); +x_95 = lean_ctor_get(x_94, 1); +lean_inc(x_95); +lean_dec(x_94); +x_79 = x_95; +goto block_91; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_96 = lean_ctor_get(x_74, 0); +lean_inc(x_96); +lean_dec(x_74); +x_97 = lean_unsigned_to_nat(1u); +x_98 = lean_nat_sub(x_96, x_97); +x_99 = l_Array_get_x3f___rarg(x_72, x_98); +lean_dec(x_98); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; +lean_dec(x_96); +lean_dec(x_72); +lean_dec(x_70); +x_100 = lean_ctor_get(x_4, 0); +lean_inc(x_100); +x_79 = x_100; +goto block_91; +} +else +{ +lean_object* x_101; lean_object* x_102; +x_101 = lean_ctor_get(x_99, 0); +lean_inc(x_101); +lean_dec(x_99); +x_102 = l_Array_get_x3f___rarg(x_72, x_96); +lean_dec(x_96); +if (lean_obj_tag(x_102) == 0) +{ +lean_object* x_103; +x_103 = l_Array_back_x3f___rarg(x_72); +lean_dec(x_72); +if (lean_obj_tag(x_103) == 0) +{ +lean_object* x_104; +lean_dec(x_101); +lean_dec(x_70); +x_104 = lean_ctor_get(x_4, 0); +lean_inc(x_104); +x_79 = x_104; +goto block_91; +} +else +{ +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_105 = lean_ctor_get(x_103, 0); +lean_inc(x_105); +lean_dec(x_103); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_101, 1); +lean_inc(x_107); +lean_dec(x_101); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_nat_abs(x_108); +lean_dec(x_108); +x_110 = lean_nat_to_int(x_109); +x_111 = lean_int_sub(x_106, x_110); +lean_dec(x_110); +lean_dec(x_106); +x_112 = lean_int_dec_lt(x_70, x_111); +lean_dec(x_111); +lean_dec(x_70); +if (x_112 == 0) +{ +lean_object* x_113; +lean_dec(x_107); +x_113 = lean_ctor_get(x_105, 1); +lean_inc(x_113); +lean_dec(x_105); +x_79 = x_113; +goto block_91; +} +else +{ +lean_dec(x_105); +x_79 = x_107; +goto block_91; +} +} +} +else +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +lean_dec(x_72); +x_114 = lean_ctor_get(x_102, 0); +lean_inc(x_114); +lean_dec(x_102); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_101, 1); +lean_inc(x_116); +lean_dec(x_101); +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_nat_abs(x_117); +lean_dec(x_117); +x_119 = lean_nat_to_int(x_118); +x_120 = lean_int_sub(x_115, x_119); +lean_dec(x_119); +lean_dec(x_115); +x_121 = lean_int_dec_lt(x_70, x_120); +lean_dec(x_120); +lean_dec(x_70); +if (x_121 == 0) +{ +lean_object* x_122; +lean_dec(x_116); +x_122 = lean_ctor_get(x_114, 1); +lean_inc(x_122); +lean_dec(x_114); +x_79 = x_122; +goto block_91; +} +else +{ +lean_dec(x_114); +x_79 = x_116; +goto block_91; +} +} +} +} +block_91: +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_80 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_79); +lean_dec(x_79); +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_int_neg(x_81); +x_83 = lean_int_mul(x_82, x_75); +lean_dec(x_82); +x_84 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_85 = lean_int_add(x_83, x_84); +lean_dec(x_83); +x_86 = lean_int_add(x_78, x_85); +lean_dec(x_85); +lean_dec(x_78); +x_87 = l_Std_Time_Duration_ofNanoseconds(x_86); +lean_dec(x_86); +lean_inc(x_87); +x_88 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_88, 0, x_87); +lean_closure_set(x_88, 1, x_81); +lean_closure_set(x_88, 2, x_75); +x_89 = lean_mk_thunk(x_88); +if (lean_is_scalar(x_5)) { + x_90 = lean_alloc_ctor(0, 4, 0); +} else { + x_90 = x_5; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_87); +lean_ctor_set(x_90, 2, x_4); +lean_ctor_set(x_90, 3, x_80); +return x_90; +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +x_123 = lean_ctor_get(x_6, 1); +x_124 = lean_ctor_get(x_6, 0); +lean_inc(x_123); +lean_inc(x_124); +lean_dec(x_6); +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +x_127 = lean_ctor_get(x_123, 2); +lean_inc(x_127); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + lean_ctor_release(x_123, 2); + lean_ctor_release(x_123, 3); + x_128 = x_123; +} else { + lean_dec_ref(x_123); + x_128 = lean_box(0); +} +if (lean_is_scalar(x_128)) { + x_129 = lean_alloc_ctor(0, 4, 0); +} else { + x_129 = x_128; +} +lean_ctor_set(x_129, 0, x_125); +lean_ctor_set(x_129, 1, x_126); +lean_ctor_set(x_129, 2, x_127); +lean_ctor_set(x_129, 3, x_2); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_124); +lean_ctor_set(x_130, 1, x_129); +x_131 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_130); +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +lean_inc(x_132); +x_133 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_133, 0, x_132); +x_134 = lean_ctor_get(x_4, 1); +lean_inc(x_134); +x_135 = lean_unsigned_to_nat(0u); +x_136 = l_Array_findIdx_x3f_loop___rarg(x_133, x_134, x_135); +x_137 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_138 = lean_int_mul(x_132, x_137); +x_139 = lean_ctor_get(x_131, 1); +lean_inc(x_139); +lean_dec(x_131); +x_140 = lean_int_add(x_138, x_139); +lean_dec(x_139); +lean_dec(x_138); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_154; +lean_dec(x_132); +x_154 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_154) == 0) +{ +lean_object* x_155; +x_155 = lean_ctor_get(x_4, 0); +lean_inc(x_155); +x_141 = x_155; +goto block_153; +} +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_ctor_get(x_154, 0); +lean_inc(x_156); +lean_dec(x_154); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_141 = x_157; +goto block_153; +} +} +else +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = lean_ctor_get(x_136, 0); +lean_inc(x_158); +lean_dec(x_136); +x_159 = lean_unsigned_to_nat(1u); +x_160 = lean_nat_sub(x_158, x_159); +x_161 = l_Array_get_x3f___rarg(x_134, x_160); +lean_dec(x_160); +if (lean_obj_tag(x_161) == 0) +{ +lean_object* x_162; +lean_dec(x_158); +lean_dec(x_134); +lean_dec(x_132); +x_162 = lean_ctor_get(x_4, 0); +lean_inc(x_162); +x_141 = x_162; +goto block_153; +} +else +{ +lean_object* x_163; lean_object* x_164; +x_163 = lean_ctor_get(x_161, 0); +lean_inc(x_163); +lean_dec(x_161); +x_164 = l_Array_get_x3f___rarg(x_134, x_158); +lean_dec(x_158); +if (lean_obj_tag(x_164) == 0) +{ +lean_object* x_165; +x_165 = l_Array_back_x3f___rarg(x_134); +lean_dec(x_134); +if (lean_obj_tag(x_165) == 0) +{ +lean_object* x_166; +lean_dec(x_163); +lean_dec(x_132); +x_166 = lean_ctor_get(x_4, 0); +lean_inc(x_166); +x_141 = x_166; +goto block_153; +} +else +{ +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +x_167 = lean_ctor_get(x_165, 0); +lean_inc(x_167); +lean_dec(x_165); +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_163, 1); +lean_inc(x_169); +lean_dec(x_163); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_nat_abs(x_170); +lean_dec(x_170); +x_172 = lean_nat_to_int(x_171); +x_173 = lean_int_sub(x_168, x_172); +lean_dec(x_172); +lean_dec(x_168); +x_174 = lean_int_dec_lt(x_132, x_173); +lean_dec(x_173); +lean_dec(x_132); +if (x_174 == 0) +{ +lean_object* x_175; +lean_dec(x_169); +x_175 = lean_ctor_get(x_167, 1); +lean_inc(x_175); +lean_dec(x_167); +x_141 = x_175; +goto block_153; +} +else +{ +lean_dec(x_167); +x_141 = x_169; +goto block_153; +} +} +} +else +{ +lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; +lean_dec(x_134); +x_176 = lean_ctor_get(x_164, 0); +lean_inc(x_176); +lean_dec(x_164); +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_163, 1); +lean_inc(x_178); +lean_dec(x_163); +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_nat_abs(x_179); +lean_dec(x_179); +x_181 = lean_nat_to_int(x_180); +x_182 = lean_int_sub(x_177, x_181); +lean_dec(x_181); +lean_dec(x_177); +x_183 = lean_int_dec_lt(x_132, x_182); +lean_dec(x_182); +lean_dec(x_132); +if (x_183 == 0) +{ +lean_object* x_184; +lean_dec(x_178); +x_184 = lean_ctor_get(x_176, 1); +lean_inc(x_184); +lean_dec(x_176); +x_141 = x_184; +goto block_153; +} +else +{ +lean_dec(x_176); +x_141 = x_178; +goto block_153; +} +} +} +} +block_153: +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_142 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_141); +lean_dec(x_141); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_int_neg(x_143); +x_145 = lean_int_mul(x_144, x_137); +lean_dec(x_144); +x_146 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_147 = lean_int_add(x_145, x_146); +lean_dec(x_145); +x_148 = lean_int_add(x_140, x_147); +lean_dec(x_147); +lean_dec(x_140); +x_149 = l_Std_Time_Duration_ofNanoseconds(x_148); +lean_dec(x_148); +lean_inc(x_149); +x_150 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_150, 0, x_149); +lean_closure_set(x_150, 1, x_143); +lean_closure_set(x_150, 2, x_137); +x_151 = lean_mk_thunk(x_150); +if (lean_is_scalar(x_5)) { + x_152 = lean_alloc_ctor(0, 4, 0); +} else { + x_152 = x_5; +} +lean_ctor_set(x_152, 0, x_151); +lean_ctor_set(x_152, 1, x_149); +lean_ctor_set(x_152, 2, x_4); +lean_ctor_set(x_152, 3, x_142); +return x_152; +} +} +} +} +LEAN_EXPORT uint8_t l_Std_Time_ZonedDateTime_inLeapYear(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +lean_dec(x_4); +x_6 = l_Std_Time_ZonedDateTime_dayOfYear___closed__1; +x_7 = lean_int_mod(x_5, x_6); +x_8 = l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1; +x_9 = lean_int_dec_eq(x_7, x_8); +lean_dec(x_7); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_5); +x_10 = 0; +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; +x_11 = l_Std_Time_ZonedDateTime_dayOfYear___closed__2; +x_12 = lean_int_mod(x_5, x_11); +x_13 = lean_int_dec_eq(x_12, x_8); +lean_dec(x_12); +x_14 = l_instDecidableNot___rarg(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = l_Std_Time_ZonedDateTime_dayOfYear___closed__3; +x_16 = lean_int_mod(x_5, x_15); +lean_dec(x_5); +x_17 = lean_int_dec_eq(x_16, x_8); +lean_dec(x_16); +return x_17; +} +else +{ +uint8_t x_18; +lean_dec(x_5); +x_18 = 1; +return x_18; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_inLeapYear___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Std_Time_ZonedDateTime_inLeapYear(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDaysSinceUNIXEpoch(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_thunk_get_own(x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +lean_dec(x_3); +x_5 = l_Std_Time_PlainDate_toDaysSinceUNIXEpoch(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_toDaysSinceUNIXEpoch___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Std_Time_ZonedDateTime_toDaysSinceUNIXEpoch(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofDaysSinceUNIXEpoch(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_4 = l_Std_Time_PlainDate_ofDaysSinceUNIXEpoch(x_1); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +x_6 = l_Std_Time_PlainDateTime_toTimestampAssumingUTC(x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_inc(x_7); +x_8 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__1___boxed), 2, 1); +lean_closure_set(x_8, 0, x_7); +x_9 = lean_ctor_get(x_3, 1); +lean_inc(x_9); +x_10 = lean_unsigned_to_nat(0u); +x_11 = l_Array_findIdx_x3f_loop___rarg(x_8, x_9, x_10); +x_12 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_13 = lean_int_mul(x_7, x_12); +x_14 = lean_ctor_get(x_6, 1); +lean_inc(x_14); +lean_dec(x_6); +x_15 = lean_int_add(x_13, x_14); +lean_dec(x_14); +lean_dec(x_13); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_29; +lean_dec(x_7); +x_29 = l_Array_back_x3f___rarg(x_9); +lean_dec(x_9); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; +x_30 = lean_ctor_get(x_3, 0); +lean_inc(x_30); +x_16 = x_30; +goto block_28; +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_16 = x_32; +goto block_28; +} +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_11, 0); +lean_inc(x_33); +lean_dec(x_11); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_sub(x_33, x_34); +x_36 = l_Array_get_x3f___rarg(x_9, x_35); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; +lean_dec(x_33); +lean_dec(x_9); +lean_dec(x_7); +x_37 = lean_ctor_get(x_3, 0); +lean_inc(x_37); +x_16 = x_37; +goto block_28; +} +else +{ +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Array_get_x3f___rarg(x_9, x_33); +lean_dec(x_33); +if (lean_obj_tag(x_39) == 0) +{ +lean_object* x_40; +x_40 = l_Array_back_x3f___rarg(x_9); +lean_dec(x_9); +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_41; +lean_dec(x_38); +lean_dec(x_7); +x_41 = lean_ctor_get(x_3, 0); +lean_inc(x_41); +x_16 = x_41; +goto block_28; +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_42 = lean_ctor_get(x_40, 0); +lean_inc(x_42); +lean_dec(x_40); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_38, 1); +lean_inc(x_44); +lean_dec(x_38); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_nat_abs(x_45); +lean_dec(x_45); +x_47 = lean_nat_to_int(x_46); +x_48 = lean_int_sub(x_43, x_47); +lean_dec(x_47); +lean_dec(x_43); +x_49 = lean_int_dec_lt(x_7, x_48); +lean_dec(x_48); +lean_dec(x_7); +if (x_49 == 0) +{ +lean_object* x_50; +lean_dec(x_44); +x_50 = lean_ctor_get(x_42, 1); +lean_inc(x_50); +lean_dec(x_42); +x_16 = x_50; +goto block_28; +} +else +{ +lean_dec(x_42); +x_16 = x_44; +goto block_28; +} +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +lean_dec(x_9); +x_51 = lean_ctor_get(x_39, 0); +lean_inc(x_51); +lean_dec(x_39); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_38, 1); +lean_inc(x_53); +lean_dec(x_38); +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_nat_abs(x_54); +lean_dec(x_54); +x_56 = lean_nat_to_int(x_55); +x_57 = lean_int_sub(x_52, x_56); +lean_dec(x_56); +lean_dec(x_52); +x_58 = lean_int_dec_lt(x_7, x_57); +lean_dec(x_57); +lean_dec(x_7); +if (x_58 == 0) +{ +lean_object* x_59; +lean_dec(x_53); +x_59 = lean_ctor_get(x_51, 1); +lean_inc(x_59); +lean_dec(x_51); +x_16 = x_59; +goto block_28; +} +else +{ +lean_dec(x_51); +x_16 = x_53; +goto block_28; +} +} +} +} +block_28: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_17 = l_Std_Time_TimeZone_LocalTimeType_getTimeZone(x_16); +lean_dec(x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_int_neg(x_18); +x_20 = lean_int_mul(x_19, x_12); +lean_dec(x_19); +x_21 = l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1; +x_22 = lean_int_add(x_20, x_21); +lean_dec(x_20); +x_23 = lean_int_add(x_15, x_22); +lean_dec(x_22); +lean_dec(x_15); +x_24 = l_Std_Time_Duration_ofNanoseconds(x_23); +lean_dec(x_23); +lean_inc(x_24); +x_25 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_ofPlainDateTime___lambda__2___boxed), 4, 3); +lean_closure_set(x_25, 0, x_24); +lean_closure_set(x_25, 1, x_18); +lean_closure_set(x_25, 2, x_12); +x_26 = lean_mk_thunk(x_25); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +lean_ctor_set(x_27, 2, x_3); +lean_ctor_set(x_27, 3, x_17); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_ofDaysSinceUNIXEpoch___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Std_Time_ZonedDateTime_ofDaysSinceUNIXEpoch(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subDays___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subWeeks___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subHours___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subMinutes___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subSeconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subMilliseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__5() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_addNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHAddOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1; +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Std_Time_ZonedDateTime_subNanoseconds___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Std_Time_ZonedDateTime_instHSubOffset__6() { +_start: +{ +lean_object* x_1; +x_1 = l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_2, 1); +x_5 = lean_ctor_get(x_4, 0); +x_6 = lean_int_neg(x_5); +x_7 = lean_ctor_get(x_4, 1); +x_8 = lean_int_neg(x_7); +x_9 = lean_ctor_get(x_3, 0); +x_10 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_11 = lean_int_mul(x_9, x_10); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_int_add(x_11, x_12); +lean_dec(x_11); +x_14 = lean_int_mul(x_6, x_10); +lean_dec(x_6); +x_15 = lean_int_add(x_14, x_8); +lean_dec(x_8); +lean_dec(x_14); +x_16 = lean_int_add(x_13, x_15); +lean_dec(x_15); +lean_dec(x_13); +x_17 = l_Std_Time_Duration_ofNanoseconds(x_16); +lean_dec(x_16); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_instHSubDuration(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddDuration(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_ZonedDateTime_addNanoseconds(x_1, x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHAddDuration___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_instHAddDuration(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_2, 0); +x_4 = l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2; +x_5 = lean_int_mul(x_3, x_4); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_int_add(x_5, x_6); +lean_dec(x_5); +x_8 = l_Std_Time_ZonedDateTime_subNanoseconds(x_1, x_7); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Std_Time_ZonedDateTime_instHSubDuration__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Std_Time_ZonedDateTime_instHSubDuration__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +lean_object* initialize_Std_Time_Zoned_DateTime(uint8_t builtin, lean_object*); +lean_object* initialize_Std_Time_Zoned_ZoneRules(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Std_Time_Zoned_ZonedDateTime(uint8_t builtin, lean_object* w) { +lean_object * res; +if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); +_G_initialized = true; +res = initialize_Std_Time_Zoned_DateTime(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Std_Time_Zoned_ZoneRules(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__1); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__2); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__3); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__4); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__5); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__6); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__7); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__8); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__9); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__10); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__11); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__12); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__13); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__14); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__15); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__16); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__17); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__18); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__19); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__20); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__21); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__22); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__23); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__24); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__25); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__26); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__27); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__28); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__29); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__30); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__31); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__32); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__33); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__34); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__35); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__36); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__37); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__38); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__39); +l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40 = _init_l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___lambda__1___closed__40); +l_Std_Time_instInhabitedZonedDateTime___closed__1 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__1); +l_Std_Time_instInhabitedZonedDateTime___closed__2 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__2(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__2); +l_Std_Time_instInhabitedZonedDateTime___closed__3 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__3(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__3); +l_Std_Time_instInhabitedZonedDateTime___closed__4 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__4(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__4); +l_Std_Time_instInhabitedZonedDateTime___closed__5 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__5(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__5); +l_Std_Time_instInhabitedZonedDateTime___closed__6 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__6(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__6); +l_Std_Time_instInhabitedZonedDateTime___closed__7 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__7(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__7); +l_Std_Time_instInhabitedZonedDateTime___closed__8 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__8(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__8); +l_Std_Time_instInhabitedZonedDateTime___closed__9 = _init_l_Std_Time_instInhabitedZonedDateTime___closed__9(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime___closed__9); +l_Std_Time_instInhabitedZonedDateTime = _init_l_Std_Time_instInhabitedZonedDateTime(); +lean_mark_persistent(l_Std_Time_instInhabitedZonedDateTime); +l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1 = _init_l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__1); +l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2 = _init_l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofTimestamp___lambda__1___closed__2); +l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1 = _init_l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofPlainDateTime___closed__1); +l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1 = _init_l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofTimestampWithZone___closed__1); +l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1 = _init_l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_ofPlainDateTimeWithZone___closed__1); +l_Std_Time_ZonedDateTime_millisecond___closed__1 = _init_l_Std_Time_ZonedDateTime_millisecond___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_millisecond___closed__1); +l_Std_Time_ZonedDateTime_dayOfYear___closed__1 = _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_dayOfYear___closed__1); +l_Std_Time_ZonedDateTime_dayOfYear___closed__2 = _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_dayOfYear___closed__2); +l_Std_Time_ZonedDateTime_dayOfYear___closed__3 = _init_l_Std_Time_ZonedDateTime_dayOfYear___closed__3(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_dayOfYear___closed__3); +l_Std_Time_ZonedDateTime_addWeeks___closed__1 = _init_l_Std_Time_ZonedDateTime_addWeeks___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_addWeeks___closed__1); +l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1 = _init_l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_addYearsRollOver___closed__1); +l_Std_Time_ZonedDateTime_addHours___closed__1 = _init_l_Std_Time_ZonedDateTime_addHours___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_addHours___closed__1); +l_Std_Time_ZonedDateTime_addMinutes___closed__1 = _init_l_Std_Time_ZonedDateTime_addMinutes___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_addMinutes___closed__1); +l_Std_Time_ZonedDateTime_addMilliseconds___closed__1 = _init_l_Std_Time_ZonedDateTime_addMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_addMilliseconds___closed__1); +l_Std_Time_ZonedDateTime_withMilliseconds___closed__1 = _init_l_Std_Time_ZonedDateTime_withMilliseconds___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_withMilliseconds___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset = _init_l_Std_Time_ZonedDateTime_instHAddOffset(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset); +l_Std_Time_ZonedDateTime_instHSubOffset___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset = _init_l_Std_Time_ZonedDateTime_instHSubOffset(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset); +l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__1___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__1); +l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__1___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__1); +l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__2___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__2 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__2); +l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__2___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__2 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__2(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__2); +l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__3___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__3 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__3(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__3); +l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__3___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__3 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__3(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__3); +l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__4___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__4 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__4(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__4); +l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__4___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__4 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__4(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__4); +l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__5___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__5 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__5(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__5); +l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__5___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__5 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__5(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__5); +l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__6___closed__1); +l_Std_Time_ZonedDateTime_instHAddOffset__6 = _init_l_Std_Time_ZonedDateTime_instHAddOffset__6(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHAddOffset__6); +l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__6___closed__1); +l_Std_Time_ZonedDateTime_instHSubOffset__6 = _init_l_Std_Time_ZonedDateTime_instHSubOffset__6(); +lean_mark_persistent(l_Std_Time_ZonedDateTime_instHSubOffset__6); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/tests/lean/1018unknowMVarIssue.lean.expected.out b/tests/lean/1018unknowMVarIssue.lean.expected.out index a7585de9af87..7e9a8746fc6b 100644 --- a/tests/lean/1018unknowMVarIssue.lean.expected.out +++ b/tests/lean/1018unknowMVarIssue.lean.expected.out @@ -30,6 +30,7 @@ a : α • a (isBinder := true) : α @ ⟨7, 9⟩-⟨7, 10⟩ • x (isBinder := true) : Fam2 α β @ ⟨7, 17⟩-⟨7, 18⟩ • CustomInfo(Lean.Elab.Term.BodyInfo) + • Partial term @ ⟨8, 2⟩-⟨10, 19⟩ @ Lean.Elab.Term.Quotation.elabMatchSyntax • match α, β, x, a with | α_1, .(α_1), Fam2.any, a => ?m x α_1 a | .(Nat), .(Nat), Fam2.nat n, a => n : β @ ⟨8, 2⟩-⟨10, 19⟩ @ Lean.Elab.Term.elabMatch diff --git a/tests/lean/StxQuot.lean.expected.out b/tests/lean/StxQuot.lean.expected.out index 1256d67be44d..ec84ae2af231 100644 --- a/tests/lean/StxQuot.lean.expected.out +++ b/tests/lean/StxQuot.lean.expected.out @@ -32,8 +32,8 @@ StxQuot.lean:19:15-19:16: error: unexpected token ']'; expected term "`a._@.UnhygienicMain._hyg.1" "(Term.explicitUniv `a._@.UnhygienicMain._hyg.1 \".{\" [(num \"0\")] \"}\")" "#[(Term.matchAlt \"|\" [[`a._@.UnhygienicMain._hyg.1]] \"=>\" (num \"1\")), (Term.matchAlt \"|\" [[(Term.hole \"_\")]] \"=>\" (num \"2\"))]" -"(Term.structInst\n \"{\"\n []\n [(Term.structInstField (Term.structInstLVal `a._@.UnhygienicMain._hyg.1 []) \":=\" `a._@.UnhygienicMain._hyg.1)]\n (Term.optEllipsis [])\n [\":\" `a._@.UnhygienicMain._hyg.1]\n \"}\")" -"(Term.structInst\n \"{\"\n []\n [(Term.structInstField (Term.structInstLVal `a._@.UnhygienicMain._hyg.1 []) \":=\" `a._@.UnhygienicMain._hyg.1)]\n (Term.optEllipsis [])\n []\n \"}\")" +"(Term.structInst\n \"{\"\n []\n (Term.structInstFields\n [(Term.structInstField (Term.structInstLVal `a._@.UnhygienicMain._hyg.1 []) \":=\" `a._@.UnhygienicMain._hyg.1)])\n (Term.optEllipsis [])\n [\":\" `a._@.UnhygienicMain._hyg.1]\n \"}\")" +"(Term.structInst\n \"{\"\n []\n (Term.structInstFields\n [(Term.structInstField (Term.structInstLVal `a._@.UnhygienicMain._hyg.1 []) \":=\" `a._@.UnhygienicMain._hyg.1)])\n (Term.optEllipsis [])\n []\n \"}\")" "(Command.section \"section\" [])" "(Command.section \"section\" [`foo._@.UnhygienicMain._hyg.1])" "(Term.match\n \"match\"\n []\n []\n [(Term.matchDiscr [] `a._@.UnhygienicMain._hyg.1)]\n \"with\"\n (Term.matchAlts\n [(Term.matchAlt \"|\" [[`a._@.UnhygienicMain._hyg.1]] \"=>\" `b._@.UnhygienicMain._hyg.1)\n (Term.matchAlt\n \"|\"\n [[(«term_+_» `a._@.UnhygienicMain._hyg.1 \"+\" (num \"1\"))]]\n \"=>\"\n («term_+_» `b._@.UnhygienicMain._hyg.1 \"+\" (num \"1\")))]))" diff --git a/tests/lean/instance1.lean.expected.out b/tests/lean/instance1.lean.expected.out index f09ccac7f632..174df899030a 100644 --- a/tests/lean/instance1.lean.expected.out +++ b/tests/lean/instance1.lean.expected.out @@ -1 +1 @@ -instance1.lean:6:24-6:29: error: fields missing: 'pure', 'bind' +instance1.lean:6:24-10:0: error: fields missing: 'pure', 'bind' diff --git a/tests/lean/interactive/1265.lean.expected.out b/tests/lean/interactive/1265.lean.expected.out index abe2bf402d4e..0f91b29507f0 100644 --- a/tests/lean/interactive/1265.lean.expected.out +++ b/tests/lean/interactive/1265.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.expandInterpolatedStr"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.expandInterpolatedStr"}}, + "cPos": 1}}, {"sortText": "1", "label": "getChar", "kind": 3, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getChar"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getChar"}}, + "cPos": 1}}, {"sortText": "2", "label": "getDocString", "kind": 3, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getDocString"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getDocString"}}, + "cPos": 1}}, {"sortText": "3", "label": "getHygieneInfo", "kind": 3, @@ -32,7 +35,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getHygieneInfo"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getHygieneInfo"}}, + "cPos": 1}}, {"sortText": "4", "label": "getId", "kind": 3, @@ -40,7 +44,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getId"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getId"}}, + "cPos": 1}}, {"sortText": "5", "label": "getName", "kind": 3, @@ -48,7 +53,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getName"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getName"}}, + "cPos": 1}}, {"sortText": "6", "label": "getNat", "kind": 3, @@ -56,7 +62,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getNat"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getNat"}}, + "cPos": 1}}, {"sortText": "7", "label": "getScientific", "kind": 3, @@ -64,7 +71,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getScientific"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getScientific"}}, + "cPos": 1}}, {"sortText": "8", "label": "getString", "kind": 3, @@ -72,7 +80,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.getString"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getString"}}, + "cPos": 1}}, {"sortText": "9", "label": "raw", "kind": 5, @@ -82,7 +91,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 0, "character": 51}}, - "id": {"const": {"declName": "Lean.TSyntax.raw"}}}}], + "id": {"const": {"declName": "Lean.TSyntax.raw"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}} @@ -94,7 +104,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.expandInterpolatedStr"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.expandInterpolatedStr"}}, + "cPos": 1}}, {"sortText": "1", "label": "getChar", "kind": 3, @@ -102,7 +113,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getChar"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getChar"}}, + "cPos": 1}}, {"sortText": "2", "label": "getDocString", "kind": 3, @@ -110,7 +122,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getDocString"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getDocString"}}, + "cPos": 1}}, {"sortText": "3", "label": "getHygieneInfo", "kind": 3, @@ -118,7 +131,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getHygieneInfo"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getHygieneInfo"}}, + "cPos": 1}}, {"sortText": "4", "label": "getId", "kind": 3, @@ -126,7 +140,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getId"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getId"}}, + "cPos": 1}}, {"sortText": "5", "label": "getName", "kind": 3, @@ -134,7 +149,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getName"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getName"}}, + "cPos": 1}}, {"sortText": "6", "label": "getNat", "kind": 3, @@ -142,7 +158,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getNat"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getNat"}}, + "cPos": 1}}, {"sortText": "7", "label": "getScientific", "kind": 3, @@ -150,7 +167,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getScientific"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getScientific"}}, + "cPos": 1}}, {"sortText": "8", "label": "getString", "kind": 3, @@ -158,7 +176,8 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.getString"}}}}, + "id": {"const": {"declName": "Lean.TSyntax.getString"}}, + "cPos": 1}}, {"sortText": "9", "label": "raw", "kind": 5, @@ -168,5 +187,6 @@ {"params": {"textDocument": {"uri": "file:///1265.lean"}, "position": {"line": 2, "character": 53}}, - "id": {"const": {"declName": "Lean.TSyntax.raw"}}}}], + "id": {"const": {"declName": "Lean.TSyntax.raw"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/1659.lean.expected.out b/tests/lean/interactive/1659.lean.expected.out index dc26f41f1c31..e98343098f4d 100644 --- a/tests/lean/interactive/1659.lean.expected.out +++ b/tests/lean/interactive/1659.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 9, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Elab.Term.elabTermEnsuringType", "kind": 21, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 9, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}}}], + "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 15, "character": 23}} @@ -28,7 +30,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 15, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "1", "label": "Term.elabTermEnsuringType", "kind": 21, @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 15, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}}}], + "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 21, "character": 23}} @@ -48,7 +52,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 21, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Elab.Tactic.elabTermEnsuringType", "kind": 21, @@ -56,7 +61,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 21, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}}}], + "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 27, "character": 23}} @@ -68,7 +74,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 27, "character": 23}}, - "id": {"const": {"declName": "elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Elab.Tactic.elabTermEnsuringType", "kind": 21, @@ -76,7 +83,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 27, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Elab.Term.elabTermEnsuringType", "kind": 21, @@ -84,7 +92,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 27, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}}}], + "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 33, "character": 23}} @@ -96,7 +105,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 33, "character": 23}}, - "id": {"const": {"declName": "elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Elab.elabTermEnsuringType", "kind": 21, @@ -104,7 +114,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 33, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Elab.Tactic.elabTermEnsuringType", "kind": 21, @@ -112,7 +123,8 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 33, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}}}, + "id": {"const": {"declName": "Lean.Elab.Tactic.elabTermEnsuringType"}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Elab.Term.elabTermEnsuringType", "kind": 21, @@ -120,5 +132,6 @@ {"params": {"textDocument": {"uri": "file:///1659.lean"}, "position": {"line": 33, "character": 23}}, - "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}}}], + "id": {"const": {"declName": "Lean.Elab.Term.elabTermEnsuringType"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/4880.lean.expected.out b/tests/lean/interactive/4880.lean.expected.out index 88d320c447ec..de96f1cf4891 100644 --- a/tests/lean/interactive/4880.lean.expected.out +++ b/tests/lean/interactive/4880.lean.expected.out @@ -5,22 +5,22 @@ "severity": 1, "range": {"start": {"line": 16, "character": 12}, - "end": {"line": 16, "character": 17}}, + "end": {"line": 17, "character": 0}}, "message": "could not synthesize default value for field 'h1' of 'B' using tactics", "fullRange": {"start": {"line": 16, "character": 12}, - "end": {"line": 16, "character": 17}}}, + "end": {"line": 20, "character": 0}}}, {"source": "Lean 4", "severity": 1, "range": {"start": {"line": 16, "character": 12}, - "end": {"line": 16, "character": 17}}, + "end": {"line": 17, "character": 0}}, "message": "failed to synthesize\n A\nAdditional diagnostic information may be available using the `set_option diagnostics true` command.", "fullRange": {"start": {"line": 16, "character": 12}, - "end": {"line": 16, "character": 17}}}, + "end": {"line": 20, "character": 0}}}, {"source": "Lean 4", "severity": 1, "range": @@ -45,8 +45,8 @@ "severity": 1, "range": {"start": {"line": 34, "character": 13}, - "end": {"line": 34, "character": 18}}, + "end": {"line": 35, "character": 0}}, "message": "unsolved goals\n⊢ A", "fullRange": {"start": {"line": 34, "character": 13}, - "end": {"line": 34, "character": 18}}}]} + "end": {"line": 36, "character": 0}}}]} diff --git a/tests/lean/interactive/533.lean.expected.out b/tests/lean/interactive/533.lean.expected.out index 72e7be4fb911..db615b173990 100644 --- a/tests/lean/interactive/533.lean.expected.out +++ b/tests/lean/interactive/533.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///533.lean"}, "position": {"line": 4, "character": 10}}, - "id": {"fvar": {"id": "_uniq.6"}}}}, + "id": {"fvar": {"id": "_uniq.6"}}, + "cPos": 0}}, {"sortText": "1", "label": "Foo", "kind": 6, @@ -16,5 +17,6 @@ {"params": {"textDocument": {"uri": "file:///533.lean"}, "position": {"line": 4, "character": 10}}, - "id": {"fvar": {"id": "_uniq.2"}}}}], + "id": {"fvar": {"id": "_uniq.2"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/863.lean.expected.out b/tests/lean/interactive/863.lean.expected.out index 6843b01e86c4..7ec3d09d6a6b 100644 --- a/tests/lean/interactive/863.lean.expected.out +++ b/tests/lean/interactive/863.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///863.lean"}, "position": {"line": 9, "character": 12}}, - "id": {"const": {"declName": "Lean.MonadRef.getRef"}}}}, + "id": {"const": {"declName": "Lean.MonadRef.getRef"}}, + "cPos": 0}}, {"sortText": "1", "label": "MonadRef.getRef", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///863.lean"}, "position": {"line": 9, "character": 12}}, - "id": {"const": {"declName": "Lean.MonadRef.getRef"}}}}], + "id": {"const": {"declName": "Lean.MonadRef.getRef"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///863.lean"}, "position": {"line": 13, "character": 12}} @@ -28,7 +30,8 @@ {"params": {"textDocument": {"uri": "file:///863.lean"}, "position": {"line": 13, "character": 12}}, - "id": {"const": {"declName": "Lean.MonadRef.getRef"}}}}, + "id": {"const": {"declName": "Lean.MonadRef.getRef"}}, + "cPos": 0}}, {"sortText": "1", "label": "MonadRef.getRef", "kind": 5, @@ -36,5 +39,6 @@ {"params": {"textDocument": {"uri": "file:///863.lean"}, "position": {"line": 13, "character": 12}}, - "id": {"const": {"declName": "Lean.MonadRef.getRef"}}}}], + "id": {"const": {"declName": "Lean.MonadRef.getRef"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/compHeader.lean.expected.out b/tests/lean/interactive/compHeader.lean.expected.out index d4ca54fdbfdd..e2418fe0a408 100644 --- a/tests/lean/interactive/compHeader.lean.expected.out +++ b/tests/lean/interactive/compHeader.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///compHeader.lean"}, "position": {"line": 2, "character": 22}}, - "id": {"fvar": {"id": "_uniq.7"}}}}, + "id": {"fvar": {"id": "_uniq.7"}}, + "cPos": 0}}, {"sortText": "1", "label": "veryLongNameForCompletion", "kind": 21, @@ -16,5 +17,6 @@ {"params": {"textDocument": {"uri": "file:///compHeader.lean"}, "position": {"line": 2, "character": 22}}, - "id": {"const": {"declName": "veryLongNameForCompletion"}}}}], + "id": {"const": {"declName": "veryLongNameForCompletion"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/compNamespace.lean.expected.out b/tests/lean/interactive/compNamespace.lean.expected.out index ba2b53585e65..46bce389fdf6 100644 --- a/tests/lean/interactive/compNamespace.lean.expected.out +++ b/tests/lean/interactive/compNamespace.lean.expected.out @@ -8,7 +8,8 @@ "data": {"params": {"textDocument": {"uri": "file:///compNamespace.lean"}, - "position": {"line": 5, "character": 12}}}}], + "position": {"line": 5, "character": 12}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///compNamespace.lean"}, "position": {"line": 9, "character": 12}} @@ -20,7 +21,8 @@ "data": {"params": {"textDocument": {"uri": "file:///compNamespace.lean"}, - "position": {"line": 9, "character": 12}}}}], + "position": {"line": 9, "character": 12}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///compNamespace.lean"}, "position": {"line": 13, "character": 11}} @@ -32,7 +34,8 @@ "data": {"params": {"textDocument": {"uri": "file:///compNamespace.lean"}, - "position": {"line": 13, "character": 11}}}}], + "position": {"line": 13, "character": 11}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///compNamespace.lean"}, "position": {"line": 16, "character": 16}} @@ -44,7 +47,8 @@ "data": {"params": {"textDocument": {"uri": "file:///compNamespace.lean"}, - "position": {"line": 16, "character": 16}}}}], + "position": {"line": 16, "character": 16}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///compNamespace.lean"}, "position": {"line": 20, "character": 12}} @@ -56,5 +60,6 @@ "data": {"params": {"textDocument": {"uri": "file:///compNamespace.lean"}, - "position": {"line": 20, "character": 12}}}}], + "position": {"line": 20, "character": 12}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion.lean.expected.out b/tests/lean/interactive/completion.lean.expected.out index 09db65135461..70876007ea5b 100644 --- a/tests/lean/interactive/completion.lean.expected.out +++ b/tests/lean/interactive/completion.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 3, "character": 22}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 5, "character": 23}} @@ -20,7 +21,8 @@ {"params": {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 5, "character": 23}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 7, "character": 28}} @@ -32,7 +34,8 @@ {"params": {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 7, "character": 28}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 9, "character": 29}} @@ -44,5 +47,6 @@ {"params": {"textDocument": {"uri": "file:///completion.lean"}, "position": {"line": 9, "character": 29}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion2.lean.expected.out b/tests/lean/interactive/completion2.lean.expected.out index 3f9aafd094c0..8788a4bbf773 100644 --- a/tests/lean/interactive/completion2.lean.expected.out +++ b/tests/lean/interactive/completion2.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 19, "character": 10}}, - "id": {"const": {"declName": "Foo.Bla.ax1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ax1"}}, + "cPos": 0}}, {"sortText": "1", "label": "ex1", "kind": 23, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 19, "character": 10}}, - "id": {"const": {"declName": "Foo.Bla.ex1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex1"}}, + "cPos": 0}}, {"sortText": "2", "label": "ex2", "kind": 23, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 19, "character": 10}}, - "id": {"const": {"declName": "Foo.Bla.ex2"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex2"}}, + "cPos": 0}}, {"sortText": "3", "label": "ex3", "kind": 23, @@ -32,7 +35,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 19, "character": 10}}, - "id": {"const": {"declName": "Foo.Bla.ex3"}}}}], + "id": {"const": {"declName": "Foo.Bla.ex3"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 25, "character": 6}} @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 25, "character": 6}}, - "id": {"const": {"declName": "Foo.Bla.ax1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ax1"}}, + "cPos": 0}}, {"sortText": "1", "label": "ex1", "kind": 23, @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 25, "character": 6}}, - "id": {"const": {"declName": "Foo.Bla.ex1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex1"}}, + "cPos": 0}}, {"sortText": "2", "label": "ex2", "kind": 23, @@ -60,7 +66,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 25, "character": 6}}, - "id": {"const": {"declName": "Foo.Bla.ex2"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex2"}}, + "cPos": 0}}, {"sortText": "3", "label": "ex3", "kind": 23, @@ -68,7 +75,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 25, "character": 6}}, - "id": {"const": {"declName": "Foo.Bla.ex3"}}}}], + "id": {"const": {"declName": "Foo.Bla.ex3"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 30, "character": 21}} @@ -80,7 +88,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 30, "character": 21}}, - "id": {"const": {"declName": "Foo.Bla.ax1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ax1"}}, + "cPos": 0}}, {"sortText": "1", "label": "ex1", "kind": 23, @@ -88,7 +97,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 30, "character": 21}}, - "id": {"const": {"declName": "Foo.Bla.ex1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex1"}}, + "cPos": 0}}, {"sortText": "2", "label": "ex2", "kind": 23, @@ -96,7 +106,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 30, "character": 21}}, - "id": {"const": {"declName": "Foo.Bla.ex2"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex2"}}, + "cPos": 0}}, {"sortText": "3", "label": "ex3", "kind": 23, @@ -104,7 +115,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 30, "character": 21}}, - "id": {"const": {"declName": "Foo.Bla.ex3"}}}}], + "id": {"const": {"declName": "Foo.Bla.ex3"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 37, "character": 22}} @@ -116,7 +128,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 37, "character": 22}}, - "id": {"const": {"declName": "Foo.Bla.ex1"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex1"}}, + "cPos": 0}}, {"sortText": "1", "label": "ex2", "kind": 23, @@ -124,7 +137,8 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 37, "character": 22}}, - "id": {"const": {"declName": "Foo.Bla.ex2"}}}}, + "id": {"const": {"declName": "Foo.Bla.ex2"}}, + "cPos": 0}}, {"sortText": "2", "label": "ex3", "kind": 23, @@ -132,5 +146,6 @@ {"params": {"textDocument": {"uri": "file:///completion2.lean"}, "position": {"line": 37, "character": 22}}, - "id": {"const": {"declName": "Foo.Bla.ex3"}}}}], + "id": {"const": {"declName": "Foo.Bla.ex3"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion3.lean.expected.out b/tests/lean/interactive/completion3.lean.expected.out index e3f61c586b8e..bca23dfea814 100644 --- a/tests/lean/interactive/completion3.lean.expected.out +++ b/tests/lean/interactive/completion3.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 7, "character": 9}}, - "id": {"const": {"declName": "S.b"}}}}, + "id": {"const": {"declName": "S.b"}}, + "cPos": 1}}, {"sortText": "1", "label": "x", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 7, "character": 9}}, - "id": {"const": {"declName": "S.x"}}}}, + "id": {"const": {"declName": "S.x"}}, + "cPos": 1}}, {"sortText": "2", "label": "y", "kind": 5, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 7, "character": 9}}, - "id": {"const": {"declName": "S.y"}}}}], + "id": {"const": {"declName": "S.y"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 12, "character": 5}} @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 12, "character": 5}}, - "id": {"const": {"declName": "S.b"}}}}, + "id": {"const": {"declName": "S.b"}}, + "cPos": 1}}, {"sortText": "1", "label": "x", "kind": 5, @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 12, "character": 5}}, - "id": {"const": {"declName": "S.x"}}}}, + "id": {"const": {"declName": "S.x"}}, + "cPos": 1}}, {"sortText": "2", "label": "y", "kind": 5, @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 12, "character": 5}}, - "id": {"const": {"declName": "S.y"}}}}], + "id": {"const": {"declName": "S.y"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 16, "character": 5}} @@ -64,7 +70,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 16, "character": 5}}, - "id": {"const": {"declName": "S.b"}}}}, + "id": {"const": {"declName": "S.b"}}, + "cPos": 1}}, {"sortText": "1", "label": "x", "kind": 5, @@ -72,7 +79,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 16, "character": 5}}, - "id": {"const": {"declName": "S.x"}}}}, + "id": {"const": {"declName": "S.x"}}, + "cPos": 1}}, {"sortText": "2", "label": "y", "kind": 5, @@ -80,7 +88,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 16, "character": 5}}, - "id": {"const": {"declName": "S.y"}}}}], + "id": {"const": {"declName": "S.y"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 20, "character": 5}} @@ -92,7 +101,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 20, "character": 5}}, - "id": {"const": {"declName": "S.b"}}}}, + "id": {"const": {"declName": "S.b"}}, + "cPos": 1}}, {"sortText": "1", "label": "x", "kind": 5, @@ -100,7 +110,8 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 20, "character": 5}}, - "id": {"const": {"declName": "S.x"}}}}, + "id": {"const": {"declName": "S.x"}}, + "cPos": 1}}, {"sortText": "2", "label": "y", "kind": 5, @@ -108,5 +119,6 @@ {"params": {"textDocument": {"uri": "file:///completion3.lean"}, "position": {"line": 20, "character": 5}}, - "id": {"const": {"declName": "S.y"}}}}], + "id": {"const": {"declName": "S.y"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion4.lean.expected.out b/tests/lean/interactive/completion4.lean.expected.out index f13560912614..822a7de2f452 100644 --- a/tests/lean/interactive/completion4.lean.expected.out +++ b/tests/lean/interactive/completion4.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 7, "character": 4}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "fn2", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 7, "character": 4}}, - "id": {"const": {"declName": "S.fn2"}}}}, + "id": {"const": {"declName": "S.fn2"}}, + "cPos": 1}}, {"sortText": "2", "label": "pred", "kind": 5, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 7, "character": 4}}, - "id": {"const": {"declName": "S.pred"}}}}], + "id": {"const": {"declName": "S.pred"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 11, "character": 10}} @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 11, "character": 10}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "fn2", "kind": 5, @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 11, "character": 10}}, - "id": {"const": {"declName": "S.fn2"}}}}, + "id": {"const": {"declName": "S.fn2"}}, + "cPos": 1}}, {"sortText": "2", "label": "pred", "kind": 5, @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 11, "character": 10}}, - "id": {"const": {"declName": "S.pred"}}}}], + "id": {"const": {"declName": "S.pred"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 16, "character": 11}} @@ -64,7 +70,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 16, "character": 11}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "fn2", "kind": 5, @@ -72,7 +79,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 16, "character": 11}}, - "id": {"const": {"declName": "S.fn2"}}}}, + "id": {"const": {"declName": "S.fn2"}}, + "cPos": 1}}, {"sortText": "2", "label": "pred", "kind": 5, @@ -80,7 +88,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 16, "character": 11}}, - "id": {"const": {"declName": "S.pred"}}}}], + "id": {"const": {"declName": "S.pred"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 20, "character": 21}} @@ -92,7 +101,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 20, "character": 21}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "fn2", "kind": 5, @@ -100,7 +110,8 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 20, "character": 21}}, - "id": {"const": {"declName": "S.fn2"}}}}, + "id": {"const": {"declName": "S.fn2"}}, + "cPos": 1}}, {"sortText": "2", "label": "pred", "kind": 5, @@ -108,5 +119,6 @@ {"params": {"textDocument": {"uri": "file:///completion4.lean"}, "position": {"line": 20, "character": 21}}, - "id": {"const": {"declName": "S.pred"}}}}], + "id": {"const": {"declName": "S.pred"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion5.lean.expected.out b/tests/lean/interactive/completion5.lean.expected.out index 152236a17a6b..91ea7851e6ab 100644 --- a/tests/lean/interactive/completion5.lean.expected.out +++ b/tests/lean/interactive/completion5.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion5.lean"}, "position": {"line": 9, "character": 15}}, - "id": {"const": {"declName": "C.b1"}}}}, + "id": {"const": {"declName": "C.b1"}}, + "cPos": 1}}, {"sortText": "1", "label": "f1", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completion5.lean"}, "position": {"line": 9, "character": 15}}, - "id": {"const": {"declName": "C.f1"}}}}, + "id": {"const": {"declName": "C.f1"}}, + "cPos": 1}}, {"sortText": "2", "label": "f2", "kind": 5, @@ -24,5 +26,6 @@ {"params": {"textDocument": {"uri": "file:///completion5.lean"}, "position": {"line": 9, "character": 15}}, - "id": {"const": {"declName": "C.f2"}}}}], + "id": {"const": {"declName": "C.f2"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion6.lean.expected.out b/tests/lean/interactive/completion6.lean.expected.out index f6cb7e7c4fb8..17fdb83289dd 100644 --- a/tests/lean/interactive/completion6.lean.expected.out +++ b/tests/lean/interactive/completion6.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "C.b1"}}}}, + "id": {"const": {"declName": "C.b1"}}, + "cPos": 1}}, {"sortText": "1", "label": "f1", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "C.f1"}}}}, + "id": {"const": {"declName": "C.f1"}}, + "cPos": 1}}, {"sortText": "2", "label": "f2", "kind": 5, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "C.f2"}}}}, + "id": {"const": {"declName": "C.f2"}}, + "cPos": 1}}, {"sortText": "3", "label": "f3", "kind": 5, @@ -32,7 +35,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "D.f3"}}}}, + "id": {"const": {"declName": "D.f3"}}, + "cPos": 1}}, {"sortText": "4", "label": "toC", "kind": 5, @@ -40,7 +44,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "D.toC"}}}}], + "id": {"const": {"declName": "D.toC"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}} @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "C.b1"}}}}, + "id": {"const": {"declName": "C.b1"}}, + "cPos": 1}}, {"sortText": "1", "label": "doubleF1", "kind": 3, @@ -60,7 +66,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "E.doubleF1"}}}}, + "id": {"const": {"declName": "E.doubleF1"}}, + "cPos": 1}}, {"sortText": "2", "label": "f1", "kind": 5, @@ -68,7 +75,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "C.f1"}}}}, + "id": {"const": {"declName": "C.f1"}}, + "cPos": 1}}, {"sortText": "3", "label": "f2", "kind": 5, @@ -76,7 +84,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "C.f2"}}}}, + "id": {"const": {"declName": "C.f2"}}, + "cPos": 1}}, {"sortText": "4", "label": "f3", "kind": 5, @@ -84,7 +93,8 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "D.f3"}}}}, + "id": {"const": {"declName": "D.f3"}}, + "cPos": 1}}, {"sortText": "5", "label": "toC", "kind": 5, @@ -92,5 +102,6 @@ {"params": {"textDocument": {"uri": "file:///completion6.lean"}, "position": {"line": 21, "character": 4}}, - "id": {"const": {"declName": "D.toC"}}}}], + "id": {"const": {"declName": "D.toC"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completion7.lean.expected.out b/tests/lean/interactive/completion7.lean.expected.out index 8b6d13ec7fef..0538b9a0b47c 100644 --- a/tests/lean/interactive/completion7.lean.expected.out +++ b/tests/lean/interactive/completion7.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completion7.lean"}, "position": {"line": 6, "character": 10}}, - "id": {"const": {"declName": "And"}}}}], + "id": {"const": {"declName": "And"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completion7.lean"}, "position": {"line": 8, "character": 11}} @@ -20,7 +21,8 @@ {"params": {"textDocument": {"uri": "file:///completion7.lean"}, "position": {"line": 8, "character": 11}}, - "id": {"const": {"declName": "And.left"}}}}, + "id": {"const": {"declName": "And.left"}}, + "cPos": 0}}, {"sortText": "1", "label": "mk", "kind": 4, @@ -28,7 +30,8 @@ {"params": {"textDocument": {"uri": "file:///completion7.lean"}, "position": {"line": 8, "character": 11}}, - "id": {"const": {"declName": "And.mk"}}}}, + "id": {"const": {"declName": "And.mk"}}, + "cPos": 0}}, {"sortText": "2", "label": "right", "kind": 5, @@ -36,5 +39,6 @@ {"params": {"textDocument": {"uri": "file:///completion7.lean"}, "position": {"line": 8, "character": 11}}, - "id": {"const": {"declName": "And.right"}}}}], + "id": {"const": {"declName": "And.right"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionAtPrint.lean.expected.out b/tests/lean/interactive/completionAtPrint.lean.expected.out index bb2fc2b59efb..61ce71e1ed67 100644 --- a/tests/lean/interactive/completionAtPrint.lean.expected.out +++ b/tests/lean/interactive/completionAtPrint.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionAtPrint.lean"}, "position": {"line": 2, "character": 25}}, - "id": {"const": {"declName": "Lean.Environment.find?"}}}}, + "id": {"const": {"declName": "Lean.Environment.find?"}}, + "cPos": 0}}, {"sortText": "1", "label": "freeRegions", "kind": 3, @@ -20,5 +21,6 @@ {"params": {"textDocument": {"uri": "file:///completionAtPrint.lean"}, "position": {"line": 2, "character": 25}}, - "id": {"const": {"declName": "Lean.Environment.freeRegions"}}}}], + "id": {"const": {"declName": "Lean.Environment.freeRegions"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionCheck.lean.expected.out b/tests/lean/interactive/completionCheck.lean.expected.out index af03d40b5fb9..8c799d1af312 100644 --- a/tests/lean/interactive/completionCheck.lean.expected.out +++ b/tests/lean/interactive/completionCheck.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionCheck.lean"}, "position": {"line": 6, "character": 33}}, - "id": {"const": {"declName": "AVerySpecificStructureName"}}}}, + "id": {"const": {"declName": "AVerySpecificStructureName"}}, + "cPos": 0}}, {"sortText": "1", "label": "AVerySpecificStructureName2", "kind": 22, @@ -16,5 +17,6 @@ {"params": {"textDocument": {"uri": "file:///completionCheck.lean"}, "position": {"line": 6, "character": 33}}, - "id": {"const": {"declName": "AVerySpecificStructureName2"}}}}], + "id": {"const": {"declName": "AVerySpecificStructureName2"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionDeprecation.lean.expected.out b/tests/lean/interactive/completionDeprecation.lean.expected.out index 9c36ff48dbe0..1550a8f6ea53 100644 --- a/tests/lean/interactive/completionDeprecation.lean.expected.out +++ b/tests/lean/interactive/completionDeprecation.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo1"}}}}, + "id": {"const": {"declName": "SomeStructure.foo1"}}, + "cPos": 0}}, {"sortText": "1", "label": "foo2", "kind": 3, @@ -17,7 +18,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo2"}}}}, + "id": {"const": {"declName": "SomeStructure.foo2"}}, + "cPos": 0}}, {"tags": [1], "sortText": "2", "label": "foo3", @@ -28,7 +30,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo3"}}}}, + "id": {"const": {"declName": "SomeStructure.foo3"}}, + "cPos": 0}}, {"tags": [1], "sortText": "3", "label": "foo4", @@ -40,7 +43,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo4"}}}}, + "id": {"const": {"declName": "SomeStructure.foo4"}}, + "cPos": 0}}, {"tags": [1], "sortText": "4", "label": "foo5", @@ -53,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo5"}}}}, + "id": {"const": {"declName": "SomeStructure.foo5"}}, + "cPos": 0}}, {"tags": [1], "sortText": "5", "label": "foo6", @@ -66,7 +71,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo6"}}}}, + "id": {"const": {"declName": "SomeStructure.foo6"}}, + "cPos": 0}}, {"tags": [1], "sortText": "6", "label": "foo7", @@ -79,7 +85,8 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo7"}}}}, + "id": {"const": {"declName": "SomeStructure.foo7"}}, + "cPos": 0}}, {"tags": [1], "sortText": "7", "label": "foo8", @@ -92,5 +99,6 @@ {"params": {"textDocument": {"uri": "file:///completionDeprecation.lean"}, "position": {"line": 25, "character": 28}}, - "id": {"const": {"declName": "SomeStructure.foo8"}}}}], + "id": {"const": {"declName": "SomeStructure.foo8"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionDocString.lean.expected.out b/tests/lean/interactive/completionDocString.lean.expected.out index 87577cd094a1..90ef754cdc1a 100644 --- a/tests/lean/interactive/completionDocString.lean.expected.out +++ b/tests/lean/interactive/completionDocString.lean.expected.out @@ -10,7 +10,8 @@ {"params": {"textDocument": {"uri": "file:///completionDocString.lean"}, "position": {"line": 0, "character": 20}}, - "id": {"const": {"declName": "Array.insertAt"}}}}, + "id": {"const": {"declName": "Array.insertAt"}}, + "cPos": 0}}, {"sortText": "1", "label": "insertAt!", "kind": 3, @@ -22,5 +23,6 @@ {"params": {"textDocument": {"uri": "file:///completionDocString.lean"}, "position": {"line": 0, "character": 20}}, - "id": {"const": {"declName": "Array.insertAt!"}}}}], + "id": {"const": {"declName": "Array.insertAt!"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionEOF.lean.expected.out b/tests/lean/interactive/completionEOF.lean.expected.out index bffabbbf7909..d03cf05aa28c 100644 --- a/tests/lean/interactive/completionEOF.lean.expected.out +++ b/tests/lean/interactive/completionEOF.lean.expected.out @@ -8,5 +8,6 @@ {"params": {"textDocument": {"uri": "file:///completionEOF.lean"}, "position": {"line": 8, "character": 9}}, - "id": {"const": {"declName": "And"}}}}], + "id": {"const": {"declName": "And"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionFallback.lean.expected.out b/tests/lean/interactive/completionFallback.lean.expected.out index 0d524243dc7b..a8841a0ad32d 100644 --- a/tests/lean/interactive/completionFallback.lean.expected.out +++ b/tests/lean/interactive/completionFallback.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.down"}}}}, + "id": {"const": {"declName": "Direction.down"}}, + "cPos": 0}}, {"sortText": "1", "label": "left", "kind": 4, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.left"}}}}, + "id": {"const": {"declName": "Direction.left"}}, + "cPos": 0}}, {"sortText": "2", "label": "noConfusionType", "kind": 3, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.noConfusionType"}}}}, + "id": {"const": {"declName": "Direction.noConfusionType"}}, + "cPos": 0}}, {"sortText": "3", "label": "right", "kind": 4, @@ -32,7 +35,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.right"}}}}, + "id": {"const": {"declName": "Direction.right"}}, + "cPos": 0}}, {"sortText": "4", "label": "toCtorIdx", "kind": 3, @@ -40,7 +44,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.toCtorIdx"}}}}, + "id": {"const": {"declName": "Direction.toCtorIdx"}}, + "cPos": 0}}, {"sortText": "5", "label": "up", "kind": 4, @@ -48,7 +53,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 14, "character": 14}}, - "id": {"const": {"declName": "Direction.up"}}}}], + "id": {"const": {"declName": "Direction.up"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 28, "character": 30}} @@ -60,7 +66,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 28, "character": 30}}, - "id": {"const": {"declName": "CustomAnd.ha"}}}}, + "id": {"const": {"declName": "CustomAnd.ha"}}, + "cPos": 0}}, {"sortText": "1", "label": "hb", "kind": 23, @@ -68,7 +75,8 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 28, "character": 30}}, - "id": {"const": {"declName": "CustomAnd.hb"}}}}, + "id": {"const": {"declName": "CustomAnd.hb"}}, + "cPos": 0}}, {"sortText": "2", "label": "mk", "kind": 4, @@ -76,5 +84,6 @@ {"params": {"textDocument": {"uri": "file:///completionFallback.lean"}, "position": {"line": 28, "character": 30}}, - "id": {"const": {"declName": "CustomAnd.mk"}}}}], + "id": {"const": {"declName": "CustomAnd.mk"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionFromExpectedType.lean.expected.out b/tests/lean/interactive/completionFromExpectedType.lean.expected.out index 938821b1d38c..14df3b78f6dd 100644 --- a/tests/lean/interactive/completionFromExpectedType.lean.expected.out +++ b/tests/lean/interactive/completionFromExpectedType.lean.expected.out @@ -8,5 +8,6 @@ {"params": {"textDocument": {"uri": "file:///completionFromExpectedType.lean"}, "position": {"line": 3, "character": 18}}, - "id": {"const": {"declName": "Foo.mk"}}}}], + "id": {"const": {"declName": "Foo.mk"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionIStr.lean.expected.out b/tests/lean/interactive/completionIStr.lean.expected.out index 2da7e5d5b6be..7869a2c34d73 100644 --- a/tests/lean/interactive/completionIStr.lean.expected.out +++ b/tests/lean/interactive/completionIStr.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionIStr.lean"}, "position": {"line": 5, "character": 34}}, - "id": {"const": {"declName": "C.b1"}}}}, + "id": {"const": {"declName": "C.b1"}}, + "cPos": 1}}, {"sortText": "1", "label": "f1", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///completionIStr.lean"}, "position": {"line": 5, "character": 34}}, - "id": {"const": {"declName": "C.f1"}}}}, + "id": {"const": {"declName": "C.f1"}}, + "cPos": 1}}, {"sortText": "2", "label": "f2", "kind": 5, @@ -24,5 +26,6 @@ {"params": {"textDocument": {"uri": "file:///completionIStr.lean"}, "position": {"line": 5, "character": 34}}, - "id": {"const": {"declName": "C.f2"}}}}], + "id": {"const": {"declName": "C.f2"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionOpenNamespaces.lean.expected.out b/tests/lean/interactive/completionOpenNamespaces.lean.expected.out index 468aaf3ced85..7fdbd1f930d1 100644 --- a/tests/lean/interactive/completionOpenNamespaces.lean.expected.out +++ b/tests/lean/interactive/completionOpenNamespaces.lean.expected.out @@ -10,6 +10,6 @@ "position": {"line": 4, "character": 68}}, "id": {"const": - {"declName": - "A.B1.verySpecificDefinitionNameOfCompletionOpenNamespaces"}}}}], + {"declName": "A.B1.verySpecificDefinitionNameOfCompletionOpenNamespaces"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionOption.lean.expected.out b/tests/lean/interactive/completionOption.lean.expected.out index 31d9a27f4b03..ffda8804a4df 100644 --- a/tests/lean/interactive/completionOption.lean.expected.out +++ b/tests/lean/interactive/completionOption.lean.expected.out @@ -17,7 +17,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -33,7 +34,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -49,7 +51,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -65,7 +68,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -81,7 +85,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -98,7 +103,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}, + "position": {"line": 1, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 1, "character": 11}, @@ -115,7 +121,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 1, "character": 17}}}}], + "position": {"line": 1, "character": 17}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionOption.lean"}, "position": {"line": 4, "character": 20}} @@ -135,7 +142,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 4, "character": 20}}}}, + "position": {"line": 4, "character": 20}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 4, "character": 11}, @@ -151,7 +159,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 4, "character": 20}}}}, + "position": {"line": 4, "character": 20}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 4, "character": 11}, @@ -168,7 +177,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 4, "character": 20}}}}], + "position": {"line": 4, "character": 20}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionOption.lean"}, "position": {"line": 7, "character": 23}} @@ -189,7 +199,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 7, "character": 23}}}}, + "position": {"line": 7, "character": 23}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 7, "character": 11}, @@ -206,7 +217,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 7, "character": 23}}}}, + "position": {"line": 7, "character": 23}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 7, "character": 11}, @@ -223,7 +235,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 7, "character": 23}}}}, + "position": {"line": 7, "character": 23}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 7, "character": 11}, @@ -240,7 +253,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 7, "character": 23}}}}], + "position": {"line": 7, "character": 23}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionOption.lean"}, "position": {"line": 10, "character": 27}} @@ -261,7 +275,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 10, "character": 27}}}}, + "position": {"line": 10, "character": 27}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 10, "character": 11}, @@ -278,7 +293,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 10, "character": 27}}}}, + "position": {"line": 10, "character": 27}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 10, "character": 11}, @@ -295,7 +311,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 10, "character": 27}}}}, + "position": {"line": 10, "character": 27}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 10, "character": 11}, @@ -312,7 +329,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 10, "character": 27}}}}], + "position": {"line": 10, "character": 27}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionOption.lean"}, "position": {"line": 13, "character": 17}} @@ -333,7 +351,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -349,7 +368,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -365,7 +385,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -381,7 +402,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -397,7 +419,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -414,7 +437,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}, + "position": {"line": 13, "character": 17}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 13, "character": 11}, @@ -431,7 +455,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 13, "character": 17}}}}], + "position": {"line": 13, "character": 17}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionOption.lean"}, "position": {"line": 16, "character": 18}} @@ -451,7 +476,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}, + "position": {"line": 16, "character": 18}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 16, "character": 11}, @@ -467,7 +493,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}, + "position": {"line": 16, "character": 18}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 16, "character": 11}, @@ -483,7 +510,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}, + "position": {"line": 16, "character": 18}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 16, "character": 11}, @@ -499,7 +527,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}, + "position": {"line": 16, "character": 18}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 16, "character": 11}, @@ -516,7 +545,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}, + "position": {"line": 16, "character": 18}}, + "cPos": 0}}, {"textEdit": {"replace": {"start": {"line": 16, "character": 11}, @@ -533,5 +563,6 @@ "data": {"params": {"textDocument": {"uri": "file:///completionOption.lean"}, - "position": {"line": 16, "character": 18}}}}], + "position": {"line": 16, "character": 18}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionPrefixIssue.lean.expected.out b/tests/lean/interactive/completionPrefixIssue.lean.expected.out index b56a773a16a6..8496fe5d798a 100644 --- a/tests/lean/interactive/completionPrefixIssue.lean.expected.out +++ b/tests/lean/interactive/completionPrefixIssue.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrefixIssue.lean"}, "position": {"line": 1, "character": 64}}, - "id": {"fvar": {"id": "_uniq.29"}}}}, + "id": {"fvar": {"id": "_uniq.29"}}, + "cPos": 0}}, {"sortText": "1", "label": "veryLongDefinitionNameVeryLongDefinitionName", "kind": 21, @@ -17,5 +18,6 @@ {"textDocument": {"uri": "file:///completionPrefixIssue.lean"}, "position": {"line": 1, "character": 64}}, "id": - {"const": {"declName": "veryLongDefinitionNameVeryLongDefinitionName"}}}}], + {"const": {"declName": "veryLongDefinitionNameVeryLongDefinitionName"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionPrivateTypes.lean.expected.out b/tests/lean/interactive/completionPrivateTypes.lean.expected.out index 15004871556b..e364ff1f7a1b 100644 --- a/tests/lean/interactive/completionPrivateTypes.lean.expected.out +++ b/tests/lean/interactive/completionPrivateTypes.lean.expected.out @@ -8,5 +8,6 @@ {"params": {"textDocument": {"uri": "file:///completionPrivateTypes.lean"}, "position": {"line": 3, "character": 26}}, - "id": {"const": {"declName": "_private.0.Foo.x"}}}}], + "id": {"const": {"declName": "_private.0.Foo.x"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionPrv.lean.expected.out b/tests/lean/interactive/completionPrv.lean.expected.out index ac03b62acc90..319c5da1fd45 100644 --- a/tests/lean/interactive/completionPrv.lean.expected.out +++ b/tests/lean/interactive/completionPrv.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 2, "character": 11}}, - "id": {"const": {"declName": "_private.0.blaBlaBoo"}}}}], + "id": {"const": {"declName": "_private.0.blaBlaBoo"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}} @@ -20,7 +21,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "_private.0.Foo.booBoo"}}}}, + "id": {"const": {"declName": "_private.0.Foo.booBoo"}}, + "cPos": 0}}, {"sortText": "1", "label": "instToBoolBool", "kind": 21, @@ -28,7 +30,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "instToBoolBool"}}}}, + "id": {"const": {"declName": "instToBoolBool"}}, + "cPos": 0}}, {"sortText": "2", "label": "BitVec.getElem_ofBoolListBE", "kind": 23, @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.getElem_ofBoolListBE"}}}}, + "id": {"const": {"declName": "BitVec.getElem_ofBoolListBE"}}, + "cPos": 0}}, {"sortText": "3", "label": "BitVec.getLsbD_ofBoolListBE", "kind": 23, @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.getLsbD_ofBoolListBE"}}}}, + "id": {"const": {"declName": "BitVec.getLsbD_ofBoolListBE"}}, + "cPos": 0}}, {"sortText": "4", "label": "BitVec.getMsbD_ofBoolListBE", "kind": 23, @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.getMsbD_ofBoolListBE"}}}}, + "id": {"const": {"declName": "BitVec.getMsbD_ofBoolListBE"}}, + "cPos": 0}}, {"sortText": "5", "label": "BitVec.ofBool_and_ofBool", "kind": 23, @@ -60,7 +66,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.ofBool_and_ofBool"}}}}, + "id": {"const": {"declName": "BitVec.ofBool_and_ofBool"}}, + "cPos": 0}}, {"sortText": "6", "label": "BitVec.ofBool_or_ofBool", "kind": 23, @@ -68,7 +75,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.ofBool_or_ofBool"}}}}, + "id": {"const": {"declName": "BitVec.ofBool_or_ofBool"}}, + "cPos": 0}}, {"sortText": "7", "label": "BitVec.ofBool_xor_ofBool", "kind": 23, @@ -76,7 +84,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.ofBool_xor_ofBool"}}}}, + "id": {"const": {"declName": "BitVec.ofBool_xor_ofBool"}}, + "cPos": 0}}, {"sortText": "8", "label": "BitVec.ofBoolListBE", "kind": 3, @@ -87,7 +96,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "BitVec.ofBoolListBE"}}}}], + "id": {"const": {"declName": "BitVec.ofBoolListBE"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 21, "character": 5}} @@ -99,7 +109,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 21, "character": 5}}, - "id": {"const": {"declName": "S.field1"}}}}, + "id": {"const": {"declName": "S.field1"}}, + "cPos": 1}}, {"sortText": "1", "label": "getInc", "kind": 3, @@ -107,7 +118,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 21, "character": 5}}, - "id": {"const": {"declName": "_private.0.S.getInc"}}}}], + "id": {"const": {"declName": "_private.0.S.getInc"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 25, "character": 4}} @@ -119,7 +131,8 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 25, "character": 4}}, - "id": {"const": {"declName": "S.field1"}}}}, + "id": {"const": {"declName": "S.field1"}}, + "cPos": 1}}, {"sortText": "1", "label": "getInc", "kind": 3, @@ -127,5 +140,6 @@ {"params": {"textDocument": {"uri": "file:///completionPrv.lean"}, "position": {"line": 25, "character": 4}}, - "id": {"const": {"declName": "_private.0.S.getInc"}}}}], + "id": {"const": {"declName": "_private.0.S.getInc"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/completionStructureInstance.lean b/tests/lean/interactive/completionStructureInstance.lean new file mode 100644 index 000000000000..ff79f641c6bb --- /dev/null +++ b/tests/lean/interactive/completionStructureInstance.lean @@ -0,0 +1,96 @@ +import Std.Data.HashSet + +structure S' where + foo : Nat + bar : Nat + +structure S extends S' where + foobar : Nat + barfoo : Nat + +example : S where -- No completions expected + --^ textDocument/completion + +example : S where -- All field completions expected + --^ textDocument/completion + +example : S where + -- All field completions expected +--^ textDocument/completion + +example : S where + f -- All field completions matching `f` expected + --^ textDocument/completion + +example : S where + foo -- All field completions matching `foo` expected + --^ textDocument/completion + +example : S where + foo := -- No completions expected + --^ textDocument/completion + +example : S where + foo := + -- No completions expected + --^ textDocument/completion + +example : S where + foo := 1 + -- All field completions expected +--^ textDocument/completion + +example : S where + foo := 1; -- All field completions expected + --^ textDocument/completion + +example : S := { } -- All field completions expected + --^ textDocument/completion + +example : S := { + -- All field completions expected +--^ textDocument/completion +} + +example : S := { + f -- All field completions matching `f` expected + --^ textDocument/completion +} + +example : S := { + foo -- All field completions matching `foo` expected + --^ textDocument/completion +} + +example : S := { + foo := + -- No completions expected + --^ textDocument/completion +} + +example : S := { + foo := 1 + -- All field completions expected +--^ textDocument/completion +} + +example : S := { + foo := 1 + -- All field completions expected +--^ textDocument/completion +} + +example : S := { foo := 1, } -- All field completions expected + --^ textDocument/completion + +example (s : S) : S := { s with } -- All field completions expected + --^ textDocument/completion + +example (s : S) : S := { s with : S } -- All field completions expected + --^ textDocument/completion + +example (s : S) : S := { s with f } -- All field completions matching `f` expected + --^ textDocument/completion + +example (aLongUniqueIdentifier : Nat) : Std.HashSet Nat := { aLongUniqueIdentifier } -- Identifier completion matching `aLongUniqueIdentifier` + --^ textDocument/completion diff --git a/tests/lean/interactive/completionStructureInstance.lean.expected.out b/tests/lean/interactive/completionStructureInstance.lean.expected.out new file mode 100644 index 000000000000..aab84e36c0f5 --- /dev/null +++ b/tests/lean/interactive/completionStructureInstance.lean.expected.out @@ -0,0 +1,612 @@ +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 10, "character": 17}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 10, "character": 17}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 10, "character": 17}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 10, "character": 17}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 10, "character": 17}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 13, "character": 18}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 13, "character": 18}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 13, "character": 18}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 13, "character": 18}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 13, "character": 18}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 17, "character": 2}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 17, "character": 2}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 17, "character": 2}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 17, "character": 2}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 17, "character": 2}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 21, "character": 3}} +{"items": + [{"sortText": "0", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 21, "character": 3}}, + "cPos": 0}}, + {"sortText": "1", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 21, "character": 3}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 25, "character": 5}} +{"items": + [{"sortText": "0", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 25, "character": 5}}, + "cPos": 0}}, + {"sortText": "1", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 25, "character": 5}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 29, "character": 9}} +{"items": [], "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 34, "character": 4}} +{"items": [], "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 39, "character": 2}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 39, "character": 2}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 39, "character": 2}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 39, "character": 2}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 39, "character": 2}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 43, "character": 12}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 43, "character": 12}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 43, "character": 12}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 43, "character": 12}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 43, "character": 12}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 46, "character": 17}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 46, "character": 17}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 46, "character": 17}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 46, "character": 17}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 46, "character": 17}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 50, "character": 2}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 50, "character": 2}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 50, "character": 2}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 50, "character": 2}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 50, "character": 2}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 55, "character": 3}} +{"items": + [{"sortText": "0", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 55, "character": 3}}, + "cPos": 1}}, + {"sortText": "1", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 55, "character": 3}}, + "cPos": 1}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 60, "character": 5}} +{"items": + [{"sortText": "0", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 60, "character": 5}}, + "cPos": 1}}, + {"sortText": "1", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 60, "character": 5}}, + "cPos": 1}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 66, "character": 4}} +{"items": [], "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 72, "character": 2}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 72, "character": 2}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 72, "character": 2}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 72, "character": 2}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 72, "character": 2}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 78, "character": 2}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 78, "character": 2}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 78, "character": 2}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 78, "character": 2}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 78, "character": 2}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 82, "character": 27}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 82, "character": 27}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 82, "character": 27}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 82, "character": 27}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 82, "character": 27}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 85, "character": 32}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 85, "character": 32}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 85, "character": 32}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 85, "character": 32}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 85, "character": 32}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 88, "character": 32}} +{"items": + [{"sortText": "0", + "label": "bar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 88, "character": 32}}, + "cPos": 0}}, + {"sortText": "1", + "label": "barfoo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 88, "character": 32}}, + "cPos": 0}}, + {"sortText": "2", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 88, "character": 32}}, + "cPos": 0}}, + {"sortText": "3", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 88, "character": 32}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 91, "character": 33}} +{"items": + [{"sortText": "0", + "label": "foo", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 91, "character": 33}}, + "cPos": 0}}, + {"sortText": "1", + "label": "foobar", + "kind": 5, + "detail": "field", + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 91, "character": 33}}, + "cPos": 0}}], + "isIncomplete": true} +{"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 94, "character": 82}} +{"items": + [{"sortText": "0", + "label": "aLongUniqueIdentifier", + "kind": 6, + "data": + {"params": + {"textDocument": {"uri": "file:///completionStructureInstance.lean"}, + "position": {"line": 94, "character": 82}}, + "id": {"fvar": {"id": "_uniq.1028"}}, + "cPos": 0}}], + "isIncomplete": true} diff --git a/tests/lean/interactive/completionTactics.lean.expected.out b/tests/lean/interactive/completionTactics.lean.expected.out index f68e2e2f77c5..dbe081fa5785 100644 --- a/tests/lean/interactive/completionTactics.lean.expected.out +++ b/tests/lean/interactive/completionTactics.lean.expected.out @@ -11,7 +11,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -22,7 +23,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -33,14 +35,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -51,7 +55,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -62,14 +67,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}, + "position": {"line": 23, "character": 21}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -77,7 +84,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 23, "character": 21}}}}], + "position": {"line": 23, "character": 21}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 26, "character": 24}} @@ -89,7 +97,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -100,7 +109,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -111,14 +121,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -129,7 +141,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -140,14 +153,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}, + "position": {"line": 26, "character": 24}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -155,7 +170,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 26, "character": 24}}}}], + "position": {"line": 26, "character": 24}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 29, "character": 25}} @@ -170,7 +186,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -181,7 +198,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -192,14 +210,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -210,7 +230,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -221,14 +242,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}, + "position": {"line": 32, "character": 26}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -236,7 +259,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 32, "character": 26}}}}], + "position": {"line": 32, "character": 26}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 35, "character": 27}} @@ -248,7 +272,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -259,7 +284,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -270,14 +296,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -288,7 +316,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -299,14 +328,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}, + "position": {"line": 35, "character": 27}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -314,7 +345,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 35, "character": 27}}}}], + "position": {"line": 35, "character": 27}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 40, "character": 7}} @@ -326,7 +358,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -337,7 +370,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -348,14 +382,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -366,7 +402,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -377,14 +414,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}, + "position": {"line": 40, "character": 7}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -392,7 +431,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 40, "character": 7}}}}], + "position": {"line": 40, "character": 7}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 44, "character": 2}} @@ -404,7 +444,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -415,7 +456,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -426,14 +468,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -444,7 +488,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -455,14 +500,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}, + "position": {"line": 44, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -470,7 +517,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 44, "character": 2}}}}], + "position": {"line": 44, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 49, "character": 2}} @@ -482,7 +530,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -493,7 +542,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -504,14 +554,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -522,7 +574,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -533,14 +586,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}, + "position": {"line": 49, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -548,7 +603,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 49, "character": 2}}}}], + "position": {"line": 49, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 53, "character": 2}} @@ -560,7 +616,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -571,7 +628,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -582,14 +640,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -600,7 +660,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -611,14 +672,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}, + "position": {"line": 53, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -626,7 +689,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 53, "character": 2}}}}], + "position": {"line": 53, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 59, "character": 4}} @@ -638,7 +702,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -649,7 +714,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -660,14 +726,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -678,7 +746,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -689,14 +758,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}, + "position": {"line": 59, "character": 4}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -704,7 +775,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 59, "character": 4}}}}], + "position": {"line": 59, "character": 4}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 64, "character": 2}} @@ -716,7 +788,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -727,7 +800,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -738,14 +812,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -756,7 +832,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -767,14 +844,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}, + "position": {"line": 64, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -782,7 +861,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 64, "character": 2}}}}], + "position": {"line": 64, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 70, "character": 4}} @@ -794,7 +874,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -805,7 +886,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -816,14 +898,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -834,7 +918,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -845,14 +930,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}, + "position": {"line": 70, "character": 4}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -860,7 +947,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 70, "character": 4}}}}], + "position": {"line": 70, "character": 4}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 76, "character": 2}} @@ -872,7 +960,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -883,7 +972,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -894,14 +984,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -912,7 +1004,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -923,14 +1016,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}, + "position": {"line": 76, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -938,7 +1033,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 76, "character": 2}}}}], + "position": {"line": 76, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 81, "character": 4}} @@ -953,7 +1049,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -964,7 +1061,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -975,14 +1073,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -993,7 +1093,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -1004,14 +1105,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}, + "position": {"line": 86, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -1019,7 +1122,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 86, "character": 2}}}}], + "position": {"line": 86, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 91, "character": 4}} @@ -1031,7 +1135,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -1042,7 +1147,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -1053,14 +1159,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -1071,7 +1179,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -1082,14 +1191,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}, + "position": {"line": 91, "character": 4}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -1097,7 +1208,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 91, "character": 4}}}}], + "position": {"line": 91, "character": 4}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 96, "character": 2}} @@ -1109,7 +1221,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -1120,7 +1233,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -1131,14 +1245,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -1149,7 +1265,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -1160,14 +1277,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}, + "position": {"line": 96, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -1175,7 +1294,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 96, "character": 2}}}}], + "position": {"line": 96, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 102, "character": 4}} @@ -1187,7 +1307,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -1198,7 +1319,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -1209,14 +1331,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -1227,7 +1351,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -1238,14 +1363,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}, + "position": {"line": 102, "character": 4}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -1253,7 +1380,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 102, "character": 4}}}}], + "position": {"line": 102, "character": 4}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///completionTactics.lean"}, "position": {"line": 108, "character": 2}} @@ -1268,7 +1396,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "1", "label": "Lean.Parser.Tactic.introMatch", "kind": 14, @@ -1279,7 +1408,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "2", "label": "Lean.Parser.Tactic.match", "kind": 14, @@ -1290,14 +1420,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "3", "label": "Lean.Parser.Tactic.nestedTactic", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "4", "label": "Lean.Parser.Tactic.open", "kind": 14, @@ -1308,7 +1440,8 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "5", "label": "Lean.Parser.Tactic.set_option", "kind": 14, @@ -1319,14 +1452,16 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "6", "label": "Lean.Parser.Tactic.unknown", "kind": 14, "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}, + "position": {"line": 112, "character": 2}}, + "cPos": 0}}, {"sortText": "7", "label": "skip", "kind": 14, @@ -1334,5 +1469,6 @@ "data": {"params": {"textDocument": {"uri": "file:///completionTactics.lean"}, - "position": {"line": 112, "character": 2}}}}], + "position": {"line": 112, "character": 2}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/dotIdCompletion.lean.expected.out b/tests/lean/interactive/dotIdCompletion.lean.expected.out index 1cca64cc656f..75efda661314 100644 --- a/tests/lean/interactive/dotIdCompletion.lean.expected.out +++ b/tests/lean/interactive/dotIdCompletion.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///dotIdCompletion.lean"}, "position": {"line": 4, "character": 5}}, - "id": {"const": {"declName": "Boo.true"}}}}, + "id": {"const": {"declName": "Boo.true"}}, + "cPos": 0}}, {"sortText": "1", "label": "truth", "kind": 4, @@ -16,5 +17,6 @@ {"params": {"textDocument": {"uri": "file:///dotIdCompletion.lean"}, "position": {"line": 4, "character": 5}}, - "id": {"const": {"declName": "Boo.truth"}}}}], + "id": {"const": {"declName": "Boo.truth"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/editCompletion.lean.expected.out b/tests/lean/interactive/editCompletion.lean.expected.out index 104ba7b4bf6e..56deb0a4c05e 100644 --- a/tests/lean/interactive/editCompletion.lean.expected.out +++ b/tests/lean/interactive/editCompletion.lean.expected.out @@ -8,5 +8,6 @@ {"params": {"textDocument": {"uri": "file:///editCompletion.lean"}, "position": {"line": 3, "character": 22}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/inWordCompletion.lean.expected.out b/tests/lean/interactive/inWordCompletion.lean.expected.out index 0c61f227a910..4859dcb72f74 100644 --- a/tests/lean/interactive/inWordCompletion.lean.expected.out +++ b/tests/lean/interactive/inWordCompletion.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 4, "character": 11}}, - "id": {"const": {"declName": "gfxabc"}}}}, + "id": {"const": {"declName": "gfxabc"}}, + "cPos": 0}}, {"sortText": "1", "label": "gfxacc", "kind": 3, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 4, "character": 11}}, - "id": {"const": {"declName": "gfxacc"}}}}, + "id": {"const": {"declName": "gfxacc"}}, + "cPos": 0}}, {"sortText": "2", "label": "gfxadc", "kind": 3, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 4, "character": 11}}, - "id": {"const": {"declName": "gfxadc"}}}}], + "id": {"const": {"declName": "gfxadc"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 12, "character": 15}} @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "Boo.gfxabc"}}}}, + "id": {"const": {"declName": "Boo.gfxabc"}}, + "cPos": 0}}, {"sortText": "1", "label": "gfxacc", "kind": 3, @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "Boo.gfxacc"}}}}, + "id": {"const": {"declName": "Boo.gfxacc"}}, + "cPos": 0}}, {"sortText": "2", "label": "gfxadc", "kind": 3, @@ -52,5 +57,6 @@ {"params": {"textDocument": {"uri": "file:///inWordCompletion.lean"}, "position": {"line": 12, "character": 15}}, - "id": {"const": {"declName": "Boo.gfxadc"}}}}], + "id": {"const": {"declName": "Boo.gfxadc"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/internalNamesIssue.lean.expected.out b/tests/lean/interactive/internalNamesIssue.lean.expected.out index d2e2dd36fd53..36d36c565f36 100644 --- a/tests/lean/interactive/internalNamesIssue.lean.expected.out +++ b/tests/lean/interactive/internalNamesIssue.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///internalNamesIssue.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "Foo.bla"}}}}, + "id": {"const": {"declName": "Foo.bla"}}, + "cPos": 0}}, {"sortText": "1", "label": "foo", "kind": 3, @@ -16,5 +17,6 @@ {"params": {"textDocument": {"uri": "file:///internalNamesIssue.lean"}, "position": {"line": 9, "character": 11}}, - "id": {"const": {"declName": "Foo.foo"}}}}], + "id": {"const": {"declName": "Foo.foo"}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/keywordCompletion.lean.expected.out b/tests/lean/interactive/keywordCompletion.lean.expected.out index 19498180d631..ca9f2d221712 100644 --- a/tests/lean/interactive/keywordCompletion.lean.expected.out +++ b/tests/lean/interactive/keywordCompletion.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, "position": {"line": 4, "character": 10}}, - "id": {"const": {"declName": "bin"}}}}, + "id": {"const": {"declName": "bin"}}, + "cPos": 0}}, {"sortText": "1", "label": "binder_predicate", "kind": 14, @@ -16,7 +17,8 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 10}}}}, + "position": {"line": 4, "character": 10}}, + "cPos": 0}}, {"sortText": "2", "label": "binop%", "kind": 14, @@ -24,7 +26,8 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 10}}}}, + "position": {"line": 4, "character": 10}}, + "cPos": 0}}, {"sortText": "3", "label": "binop_lazy%", "kind": 14, @@ -32,7 +35,8 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 10}}}}, + "position": {"line": 4, "character": 10}}, + "cPos": 0}}, {"sortText": "4", "label": "binrel%", "kind": 14, @@ -40,7 +44,8 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 10}}}}, + "position": {"line": 4, "character": 10}}, + "cPos": 0}}, {"sortText": "5", "label": "binrel_no_prop%", "kind": 14, @@ -48,7 +53,8 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 10}}}}], + "position": {"line": 4, "character": 10}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///keywordCompletion.lean"}, "position": {"line": 4, "character": 13}} @@ -60,5 +66,6 @@ "data": {"params": {"textDocument": {"uri": "file:///keywordCompletion.lean"}, - "position": {"line": 4, "character": 13}}}}], + "position": {"line": 4, "character": 13}}, + "cPos": 0}}], "isIncomplete": true} diff --git a/tests/lean/interactive/match.lean.expected.out b/tests/lean/interactive/match.lean.expected.out index 4965eee71f15..3c149c1dd53e 100644 --- a/tests/lean/interactive/match.lean.expected.out +++ b/tests/lean/interactive/match.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 6, "character": 11}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "name", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 6, "character": 11}}, - "id": {"const": {"declName": "S.name"}}}}, + "id": {"const": {"declName": "S.name"}}, + "cPos": 1}}, {"sortText": "2", "label": "value", "kind": 5, @@ -24,7 +26,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 6, "character": 11}}, - "id": {"const": {"declName": "S.value"}}}}], + "id": {"const": {"declName": "S.value"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 10, "character": 10}} @@ -36,7 +39,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 10, "character": 10}}, - "id": {"const": {"declName": "S.fn1"}}}}, + "id": {"const": {"declName": "S.fn1"}}, + "cPos": 1}}, {"sortText": "1", "label": "name", "kind": 5, @@ -44,7 +48,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 10, "character": 10}}, - "id": {"const": {"declName": "S.name"}}}}, + "id": {"const": {"declName": "S.name"}}, + "cPos": 1}}, {"sortText": "2", "label": "value", "kind": 5, @@ -52,7 +57,8 @@ {"params": {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 10, "character": 10}}, - "id": {"const": {"declName": "S.value"}}}}], + "id": {"const": {"declName": "S.value"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///match.lean"}, "position": {"line": 14, "character": 2}} diff --git a/tests/lean/interactive/matchStxCompletion.lean.expected.out b/tests/lean/interactive/matchStxCompletion.lean.expected.out index 6e9e15b0974c..86d0747ddd39 100644 --- a/tests/lean/interactive/matchStxCompletion.lean.expected.out +++ b/tests/lean/interactive/matchStxCompletion.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///matchStxCompletion.lean"}, "position": {"line": 8, "character": 9}}, - "id": {"const": {"declName": "C.b1"}}}}, + "id": {"const": {"declName": "C.b1"}}, + "cPos": 1}}, {"sortText": "1", "label": "f1", "kind": 5, @@ -16,7 +17,8 @@ {"params": {"textDocument": {"uri": "file:///matchStxCompletion.lean"}, "position": {"line": 8, "character": 9}}, - "id": {"const": {"declName": "C.f1"}}}}, + "id": {"const": {"declName": "C.f1"}}, + "cPos": 1}}, {"sortText": "2", "label": "f2", "kind": 5, @@ -24,5 +26,6 @@ {"params": {"textDocument": {"uri": "file:///matchStxCompletion.lean"}, "position": {"line": 8, "character": 9}}, - "id": {"const": {"declName": "C.f2"}}}}], + "id": {"const": {"declName": "C.f2"}}, + "cPos": 1}}], "isIncomplete": true} diff --git a/tests/lean/interactive/travellingCompletions.lean.expected.out b/tests/lean/interactive/travellingCompletions.lean.expected.out index 15c1f442e7d0..b04e4b81cdb9 100644 --- a/tests/lean/interactive/travellingCompletions.lean.expected.out +++ b/tests/lean/interactive/travellingCompletions.lean.expected.out @@ -8,7 +8,8 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 7, "character": 33}}, - "id": {"const": {"declName": "aaaaaaaa"}}}}], + "id": {"const": {"declName": "aaaaaaaa"}}, + "cPos": 0}}], "isIncomplete": true} {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 20, "character": 20}} @@ -20,7 +21,8 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 20, "character": 20}}, - "id": {"const": {"declName": "Bar.foobar"}}}}], + "id": {"const": {"declName": "Bar.foobar"}}, + "cPos": 2}}], "isIncomplete": true} {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 24, "character": 16}} @@ -32,7 +34,8 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 24, "character": 16}}, - "id": {"const": {"declName": "Bar.foobar"}}}}], + "id": {"const": {"declName": "Bar.foobar"}}, + "cPos": 2}}], "isIncomplete": true} {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 38, "character": 45}} @@ -44,7 +47,8 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 38, "character": 45}}, - "id": {"const": {"declName": "Prod.continuousAdd"}}}}, + "id": {"const": {"declName": "Prod.continuousAdd"}}, + "cPos": 1}}, {"sortText": "1", "label": "continuousSMul", "kind": 23, @@ -52,7 +56,8 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 38, "character": 45}}, - "id": {"const": {"declName": "Prod.continuousSMul"}}}}], + "id": {"const": {"declName": "Prod.continuousSMul"}}, + "cPos": 1}}], "isIncomplete": true} {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 42, "character": 25}} @@ -68,5 +73,6 @@ {"params": {"textDocument": {"uri": "file:///travellingCompletions.lean"}, "position": {"line": 42, "character": 25}}, - "id": {"const": {"declName": "True.intro"}}}}], + "id": {"const": {"declName": "True.intro"}}, + "cPos": 1}}], "isIncomplete": true}